# HG changeset patch # User kleing # Date 1432797957 -36000 # Node ID 3f429b7d8eb5c1e7d2321b03a37f22072918802a # Parent 00c06f1315d0fd39497bca754a0ab65fcd04e46a modernized (slightly) type compiler in MicroJava diff -r 00c06f1315d0 -r 3f429b7d8eb5 src/HOL/MicroJava/Comp/AuxLemmas.thy --- a/src/HOL/MicroJava/Comp/AuxLemmas.thy Tue May 26 21:58:04 2015 +0100 +++ b/src/HOL/MicroJava/Comp/AuxLemmas.thy Thu May 28 17:25:57 2015 +1000 @@ -16,30 +16,30 @@ lemma app_nth_greater_len [simp]: - "length pre \ ind \ (pre @ a # post) ! (Suc ind) = (pre @ post) ! ind" -apply (induct pre arbitrary: ind) -apply auto -apply (case_tac ind) -apply auto -done + "length pre \ ind \ (pre @ a # post) ! (Suc ind) = (pre @ post) ! ind" + apply (induct pre arbitrary: ind) + apply clarsimp + apply (case_tac ind) + apply auto + done -lemma length_takeWhile: "v \ set xs \ length (takeWhile (%z. z~=v) xs) < length xs" -by (induct xs) auto +lemma length_takeWhile: "v \ set xs \ length (takeWhile (\z. z \ v) xs) < length xs" + by (induct xs) auto lemma nth_length_takeWhile [simp]: "v \ set xs \ xs ! (length (takeWhile (%z. z~=v) xs)) = v" -by (induct xs) auto + by (induct xs) auto lemma map_list_update [simp]: "\ x \ set xs; distinct xs\ \ - (map f xs) [length (takeWhile (\z. z \ x) xs) := v] = - map (f(x:=v)) xs" -apply (induct xs) -apply simp -apply (case_tac "x=a") -apply auto -done + (map f xs) [length (takeWhile (\z. z \ x) xs) := v] = map (f(x:=v)) xs" + apply (induct xs) + apply simp + apply (rename_tac a xs) + apply (case_tac "x=a") + apply auto + done (**********************************************************************) @@ -47,13 +47,13 @@ (**********************************************************************) -lemma split_compose: "(split f) \ (\ (a,b). ((fa a), (fb b))) = - (\ (a,b). (f (fa a) (fb b)))" -by (simp add: split_def o_def) +lemma split_compose: + "(split f) \ (\ (a,b). ((fa a), (fb b))) = (\ (a,b). (f (fa a) (fb b)))" + by (simp add: split_def o_def) -lemma split_iter: "(\ (a,b,c). ((g1 a), (g2 b), (g3 c))) = - (\ (a,p). ((g1 a), (\ (b, c). ((g2 b), (g3 c))) p))" -by (simp add: split_def o_def) +lemma split_iter: + "(\ (a,b,c). ((g1 a), (g2 b), (g3 c))) = (\ (a,p). ((g1 a), (\ (b, c). ((g2 b), (g3 c))) p))" + by (simp add: split_def o_def) (**********************************************************************) @@ -67,64 +67,51 @@ (**********************************************************************) lemma the_map_upd: "(the \ f(x\v)) = (the \ f)(x:=v)" -by (simp add: fun_eq_iff) + by (simp add: fun_eq_iff) lemma map_of_in_set: "(map_of xs x = None) = (x \ set (map fst xs))" -by (induct xs, auto) + by (induct xs, auto) lemma map_map_upd [simp]: "y \ set xs \ map (the \ f(y\v)) xs = map (the \ f) xs" -by (simp add: the_map_upd) + by (simp add: the_map_upd) lemma map_map_upds [simp]: "(\y\set ys. y \ set xs) \ map (the \ f(ys[\]vs)) xs = map (the \ f) xs" -apply (induct xs arbitrary: f vs) - apply simp -apply fastforce -done + by (induct xs arbitrary: f vs) auto lemma map_upds_distinct [simp]: "distinct ys \ length ys = length vs \ map (the \ f(ys[\]vs)) ys = vs" -apply (induct ys arbitrary: f vs) -apply simp -apply (case_tac vs) -apply simp_all -done + apply (induct ys arbitrary: f vs) + apply simp + apply (case_tac vs) + apply simp_all + done lemma map_of_map_as_map_upd: "distinct (map f zs) \ map_of (map (\ p. (f p, g p)) zs) = empty (map f zs [\] map g zs)" by (induct zs) auto (* In analogy to Map.map_of_SomeD *) -lemma map_upds_SomeD [rule_format (no_asm)]: - "\ m ys. (m(xs[\]ys)) k = Some y \ k \ (set xs) \ (m k = Some y)" -apply (induct xs) -apply simp -apply auto -apply(case_tac ys) -apply auto -done +lemma map_upds_SomeD: + "(m(xs[\]ys)) k = Some y \ k \ (set xs) \ (m k = Some y)" + apply (induct xs arbitrary: m ys) + apply simp + apply (case_tac ys) + apply fastforce+ + done lemma map_of_upds_SomeD: "(map_of m (xs[\]ys)) k = Some y \ k \ (set (xs @ map fst m))" -by (auto dest: map_upds_SomeD map_of_SomeD fst_in_set_lemma) - + by (auto dest: map_upds_SomeD map_of_SomeD fst_in_set_lemma) -lemma map_of_map_prop [rule_format (no_asm)]: - "(map_of (map f xs) k = Some v) \ - (\ x \ set xs. (P1 x)) \ - (\ x. (P1 x) \ (P2 (f x))) \ - (P2(k, v))" -by (induct xs,auto) +lemma map_of_map_prop: + "\map_of (map f xs) k = Some v; \x \ set xs. P1 x; \x. P1 x \ P2 (f x)\ \ P2 (k, v)" + by (induct xs) (auto split: split_if_asm) -lemma map_of_map2: "\ x \ set xs. (fst (f x)) = (fst x) \ +lemma map_of_map2: "\x \ set xs. (fst (f x)) = (fst x) \ map_of (map f xs) a = map_option (\ b. (snd (f (a, b)))) (map_of xs a)" -by (induct xs, auto) - -lemma map_option_map_of [simp]: "(map_option f (map_of xs k) = None) = ((map_of xs k) = None)" -by (simp add: map_option_case split: option.split) - - + by (induct xs, auto) end diff -r 00c06f1315d0 -r 3f429b7d8eb5 src/HOL/MicroJava/Comp/CorrComp.thy --- a/src/HOL/MicroJava/Comp/CorrComp.thy Tue May 26 21:58:04 2015 +0100 +++ b/src/HOL/MicroJava/Comp/CorrComp.thy Thu May 28 17:25:57 2015 +1000 @@ -16,7 +16,7 @@ "(G \ xs -ex\val-> xs' \ gx xs' = None \ gx xs = None) \ (G \ xs -exs[\]vals-> xs' \ gx xs' = None \ gx xs = None) \ (G \ xs -st-> xs' \ gx xs' = None \ gx xs = None)" -by (induct rule: eval_evals_exec.induct) auto + by (induct rule: eval_evals_exec.induct) auto (* instance of eval_evals_exec_xcpt for eval *) @@ -49,22 +49,21 @@ (**********************************************************************) theorem exec_all_trans: "\(exec_all G s0 s1); (exec_all G s1 s2)\ \ (exec_all G s0 s2)" -apply (auto simp: exec_all_def elim: Transitive_Closure.rtrancl_trans) -done + by (auto simp: exec_all_def elim: Transitive_Closure.rtrancl_trans) theorem exec_all_refl: "exec_all G s s" -by (simp only: exec_all_def) + by (simp only: exec_all_def) theorem exec_instr_in_exec_all: "\ exec_instr i G hp stk lvars C S pc frs = (None, hp', frs'); gis (gmb G C S) ! pc = i\ \ G \ (None, hp, (stk, lvars, C, S, pc) # frs) \jvm\ (None, hp', frs')" -apply (simp only: exec_all_def) -apply (rule rtrancl_refl [THEN rtrancl_into_rtrancl]) -apply (simp add: gis_def gmb_def) -apply (case_tac frs', simp+) -done + apply (simp only: exec_all_def) + apply (rule rtrancl_refl [THEN rtrancl_into_rtrancl]) + apply (simp add: gis_def gmb_def) + apply (case_tac frs', simp+) + done theorem exec_all_one_step: " \ gis (gmb G C S) = pre @ (i # post); pc0 = length pre; @@ -73,10 +72,10 @@ \ G \ (None, hp0, (stk0,lvars0,C,S, pc0)#frs) \jvm\ (None, hp1, (stk1,lvars1,C,S, Suc pc0)#frs)" -apply (unfold exec_all_def) -apply (rule r_into_rtrancl) -apply (simp add: gis_def gmb_def split_beta) -done + apply (unfold exec_all_def) + apply (rule r_into_rtrancl) + apply (simp add: gis_def gmb_def split_beta) + done (***********************************************************************) @@ -88,7 +87,7 @@ bool" ("{_,_,_} \ {_, _, _} >- _ \ {_, _, _}" [61,61,61,61,61,61,90,61,61,61]60) where "{G,C,S} \ {hp0, os0, lvars0} >- instrs \ {hp1, os1, lvars1} == - \ pre post frs. + \pre post frs. (gis (gmb G C S) = pre @ instrs @ post) \ G \ (None,hp0,(os0,lvars0,C,S,length pre)#frs) \jvm\ (None,hp1,(os1,lvars1,C,S,(length pre) + (length instrs))#frs)" @@ -96,7 +95,7 @@ lemma progression_call: - "\ \ pc frs. + "\ \pc frs. exec_instr instr G hp0 os0 lvars0 C S pc frs = (None, hp', (os', lvars', C', S', 0) # (fr pc) # frs) \ gis (gmb G C' S') = instrs' @ [Return] \ @@ -105,22 +104,23 @@ ((fr pc) # frs) = (None, hp2, (os2, lvars2, C, S, Suc pc) # frs) \ \ {G, C, S} \ {hp0, os0, lvars0} >-[instr]\ {hp2,os2,lvars2}" -apply (simp only: progression_def) -apply (intro strip) -apply (drule_tac x="length pre" in spec) -apply (drule_tac x="frs" in spec) -apply clarify -apply (rule exec_all_trans) -apply (rule exec_instr_in_exec_all) apply simp -apply simp -apply (rule exec_all_trans) -apply (simp only: append_Nil) -apply (drule_tac x="[]" in spec) -apply (simp only: list.simps list.size) -apply blast -apply (rule exec_instr_in_exec_all) -apply auto -done + apply (simp only: progression_def) + apply (intro strip) + apply (drule_tac x="length pre" in spec) + apply (drule_tac x="frs" in spec) + apply clarify + apply (rule exec_all_trans) + apply (rule exec_instr_in_exec_all) + apply simp + apply simp + apply (rule exec_all_trans) + apply (simp only: append_Nil) + apply (drule_tac x="[]" in spec) + apply (simp only: list.simps list.size) + apply blast + apply (rule exec_instr_in_exec_all) + apply auto + done lemma progression_transitive: @@ -129,57 +129,57 @@ {G, C, S} \ {hp1, os1, lvars1} >- instrs1 \ {hp2, os2, lvars2} \ \ {G, C, S} \ {hp0, os0, lvars0} >- instrs_comb \ {hp2, os2, lvars2}" -apply (simp only: progression_def) -apply (intro strip) -apply (rule_tac ?s1.0 = "Norm (hp1, (os1, lvars1, C, S, - length pre + length instrs0) # frs)" - in exec_all_trans) -apply (simp only: append_assoc) -apply (erule thin_rl, erule thin_rl) -apply (drule_tac x="pre @ instrs0" in spec) -apply (simp add: add.assoc) -done + apply (simp only: progression_def) + apply (intro strip) + apply (rule_tac ?s1.0 = "Norm (hp1, (os1, lvars1, C, S, + length pre + length instrs0) # frs)" + in exec_all_trans) + apply (simp only: append_assoc) + apply (erule thin_rl, erule thin_rl) + apply (drule_tac x="pre @ instrs0" in spec) + apply (simp add: add.assoc) + done lemma progression_refl: "{G, C, S} \ {hp0, os0, lvars0} >- [] \ {hp0, os0, lvars0}" -apply (simp add: progression_def) -apply (intro strip) -apply (rule exec_all_refl) -done + apply (simp add: progression_def) + apply (intro strip) + apply (rule exec_all_refl) + done lemma progression_one_step: " - \ pc frs. + \pc frs. (exec_instr i G hp0 os0 lvars0 C S pc frs) = (None, hp1, (os1,lvars1,C,S, Suc pc)#frs) \ {G, C, S} \ {hp0, os0, lvars0} >- [i] \ {hp1, os1, lvars1}" -apply (unfold progression_def) -apply (intro strip) -apply simp -apply (rule exec_all_one_step) -apply auto -done + apply (unfold progression_def) + apply (intro strip) + apply simp + apply (rule exec_all_one_step) + apply auto + done (*****) definition jump_fwd :: "jvm_prog \ cname \ sig \ aheap \ locvars \ opstack \ opstack \ instr \ bytecode \ bool" where "jump_fwd G C S hp lvars os0 os1 instr instrs == - \ pre post frs. + \pre post frs. (gis (gmb G C S) = pre @ instr # instrs @ post) \ exec_all G (None,hp,(os0,lvars,C,S, length pre)#frs) (None,hp,(os1,lvars,C,S,(length pre) + (length instrs) + 1)#frs)" lemma jump_fwd_one_step: - "\ pc frs. + "\pc frs. exec_instr instr G hp os0 lvars C S pc frs = (None, hp, (os1, lvars, C, S, pc + (length instrs) + 1)#frs) \ jump_fwd G C S hp lvars os0 os1 instr instrs" -apply (unfold jump_fwd_def) -apply (intro strip) -apply (rule exec_instr_in_exec_all) -apply auto -done + apply (unfold jump_fwd_def) + apply (intro strip) + apply (rule exec_instr_in_exec_all) + apply auto + done lemma jump_fwd_progression_aux: @@ -187,17 +187,17 @@ jump_fwd G C S hp lvars os0 os1 instr instrs0; {G, C, S} \ {hp, os1, lvars} >- instrs1 \ {hp2, os2, lvars2} \ \ {G, C, S} \ {hp, os0, lvars} >- instrs_comb \ {hp2, os2, lvars2}" -apply (simp only: progression_def jump_fwd_def) -apply (intro strip) -apply (rule_tac ?s1.0 = "Norm(hp, (os1, lvars, C, S, length pre + length instrs0 + 1) # frs)" in exec_all_trans) -apply (simp only: append_assoc) -apply (subgoal_tac "pre @ (instr # instrs0 @ instrs1) @ post = pre @ instr # instrs0 @ (instrs1 @ post)") -apply blast -apply simp -apply (erule thin_rl, erule thin_rl) -apply (drule_tac x="pre @ instr # instrs0" in spec) -apply (simp add: add.assoc) -done + apply (simp only: progression_def jump_fwd_def) + apply (intro strip) + apply (rule_tac ?s1.0 = "Norm(hp, (os1, lvars, C, S, length pre + length instrs0 + 1) # frs)" in exec_all_trans) + apply (simp only: append_assoc) + apply (subgoal_tac "pre @ (instr # instrs0 @ instrs1) @ post = pre @ instr # instrs0 @ (instrs1 @ post)") + apply blast + apply simp + apply (erule thin_rl, erule thin_rl) + apply (drule_tac x="pre @ instr # instrs0" in spec) + apply (simp add: add.assoc) + done lemma jump_fwd_progression: @@ -207,10 +207,10 @@ (None, hp, (os1, lvars, C, S, pc + (length instrs0) + 1)#frs); {G, C, S} \ {hp, os1, lvars} >- instrs1 \ {hp2, os2, lvars2} \ \ {G, C, S} \ {hp, os0, lvars} >- instrs_comb \ {hp2, os2, lvars2}" -apply (rule jump_fwd_progression_aux) -apply assumption -apply (rule jump_fwd_one_step) apply assumption+ -done + apply (rule jump_fwd_progression_aux) + apply assumption + apply (rule jump_fwd_one_step, assumption+) + done (* note: instrs and instr reversed wrt. jump_fwd *) @@ -230,11 +230,11 @@ (None, hp, (os1, lvars, C, S, pc)#frs) \ jump_bwd G C S hp lvars os0 os1 instrs instr" -apply (unfold jump_bwd_def) -apply (intro strip) -apply (rule exec_instr_in_exec_all) -apply auto -done + apply (unfold jump_bwd_def) + apply (intro strip) + apply (rule exec_instr_in_exec_all) + apply auto + done lemma jump_bwd_progression: "\ instrs_comb = instrs @ [instr]; @@ -242,14 +242,14 @@ jump_bwd G C S hp1 lvars1 os1 os2 instrs instr; {G, C, S} \ {hp1, os2, lvars1} >- instrs_comb \ {hp3, os3, lvars3} \ \ {G, C, S} \ {hp0, os0, lvars0} >- instrs_comb \ {hp3, os3, lvars3}" -apply (simp only: progression_def jump_bwd_def) -apply (intro strip) -apply (rule exec_all_trans, force) -apply (rule exec_all_trans, force) -apply (rule exec_all_trans, force) -apply simp -apply (rule exec_all_refl) -done + apply (simp only: progression_def jump_bwd_def) + apply (intro strip) + apply (rule exec_all_trans, force) + apply (rule exec_all_trans, force) + apply (rule exec_all_trans, force) + apply simp + apply (rule exec_all_refl) + done (**********************************************************************) @@ -281,50 +281,52 @@ \ \ D. method (G, C) S = Some (D, rT, mb) \ (P D S (rT,mb))" -apply (frule wf_prog_ws_prog [THEN wf_subcls1]) -apply (rule subcls1_induct, assumption, assumption) + apply (frule wf_prog_ws_prog [THEN wf_subcls1]) + apply (rule subcls1_induct, assumption, assumption) -apply (intro strip) -apply ((drule spec)+, drule_tac x="Object" in bspec) -apply (simp add: wf_prog_def ws_prog_def wf_syscls_def) -apply (subgoal_tac "D=Object") apply simp -apply (drule mp) -apply (frule_tac C=Object in method_wf_mdecl) - apply simp - apply assumption apply simp apply assumption apply simp + apply (intro strip) + apply ((drule spec)+, drule_tac x="Object" in bspec) + apply (simp add: wf_prog_def ws_prog_def wf_syscls_def) + apply (subgoal_tac "D=Object") apply simp + apply (drule mp) + apply (frule_tac C=Object in method_wf_mdecl) + apply simp + apply assumption + apply simp + apply assumption + apply simp -apply (simplesubst method_rec) apply simp -apply force -apply simp -apply (simp only: map_add_def) -apply (split option.split) -apply (rule conjI) -apply force -apply (intro strip) -apply (frule_tac - ?P1.0 = "wf_mdecl wf_mb G Ca" and - ?P2.0 = "%(S, (Da, rT, mb)). P Da S (rT, mb)" in map_of_map_prop) -apply (force simp: wf_cdecl_def) + apply (simplesubst method_rec) + apply simp + apply force + apply simp + apply (simp only: map_add_def) + apply (split option.split) + apply (rule conjI) + apply force + apply (intro strip) + apply (frule_tac ?P1.0 = "wf_mdecl wf_mb G Ca" and + ?P2.0 = "%(S, (Da, rT, mb)). P Da S (rT, mb)" in map_of_map_prop) + apply (force simp: wf_cdecl_def) -apply clarify + apply clarify -apply (simp only: class_def) -apply (drule map_of_SomeD)+ -apply (frule_tac A="set G" and f=fst in imageI, simp) -apply blast -apply simp -done + apply (simp only: class_def) + apply (drule map_of_SomeD)+ + apply (frule_tac A="set G" and f=fst in imageI, simp) + apply blast + apply simp + done lemma method_preserves_length: "\ wf_java_prog G; is_class G C; method (G, C) (mn,pTs) = Some (D, rT, pns, lvars, blk, res)\ \ length pns = length pTs" -apply (frule_tac - P="%D (mn,pTs) (rT, pns, lvars, blk, res). length pns = length pTs" - in method_preserves) -apply (auto simp: wf_mdecl_def wf_java_mdecl_def) -done + apply (frule_tac P="%D (mn,pTs) (rT, pns, lvars, blk, res). length pns = length pTs" + in method_preserves) + apply (auto simp: wf_mdecl_def wf_java_mdecl_def) + done (**********************************************************************) @@ -338,91 +340,84 @@ "wtpd_stmt E c == (E\c \)" lemma wtpd_expr_newc: "wtpd_expr E (NewC C) \ is_class (prg E) C" -by (simp only: wtpd_expr_def, erule exE, erule ty_expr.cases, auto) + by (simp only: wtpd_expr_def, erule exE, erule ty_expr.cases, auto) lemma wtpd_expr_cast: "wtpd_expr E (Cast cn e) \ (wtpd_expr E e)" -by (simp only: wtpd_expr_def, erule exE, erule ty_expr.cases, auto) + by (simp only: wtpd_expr_def, erule exE, erule ty_expr.cases, auto) -lemma wtpd_expr_lacc: "\ wtpd_expr (env_of_jmb G C S) (LAcc vn); - class_sig_defined G C S \ +lemma wtpd_expr_lacc: + "\ wtpd_expr (env_of_jmb G C S) (LAcc vn); class_sig_defined G C S \ \ vn \ set (gjmb_plns (gmb G C S)) \ vn = This" -apply (simp only: wtpd_expr_def env_of_jmb_def class_sig_defined_def galldefs) -apply clarify -apply (case_tac S) -apply simp -apply (erule ty_expr.cases) -apply (auto dest: map_upds_SomeD map_of_SomeD fst_in_set_lemma) -apply (drule map_upds_SomeD) -apply (erule disjE) - apply assumption - apply (drule map_of_SomeD) apply (auto dest: fst_in_set_lemma) -done + apply (clarsimp simp: wtpd_expr_def env_of_jmb_def class_sig_defined_def galldefs) + apply (case_tac S) + apply (erule ty_expr.cases; fastforce dest: map_upds_SomeD map_of_SomeD fst_in_set_lemma) + done lemma wtpd_expr_lass: "wtpd_expr E (vn::=e) \ (vn \ This) & (wtpd_expr E (LAcc vn)) & (wtpd_expr E e)" -by (simp only: wtpd_expr_def, erule exE, erule ty_expr.cases, auto) + by (simp only: wtpd_expr_def, erule exE, erule ty_expr.cases, auto) lemma wtpd_expr_facc: "wtpd_expr E ({fd}a..fn) \ (wtpd_expr E a)" -by (simp only: wtpd_expr_def, erule exE, erule ty_expr.cases, auto) + by (simp only: wtpd_expr_def, erule exE, erule ty_expr.cases, auto) lemma wtpd_expr_fass: "wtpd_expr E ({fd}a..fn:=v) \ (wtpd_expr E ({fd}a..fn)) & (wtpd_expr E v)" -by (simp only: wtpd_expr_def, erule exE, erule ty_expr.cases, auto) + by (simp only: wtpd_expr_def, erule exE, erule ty_expr.cases, auto) lemma wtpd_expr_binop: "wtpd_expr E (BinOp bop e1 e2) \ (wtpd_expr E e1) & (wtpd_expr E e2)" -by (simp only: wtpd_expr_def, erule exE, erule ty_expr.cases, auto) + by (simp only: wtpd_expr_def, erule exE, erule ty_expr.cases, auto) lemma wtpd_exprs_cons: "wtpd_exprs E (e # es) \ (wtpd_expr E e) & (wtpd_exprs E es)" -by (simp only: wtpd_exprs_def wtpd_expr_def, erule exE, erule ty_exprs.cases, auto) + by (simp only: wtpd_exprs_def wtpd_expr_def, erule exE, erule ty_exprs.cases, auto) lemma wtpd_stmt_expr: "wtpd_stmt E (Expr e) \ (wtpd_expr E e)" -by (simp only: wtpd_stmt_def wtpd_expr_def, erule wt_stmt.cases, auto) + by (simp only: wtpd_stmt_def wtpd_expr_def, erule wt_stmt.cases, auto) lemma wtpd_stmt_comp: "wtpd_stmt E (s1;; s2) \ (wtpd_stmt E s1) & (wtpd_stmt E s2)" -by (simp only: wtpd_stmt_def wtpd_expr_def, erule wt_stmt.cases, auto) + by (simp only: wtpd_stmt_def wtpd_expr_def, erule wt_stmt.cases, auto) lemma wtpd_stmt_cond: "wtpd_stmt E (If(e) s1 Else s2) \ (wtpd_expr E e) & (wtpd_stmt E s1) & (wtpd_stmt E s2) & (E\e::PrimT Boolean)" -by (simp only: wtpd_stmt_def wtpd_expr_def, erule wt_stmt.cases, auto) + by (simp only: wtpd_stmt_def wtpd_expr_def, erule wt_stmt.cases, auto) lemma wtpd_stmt_loop: "wtpd_stmt E (While(e) s) \ (wtpd_expr E e) & (wtpd_stmt E s) & (E\e::PrimT Boolean)" -by (simp only: wtpd_stmt_def wtpd_expr_def, erule wt_stmt.cases, auto) + by (simp only: wtpd_stmt_def wtpd_expr_def, erule wt_stmt.cases, auto) lemma wtpd_expr_call: "wtpd_expr E ({C}a..mn({pTs'}ps)) \ (wtpd_expr E a) & (wtpd_exprs E ps) & (length ps = length pTs') & (E\a::Class C) & (\ pTs md rT. E\ps[::]pTs & max_spec (prg E) C (mn, pTs) = {((md,rT),pTs')})" -apply (simp only: wtpd_expr_def wtpd_exprs_def) -apply (erule exE) -apply (ind_cases "E \ {C}a..mn( {pTs'}ps) :: T" for T) -apply (auto simp: max_spec_preserves_length) -done + apply (simp only: wtpd_expr_def wtpd_exprs_def) + apply (erule exE) + apply (ind_cases "E \ {C}a..mn( {pTs'}ps) :: T" for T) + apply (auto simp: max_spec_preserves_length) + done lemma wtpd_blk: "\ method (G, D) (md, pTs) = Some (D, rT, (pns, lvars, blk, res)); wf_prog wf_java_mdecl G; is_class G D \ \ wtpd_stmt (env_of_jmb G D (md, pTs)) blk" -apply (simp add: wtpd_stmt_def env_of_jmb_def) -apply (frule_tac P="%D (md, pTs) (rT, (pns, lvars, blk, res)). (G, map_of lvars(pns[\]pTs)(This\Class D)) \ blk \ " in method_preserves) -apply (auto simp: wf_mdecl_def wf_java_mdecl_def) -done + apply (simp add: wtpd_stmt_def env_of_jmb_def) + apply (frule_tac P="%D (md, pTs) (rT, (pns, lvars, blk, res)). (G, map_of lvars(pns[\]pTs)(This\Class D)) \ blk \ " in method_preserves) + apply (auto simp: wf_mdecl_def wf_java_mdecl_def) + done lemma wtpd_res: "\ method (G, D) (md, pTs) = Some (D, rT, (pns, lvars, blk, res)); wf_prog wf_java_mdecl G; is_class G D \ \ wtpd_expr (env_of_jmb G D (md, pTs)) res" -apply (simp add: wtpd_expr_def env_of_jmb_def) -apply (frule_tac P="%D (md, pTs) (rT, (pns, lvars, blk, res)). \T. (G, map_of lvars(pns[\]pTs)(This\Class D)) \ res :: T " in method_preserves) -apply (auto simp: wf_mdecl_def wf_java_mdecl_def) -done + apply (simp add: wtpd_expr_def env_of_jmb_def) + apply (frule_tac P="%D (md, pTs) (rT, (pns, lvars, blk, res)). \T. (G, map_of lvars(pns[\]pTs)(This\Class D)) \ res :: T " in method_preserves) + apply (auto simp: wf_mdecl_def wf_java_mdecl_def) + done (**********************************************************************) @@ -431,16 +426,15 @@ (* Is there a more elegant proof? *) lemma evals_preserves_length: "G\ xs -es[\]vs-> (None, s) \ length es = length vs" -apply (subgoal_tac - "\ xs'. (G \ xk -xj\xi-> xh \ True) & - (G\ xs -es[\]vs-> xs' \ (\ s. (xs' = (None, s))) \ - length es = length vs) & - (G \ xc -xb-> xa \ True)") -apply blast -apply (rule allI) -apply (rule Eval.eval_evals_exec.induct) -apply auto -done + apply (subgoal_tac + "\ xs'. (G \ xk -xj\xi-> xh \ True) & + (G\ xs -es[\]vs-> xs' \ (\ s. (xs' = (None, s))) \ + length es = length vs) & + (G \ xc -xb-> xa \ True)") + apply blast + apply (rule allI) + apply (rule Eval.eval_evals_exec.induct; simp) + done (***********************************************************************) @@ -451,21 +445,22 @@ {hp, (v2 # v1 # os), lvars} >- [Ifcmpeq 3, LitPush (Bool False), Goto 2, LitPush (Bool True)] \ {hp, (Bool (v1 = v2) # os), lvars}" -apply (case_tac "v1 = v2") + apply (case_tac "v1 = v2") -(* case v1 = v2 *) -apply (rule_tac ?instrs1.0 = "[LitPush (Bool True)]" in jump_fwd_progression) -apply (auto simp: nat_add_distrib) -apply (rule progression_one_step) apply simp + (* case v1 = v2 *) + apply (rule_tac ?instrs1.0 = "[LitPush (Bool True)]" in jump_fwd_progression) + apply (auto simp: nat_add_distrib) + apply (rule progression_one_step) + apply simp -(* case v1 \ v2 *) -apply (rule progression_one_step [THEN HOL.refl [THEN progression_transitive], simplified]) -apply auto -apply (rule progression_one_step [THEN HOL.refl [THEN progression_transitive], simplified]) -apply auto -apply (rule_tac ?instrs1.0 = "[]" in jump_fwd_progression) -apply (auto simp: nat_add_distrib intro: progression_refl) -done + (* case v1 \ v2 *) + apply (rule progression_one_step [THEN HOL.refl [THEN progression_transitive], simplified]) + apply auto + apply (rule progression_one_step [THEN HOL.refl [THEN progression_transitive], simplified]) + apply auto + apply (rule_tac ?instrs1.0 = "[]" in jump_fwd_progression) + apply (auto simp: nat_add_distrib intro: progression_refl) + done (**********************************************************************) @@ -475,28 +470,27 @@ declare split_paired_All [simp del] split_paired_Ex [simp del] -lemma distinct_method: "\ wf_java_prog G; is_class G C; - method (G, C) S = Some (D, rT, pns, lvars, blk, res) \ \ +lemma distinct_method: + "\ wf_java_prog G; is_class G C; method (G, C) S = Some (D, rT, pns, lvars, blk, res) \ \ distinct (gjmb_plns (gmb G C S))" -apply (frule method_wf_mdecl [THEN conjunct2], assumption, assumption) -apply (case_tac S) -apply (simp add: wf_mdecl_def wf_java_mdecl_def galldefs) -apply (simp add: unique_def map_of_in_set) -apply blast -done + apply (frule method_wf_mdecl [THEN conjunct2], assumption, assumption) + apply (case_tac S) + apply (simp add: wf_mdecl_def wf_java_mdecl_def galldefs) + apply (simp add: unique_def map_of_in_set) + apply blast + done lemma distinct_method_if_class_sig_defined : - "\ wf_java_prog G; class_sig_defined G C S \ \ - distinct (gjmb_plns (gmb G C S))" -by (auto intro: distinct_method simp: class_sig_defined_def) + "\ wf_java_prog G; class_sig_defined G C S \ \ distinct (gjmb_plns (gmb G C S))" + by (auto intro: distinct_method simp: class_sig_defined_def) lemma method_yields_wf_java_mdecl: "\ wf_java_prog G; is_class G C; method (G, C) S = Some (D, rT, pns, lvars, blk, res) \ \ wf_java_mdecl G D (S,rT,(pns,lvars,blk,res))" -apply (frule method_wf_mdecl) -apply (auto simp: wf_mdecl_def) -done + apply (frule method_wf_mdecl) + apply (auto simp: wf_mdecl_def) + done (**********************************************************************) @@ -511,28 +505,30 @@ {h, os, (prfx @ lvals)} >- (concat (map (compInit (pns, lvars0, blk, res)) lvars)) \ {h, os, (prfx @ (map (\p. (default_val (snd p))) lvars))}))" -apply simp -apply (induct lvars) -apply (clarsimp, rule progression_refl) -apply (intro strip) -apply (case_tac lvals) apply simp -apply (simp (no_asm_simp) ) + apply simp + apply (induct lvars) + apply (clarsimp, rule progression_refl) + apply (intro strip) + apply (case_tac lvals) + apply simp + apply (simp (no_asm_simp) ) -apply (rule_tac ?lvars1.0 = "(prfx @ [default_val (snd a)]) @ list" in progression_transitive, rule HOL.refl) -apply (case_tac a) apply (simp (no_asm_simp) add: compInit_def) -apply (rule_tac ?instrs0.0 = "[load_default_val b]" in progression_transitive, simp) -apply (rule progression_one_step) -apply (simp (no_asm_simp) add: load_default_val_def) + apply (rule_tac ?lvars1.0 = "(prfx @ [default_val (snd a)]) @ list" in progression_transitive, rule HOL.refl) + apply (case_tac a) + apply (simp (no_asm_simp) add: compInit_def) + apply (rule_tac ?instrs0.0 = "[load_default_val b]" in progression_transitive, simp) + apply (rule progression_one_step) + apply (simp (no_asm_simp) add: load_default_val_def) -apply (rule progression_one_step) -apply (simp (no_asm_simp)) -apply (rule conjI, simp)+ -apply (simp add: index_of_var2) -apply (drule_tac x="zs @ [a]" in spec) (* instantiate zs *) -apply (drule mp, simp) -apply (drule_tac x="(prfx @ [default_val (snd a)])" in spec) (* instantiate prfx *) -apply auto -done + apply (rule progression_one_step) + apply (simp (no_asm_simp)) + apply (rule conjI, simp)+ + apply (simp add: index_of_var2) + apply (drule_tac x="zs @ [a]" in spec) (* instantiate zs *) + apply (drule mp, simp) + apply (drule_tac x="(prfx @ [default_val (snd a)])" in spec) (* instantiate prfx *) + apply auto + done lemma progression_lvar_init [rule_format (no_asm)]: "\ wf_java_prog G; is_class G C; @@ -544,147 +540,142 @@ {h, os, (a' # pvs @ lvals)} >- (compInitLvars (pns, lvars, blk, res) lvars) \ {h, os, (locvars_xstate G C S (Norm (h, init_vars lvars(pns[\]pvs)(This\a'))))})" -apply (simp only: compInitLvars_def) -apply (frule method_yields_wf_java_mdecl, assumption, assumption) + apply (simp only: compInitLvars_def) + apply (frule method_yields_wf_java_mdecl, assumption, assumption) -apply (simp only: wf_java_mdecl_def) -apply (subgoal_tac "(\y\set pns. y \ set (map fst lvars))") -apply (simp add: init_vars_def locvars_xstate_def locvars_locals_def galldefs unique_def split_def map_of_map_as_map_upd del: map_map) -apply (intro strip) -apply (simp (no_asm_simp) only: append_Cons [symmetric]) -apply (rule progression_lvar_init_aux) -apply (auto simp: unique_def map_of_in_set disjoint_varnames_def) -done + apply (simp only: wf_java_mdecl_def) + apply (subgoal_tac "(\y\set pns. y \ set (map fst lvars))") + apply (simp add: init_vars_def locvars_xstate_def locvars_locals_def galldefs unique_def split_def map_of_map_as_map_upd del: map_map) + apply (intro strip) + apply (simp (no_asm_simp) only: append_Cons [symmetric]) + apply (rule progression_lvar_init_aux) + apply (auto simp: unique_def map_of_in_set disjoint_varnames_def) + done (**********************************************************************) -lemma state_ok_eval: "\xs::\E; wf_java_prog (prg E); wtpd_expr E e; - (prg E)\xs -e\v -> xs'\ \ xs'::\E" -apply (simp only: wtpd_expr_def) -apply (erule exE) -apply (case_tac xs', case_tac xs) -apply (auto intro: eval_type_sound [THEN conjunct1]) -done +lemma state_ok_eval: + "\xs::\E; wf_java_prog (prg E); wtpd_expr E e; (prg E)\xs -e\v -> xs'\ \ xs'::\E" + apply (simp only: wtpd_expr_def) + apply (erule exE) + apply (case_tac xs', case_tac xs) + apply (auto intro: eval_type_sound [THEN conjunct1]) + done -lemma state_ok_evals: "\xs::\E; wf_java_prog (prg E); wtpd_exprs E es; - prg E \ xs -es[\]vs-> xs'\ \ xs'::\E" -apply (simp only: wtpd_exprs_def) -apply (erule exE) -apply (case_tac xs) apply (case_tac xs') -apply (auto intro: evals_type_sound [THEN conjunct1]) -done +lemma state_ok_evals: + "\xs::\E; wf_java_prog (prg E); wtpd_exprs E es; prg E \ xs -es[\]vs-> xs'\ \ xs'::\E" + apply (simp only: wtpd_exprs_def) + apply (erule exE) + apply (case_tac xs) + apply (case_tac xs') + apply (auto intro: evals_type_sound [THEN conjunct1]) + done -lemma state_ok_exec: "\xs::\E; wf_java_prog (prg E); wtpd_stmt E st; - prg E \ xs -st-> xs'\ \ xs'::\E" -apply (simp only: wtpd_stmt_def) -apply (case_tac xs', case_tac xs) -apply (auto dest: exec_type_sound) -done - +lemma state_ok_exec: + "\xs::\E; wf_java_prog (prg E); wtpd_stmt E st; prg E \ xs -st-> xs'\ \ xs'::\E" + apply (simp only: wtpd_stmt_def) + apply (case_tac xs', case_tac xs) + apply (auto dest: exec_type_sound) + done lemma state_ok_init: "\ wf_java_prog G; (x, h, l)::\(env_of_jmb G C S); is_class G dynT; method (G, dynT) (mn, pTs) = Some (md, rT, pns, lvars, blk, res); list_all2 (conf G h) pvs pTs; G,h \ a' ::\ Class md\ -\ -(np a' x, h, init_vars lvars(pns[\]pvs)(This\a'))::\(env_of_jmb G md (mn, pTs))" -apply (frule wf_prog_ws_prog) -apply (frule method_in_md [THEN conjunct2], assumption+) -apply (frule method_yields_wf_java_mdecl, assumption, assumption) -apply (simp add: env_of_jmb_def gs_def conforms_def split_beta) -apply (simp add: wf_java_mdecl_def) -apply (rule conjI) -apply (rule lconf_ext) -apply (rule lconf_ext_list) -apply (rule lconf_init_vars) -apply (auto dest: Ball_set_table) -apply (simp add: np_def xconf_raise_if) -done + \ + (np a' x, h, init_vars lvars(pns[\]pvs)(This\a'))::\(env_of_jmb G md (mn, pTs))" + apply (frule wf_prog_ws_prog) + apply (frule method_in_md [THEN conjunct2], assumption+) + apply (frule method_yields_wf_java_mdecl, assumption, assumption) + apply (simp add: env_of_jmb_def gs_def conforms_def split_beta) + apply (simp add: wf_java_mdecl_def) + apply (rule conjI) + apply (rule lconf_ext) + apply (rule lconf_ext_list) + apply (rule lconf_init_vars) + apply (auto dest: Ball_set_table) + apply (simp add: np_def xconf_raise_if) + done lemma ty_exprs_list_all2 [rule_format (no_asm)]: "(\ Ts. (E \ es [::] Ts) = list_all2 (\e T. E \ e :: T) es Ts)" -apply (rule list.induct) -apply simp -apply (rule allI) -apply (rule iffI) - apply (ind_cases "E \ [] [::] Ts" for Ts, assumption) - apply simp apply (rule WellType.Nil) -apply (simp add: list_all2_Cons1) -apply (rule allI) -apply (rule iffI) - apply (rename_tac a exs Ts) - apply (ind_cases "E \ a # exs [::] Ts" for a exs Ts) apply blast + apply (rule list.induct) + apply simp + apply (rule allI) + apply (rule iffI) + apply (ind_cases "E \ [] [::] Ts" for Ts, assumption) + apply simp apply (rule WellType.Nil) + apply (simp add: list_all2_Cons1) + apply (rule allI) + apply (rule iffI) + apply (rename_tac a exs Ts) + apply (ind_cases "E \ a # exs [::] Ts" for a exs Ts) apply blast apply (auto intro: WellType.Cons) -done + done lemma conf_bool: "G,h \ v::\PrimT Boolean \ \ b. v = Bool b" -apply (simp add: conf_def) -apply (erule exE) -apply (case_tac v) -apply (auto elim: widen.cases) -done - - -lemma class_expr_is_class: "\E \ e :: Class C; ws_prog (prg E)\ - \ is_class (prg E) C" -by (case_tac E, auto dest: ty_expr_is_type) + apply (simp add: conf_def) + apply (erule exE) + apply (case_tac v) + apply (auto elim: widen.cases) + done lemma max_spec_widen: "max_spec G C (mn, pTs) = {((md,rT),pTs')} \ list_all2 (\ T T'. G \ T \ T') pTs pTs'" -by (blast dest: singleton_in_set max_spec2appl_meths appl_methsD) + by (blast dest: singleton_in_set max_spec2appl_meths appl_methsD) lemma eval_conf: "\G \ s -e\v-> s'; wf_java_prog G; s::\E; E\e::T; gx s' = None; prg E = G \ \ G,gh s'\v::\T" -apply (simp add: gh_def) -apply (rule_tac x3="fst s" and s3="snd s"and x'3="fst s'" - in eval_type_sound [THEN conjunct2 [THEN conjunct1 [THEN mp]], simplified]) -apply assumption+ -apply (simp (no_asm_use) add: surjective_pairing [symmetric]) -apply (simp only: surjective_pairing [symmetric]) -apply (auto simp add: gs_def gx_def) -done + apply (simp add: gh_def) + apply (rule_tac x3="fst s" and s3="snd s"and x'3="fst s'" + in eval_type_sound [THEN conjunct2 [THEN conjunct1 [THEN mp]], simplified]) + apply assumption+ + apply (simp (no_asm_use)) + apply (simp only: surjective_pairing [symmetric]) + apply (auto simp add: gs_def gx_def) + done lemma evals_preserves_conf: "\ G\ s -es[\]vs-> s'; G,gh s \ t ::\ T; E \es[::]Ts; wf_java_prog G; s::\E; prg E = G \ \ G,gh s' \ t ::\ T" -apply (subgoal_tac "gh s\| gh s'") -apply (frule conf_hext, assumption, assumption) -apply (frule eval_evals_exec_type_sound [THEN conjunct2 [THEN conjunct1 [THEN mp]]]) -apply (subgoal_tac "G \ (gx s, (gh s, gl s)) -es[\]vs-> (gx s', (gh s', gl s'))") -apply assumption -apply (simp add: gx_def gh_def gl_def surjective_pairing [symmetric]) -apply (case_tac E) -apply (simp add: gx_def gs_def gh_def gl_def surjective_pairing [symmetric]) -done + apply (subgoal_tac "gh s\| gh s'") + apply (frule conf_hext, assumption, assumption) + apply (frule eval_evals_exec_type_sound [THEN conjunct2 [THEN conjunct1 [THEN mp]]]) + apply (subgoal_tac "G \ (gx s, (gh s, gl s)) -es[\]vs-> (gx s', (gh s', gl s'))") + apply assumption + apply (simp add: gx_def gh_def gl_def) + apply (case_tac E) + apply (simp add: gx_def gs_def gh_def gl_def) + done -lemma eval_of_class: "\ G \ s -e\a'-> s'; E \ e :: Class C; - wf_java_prog G; s::\E; gx s'=None; a' \ Null; G=prg E\ +lemma eval_of_class: + "\ G \ s -e\a'-> s'; E \ e :: Class C; wf_java_prog G; s::\E; gx s'=None; a' \ Null; G=prg E\ \ (\ lc. a' = Addr lc)" -apply (case_tac s, case_tac s', simp) -apply (frule eval_type_sound, (simp add: gs_def)+) -apply (case_tac a') -apply (auto simp: conf_def) -done + apply (case_tac s, case_tac s', simp) + apply (frule eval_type_sound, (simp add: gs_def)+) + apply (case_tac a') + apply (auto simp: conf_def) + done lemma dynT_subcls: "\ a' \ Null; G,h\a'::\ Class C; dynT = fst (the (h (the_Addr a'))); is_class G dynT; ws_prog G \ \ G\dynT \C C" -apply (case_tac "C = Object") -apply (simp, rule subcls_C_Object, assumption+) -apply simp -apply (frule non_np_objD, auto) -done + apply (case_tac "C = Object") + apply (simp, rule subcls_C_Object, assumption+) + apply simp + apply (frule non_np_objD, auto) + done lemma method_defined: "\ @@ -693,15 +684,13 @@ a' \ Null; G,h\a'::\ Class C; a = the_Addr a'; \pTsa md rT. max_spec G C (mn, pTsa) = {((md, rT), pTs)} \ \ (method (G, dynT) (mn, pTs)) = Some m" -apply (erule exE)+ -apply (drule singleton_in_set, drule max_spec2appl_meths) -apply (simp add: appl_methds_def) -apply ((erule exE)+, (erule conjE)+, (erule exE)+) -apply (drule widen_methd) -apply assumption -apply (rule dynT_subcls) apply assumption+ apply simp apply simp -apply (erule exE)+ apply simp -done + apply (erule exE)+ + apply (drule singleton_in_set, drule max_spec2appl_meths) + apply (simp add: appl_methds_def) + apply (elim exE conjE) + apply (drule widen_methd) + apply (auto intro!: dynT_subcls) + done @@ -744,476 +733,515 @@ {gh xs, os, (locvars_xstate G CL S xs)} >- (compStmt (gmb G CL S) st) \ {gh xs', os, (locvars_xstate G CL S xs')})))" -apply (rule Eval.eval_evals_exec.induct) + apply (rule Eval.eval_evals_exec.induct) -(* case XcptE *) -apply simp + (* case XcptE *) + apply simp -(* case NewC *) -apply clarify -apply (frule wf_prog_ws_prog [THEN wf_subcls1]) (* establish wf ((subcls1 G)^-1) *) -apply (simp add: c_hupd_hp_invariant) -apply (rule progression_one_step) -apply (rotate_tac 1, drule sym) (* reverse equation (a, None) = new_Addr (fst s) *) -apply (simp add: locvars_xstate_def locvars_locals_def comp_fields) + (* case NewC *) + apply clarify + apply (frule wf_prog_ws_prog [THEN wf_subcls1]) (* establish wf ((subcls1 G)^-1) *) + apply (simp add: c_hupd_hp_invariant) + apply (rule progression_one_step) + apply (rotate_tac 1, drule sym) (* reverse equation (a, None) = new_Addr (fst s) *) + apply (simp add: locvars_xstate_def locvars_locals_def comp_fields) -(* case Cast *) -apply (intro allI impI) -apply simp -apply (frule raise_if_NoneD) -apply (frule wtpd_expr_cast) -apply simp -apply (rule_tac ?instrs0.0 = "(compExpr (gmb G CL S) e)" in progression_transitive, simp) -apply blast -apply (rule progression_one_step) -apply (simp add: raise_system_xcpt_def gh_def comp_cast_ok) + (* case Cast *) + apply (intro allI impI) + apply simp + apply (frule raise_if_NoneD) + apply (frule wtpd_expr_cast) + apply simp + apply (rule_tac ?instrs0.0 = "(compExpr (gmb G CL S) e)" in progression_transitive, simp) + apply blast + apply (rule progression_one_step) + apply (simp add: raise_system_xcpt_def gh_def comp_cast_ok) -(* case Lit *) -apply simp -apply (intro strip) -apply (rule progression_one_step) -apply simp + (* case Lit *) + apply simp + apply (intro strip) + apply (rule progression_one_step) + apply simp + + + (* case BinOp *) + apply (intro allI impI) + apply (frule_tac xs=s1 in eval_xcpt, assumption) (* establish (gx s1 = None) *) + apply (frule wtpd_expr_binop) + (* establish (s1::\ \) *) + apply (frule_tac e=e1 in state_ok_eval) apply (simp (no_asm_simp)) + apply simp + apply (simp (no_asm_use) only: env_of_jmb_fst) -(* case BinOp *) -apply (intro allI impI) -apply (frule_tac xs=s1 in eval_xcpt, assumption) (* establish (gx s1 = None) *) -apply (frule wtpd_expr_binop) -(* establish (s1::\ \) *) -apply (frule_tac e=e1 in state_ok_eval) apply (simp (no_asm_simp)) apply simp apply (simp (no_asm_use) only: env_of_jmb_fst) - - -apply (simp (no_asm_use) only: compExpr.simps compExprs.simps) -apply (rule_tac ?instrs0.0 = "compExpr (gmb G CL S) e1" in progression_transitive, simp) apply blast -apply (rule_tac ?instrs0.0 = "compExpr (gmb G CL S) e2" in progression_transitive, simp) apply blast -apply (case_tac bop) - (*subcase bop = Eq *) apply simp apply (rule progression_Eq) - (*subcase bop = Add *) apply simp apply (rule progression_one_step) apply simp + apply (simp (no_asm_use) only: compExpr.simps compExprs.simps) + apply (rule_tac ?instrs0.0 = "compExpr (gmb G CL S) e1" in progression_transitive, simp) apply blast + apply (rule_tac ?instrs0.0 = "compExpr (gmb G CL S) e2" in progression_transitive, simp) apply blast + apply (case_tac bop) + (*subcase bop = Eq *) + apply simp + apply (rule progression_Eq) + (*subcase bop = Add *) + apply simp + apply (rule progression_one_step) + apply simp -(* case LAcc *) -apply simp -apply (intro strip) -apply (rule progression_one_step) -apply (simp add: locvars_xstate_def locvars_locals_def) -apply (frule wtpd_expr_lacc) -apply assumption -apply (simp add: gl_def) -apply (erule select_at_index) + (* case LAcc *) + apply simp + apply (intro strip) + apply (rule progression_one_step) + apply (simp add: locvars_xstate_def locvars_locals_def) + apply (frule wtpd_expr_lacc) + apply assumption + apply (simp add: gl_def) + apply (erule select_at_index) -(* case LAss *) -apply (intro allI impI) -apply (frule wtpd_expr_lass, erule conjE, erule conjE) -apply simp + (* case LAss *) + apply (intro allI impI) + apply (frule wtpd_expr_lass, erule conjE, erule conjE) + apply simp -apply (rule_tac ?instrs0.0 = "(compExpr (gmb G CL S) e)" in progression_transitive, rule HOL.refl) -apply blast + apply (rule_tac ?instrs0.0 = "(compExpr (gmb G CL S) e)" in progression_transitive, rule HOL.refl) + apply blast -apply (rule_tac ?instrs0.0 = "[Dup]" in progression_transitive, simp) -apply (rule progression_one_step) -apply (simp add: gh_def) -apply simp -apply (rule progression_one_step) -apply (simp add: gh_def) -(* the following falls out of the general scheme *) -apply (frule wtpd_expr_lacc) apply assumption -apply (rule update_at_index) -apply (rule distinct_method_if_class_sig_defined) apply assumption -apply assumption apply simp apply assumption + apply (rule_tac ?instrs0.0 = "[Dup]" in progression_transitive, simp) + apply (rule progression_one_step) + apply (simp add: gh_def) + apply simp + apply (rule progression_one_step) + apply (simp add: gh_def) + (* the following falls out of the general scheme *) + apply (frule wtpd_expr_lacc) apply assumption + apply (rule update_at_index) + apply (rule distinct_method_if_class_sig_defined) + apply assumption + apply assumption + apply simp + apply assumption -(* case FAcc *) -apply (intro allI impI) - (* establish x1 = None \ a' \ Null *) -apply (simp (no_asm_use) only: gx_conv, frule np_NoneD) -apply (frule wtpd_expr_facc) + (* case FAcc *) + apply (intro allI impI) + (* establish x1 = None \ a' \ Null *) + apply (simp (no_asm_use) only: gx_conv, frule np_NoneD) + apply (frule wtpd_expr_facc) -apply (simp (no_asm_use) only: compExpr.simps compExprs.simps) -apply (rule_tac ?instrs0.0 = "(compExpr (gmb G CL S) e)" in progression_transitive, rule HOL.refl) -apply blast -apply (rule progression_one_step) -apply (simp add: gh_def) -apply (case_tac "(the (fst s1 (the_Addr a')))") -apply (simp add: raise_system_xcpt_def) + apply (simp (no_asm_use) only: compExpr.simps compExprs.simps) + apply (rule_tac ?instrs0.0 = "(compExpr (gmb G CL S) e)" in progression_transitive, rule HOL.refl) + apply blast + apply (rule progression_one_step) + apply (simp add: gh_def) + apply (case_tac "(the (fst s1 (the_Addr a')))") + apply (simp add: raise_system_xcpt_def) -(* case FAss *) -apply (intro allI impI) -apply (frule wtpd_expr_fass) apply (erule conjE) apply (frule wtpd_expr_facc) -apply (simp only: c_hupd_xcpt_invariant) (* establish x2 = None *) - (* establish x1 = None and a' \ Null *) -apply (frule_tac xs="(np a' x1, s1)" in eval_xcpt) -apply (simp only: gx_conv, simp only: gx_conv, frule np_NoneD, erule conjE) + (* case FAss *) + apply (intro allI impI) + apply (frule wtpd_expr_fass) apply (erule conjE) apply (frule wtpd_expr_facc) + apply (simp only: c_hupd_xcpt_invariant) (* establish x2 = None *) + (* establish x1 = None and a' \ Null *) + apply (frule_tac xs="(np a' x1, s1)" in eval_xcpt) + apply (simp only: gx_conv, simp only: gx_conv, frule np_NoneD, erule conjE) - (* establish ((Norm s1)::\ \) *) -apply (frule_tac e=e1 in state_ok_eval) apply (simp (no_asm_simp) only: env_of_jmb_fst) - apply assumption apply (simp (no_asm_use) only: env_of_jmb_fst) + (* establish ((Norm s1)::\ \) *) + apply (frule_tac e=e1 in state_ok_eval) + apply (simp (no_asm_simp) only: env_of_jmb_fst) + apply assumption + apply (simp (no_asm_use) only: env_of_jmb_fst) -apply (simp only: compExpr.simps compExprs.simps) + apply (simp only: compExpr.simps compExprs.simps) -apply (rule_tac ?instrs0.0 = "(compExpr (gmb G CL S) e1)" in progression_transitive, rule HOL.refl) -apply fast (* blast does not seem to work - why? *) -apply (rule_tac ?instrs0.0 = "(compExpr (gmb G CL S) e2)" in progression_transitive, rule HOL.refl) -apply fast -apply (rule_tac ?instrs0.0 = "[Dup_x1]" and ?instrs1.0 = "[Putfield fn T]" in progression_transitive, simp) + apply (rule_tac ?instrs0.0 = "(compExpr (gmb G CL S) e1)" in progression_transitive, rule HOL.refl) + apply fast (* blast does not seem to work - why? *) + apply (rule_tac ?instrs0.0 = "(compExpr (gmb G CL S) e2)" in progression_transitive, rule HOL.refl) + apply fast + apply (rule_tac ?instrs0.0 = "[Dup_x1]" and ?instrs1.0 = "[Putfield fn T]" in progression_transitive, simp) - (* Dup_x1 *) - apply (rule progression_one_step) - apply (simp add: gh_def) + (* Dup_x1 *) + apply (rule progression_one_step) + apply (simp add: gh_def) - (* Putfield \ still looks nasty*) - apply (rule progression_one_step) - apply simp - apply (case_tac "(the (fst s2 (the_Addr a')))") - apply (simp add: c_hupd_hp_invariant) - apply (case_tac s2) - apply (simp add: c_hupd_conv raise_system_xcpt_def) - apply (rule locvars_xstate_par_dep, rule HOL.refl) + (* Putfield \ still looks nasty*) + apply (rule progression_one_step) + apply simp + apply (case_tac "(the (fst s2 (the_Addr a')))") + apply (simp add: c_hupd_hp_invariant) + apply (case_tac s2) + apply (simp add: c_hupd_conv raise_system_xcpt_def) + apply (rule locvars_xstate_par_dep, rule HOL.refl) -defer (* method call *) + defer (* method call *) -(* case XcptEs *) -apply simp + (* case XcptEs *) + apply simp -(* case Nil *) -apply simp -apply (intro strip) -apply (rule progression_refl) + (* case Nil *) + apply simp + apply (intro strip) + apply (rule progression_refl) -(* case Cons *) -apply (intro allI impI) -apply (frule_tac xs=s1 in evals_xcpt, simp only: gx_conv) (* establish gx s1 = None *) -apply (frule wtpd_exprs_cons) - (* establish ((Norm s0)::\ \) *) -apply (frule_tac e=e in state_ok_eval) apply (simp (no_asm_simp) only: env_of_jmb_fst) apply simp apply (simp (no_asm_use) only: env_of_jmb_fst) + (* case Cons *) + apply (intro allI impI) + apply (frule_tac xs=s1 in evals_xcpt, simp only: gx_conv) (* establish gx s1 = None *) + apply (frule wtpd_exprs_cons) + (* establish ((Norm s0)::\ \) *) + apply (frule_tac e=e in state_ok_eval) + apply (simp (no_asm_simp) only: env_of_jmb_fst) + apply simp + apply (simp (no_asm_use) only: env_of_jmb_fst) -apply simp -apply (rule_tac ?instrs0.0 = "(compExpr (gmb G CL S) e)" in progression_transitive, rule HOL.refl) -apply fast -apply fast + apply simp + apply (rule_tac ?instrs0.0 = "(compExpr (gmb G CL S) e)" in progression_transitive, rule HOL.refl) + apply fast + apply fast -(* case Statement: exception *) -apply simp + (* case Statement: exception *) + apply simp -(* case Skip *) -apply (intro allI impI) -apply simp -apply (rule progression_refl) + (* case Skip *) + apply (intro allI impI) + apply simp + apply (rule progression_refl) -(* case Expr *) -apply (intro allI impI) -apply (frule wtpd_stmt_expr) -apply simp -apply (rule_tac ?instrs0.0 = "(compExpr (gmb G CL S) e)" in progression_transitive, rule HOL.refl) -apply fast -apply (rule progression_one_step) -apply simp + (* case Expr *) + apply (intro allI impI) + apply (frule wtpd_stmt_expr) + apply simp + apply (rule_tac ?instrs0.0 = "(compExpr (gmb G CL S) e)" in progression_transitive, rule HOL.refl) + apply fast + apply (rule progression_one_step) + apply simp -(* case Comp *) -apply (intro allI impI) -apply (frule_tac xs=s1 in exec_xcpt, assumption) (* establish (gx s1 = None) *) -apply (frule wtpd_stmt_comp) + (* case Comp *) + apply (intro allI impI) + apply (frule_tac xs=s1 in exec_xcpt, assumption) (* establish (gx s1 = None) *) + apply (frule wtpd_stmt_comp) - (* establish (s1::\ \) *) -apply (frule_tac st=c1 in state_ok_exec) apply (simp (no_asm_simp) only: env_of_jmb_fst) apply simp apply (simp (no_asm_use) only: env_of_jmb_fst) + (* establish (s1::\ \) *) + apply (frule_tac st=c1 in state_ok_exec) + apply (simp (no_asm_simp) only: env_of_jmb_fst) + apply simp apply (simp (no_asm_use) only: env_of_jmb_fst) -apply simp -apply (rule_tac ?instrs0.0 = "(compStmt (gmb G CL S) c1)" in progression_transitive, rule HOL.refl) -apply fast -apply fast + apply simp + apply (rule_tac ?instrs0.0 = "(compStmt (gmb G CL S) c1)" in progression_transitive, rule HOL.refl) + apply fast + apply fast -(* case Cond *) -apply (intro allI impI) -apply (frule_tac xs=s1 in exec_xcpt, assumption) (* establish (gx s1 = None) *) -apply (frule wtpd_stmt_cond, (erule conjE)+) -(* establish (s1::\ \) *) -apply (frule_tac e=e in state_ok_eval) -apply (simp (no_asm_simp) only: env_of_jmb_fst) -apply assumption -apply (simp (no_asm_use) only: env_of_jmb_fst) -(* establish G,gh s1\v::\PrimT Boolean *) -apply (frule eval_conf, assumption+, rule env_of_jmb_fst) -apply (frule conf_bool) (* establish \b. v = Bool b *) -apply (erule exE) + (* case Cond *) + apply (intro allI impI) + apply (frule_tac xs=s1 in exec_xcpt, assumption) (* establish (gx s1 = None) *) + apply (frule wtpd_stmt_cond, (erule conjE)+) + (* establish (s1::\ \) *) + apply (frule_tac e=e in state_ok_eval) + apply (simp (no_asm_simp) only: env_of_jmb_fst) + apply assumption + apply (simp (no_asm_use) only: env_of_jmb_fst) + (* establish G,gh s1\v::\PrimT Boolean *) + apply (frule eval_conf, assumption+, rule env_of_jmb_fst) + apply (frule conf_bool) (* establish \b. v = Bool b *) + apply (erule exE) -apply simp -apply (rule_tac ?instrs0.0 = "[LitPush (Bool False)]" in progression_transitive, simp (no_asm_simp)) -apply (rule progression_one_step, simp) - -apply (rule_tac ?instrs0.0 = "compExpr (gmb G CL S) e" in progression_transitive, rule HOL.refl) -apply fast + apply simp + apply (rule_tac ?instrs0.0 = "[LitPush (Bool False)]" in progression_transitive, simp (no_asm_simp)) + apply (rule progression_one_step, simp) -apply (case_tac b) - (* case b= True *) -apply simp -apply (rule_tac ?instrs0.0 = "[Ifcmpeq (2 + int (length (compStmt (gmb G CL S) c1)))]" in progression_transitive, simp) -apply (rule progression_one_step) apply simp -apply (rule_tac ?instrs0.0 = "(compStmt (gmb G CL S) c1)" in progression_transitive, simp) -apply fast -apply (rule_tac ?instrs1.0 = "[]" in jump_fwd_progression) -apply (simp) -apply simp apply (rule conjI, simp) apply (simp add: nat_add_distrib) -apply (rule progression_refl) + apply (rule_tac ?instrs0.0 = "compExpr (gmb G CL S) e" in progression_transitive, rule HOL.refl) + apply fast - (* case b= False *) -apply simp -apply (rule_tac ?instrs1.0 = "compStmt (gmb G CL S) c2" in jump_fwd_progression) -apply (simp) -apply (simp, rule conjI, rule HOL.refl, simp add: nat_add_distrib) -apply fast + apply (case_tac b) + (* case b= True *) + apply simp + apply (rule_tac ?instrs0.0 = "[Ifcmpeq (2 + int (length (compStmt (gmb G CL S) c1)))]" in progression_transitive, simp) + apply (rule progression_one_step) + apply simp + apply (rule_tac ?instrs0.0 = "(compStmt (gmb G CL S) c1)" in progression_transitive, simp) + apply fast + apply (rule_tac ?instrs1.0 = "[]" in jump_fwd_progression) + apply (simp) + apply simp + apply (rule conjI, simp) + apply (simp add: nat_add_distrib) + apply (rule progression_refl) -(* case exit Loop *) -apply (intro allI impI) -apply (frule wtpd_stmt_loop, (erule conjE)+) - -(* establish G,gh s1\v::\PrimT Boolean *) -apply (frule eval_conf, assumption+, rule env_of_jmb_fst) -apply (frule conf_bool) (* establish \b. v = Bool b *) -apply (erule exE) -apply (case_tac b) + (* case b= False *) + apply simp + apply (rule_tac ?instrs1.0 = "compStmt (gmb G CL S) c2" in jump_fwd_progression) + apply (simp) + apply (simp, rule conjI, rule HOL.refl, simp add: nat_add_distrib) + apply fast - (* case b= True \ contradiction *) -apply simp + (* case exit Loop *) + apply (intro allI impI) + apply (frule wtpd_stmt_loop, (erule conjE)+) - (* case b= False *) -apply simp + (* establish G,gh s1\v::\PrimT Boolean *) + apply (frule eval_conf, assumption+, rule env_of_jmb_fst) + apply (frule conf_bool) (* establish \b. v = Bool b *) + apply (erule exE) + apply (case_tac b) + + (* case b= True \ contradiction *) + apply simp -apply (rule_tac ?instrs0.0 = "[LitPush (Bool False)]" in progression_transitive, simp (no_asm_simp)) -apply (rule progression_one_step) - apply simp + (* case b= False *) + apply simp + + apply (rule_tac ?instrs0.0 = "[LitPush (Bool False)]" in progression_transitive, simp (no_asm_simp)) + apply (rule progression_one_step) + apply simp -apply (rule_tac ?instrs0.0 = "compExpr (gmb G CL S) e" in progression_transitive, rule HOL.refl) -apply fast -apply (rule_tac ?instrs1.0 = "[]" in jump_fwd_progression) -apply (simp) -apply (simp, rule conjI, rule HOL.refl, simp add: nat_add_distrib) -apply (rule progression_refl) + apply (rule_tac ?instrs0.0 = "compExpr (gmb G CL S) e" in progression_transitive, rule HOL.refl) + apply fast + apply (rule_tac ?instrs1.0 = "[]" in jump_fwd_progression) + apply (simp) + apply (simp, rule conjI, rule HOL.refl, simp add: nat_add_distrib) + apply (rule progression_refl) -(* case continue Loop *) -apply (intro allI impI) -apply (frule_tac xs=s2 in exec_xcpt, assumption) (* establish (gx s2 = None) *) -apply (frule_tac xs=s1 in exec_xcpt, assumption) (* establish (gx s1 = None) *) -apply (frule wtpd_stmt_loop, (erule conjE)+) - -(* establish (s1::\ \) *) -apply (frule_tac e=e in state_ok_eval) apply (simp (no_asm_simp) only: env_of_jmb_fst) apply simp apply (simp (no_asm_use) only: env_of_jmb_fst) -(* establish (s2::\ \) *) -apply (frule_tac xs=s1 and st=c in state_ok_exec) -apply (simp (no_asm_simp) only: env_of_jmb_fst) apply assumption apply (simp (no_asm_use) only: env_of_jmb_fst) + (* case continue Loop *) + apply (intro allI impI) + apply (frule_tac xs=s2 in exec_xcpt, assumption) (* establish (gx s2 = None) *) + apply (frule_tac xs=s1 in exec_xcpt, assumption) (* establish (gx s1 = None) *) + apply (frule wtpd_stmt_loop, (erule conjE)+) -(* establish G,gh s1\v::\PrimT Boolean *) -apply (frule eval_conf, assumption+, rule env_of_jmb_fst) -apply (frule conf_bool) (* establish \b. v = Bool b *) -apply (erule exE) + (* establish (s1::\ \) *) + apply (frule_tac e=e in state_ok_eval) + apply (simp (no_asm_simp) only: env_of_jmb_fst) + apply simp + apply (simp (no_asm_use) only: env_of_jmb_fst) + (* establish (s2::\ \) *) + apply (frule_tac xs=s1 and st=c in state_ok_exec) + apply (simp (no_asm_simp) only: env_of_jmb_fst) + apply assumption + apply (simp (no_asm_use) only: env_of_jmb_fst) -apply simp -apply (rule jump_bwd_progression) -apply simp + (* establish G,gh s1\v::\PrimT Boolean *) + apply (frule eval_conf, assumption+, rule env_of_jmb_fst) + apply (frule conf_bool) (* establish \b. v = Bool b *) + apply (erule exE) -apply (rule_tac ?instrs0.0 = "[LitPush (Bool False)]" in progression_transitive, simp (no_asm_simp)) -apply (rule progression_one_step) - apply simp + apply simp + apply (rule jump_bwd_progression) + apply simp + + apply (rule_tac ?instrs0.0 = "[LitPush (Bool False)]" in progression_transitive, simp (no_asm_simp)) + apply (rule progression_one_step) + apply simp -apply (rule_tac ?instrs0.0 = "compExpr (gmb G CL S) e" in progression_transitive, rule HOL.refl) -apply fast + apply (rule_tac ?instrs0.0 = "compExpr (gmb G CL S) e" in progression_transitive, rule HOL.refl) + apply fast -apply (case_tac b) - (* case b= True *) -apply simp + apply (case_tac b) + (* case b= True *) + apply simp -apply (rule_tac ?instrs0.0 = "[Ifcmpeq (2 + int (length (compStmt (gmb G CL S) c)))]" in progression_transitive, simp) -apply (rule progression_one_step) apply simp -apply fast + apply (rule_tac ?instrs0.0 = "[Ifcmpeq (2 + int (length (compStmt (gmb G CL S) c)))]" in progression_transitive, simp) + apply (rule progression_one_step) + apply simp + apply fast - (* case b= False \ contradiction*) -apply simp + (* case b= False \ contradiction*) + apply simp -apply (rule jump_bwd_one_step) -apply simp -apply blast + apply (rule jump_bwd_one_step) + apply simp + apply blast -(*****) -(* case method call *) + (*****) + (* case method call *) -apply (intro allI impI) + apply (intro allI impI) -apply (frule_tac xs=s3 in eval_xcpt, simp only: gx_conv) (* establish gx s3 = None *) -apply (frule exec_xcpt, assumption, simp (no_asm_use) only: gx_conv, frule np_NoneD) (* establish x = None \ a' \ Null *) -apply (frule evals_xcpt, simp only: gx_conv) (* establish gx s1 = None *) + apply (frule_tac xs=s3 in eval_xcpt, simp only: gx_conv) (* establish gx s3 = None *) + apply (frule exec_xcpt, assumption, simp (no_asm_use) only: gx_conv, frule np_NoneD) (* establish x = None \ a' \ Null *) + apply (frule evals_xcpt, simp only: gx_conv) (* establish gx s1 = None *) -apply (frule wtpd_expr_call, (erule conjE)+) + apply (frule wtpd_expr_call, (erule conjE)+) -(* assumptions about state_ok and is_class *) + (* assumptions about state_ok and is_class *) -(* establish s1::\ (env_of_jmb G CL S) *) -apply (frule_tac xs="Norm s0" and e=e in state_ok_eval) -apply (simp (no_asm_simp) only: env_of_jmb_fst, assumption, simp (no_asm_use) only: env_of_jmb_fst) + (* establish s1::\ (env_of_jmb G CL S) *) + apply (frule_tac xs="Norm s0" and e=e in state_ok_eval) + apply (simp (no_asm_simp) only: env_of_jmb_fst, assumption, simp (no_asm_use) only: env_of_jmb_fst) -(* establish (x, h, l)::\(env_of_jmb G CL S) *) -apply (frule_tac xs=s1 and xs'="(x, h, l)" in state_ok_evals) -apply (simp (no_asm_simp) only: env_of_jmb_fst, assumption, simp only: env_of_jmb_fst) + (* establish (x, h, l)::\(env_of_jmb G CL S) *) + apply (frule_tac xs=s1 and xs'="(x, h, l)" in state_ok_evals) + apply (simp (no_asm_simp) only: env_of_jmb_fst, assumption, simp only: env_of_jmb_fst) -(* establish \ lc. a' = Addr lc *) -apply (frule (5) eval_of_class, rule env_of_jmb_fst [symmetric]) -apply (subgoal_tac "G,h \ a' ::\ Class C") -apply (subgoal_tac "is_class G dynT") + (* establish \ lc. a' = Addr lc *) + apply (frule (5) eval_of_class, rule env_of_jmb_fst [symmetric]) + apply (subgoal_tac "G,h \ a' ::\ Class C") + apply (subgoal_tac "is_class G dynT") -(* establish method (G, dynT) (mn, pTs) = Some(md, rT, pns, lvars, blk, res) *) -apply (drule method_defined, assumption+) -apply (simp only: env_of_jmb_fst) -apply ((erule exE)+, erule conjE, (rule exI)+, assumption) + (* establish method (G, dynT) (mn, pTs) = Some(md, rT, pns, lvars, blk, res) *) + apply (drule method_defined, assumption+) + apply (simp only: env_of_jmb_fst) + apply ((erule exE)+, erule conjE, (rule exI)+, assumption) -apply (subgoal_tac "is_class G md") -apply (subgoal_tac "G\Class dynT \ Class md") -apply (subgoal_tac " method (G, md) (mn, pTs) = Some (md, rT, pns, lvars, blk, res)") -apply (subgoal_tac "list_all2 (conf G h) pvs pTs") + apply (subgoal_tac "is_class G md") + apply (subgoal_tac "G\Class dynT \ Class md") + apply (subgoal_tac " method (G, md) (mn, pTs) = Some (md, rT, pns, lvars, blk, res)") + apply (subgoal_tac "list_all2 (conf G h) pvs pTs") -(* establish (np a' x, h, init_vars lvars(pns[\]pvs)(This\a'))::\(env_of_jmb G md (mn, pTs)) *) -apply (subgoal_tac "G,h \ a' ::\ Class dynT") -apply (frule (2) conf_widen) -apply (frule state_ok_init, assumption+) + (* establish (np a' x, h, init_vars lvars(pns[\]pvs)(This\a'))::\(env_of_jmb G md (mn, pTs)) *) + apply (subgoal_tac "G,h \ a' ::\ Class dynT") + apply (frule (2) conf_widen) + apply (frule state_ok_init, assumption+) -apply (subgoal_tac "class_sig_defined G md (mn, pTs)") -apply (frule wtpd_blk, assumption, assumption) -apply (frule wtpd_res, assumption, assumption) -apply (subgoal_tac "s3::\(env_of_jmb G md (mn, pTs))") + apply (subgoal_tac "class_sig_defined G md (mn, pTs)") + apply (frule wtpd_blk, assumption, assumption) + apply (frule wtpd_res, assumption, assumption) + apply (subgoal_tac "s3::\(env_of_jmb G md (mn, pTs))") -apply (subgoal_tac "method (TranslComp.comp G, md) (mn, pTs) = - Some (md, rT, snd (snd (compMethod G md ((mn, pTs), rT, pns, lvars, blk, res))))") -prefer 2 apply (simp add: wf_prog_ws_prog [THEN comp_method]) -apply (subgoal_tac "method (TranslComp.comp G, dynT) (mn, pTs) = - Some (md, rT, snd (snd (compMethod G md ((mn, pTs), rT, pns, lvars, blk, res))))") -prefer 2 apply (simp add: wf_prog_ws_prog [THEN comp_method]) - apply (simp only: fst_conv snd_conv) + apply (subgoal_tac "method (TranslComp.comp G, md) (mn, pTs) = + Some (md, rT, snd (snd (compMethod G md ((mn, pTs), rT, pns, lvars, blk, res))))") + prefer 2 + apply (simp add: wf_prog_ws_prog [THEN comp_method]) + apply (subgoal_tac "method (TranslComp.comp G, dynT) (mn, pTs) = + Some (md, rT, snd (snd (compMethod G md ((mn, pTs), rT, pns, lvars, blk, res))))") + prefer 2 + apply (simp add: wf_prog_ws_prog [THEN comp_method]) + apply (simp only: fst_conv snd_conv) -(* establish length pns = length pTs *) -apply (frule method_preserves_length, assumption, assumption) -(* establish length pvs = length ps *) -apply (frule evals_preserves_length [symmetric]) + (* establish length pns = length pTs *) + apply (frule method_preserves_length, assumption, assumption) + (* establish length pvs = length ps *) + apply (frule evals_preserves_length [symmetric]) -(** start evaluating subexpressions **) -apply (simp (no_asm_use) only: compExpr.simps compExprs.simps) - - (* evaluate e *) -apply (rule_tac ?instrs0.0 = "(compExpr (gmb G CL S) e)" in progression_transitive, rule HOL.refl) -apply fast + (** start evaluating subexpressions **) + apply (simp (no_asm_use) only: compExpr.simps compExprs.simps) - (* evaluate parameters *) -apply (rule_tac ?instrs0.0 = "compExprs (gmb G CL S) ps" in progression_transitive, rule HOL.refl) -apply fast + (* evaluate e *) + apply (rule_tac ?instrs0.0 = "(compExpr (gmb G CL S) e)" in progression_transitive, rule HOL.refl) + apply fast + + (* evaluate parameters *) + apply (rule_tac ?instrs0.0 = "compExprs (gmb G CL S) ps" in progression_transitive, rule HOL.refl) + apply fast - (* invokation *) -apply (rule progression_call) -apply (intro allI impI conjI) - (* execute Invoke statement *) -apply (simp (no_asm_use) only: exec_instr.simps) -apply (erule thin_rl, erule thin_rl, erule thin_rl) -apply (simp add: compMethod_def raise_system_xcpt_def) + (* invokation *) + apply (rule progression_call) + apply (intro allI impI conjI) + (* execute Invoke statement *) + apply (simp (no_asm_use) only: exec_instr.simps) + apply (erule thin_rl, erule thin_rl, erule thin_rl) + apply (simp add: compMethod_def raise_system_xcpt_def) - (* get instructions of invoked method *) -apply (simp (no_asm_simp) add: gis_def gmb_def compMethod_def) + (* get instructions of invoked method *) + apply (simp (no_asm_simp) add: gis_def gmb_def compMethod_def) - (* var. initialization *) -apply (rule_tac ?instrs0.0 = "(compInitLvars (pns, lvars, blk, res) lvars)" in progression_transitive, rule HOL.refl) -apply (rule_tac C=md in progression_lvar_init, assumption, assumption, assumption) -apply (simp (no_asm_simp)) (* length pns = length pvs *) -apply (simp (no_asm_simp)) (* length lvars = length (replicate (length lvars) undefined) *) + (* var. initialization *) + apply (rule_tac ?instrs0.0 = "(compInitLvars (pns, lvars, blk, res) lvars)" in progression_transitive, rule HOL.refl) + apply (rule_tac C=md in progression_lvar_init, assumption, assumption, assumption) + apply (simp (no_asm_simp)) (* length pns = length pvs *) + apply (simp (no_asm_simp)) (* length lvars = length (replicate (length lvars) undefined) *) - (* body statement *) -apply (rule_tac ?instrs0.0 = "compStmt (pns, lvars, blk, res) blk" in progression_transitive, rule HOL.refl) -apply (subgoal_tac "(pns, lvars, blk, res) = gmb G md (mn, pTs)") -apply (simp (no_asm_simp)) -apply (simp only: gh_conv) -apply (drule mp [OF _ TrueI])+ -apply (erule allE, erule allE, erule allE, erule impE, assumption)+ -apply ((erule impE, assumption)+, assumption) + (* body statement *) + apply (rule_tac ?instrs0.0 = "compStmt (pns, lvars, blk, res) blk" in progression_transitive, rule HOL.refl) + apply (subgoal_tac "(pns, lvars, blk, res) = gmb G md (mn, pTs)") + apply (simp (no_asm_simp)) + apply (simp only: gh_conv) + apply (drule mp [OF _ TrueI])+ + apply (erule allE, erule allE, erule allE, erule impE, assumption)+ + apply ((erule impE, assumption)+, assumption) -apply (simp (no_asm_use)) -apply (simp (no_asm_simp) add: gmb_def) + apply (simp (no_asm_use)) + apply (simp (no_asm_simp) add: gmb_def) - (* return expression *) -apply (subgoal_tac "(pns, lvars, blk, res) = gmb G md (mn, pTs)") -apply (simp (no_asm_simp)) -apply (simp only: gh_conv) -apply ((drule mp, rule TrueI)+, (drule spec)+, (drule mp, assumption)+, assumption) -apply (simp (no_asm_use)) -apply (simp (no_asm_simp) add: gmb_def) + (* return expression *) + apply (subgoal_tac "(pns, lvars, blk, res) = gmb G md (mn, pTs)") + apply (simp (no_asm_simp)) + apply (simp only: gh_conv) + apply ((drule mp, rule TrueI)+, (drule spec)+, (drule mp, assumption)+, assumption) + apply (simp (no_asm_use)) + apply (simp (no_asm_simp) add: gmb_def) - (* execute return statement *) -apply (simp (no_asm_use) add: gh_def locvars_xstate_def gl_def del: drop_append) -apply (subgoal_tac "rev pvs @ a' # os = (rev (a' # pvs)) @ os") -apply (simp only: drop_append) -apply (simp (no_asm_simp)) -apply (simp (no_asm_simp)) + (* execute return statement *) + apply (simp (no_asm_use) add: gh_def locvars_xstate_def gl_def del: drop_append) + apply (subgoal_tac "rev pvs @ a' # os = (rev (a' # pvs)) @ os") + apply (simp only: drop_append) + apply (simp (no_asm_simp)) + apply (simp (no_asm_simp)) -(* show s3::\\ *) -apply (rule_tac xs = "(np a' x, h, init_vars lvars(pns[\]pvs)(This\a'))" and st=blk in state_ok_exec) -apply assumption apply (simp (no_asm_simp) only: env_of_jmb_fst) -apply assumption apply (simp (no_asm_use) only: env_of_jmb_fst) + (* show s3::\\ *) + apply (rule_tac xs = "(np a' x, h, init_vars lvars(pns[\]pvs)(This\a'))" and st=blk in state_ok_exec) + apply assumption + apply (simp (no_asm_simp) only: env_of_jmb_fst) + apply assumption + apply (simp (no_asm_use) only: env_of_jmb_fst) -(* show class_sig_defined G md (mn, pTs) *) -apply (simp (no_asm_simp) add: class_sig_defined_def) + (* show class_sig_defined G md (mn, pTs) *) + apply (simp (no_asm_simp) add: class_sig_defined_def) -(* show G,h \ a' ::\ Class dynT *) -apply (frule non_npD) apply assumption -apply (erule exE)+ apply simp -apply (rule conf_obj_AddrI) apply simp -apply (rule widen_Class_Class [THEN iffD1], rule widen.refl) + (* show G,h \ a' ::\ Class dynT *) + apply (frule non_npD) + apply assumption + apply (erule exE)+ + apply simp + apply (rule conf_obj_AddrI) + apply simp + apply (rule widen_Class_Class [THEN iffD1], rule widen.refl) - (* show list_all2 (conf G h) pvs pTs *) -apply (erule exE)+ apply (erule conjE)+ -apply (rule_tac Ts="pTsa" in conf_list_gext_widen) apply assumption -apply (subgoal_tac "G \ (gx s1, gs s1) -ps[\]pvs-> (x, h, l)") -apply (frule_tac E="env_of_jmb G CL S" in evals_type_sound) -apply assumption+ -apply (simp only: env_of_jmb_fst) -apply (simp add: conforms_def xconf_def gs_def) -apply simp -apply (simp (no_asm_use) only: gx_def gs_def surjective_pairing [symmetric]) -apply (simp (no_asm_use) only: ty_exprs_list_all2) apply simp -apply simp -apply (simp (no_asm_use) only: gx_def gs_def surjective_pairing [symmetric]) - (* list_all2 (\T T'. G \ T \ T') pTsa pTs *) - apply (rule max_spec_widen, simp only: env_of_jmb_fst) + (* show list_all2 (conf G h) pvs pTs *) + apply (erule exE)+ apply (erule conjE)+ + apply (rule_tac Ts="pTsa" in conf_list_gext_widen) + apply assumption + apply (subgoal_tac "G \ (gx s1, gs s1) -ps[\]pvs-> (x, h, l)") + apply (frule_tac E="env_of_jmb G CL S" in evals_type_sound) + apply assumption+ + apply (simp only: env_of_jmb_fst) + apply (simp add: conforms_def xconf_def gs_def) + apply simp + apply (simp (no_asm_use) only: gx_def gs_def surjective_pairing [symmetric]) + apply (simp (no_asm_use) only: ty_exprs_list_all2) + apply simp + apply simp + apply (simp (no_asm_use) only: gx_def gs_def surjective_pairing [symmetric]) + (* list_all2 (\T T'. G \ T \ T') pTsa pTs *) + apply (rule max_spec_widen, simp only: env_of_jmb_fst) -(* show method (G, md) (mn, pTs) = Some (md, rT, pns, lvars, blk, res) *) -apply (frule wf_prog_ws_prog [THEN method_in_md [THEN conjunct2]], assumption+) + (* show method (G, md) (mn, pTs) = Some (md, rT, pns, lvars, blk, res) *) + apply (frule wf_prog_ws_prog [THEN method_in_md [THEN conjunct2]], assumption+) - (* show G\Class dynT \ Class md *) -apply (simp (no_asm_use) only: widen_Class_Class) -apply (rule method_wf_mdecl [THEN conjunct1], assumption+) + (* show G\Class dynT \ Class md *) + apply (simp (no_asm_use) only: widen_Class_Class) + apply (rule method_wf_mdecl [THEN conjunct1], assumption+) - (* is_class G md *) -apply (rule wf_prog_ws_prog [THEN method_in_md [THEN conjunct1]], assumption+) + (* is_class G md *) + apply (rule wf_prog_ws_prog [THEN method_in_md [THEN conjunct1]], assumption+) - (* show is_class G dynT *) -apply (frule non_npD) apply assumption -apply (erule exE)+ apply (erule conjE)+ -apply simp -apply (rule subcls_is_class2) apply assumption -apply (frule class_expr_is_class) apply (simp only: env_of_jmb_fst) -apply (rule wf_prog_ws_prog, assumption) -apply (simp only: env_of_jmb_fst) + (* show is_class G dynT *) + apply (frule non_npD) + apply assumption + apply (erule exE)+ + apply (erule conjE)+ + apply simp + apply (rule subcls_is_class2) + apply assumption + apply (frule expr_class_is_class [rotated]) + apply (simp only: env_of_jmb_fst) + apply (rule wf_prog_ws_prog, assumption) + apply (simp only: env_of_jmb_fst) - (* show G,h \ a' ::\ Class C *) -apply (simp only: wtpd_exprs_def, erule exE) -apply (frule evals_preserves_conf) -apply (rule eval_conf, assumption+) -apply (rule env_of_jmb_fst, assumption+) -apply (rule env_of_jmb_fst) -apply (simp only: gh_conv) -done + (* show G,h \ a' ::\ Class C *) + apply (simp only: wtpd_exprs_def, erule exE) + apply (frule evals_preserves_conf) + apply (rule eval_conf, assumption+) + apply (rule env_of_jmb_fst, assumption+) + apply (rule env_of_jmb_fst) + apply (simp only: gh_conv) + done theorem compiler_correctness_eval: " @@ -1226,9 +1254,9 @@ {hp, os, (locvars_locals G C S loc)} >- (compExpr (gmb G C S) ex) \ {hp', val#os, (locvars_locals G C S loc')}" -apply (frule compiler_correctness [THEN conjunct1]) -apply (auto simp: gh_def gx_def gs_def gl_def locvars_xstate_def) -done + apply (frule compiler_correctness [THEN conjunct1]) + apply (auto simp: gh_def gx_def gs_def gl_def locvars_xstate_def) + done theorem compiler_correctness_exec: " \ G \ Norm (hp, loc) -st-> Norm (hp', loc'); @@ -1240,9 +1268,9 @@ {hp, os, (locvars_locals G C S loc)} >- (compStmt (gmb G C S) st) \ {hp', os, (locvars_locals G C S loc')}" -apply (frule compiler_correctness [THEN conjunct2 [THEN conjunct2]]) -apply (auto simp: gh_def gx_def gs_def gl_def locvars_xstate_def) -done + apply (frule compiler_correctness [THEN conjunct2 [THEN conjunct2]]) + apply (auto simp: gh_def gx_def gs_def gl_def locvars_xstate_def) + done (**********************************************************************) diff -r 00c06f1315d0 -r 3f429b7d8eb5 src/HOL/MicroJava/Comp/CorrCompTp.thy --- a/src/HOL/MicroJava/Comp/CorrCompTp.thy Tue May 26 21:58:04 2015 +0100 +++ b/src/HOL/MicroJava/Comp/CorrCompTp.thy Thu May 28 17:25:57 2015 +1000 @@ -23,174 +23,171 @@ let (mn, pTs) = S in (G,map_of lvars(pns[\]pTs)(This\Class C))" lemma local_env_fst [simp]: "fst (local_env G C S pns lvars) = G" -by (simp add: local_env_def split_beta) - - -lemma wt_class_expr_is_class: "\ ws_prog G; E \ expr :: Class cname; - E = local_env G C (mn, pTs) pns lvars\ + by (simp add: local_env_def split_beta) + + +lemma wt_class_expr_is_class: + "\ ws_prog G; E \ expr :: Class cname; E = local_env G C (mn, pTs) pns lvars\ \ is_class G cname " -apply (subgoal_tac "((fst E), (snd E)) \ expr :: Class cname") -apply (frule ty_expr_is_type) apply simp -apply simp apply (simp (no_asm_use)) -done + apply (subgoal_tac "((fst E), (snd E)) \ expr :: Class cname") + apply (frule ty_expr_is_type) + apply simp + apply simp + apply (simp (no_asm_use)) + done (********************************************************************************) subsection "index" -lemma local_env_snd: " - snd (local_env G C (mn, pTs) pns lvars) = map_of lvars(pns[\]pTs)(This\Class C)" -by (simp add: local_env_def) - - - -lemma index_in_bounds: " length pns = length pTs \ +lemma local_env_snd: + "snd (local_env G C (mn, pTs) pns lvars) = map_of lvars(pns[\]pTs)(This\Class C)" + by (simp add: local_env_def) + + + +lemma index_in_bounds: + "length pns = length pTs \ snd (local_env G C (mn, pTs) pns lvars) vname = Some T \ index (pns, lvars, blk, res) vname < length (inited_LT C pTs lvars)" -apply (simp add: local_env_snd index_def split_beta) -apply (case_tac "vname = This") -apply (simp add: inited_LT_def) -apply simp -apply (drule map_of_upds_SomeD) -apply (drule length_takeWhile) -apply (simp add: inited_LT_def) -done - - -lemma map_upds_append [rule_format (no_asm)]: - "\ x1s m. (length k1s = length x1s - \ m(k1s[\]x1s)(k2s[\]x2s) = m ((k1s@k2s)[\](x1s@x2s)))" -apply (induct k1s) -apply simp -apply (intro strip) -apply (subgoal_tac "\ x xr. x1s = x # xr") -apply clarify -apply simp + apply (simp add: local_env_snd index_def split_beta) + apply (case_tac "vname = This") + apply (simp add: inited_LT_def) + apply simp + apply (drule map_of_upds_SomeD) + apply (drule length_takeWhile) + apply (simp add: inited_LT_def) + done + + +lemma map_upds_append: + "length k1s = length x1s \ m(k1s[\]x1s)(k2s[\]x2s) = m ((k1s@k2s)[\](x1s@x2s))" + apply (induct k1s arbitrary: x1s m) + apply simp + apply (subgoal_tac "\x xr. x1s = x # xr") + apply clarsimp (* subgoal *) -apply (case_tac x1s) -apply auto -done - - -lemma map_of_append [rule_format]: - "\ ys. (map_of ((rev xs) @ ys) = (map_of ys) ((map fst xs) [\] (map snd xs)))" -apply (induct xs) -apply simp -apply (rule allI) -apply (drule_tac x="a # ys" in spec) -apply (simp only: rev.simps append_assoc append_Cons append_Nil - list.map map_of.simps map_upds_Cons list.sel) -done + apply (case_tac x1s) + apply auto + done + + +lemma map_of_append: + "map_of ((rev xs) @ ys) = (map_of ys) ((map fst xs) [\] (map snd xs))" + apply (induct xs arbitrary: ys) + apply simp + apply (rename_tac a xs ys) + apply (drule_tac x="a # ys" in meta_spec) + apply (simp only: rev.simps append_assoc append_Cons append_Nil + list.map map_of.simps map_upds_Cons list.sel) + done lemma map_of_as_map_upds: "map_of (rev xs) = empty ((map fst xs) [\] (map snd xs))" -by (rule map_of_append [of _ "[]", simplified]) + by (rule map_of_append [of _ "[]", simplified]) lemma map_of_rev: "unique xs \ map_of (rev xs) = map_of xs" -apply (induct xs) -apply simp -apply (simp add: unique_def map_of_append map_of_as_map_upds [symmetric] - Map.map_of_append[symmetric] del:Map.map_of_append) -done - -lemma map_upds_rev [rule_format]: "\ xs. (distinct ks \ length ks = length xs - \ m (rev ks [\] rev xs) = m (ks [\] xs))" -apply (induct ks) -apply simp -apply (intro strip) -apply (subgoal_tac "\ x xr. xs = x # xr") -apply clarify -apply (drule_tac x=xr in spec) -apply (simp add: map_upds_append [symmetric]) - (* subgoal *) -apply (case_tac xs) -apply auto -done + apply (induct xs) + apply simp + apply (simp add: unique_def map_of_append map_of_as_map_upds [symmetric] + Map.map_of_append[symmetric] del:Map.map_of_append) + done + +lemma map_upds_rev: + "\ distinct ks; length ks = length xs \ \ m (rev ks [\] rev xs) = m (ks [\] xs)" + apply (induct ks arbitrary: xs) + apply simp + apply (subgoal_tac "\x xr. xs = x # xr") + apply clarify + apply (drule_tac x=xr in meta_spec) + apply (simp add: map_upds_append [symmetric]) + apply (case_tac xs, auto) + done lemma map_upds_takeWhile [rule_format]: - "\ ks. (empty(rev ks[\]rev xs)) k = Some x \ length ks = length xs \ + "\ks. (empty(rev ks[\]rev xs)) k = Some x \ length ks = length xs \ xs ! length (takeWhile (\z. z \ k) ks) = x" -apply (induct xs) - apply simp -apply (intro strip) -apply (subgoal_tac "\ k' kr. ks = k' # kr") - apply (clarify) - apply (drule_tac x=kr in spec) - apply (simp only: rev.simps) - apply (subgoal_tac "(empty(rev kr @ [k'][\]rev xs @ [a])) = empty (rev kr[\]rev xs)([k'][\][a])") - apply (simp split:split_if_asm) - apply (simp add: map_upds_append [symmetric]) -apply (case_tac ks) - apply auto -done - - -lemma local_env_inited_LT: "\ snd (local_env G C (mn, pTs) pns lvars) vname = Some T; + apply (induct xs) + apply simp + apply (intro strip) + apply (subgoal_tac "\k' kr. ks = k' # kr") + apply (clarify) + apply (drule_tac x=kr in spec) + apply (simp only: rev.simps) + apply (subgoal_tac "(empty(rev kr @ [k'][\]rev xs @ [a])) = empty (rev kr[\]rev xs)([k'][\][a])") + apply (simp split:split_if_asm) + apply (simp add: map_upds_append [symmetric]) + apply (case_tac ks) + apply auto + done + + +lemma local_env_inited_LT: + "\ snd (local_env G C (mn, pTs) pns lvars) vname = Some T; length pns = length pTs; distinct pns; unique lvars \ \ (inited_LT C pTs lvars ! index (pns, lvars, blk, res) vname) = OK T" -apply (simp add: local_env_snd index_def) -apply (case_tac "vname = This") -apply (simp add: inited_LT_def) -apply (simp add: inited_LT_def) -apply (simp (no_asm_simp) only: map_map [symmetric] map_append [symmetric] list.map [symmetric]) -apply (subgoal_tac "length (takeWhile (\z. z \ vname) (pns @ map fst lvars)) < length (pTs @ map snd lvars)") -apply (simp (no_asm_simp) only: List.nth_map ok_val.simps) -apply (subgoal_tac "map_of lvars = map_of (map (\ p. (fst p, snd p)) lvars)") -apply (simp only:) -apply (subgoal_tac "distinct (map fst lvars)") -apply (frule_tac g=snd in AuxLemmas.map_of_map_as_map_upd) -apply (simp only:) -apply (simp add: map_upds_append) -apply (frule map_upds_SomeD) -apply (rule map_upds_takeWhile) -apply (simp (no_asm_simp)) -apply (simp add: map_upds_append [symmetric]) -apply (simp add: map_upds_rev) - - (* show length (pns @ map fst lvars) = length (pTs @ map snd lvars) *) -apply simp - - (* show distinct (map fst lvars) *) -apply (simp only: unique_def Fun.comp_def) - - (* show map_of lvars = map_of (map (\p. (fst p, snd p)) lvars) *) -apply simp + apply (simp add: local_env_snd index_def) + apply (case_tac "vname = This") + apply (simp add: inited_LT_def) + apply (simp add: inited_LT_def) + apply (simp (no_asm_simp) only: map_map [symmetric] map_append [symmetric] list.map [symmetric]) + apply (subgoal_tac "length (takeWhile (\z. z \ vname) (pns @ map fst lvars)) < length (pTs @ map snd lvars)") + apply (simp (no_asm_simp) only: List.nth_map ok_val.simps) + apply (subgoal_tac "map_of lvars = map_of (map (\ p. (fst p, snd p)) lvars)") + apply (simp only:) + apply (subgoal_tac "distinct (map fst lvars)") + apply (frule_tac g=snd in AuxLemmas.map_of_map_as_map_upd) + apply (simp only:) + apply (simp add: map_upds_append) + apply (frule map_upds_SomeD) + apply (rule map_upds_takeWhile) + apply (simp (no_asm_simp)) + apply (simp add: map_upds_append [symmetric]) + apply (simp add: map_upds_rev) + + (* show length (pns @ map fst lvars) = length (pTs @ map snd lvars) *) + apply simp + + (* show distinct (map fst lvars) *) + apply (simp only: unique_def Fun.comp_def) + + (* show map_of lvars = map_of (map (\p. (fst p, snd p)) lvars) *) + apply simp (* show length (takeWhile (\z. z \ vname) (pns @ map fst lvars)) < length (pTs @ map snd lvars) *) -apply (drule map_of_upds_SomeD) -apply (drule length_takeWhile) -apply simp -done - - -lemma inited_LT_at_index_no_err: " i < length (inited_LT C pTs lvars) - \ inited_LT C pTs lvars ! i \ Err" -apply (simp only: inited_LT_def) -apply (simp only: map_map [symmetric] map_append [symmetric] list.map [symmetric] length_map) -apply (simp only: nth_map) -apply simp -done + apply (drule map_of_upds_SomeD) + apply (drule length_takeWhile) + apply simp + done + + +lemma inited_LT_at_index_no_err: + "i < length (inited_LT C pTs lvars) \ inited_LT C pTs lvars ! i \ Err" + apply (simp only: inited_LT_def) + apply (simp only: map_map [symmetric] map_append [symmetric] list.map [symmetric] length_map) + apply (simp only: nth_map) + apply simp + done lemma sup_loc_update_index: " \ G \ T \ T'; is_type G T'; length pns = length pTs; distinct pns; unique lvars; snd (local_env G C (mn, pTs) pns lvars) vname = Some T' \ \ - comp G \ - inited_LT C pTs lvars [index (pns, lvars, blk, res) vname := OK T] <=l - inited_LT C pTs lvars" -apply (subgoal_tac " index (pns, lvars, blk, res) vname < length (inited_LT C pTs lvars)") -apply (frule_tac blk=blk and res=res in local_env_inited_LT, assumption+) -apply (rule sup_loc_trans) -apply (rule_tac b="OK T'" in sup_loc_update) -apply (simp add: comp_widen) -apply assumption -apply (rule sup_loc_refl) -apply (simp add: list_update_same_conv [THEN iffD2]) + comp G \ inited_LT C pTs lvars [index (pns, lvars, blk, res) vname := OK T] <=l + inited_LT C pTs lvars" + apply (subgoal_tac " index (pns, lvars, blk, res) vname < length (inited_LT C pTs lvars)") + apply (frule_tac blk=blk and res=res in local_env_inited_LT, assumption+) + apply (rule sup_loc_trans) + apply (rule_tac b="OK T'" in sup_loc_update) + apply (simp add: comp_widen) + apply assumption + apply (rule sup_loc_refl) + apply (simp add: list_update_same_conv [THEN iffD2]) (* subgoal *) -apply (rule index_in_bounds) -apply simp+ -done + apply (rule index_in_bounds) + apply simp+ + done (********************************************************************************) @@ -198,147 +195,147 @@ subsection "Preservation of ST and LT by compTpExpr / compTpStmt" lemma sttp_of_comb_nil [simp]: "sttp_of (comb_nil sttp) = sttp" -by (simp add: comb_nil_def) + by (simp add: comb_nil_def) lemma mt_of_comb_nil [simp]: "mt_of (comb_nil sttp) = []" -by (simp add: comb_nil_def) + by (simp add: comb_nil_def) lemma sttp_of_comb [simp]: "sttp_of ((f1 \ f2) sttp) = sttp_of (f2 (sttp_of (f1 sttp)))" -apply (case_tac "f1 sttp") -apply (case_tac "(f2 (sttp_of (f1 sttp)))") -apply (simp add: comb_def) -done + apply (case_tac "f1 sttp") + apply (case_tac "(f2 (sttp_of (f1 sttp)))") + apply (simp add: comb_def) + done lemma mt_of_comb: "(mt_of ((f1 \ f2) sttp)) = (mt_of (f1 sttp)) @ (mt_of (f2 (sttp_of (f1 sttp))))" -by (simp add: comb_def split_beta) + by (simp add: comb_def split_beta) lemma mt_of_comb_length [simp]: "\ n1 = length (mt_of (f1 sttp)); n1 \ n \ \ (mt_of ((f1 \ f2) sttp) ! n) = (mt_of (f2 (sttp_of (f1 sttp))) ! (n - n1))" -by (simp add: comb_def nth_append split_beta) + by (simp add: comb_def nth_append split_beta) lemma compTpExpr_Exprs_LT_ST: " - \jmb = (pns, lvars, blk, res); + \jmb = (pns, lvars, blk, res); wf_prog wf_java_mdecl G; wf_java_mdecl G C ((mn, pTs), rT, jmb); E = local_env G C (mn, pTs) pns lvars \ - \ - (\ ST LT T. + \ + (\ ST LT T. E \ ex :: T \ is_inited_LT C pTs lvars LT \ sttp_of (compTpExpr jmb G ex (ST, LT)) = (T # ST, LT)) - \ - (\ ST LT Ts. + \ + (\ ST LT Ts. E \ exs [::] Ts \ is_inited_LT C pTs lvars LT \ sttp_of (compTpExprs jmb G exs (ST, LT)) = ((rev Ts) @ ST, LT))" -apply (rule compat_expr_expr_list.induct) - -(* expresssions *) - -(* NewC *) -apply (intro strip) -apply (drule NewC_invers) -apply (simp add: pushST_def) - -(* Cast *) -apply (intro strip) -apply (drule Cast_invers, clarify) -apply ((drule_tac x=ST in spec), (drule spec)+, (drule mp, assumption)+) -apply (simp add: replST_def split_beta) - -(* Lit *) -apply (intro strip) -apply (drule Lit_invers) -apply (simp add: pushST_def) - -(* BinOp *) -apply (intro strip) -apply (drule BinOp_invers, clarify) -apply (drule_tac x=ST in spec) -apply (drule_tac x="Ta # ST" in spec) -apply ((drule spec)+, (drule mp, assumption)+) - apply (rename_tac binop x2 x3 ST LT T Ta, case_tac binop) - apply (simp (no_asm_simp)) - apply (simp (no_asm_simp) add: popST_def pushST_def) - apply (simp) - apply (simp (no_asm_simp) add: replST_def) - - -(* LAcc *) -apply (intro strip) -apply (drule LAcc_invers) -apply (simp add: pushST_def is_inited_LT_def) -apply (simp add: wf_prog_def) -apply (frule wf_java_mdecl_disjoint_varnames) - apply (simp add: disjoint_varnames_def) -apply (frule wf_java_mdecl_length_pTs_pns) -apply (erule conjE)+ -apply (simp (no_asm_simp) add: local_env_inited_LT) - -(* LAss *) -apply (intro strip) -apply (drule LAss_invers, clarify) -apply (drule LAcc_invers) -apply ((drule_tac x=ST in spec), (drule spec)+, (drule mp, assumption)+) -apply (simp add: popST_def dupST_def) - -(* FAcc *) -apply (intro strip) -apply (drule FAcc_invers, clarify) -apply ((drule_tac x=ST in spec), (drule spec)+, (drule mp, assumption)+) -apply (simp add: replST_def) - - (* show snd (the (field (G, cname) vname)) = T *) - apply (subgoal_tac "is_class G Ca") - apply (rename_tac cname x2 vname ST LT T Ca, subgoal_tac "is_class G cname \ field (G, cname) vname = Some (cname, T)") + apply (rule compat_expr_expr_list.induct) + + (* expresssions *) + + (* NewC *) + apply (intro strip) + apply (drule NewC_invers) + apply (simp add: pushST_def) + + (* Cast *) + apply (intro strip) + apply (drule Cast_invers, clarify) + apply ((drule_tac x=ST in spec), (drule spec)+, (drule mp, assumption)+) + apply (simp add: replST_def split_beta) + + (* Lit *) + apply (intro strip) + apply (drule Lit_invers) + apply (simp add: pushST_def) + + (* BinOp *) + apply (intro strip) + apply (drule BinOp_invers, clarify) + apply (drule_tac x=ST in spec) + apply (drule_tac x="Ta # ST" in spec) + apply ((drule spec)+, (drule mp, assumption)+) + apply (rename_tac binop x2 x3 ST LT T Ta, case_tac binop) + apply (simp (no_asm_simp)) + apply (simp (no_asm_simp) add: popST_def pushST_def) + apply (simp) + apply (simp (no_asm_simp) add: replST_def) + + + (* LAcc *) + apply (intro strip) + apply (drule LAcc_invers) + apply (simp add: pushST_def is_inited_LT_def) + apply (simp add: wf_prog_def) + apply (frule wf_java_mdecl_disjoint_varnames) + apply (simp add: disjoint_varnames_def) + apply (frule wf_java_mdecl_length_pTs_pns) + apply (erule conjE)+ + apply (simp (no_asm_simp) add: local_env_inited_LT) + + (* LAss *) + apply (intro strip) + apply (drule LAss_invers, clarify) + apply (drule LAcc_invers) + apply ((drule_tac x=ST in spec), (drule spec)+, (drule mp, assumption)+) + apply (simp add: popST_def dupST_def) + + (* FAcc *) + apply (intro strip) + apply (drule FAcc_invers, clarify) + apply ((drule_tac x=ST in spec), (drule spec)+, (drule mp, assumption)+) + apply (simp add: replST_def) + + (* show snd (the (field (G, cname) vname)) = T *) + apply (subgoal_tac "is_class G Ca") + apply (rename_tac cname x2 vname ST LT T Ca, subgoal_tac "is_class G cname \ field (G, cname) vname = Some (cname, T)") + apply simp + + (* show is_class G cname \ field (G, cname) vname = Some (cname, T) *) + apply (rule field_in_fd) apply assumption+ + (* show is_class G Ca *) + apply (fast intro: wt_class_expr_is_class) + + (* FAss *) + apply (intro strip) + apply (drule FAss_invers, clarify) + apply (drule FAcc_invers, clarify) + apply (drule_tac x=ST in spec) + apply (drule_tac x="Class Ca # ST" in spec) + apply ((drule spec)+, (drule mp, assumption)+) + apply (simp add: popST_def dup_x1ST_def) + + + (* Call *) + apply (intro strip) + apply (drule Call_invers, clarify) + apply (drule_tac x=ST in spec) + apply (rename_tac cname x2 x3 x4 x5 ST LT T pTsa md, drule_tac x="Class cname # ST" in spec) + apply ((drule spec)+, (drule mp, assumption)+) + apply (simp add: replST_def) + + + (* expression lists *) + (* nil *) + + apply (intro strip) + apply (drule Nil_invers) + apply (simp add: comb_nil_def) + + (* cons *) + + apply (intro strip) + apply (drule Cons_invers, clarify) + apply (drule_tac x=ST in spec) + apply (drule_tac x="T # ST" in spec) + apply ((drule spec)+, (drule mp, assumption)+) apply simp - (* show is_class G cname \ field (G, cname) vname = Some (cname, T) *) - apply (rule field_in_fd) apply assumption+ - (* show is_class G Ca *) - apply (fast intro: wt_class_expr_is_class) - -(* FAss *) -apply (intro strip) -apply (drule FAss_invers, clarify) -apply (drule FAcc_invers, clarify) -apply (drule_tac x=ST in spec) -apply (drule_tac x="Class Ca # ST" in spec) -apply ((drule spec)+, (drule mp, assumption)+) -apply (simp add: popST_def dup_x1ST_def) - - -(* Call *) -apply (intro strip) -apply (drule Call_invers, clarify) -apply (drule_tac x=ST in spec) -apply (rename_tac cname x2 x3 x4 x5 ST LT T pTsa md, drule_tac x="Class cname # ST" in spec) -apply ((drule spec)+, (drule mp, assumption)+) -apply (simp add: replST_def) - - -(* expression lists *) -(* nil *) - -apply (intro strip) -apply (drule Nil_invers) -apply (simp add: comb_nil_def) - -(* cons *) - -apply (intro strip) -apply (drule Cons_invers, clarify) -apply (drule_tac x=ST in spec) -apply (drule_tac x="T # ST" in spec) -apply ((drule spec)+, (drule mp, assumption)+) -apply simp - -done + done @@ -358,50 +355,50 @@ (is_inited_LT C pTs lvars LT) \ sttp_of (compTpStmt jmb G s (ST, LT)) = (ST, LT))" -apply (rule stmt.induct) - -(* Skip *) -apply (intro strip) -apply simp - -(* Expr *) -apply (intro strip) -apply (drule Expr_invers, erule exE) -apply (simp (no_asm_simp) add: compTpExpr_LT_ST) -apply (frule_tac ST=ST in compTpExpr_LT_ST, assumption+) -apply (simp add: popST_def) - -(* Comp *) -apply (intro strip) -apply (drule Comp_invers, clarify) -apply (simp (no_asm_use)) -apply simp - -(* Cond *) -apply (intro strip) -apply (drule Cond_invers) -apply (erule conjE)+ -apply (drule_tac x=ST in spec) -apply (drule_tac x=ST in spec) -apply (drule spec)+ apply (drule mp, assumption)+ -apply (drule_tac ST="PrimT Boolean # ST" in compTpExpr_LT_ST, assumption+) -apply (simp add: popST_def pushST_def nochangeST_def) - -(* Loop *) -apply (intro strip) -apply (drule Loop_invers) -apply (erule conjE)+ -apply (drule_tac x=ST in spec) -apply (drule spec)+ apply (drule mp, assumption)+ -apply (drule_tac ST="PrimT Boolean # ST" in compTpExpr_LT_ST, assumption+) -apply (simp add: popST_def pushST_def nochangeST_def) -done + apply (rule stmt.induct) + + (* Skip *) + apply (intro strip) + apply simp + + (* Expr *) + apply (intro strip) + apply (drule Expr_invers, erule exE) + apply (simp (no_asm_simp) add: compTpExpr_LT_ST) + apply (frule_tac ST=ST in compTpExpr_LT_ST, assumption+) + apply (simp add: popST_def) + + (* Comp *) + apply (intro strip) + apply (drule Comp_invers, clarify) + apply (simp (no_asm_use)) + apply simp + + (* Cond *) + apply (intro strip) + apply (drule Cond_invers) + apply (erule conjE)+ + apply (drule_tac x=ST in spec) + apply (drule_tac x=ST in spec) + apply (drule spec)+ apply (drule mp, assumption)+ + apply (drule_tac ST="PrimT Boolean # ST" in compTpExpr_LT_ST, assumption+) + apply (simp add: popST_def pushST_def nochangeST_def) + + (* Loop *) + apply (intro strip) + apply (drule Loop_invers) + apply (erule conjE)+ + apply (drule_tac x=ST in spec) + apply (drule spec)+ apply (drule mp, assumption)+ + apply (drule_tac ST="PrimT Boolean # ST" in compTpExpr_LT_ST, assumption+) + apply (simp add: popST_def pushST_def nochangeST_def) + done lemma compTpInit_LT_ST: " sttp_of (compTpInit jmb (vn,ty) (ST, LT)) = (ST, LT[(index jmb vn):= OK ty])" -by (simp add: compTpInit_def storeST_def pushST_def) + by (simp add: compTpInit_def storeST_def pushST_def) lemma compTpInitLvars_LT_ST_aux [rule_format (no_asm)]: @@ -413,125 +410,124 @@ \ sttp_of (compTpInitLvars jmb lvars (ST, pre @ replicate (length lvars) Err)) = (ST, pre @ map (Fun.comp OK snd) lvars)" -apply (induct lvars) -apply simp - -apply (intro strip) -apply (subgoal_tac "\ vn ty. a = (vn, ty)") - prefer 2 apply (simp (no_asm_simp)) + apply (induct lvars) + apply simp + + apply (intro strip) + apply (subgoal_tac "\vn ty. a = (vn, ty)") + prefer 2 + apply (simp (no_asm_simp)) apply ((erule exE)+, simp (no_asm_simp)) -apply (drule_tac x="pre @ [OK ty]" in spec) -apply (drule_tac x="lvars_pre @ [a]" in spec) -apply (drule_tac x="lvars0" in spec) -apply (simp add: compTpInit_LT_ST index_of_var2) -done + apply (drule_tac x="pre @ [OK ty]" in spec) + apply (drule_tac x="lvars_pre @ [a]" in spec) + apply (drule_tac x="lvars0" in spec) + apply (simp add: compTpInit_LT_ST index_of_var2) + done lemma compTpInitLvars_LT_ST: "\ jmb = (pns, lvars, blk, res); wf_java_mdecl G C ((mn, pTs), rT, jmb) \ -\(sttp_of (compTpInitLvars jmb lvars (ST, start_LT C pTs (length lvars)))) - = (ST, inited_LT C pTs lvars)" -apply (simp add: start_LT_def inited_LT_def) -apply (simp only: append_Cons [symmetric]) -apply (rule compTpInitLvars_LT_ST_aux) -apply (auto dest: wf_java_mdecl_length_pTs_pns wf_java_mdecl_disjoint_varnames) -done + \ sttp_of (compTpInitLvars jmb lvars (ST, start_LT C pTs (length lvars))) + = (ST, inited_LT C pTs lvars)" + apply (simp add: start_LT_def inited_LT_def) + apply (simp only: append_Cons [symmetric]) + apply (rule compTpInitLvars_LT_ST_aux) + apply (auto dest: wf_java_mdecl_length_pTs_pns wf_java_mdecl_disjoint_varnames) + done (********************************************************************************) lemma max_of_list_elem: "x \ set xs \ x \ (max_of_list xs)" -by (induct xs, auto intro: max.cobounded1 simp: le_max_iff_disj max_of_list_def) + by (induct xs, auto intro: max.cobounded1 simp: le_max_iff_disj max_of_list_def) lemma max_of_list_sublist: "set xs \ set ys \ (max_of_list xs) \ (max_of_list ys)" -by (induct xs, auto dest: max_of_list_elem simp: max_of_list_def) + by (induct xs, auto dest: max_of_list_elem simp: max_of_list_def) lemma max_of_list_append [simp]: "max_of_list (xs @ ys) = max (max_of_list xs) (max_of_list ys)" -apply (simp add: max_of_list_def) -apply (induct xs) -apply simp -using [[linarith_split_limit = 0]] -apply simp -apply arith -done + apply (simp add: max_of_list_def) + apply (induct xs) + apply simp + using [[linarith_split_limit = 0]] + apply simp + apply arith + done lemma app_mono_mxs: "\ app i G mxs rT pc et s; mxs \ mxs' \ \ app i G mxs' rT pc et s" -apply (case_tac s) -apply (simp add: app_def) -apply (case_tac i) -apply (auto intro: less_trans) -done + apply (case_tac s) + apply (simp add: app_def) + apply (case_tac i, auto intro: less_trans) + done lemma err_mono [simp]: "A \ B \ err A \ err B" -by (auto simp: err_def) + by (auto simp: err_def) lemma opt_mono [simp]: "A \ B \ opt A \ opt B" -by (auto simp: opt_def) + by (auto simp: opt_def) lemma states_mono: "\ mxs \ mxs' \ \ states G mxs mxr \ states G mxs' mxr" -apply (simp add: states_def JVMType.sl_def) -apply (simp add: Product.esl_def stk_esl_def reg_sl_def - upto_esl_def Listn.sl_def Err.sl_def JType.esl_def) -apply (simp add: Err.esl_def Err.le_def Listn.le_def) -apply (simp add: Product.le_def Product.sup_def Err.sup_def) -apply (simp add: Opt.esl_def Listn.sup_def) -apply (rule err_mono) -apply (rule opt_mono) -apply (rule Sigma_mono) -apply (rule Union_mono) -apply auto -done - - -lemma check_type_mono: "\ check_type G mxs mxr s; mxs \ mxs' \ - \ check_type G mxs' mxr s" -apply (simp add: check_type_def) -apply (frule_tac G=G and mxr=mxr in states_mono) -apply auto -done - - - (* wt is preserved when adding instructions/state-types at the end *) + apply (simp add: states_def JVMType.sl_def) + apply (simp add: Product.esl_def stk_esl_def reg_sl_def + upto_esl_def Listn.sl_def Err.sl_def JType.esl_def) + apply (simp add: Err.esl_def Err.le_def Listn.le_def) + apply (simp add: Product.le_def Product.sup_def Err.sup_def) + apply (simp add: Opt.esl_def Listn.sup_def) + apply (rule err_mono) + apply (rule opt_mono) + apply (rule Sigma_mono) + apply (rule Union_mono) + apply auto + done + + +lemma check_type_mono: + "\ check_type G mxs mxr s; mxs \ mxs' \ \ check_type G mxs' mxr s" + apply (simp add: check_type_def) + apply (frule_tac G=G and mxr=mxr in states_mono) + apply auto + done + + +(* wt is preserved when adding instructions/state-types at the end *) lemma wt_instr_prefix: " - \ wt_instr_altern (bc ! pc) cG rT mt mxs mxr max_pc et pc; - bc' = bc @ bc_post; mt' = mt @ mt_post; - mxs \ mxs'; max_pc \ max_pc'; - pc < length bc; pc < length mt; - max_pc = (length mt)\ -\ wt_instr_altern (bc' ! pc) cG rT mt' mxs' mxr max_pc' et pc" -apply (simp add: wt_instr_altern_def nth_append) -apply (auto intro: app_mono_mxs check_type_mono) -done + \ wt_instr_altern (bc ! pc) cG rT mt mxs mxr max_pc et pc; + bc' = bc @ bc_post; mt' = mt @ mt_post; + mxs \ mxs'; max_pc \ max_pc'; + pc < length bc; pc < length mt; + max_pc = (length mt)\ + \ wt_instr_altern (bc' ! pc) cG rT mt' mxs' mxr max_pc' et pc" + apply (simp add: wt_instr_altern_def nth_append) + apply (auto intro: app_mono_mxs check_type_mono) + done (************************************************************************) -lemma pc_succs_shift: "pc'\set (succs i (pc'' + n)) - \ ((pc' - n) \set (succs i pc''))" -apply (induct i) -apply simp_all -apply arith -done - - -lemma pc_succs_le: "\ pc' \ set (succs i (pc'' + n)); - \ b. ((i = (Goto b) \ i=(Ifcmpeq b)) \ 0 \ (int pc'' + b)) \ +lemma pc_succs_shift: + "pc'\set (succs i (pc'' + n)) \ ((pc' - n) \set (succs i pc''))" + apply (induct i, simp_all) + apply arith + done + + +lemma pc_succs_le: + "\ pc' \ set (succs i (pc'' + n)); + \b. ((i = (Goto b) \ i=(Ifcmpeq b)) \ 0 \ (int pc'' + b)) \ \ n \ pc'" -apply (induct i) -apply simp_all -apply arith -done + apply (induct i, simp_all) + apply arith + done (**********************************************************************) @@ -547,37 +543,36 @@ lemma match_xcentry_offset [simp]: " match_exception_entry G cn (pc + n) (offset_xcentry n ee) = match_exception_entry G cn pc ee" -by (simp add: match_exception_entry_def offset_xcentry_def split_beta) + by (simp add: match_exception_entry_def offset_xcentry_def split_beta) lemma match_xctable_offset: " (match_exception_table G cn (pc + n) (offset_xctable n et)) = (map_option (\ pc'. pc' + n) (match_exception_table G cn pc et))" -apply (induct et) -apply (simp add: offset_xctable_def)+ -apply (case_tac "match_exception_entry G cn pc a") -apply (simp add: offset_xcentry_def split_beta)+ -done + apply (induct et) + apply (simp add: offset_xctable_def)+ + apply (case_tac "match_exception_entry G cn pc a") + apply (simp add: offset_xcentry_def split_beta)+ + done lemma match_offset [simp]: " match G cn (pc + n) (offset_xctable n et) = match G cn pc et" -apply (induct et) -apply (simp add: offset_xctable_def)+ -done + apply (induct et) + apply (simp add: offset_xctable_def)+ + done lemma match_any_offset [simp]: " match_any G (pc + n) (offset_xctable n et) = match_any G pc et" -apply (induct et) -apply (simp add: offset_xctable_def offset_xcentry_def split_beta)+ -done + apply (induct et) + apply (simp add: offset_xctable_def offset_xcentry_def split_beta)+ + done lemma app_mono_pc: "\ app i G mxs rT pc et s; pc'= pc + n \ \ app i G mxs rT pc' (offset_xctable n et) s" -apply (case_tac s) -apply (simp add: app_def) -apply (case_tac i) -apply (auto) -done + apply (case_tac s) + apply (simp add: app_def) + apply (case_tac i, auto) + done (**********************************************************************) @@ -591,18 +586,15 @@ (* move into Effect.thy *) lemma xcpt_names_Nil [simp]: "(xcpt_names (i, G, pc, [])) = []" -by (induct i, simp_all) + by (induct i, simp_all) lemma xcpt_eff_Nil [simp]: "(xcpt_eff i G pc s []) = []" -by (simp add: xcpt_eff_def) + by (simp add: xcpt_eff_def) lemma app_jumps_lem: "\ app i cG mxs rT pc empty_et s; s=(Some st) \ \ \ b. ((i = (Goto b) \ i=(Ifcmpeq b)) \ 0 \ (int pc + b))" -apply (simp only:) -apply (induct i) -apply auto -done + by (induct i) auto (* wt is preserved when adding instructions/state-types to the front *) @@ -614,67 +606,71 @@ length mt_pre \ pc; pc < length (mt_pre @ mt); mxs \ mxs'; max_pc + length mt_pre \ max_pc' \ \ wt_instr_altern (bc' ! pc) cG rT mt' mxs' mxr max_pc' empty_et pc" - -apply (simp add: wt_instr_altern_def) -apply (subgoal_tac "\ pc''. pc = pc'' + length mt_pre", erule exE) -prefer 2 apply (rule_tac x="pc - length mt_pre" in exI, arith) - -apply (drule_tac x=pc'' in spec) -apply (drule mp) apply arith (* show pc'' < length mt *) -apply clarify - -apply (rule conjI) - (* app *) - apply (simp add: nth_append) - apply (rule app_mono_mxs) - apply (frule app_mono_pc) apply (rule HOL.refl) apply (simp add: offset_xctable_def) - apply assumption+ + apply (simp add: wt_instr_altern_def) + apply (subgoal_tac "\ pc''. pc = pc'' + length mt_pre", erule exE) + prefer 2 + apply (rule_tac x="pc - length mt_pre" in exI, arith) + + apply (drule_tac x=pc'' in spec) + apply (drule mp) + apply arith (* show pc'' < length mt *) + apply clarify + + apply (rule conjI) + (* app *) + apply (simp add: nth_append) + apply (rule app_mono_mxs) + apply (frule app_mono_pc) + apply (rule HOL.refl) + apply (simp add: offset_xctable_def) + apply assumption+ (* check_type *) -apply (rule conjI) -apply (simp add: nth_append) -apply (rule check_type_mono) apply assumption+ + apply (rule conjI) + apply (simp add: nth_append) + apply (rule check_type_mono) + apply assumption+ (* ..eff.. *) -apply (intro ballI) -apply (subgoal_tac "\ pc' s'. x = (pc', s')", (erule exE)+, simp) - -apply (case_tac s') - (* s' = None *) -apply (simp add: eff_def nth_append norm_eff_def) -apply (frule_tac x="(pc', None)" and f=fst and b=pc' in rev_image_eqI) - apply (simp (no_asm_simp)) - apply (simp add: image_comp Fun.comp_def) - apply (frule pc_succs_shift) -apply (drule bspec, assumption) -apply arith - - (* s' = Some a *) -apply (drule_tac x="(pc' - length mt_pre, s')" in bspec) - - (* show (pc' - length mt_pre, s') \ set (eff \) *) -apply (simp add: eff_def) + apply (intro ballI) + apply (subgoal_tac "\ pc' s'. x = (pc', s')", (erule exE)+, simp) + + apply (case_tac s') + (* s' = None *) + apply (simp add: eff_def nth_append norm_eff_def) + apply (frule_tac x="(pc', None)" and f=fst and b=pc' in rev_image_eqI) + apply (simp (no_asm_simp)) + apply (simp add: image_comp Fun.comp_def) + apply (frule pc_succs_shift) + apply (drule bspec, assumption) + apply arith + + (* s' = Some a *) + apply (drule_tac x="(pc' - length mt_pre, s')" in bspec) + + (* show (pc' - length mt_pre, s') \ set (eff \) *) + apply (simp add: eff_def) (* norm_eff *) - apply (clarsimp simp: nth_append pc_succs_shift) - - (* show P x of bspec *) -apply simp - apply (subgoal_tac "length mt_pre \ pc'") - apply (simp add: nth_append) apply arith - - (* subgoals *) -apply (simp add: eff_def xcpt_eff_def) -apply (clarsimp) -apply (rule pc_succs_le) apply assumption+ -apply (subgoal_tac "\ st. mt ! pc'' = Some st", erule exE) - apply (rule_tac s="Some st" and st=st and cG=cG and mxs=mxs and rT=rT - in app_jumps_lem) - apply (simp add: nth_append)+ - (* subgoal \ st. mt ! pc'' = Some st *) - apply (simp add: norm_eff_def map_option_case nth_append) - apply (case_tac "mt ! pc''") -apply simp+ -done + apply (clarsimp simp: nth_append pc_succs_shift) + + (* show P x of bspec *) + apply simp + apply (subgoal_tac "length mt_pre \ pc'") + apply (simp add: nth_append) + apply arith + + (* subgoals *) + apply (simp add: eff_def xcpt_eff_def) + apply (clarsimp) + apply (rule pc_succs_le, assumption+) + apply (subgoal_tac "\ st. mt ! pc'' = Some st", erule exE) + apply (rule_tac s="Some st" and st=st and cG=cG and mxs=mxs and rT=rT in app_jumps_lem) + apply (simp add: nth_append)+ + (* subgoal \ st. mt ! pc'' = Some st *) + apply (simp add: norm_eff_def map_option_case nth_append) + apply (case_tac "mt ! pc''") + apply simp+ + done (**********************************************************************) @@ -688,100 +684,92 @@ "start_sttp_resp f == (f = comb_nil) \ (start_sttp_resp_cons f)" lemma start_sttp_resp_comb_nil [simp]: "start_sttp_resp comb_nil" -by (simp add: start_sttp_resp_def) + by (simp add: start_sttp_resp_def) lemma start_sttp_resp_cons_comb_cons [simp]: "start_sttp_resp_cons f \ start_sttp_resp_cons (f \ f')" -apply (simp add: start_sttp_resp_cons_def comb_def split_beta) -apply (rule allI) -apply (drule_tac x=sttp in spec) -apply auto -done + apply (simp add: start_sttp_resp_cons_def comb_def split_beta) + apply (rule allI) + apply (drule_tac x=sttp in spec) + apply auto + done lemma start_sttp_resp_cons_comb_cons_r: "\ start_sttp_resp f; start_sttp_resp_cons f'\ \ start_sttp_resp_cons (f \ f')" -apply (simp add: start_sttp_resp_def) -apply (erule disjE) -apply simp+ -done + by (auto simp: start_sttp_resp_def) lemma start_sttp_resp_cons_comb [simp]: "start_sttp_resp_cons f \ start_sttp_resp (f \ f')" -by (simp add: start_sttp_resp_def) + by (simp add: start_sttp_resp_def) lemma start_sttp_resp_comb: "\ start_sttp_resp f; start_sttp_resp f' \ \ start_sttp_resp (f \ f')" -apply (simp add: start_sttp_resp_def) -apply (erule disjE) -apply simp -apply (erule disjE) -apply simp+ -done + by (auto simp: start_sttp_resp_def) lemma start_sttp_resp_cons_nochangeST [simp]: "start_sttp_resp_cons nochangeST" -by (simp add: start_sttp_resp_cons_def nochangeST_def) + by (simp add: start_sttp_resp_cons_def nochangeST_def) lemma start_sttp_resp_cons_pushST [simp]: "start_sttp_resp_cons (pushST Ts)" -by (simp add: start_sttp_resp_cons_def pushST_def split_beta) + by (simp add: start_sttp_resp_cons_def pushST_def split_beta) lemma start_sttp_resp_cons_dupST [simp]: "start_sttp_resp_cons dupST" -by (simp add: start_sttp_resp_cons_def dupST_def split_beta) + by (simp add: start_sttp_resp_cons_def dupST_def split_beta) lemma start_sttp_resp_cons_dup_x1ST [simp]: "start_sttp_resp_cons dup_x1ST" -by (simp add: start_sttp_resp_cons_def dup_x1ST_def split_beta) + by (simp add: start_sttp_resp_cons_def dup_x1ST_def split_beta) lemma start_sttp_resp_cons_popST [simp]: "start_sttp_resp_cons (popST n)" -by (simp add: start_sttp_resp_cons_def popST_def split_beta) + by (simp add: start_sttp_resp_cons_def popST_def split_beta) lemma start_sttp_resp_cons_replST [simp]: "start_sttp_resp_cons (replST n tp)" -by (simp add: start_sttp_resp_cons_def replST_def split_beta) + by (simp add: start_sttp_resp_cons_def replST_def split_beta) lemma start_sttp_resp_cons_storeST [simp]: "start_sttp_resp_cons (storeST i tp)" -by (simp add: start_sttp_resp_cons_def storeST_def split_beta) + by (simp add: start_sttp_resp_cons_def storeST_def split_beta) lemma start_sttp_resp_cons_compTpExpr [simp]: "start_sttp_resp_cons (compTpExpr jmb G ex)" -apply (induct ex) -apply simp+ -apply (simp add: start_sttp_resp_cons_def comb_def pushST_def split_beta) (* LAcc *) -apply simp+ -done + apply (induct ex) + apply simp+ + apply (simp add: start_sttp_resp_cons_def comb_def pushST_def split_beta) (* LAcc *) + apply simp+ + done lemma start_sttp_resp_cons_compTpInit [simp]: "start_sttp_resp_cons (compTpInit jmb lv)" -by (simp add: compTpInit_def split_beta) + by (simp add: compTpInit_def split_beta) lemma start_sttp_resp_nochangeST [simp]: "start_sttp_resp nochangeST" -by (simp add: start_sttp_resp_def) + by (simp add: start_sttp_resp_def) lemma start_sttp_resp_pushST [simp]: "start_sttp_resp (pushST Ts)" -by (simp add: start_sttp_resp_def) + by (simp add: start_sttp_resp_def) lemma start_sttp_resp_dupST [simp]: "start_sttp_resp dupST" -by (simp add: start_sttp_resp_def) + by (simp add: start_sttp_resp_def) lemma start_sttp_resp_dup_x1ST [simp]: "start_sttp_resp dup_x1ST" -by (simp add: start_sttp_resp_def) + by (simp add: start_sttp_resp_def) lemma start_sttp_resp_popST [simp]: "start_sttp_resp (popST n)" -by (simp add: start_sttp_resp_def) + by (simp add: start_sttp_resp_def) lemma start_sttp_resp_replST [simp]: "start_sttp_resp (replST n tp)" -by (simp add: start_sttp_resp_def) + by (simp add: start_sttp_resp_def) lemma start_sttp_resp_storeST [simp]: "start_sttp_resp (storeST i tp)" -by (simp add: start_sttp_resp_def) + by (simp add: start_sttp_resp_def) lemma start_sttp_resp_compTpExpr [simp]: "start_sttp_resp (compTpExpr jmb G ex)" -by (simp add: start_sttp_resp_def) + by (simp add: start_sttp_resp_def) lemma start_sttp_resp_compTpExprs [simp]: "start_sttp_resp (compTpExprs jmb G exs)" -by (induct exs, (simp add: start_sttp_resp_comb)+) + by (induct exs, (simp add: start_sttp_resp_comb)+) lemma start_sttp_resp_compTpStmt [simp]: "start_sttp_resp (compTpStmt jmb G s)" -by (induct s, (simp add: start_sttp_resp_comb)+) + by (induct s, (simp add: start_sttp_resp_comb)+) lemma start_sttp_resp_compTpInitLvars [simp]: "start_sttp_resp (compTpInitLvars jmb lvars)" -by (induct lvars, simp+) + by (induct lvars, simp+) @@ -792,64 +780,59 @@ lemma length_comb [simp]: "length (mt_of ((f1 \ f2) sttp)) = length (mt_of (f1 sttp)) + length (mt_of (f2 (sttp_of (f1 sttp))))" -by (simp add: comb_def split_beta) + by (simp add: comb_def split_beta) lemma length_comb_nil [simp]: "length (mt_of (comb_nil sttp)) = 0" -by (simp add: comb_nil_def) + by (simp add: comb_nil_def) lemma length_nochangeST [simp]: "length (mt_of (nochangeST sttp)) = 1" -by (simp add: nochangeST_def) + by (simp add: nochangeST_def) lemma length_pushST [simp]: "length (mt_of (pushST Ts sttp)) = 1" -by (simp add: pushST_def split_beta) + by (simp add: pushST_def split_beta) lemma length_dupST [simp]: "length (mt_of (dupST sttp)) = 1" -by (simp add: dupST_def split_beta) + by (simp add: dupST_def split_beta) lemma length_dup_x1ST [simp]: "length (mt_of (dup_x1ST sttp)) = 1" -by (simp add: dup_x1ST_def split_beta) + by (simp add: dup_x1ST_def split_beta) lemma length_popST [simp]: "length (mt_of (popST n sttp)) = 1" -by (simp add: popST_def split_beta) + by (simp add: popST_def split_beta) lemma length_replST [simp]: "length (mt_of (replST n tp sttp)) = 1" -by (simp add: replST_def split_beta) + by (simp add: replST_def split_beta) lemma length_storeST [simp]: "length (mt_of (storeST i tp sttp)) = 1" -by (simp add: storeST_def split_beta) + by (simp add: storeST_def split_beta) lemma length_compTpExpr_Exprs [rule_format]: " - (\ sttp. (length (mt_of (compTpExpr jmb G ex sttp)) = length (compExpr jmb ex))) - \ (\ sttp. (length (mt_of (compTpExprs jmb G exs sttp)) = length (compExprs jmb exs)))" -apply (rule compat_expr_expr_list.induct) -apply simp+ -apply (rename_tac binop a b, case_tac binop) -apply simp+ -apply (simp add: pushST_def split_beta) -apply simp+ -done + (\sttp. (length (mt_of (compTpExpr jmb G ex sttp)) = length (compExpr jmb ex))) + \ (\sttp. (length (mt_of (compTpExprs jmb G exs sttp)) = length (compExprs jmb exs)))" + apply (rule compat_expr_expr_list.induct) + apply (simp_all)[3] + apply (rename_tac binop a b, case_tac binop) + apply (auto simp add: pushST_def split_beta) + done lemma length_compTpExpr: "length (mt_of (compTpExpr jmb G ex sttp)) = length (compExpr jmb ex)" -by (rule length_compTpExpr_Exprs [THEN conjunct1 [THEN spec]]) + by (rule length_compTpExpr_Exprs [THEN conjunct1 [THEN spec]]) lemma length_compTpExprs: "length (mt_of (compTpExprs jmb G exs sttp)) = length (compExprs jmb exs)" -by (rule length_compTpExpr_Exprs [THEN conjunct2 [THEN spec]]) + by (rule length_compTpExpr_Exprs [THEN conjunct2 [THEN spec]]) lemma length_compTpStmt [rule_format]: " (\ sttp. (length (mt_of (compTpStmt jmb G s sttp)) = length (compStmt jmb s)))" -apply (rule stmt.induct) -apply (simp add: length_compTpExpr)+ -done - + by (rule stmt.induct) (auto simp: length_compTpExpr) lemma length_compTpInit: "length (mt_of (compTpInit jmb lv sttp)) = length (compInit jmb lv)" -by (simp add: compTpInit_def compInit_def split_beta) + by (simp add: compTpInit_def compInit_def split_beta) lemma length_compTpInitLvars [rule_format]: "\ sttp. length (mt_of (compTpInitLvars jmb lvars sttp)) = length (compInitLvars jmb lvars)" -by (induct lvars, (simp add: compInitLvars_def length_compTpInit split_beta)+) + by (induct lvars, (simp add: compInitLvars_def length_compTpInit split_beta)+) (* ********************************************************************** *) @@ -867,20 +850,20 @@ lemma states_lower: "\ OK (Some (ST, LT)) \ states cG mxs mxr; length ST \ mxs\ \ OK (Some (ST, LT)) \ states cG (length ST) mxr" -apply (simp add: states_def JVMType.sl_def) -apply (simp add: Product.esl_def stk_esl_def reg_sl_def upto_esl_def Listn.sl_def Err.sl_def - JType.esl_def) -apply (simp add: Err.esl_def Err.le_def Listn.le_def) -apply (simp add: Product.le_def Product.sup_def Err.sup_def) -apply (simp add: Opt.esl_def Listn.sup_def) -apply clarify -apply auto -done + apply (simp add: states_def JVMType.sl_def) + apply (simp add: Product.esl_def stk_esl_def reg_sl_def upto_esl_def Listn.sl_def Err.sl_def + JType.esl_def) + apply (simp add: Err.esl_def Err.le_def Listn.le_def) + apply (simp add: Product.le_def Product.sup_def Err.sup_def) + apply (simp add: Opt.esl_def Listn.sup_def) + apply clarify + apply auto + done lemma check_type_lower: "\ check_type cG mxs mxr (OK (Some (ST, LT))); length ST \ mxs\ \check_type cG (length ST) mxr (OK (Some (ST, LT)))" -by (simp add: check_type_def states_lower) + by (simp add: check_type_def states_lower) (* ******************************************************************* *) @@ -906,85 +889,86 @@ bc_mt_corresp bc2 f2 (sttp_of (f1 sttp0)) cG rT mxr (length bc2); start_sttp_resp f2\ \ bc_mt_corresp bc' (f1 \ f2) sttp0 cG rT mxr l'" -apply (subgoal_tac "\ mt1 sttp1. (f1 sttp0) = (mt1, sttp1)", (erule exE)+) -apply (subgoal_tac "\ mt2 sttp2. (f2 sttp1) = (mt2, sttp2)", (erule exE)+) - - (* unfold start_sttp_resp and make case distinction *) -apply (simp only: start_sttp_resp_def) -apply (erule disjE) - (* case f2 = comb_nil *) -apply (simp add: bc_mt_corresp_def comb_nil_def start_sttp_resp_cons_def) -apply (erule conjE)+ -apply (intro strip) -apply simp - - (* case start_sttp_resp_cons f2 *) -apply (simp add: bc_mt_corresp_def comb_def start_sttp_resp_cons_def del: all_simps) -apply (intro strip) -apply (erule conjE)+ -apply (drule mp, assumption) -apply (subgoal_tac "check_type cG (length (fst sttp1)) mxr (OK (Some sttp1))") -apply (erule conjE)+ -apply (drule mp, assumption) -apply (erule conjE)+ - -apply (rule conjI) - (* show wt_instr \ *) - -apply (drule_tac x=sttp1 in spec, simp) -apply (erule exE) -apply (intro strip) -apply (case_tac "pc < length mt1") - - (* case pc < length mt1 *) -apply (drule spec, drule mp, simp) -apply simp -apply (rule_tac mt="mt1 @ [Some sttp1]" in wt_instr_prefix) - apply assumption+ apply (rule HOL.refl) - apply (simp (no_asm_simp)) - apply (simp (no_asm_simp) add: max_ssize_def) - apply (simp add: max_of_list_def ac_simps) - apply arith - apply (simp (no_asm_simp))+ - - (* case pc \ length mt1 *) -apply (rule_tac bc=bc2 and mt=mt2 and bc_post="[]" and mt_post="[Some sttp2]" - and mxr=mxr - in wt_instr_offset) -apply simp -apply (simp (no_asm_simp))+ -apply simp -apply (simp add: max_ssize_def) apply (simp (no_asm_simp)) - - (* show check_type \ *) -apply (subgoal_tac "((mt2 @ [Some sttp2]) ! length bc2) = Some sttp2") -apply (simp only:) -apply (rule check_type_mono) apply assumption -apply (simp (no_asm_simp) add: max_ssize_def ac_simps) -apply (simp add: nth_append) - -apply (erule conjE)+ -apply (case_tac sttp1) -apply (simp add: check_type_def) -apply (rule states_lower, assumption) -apply (simp (no_asm_simp) add: max_ssize_def) -apply (simp (no_asm_simp) add: max_of_list_def ssize_sto_def) -apply (simp (no_asm_simp))+ -done - - -lemma bc_mt_corresp_zero [simp]: "\ length (mt_of (f sttp)) = length bc; start_sttp_resp f\ + apply (subgoal_tac "\mt1 sttp1. (f1 sttp0) = (mt1, sttp1)", (erule exE)+) + apply (subgoal_tac "\mt2 sttp2. (f2 sttp1) = (mt2, sttp2)", (erule exE)+) + + (* unfold start_sttp_resp and make case distinction *) + apply (simp only: start_sttp_resp_def) + apply (erule disjE) + (* case f2 = comb_nil *) + apply (simp add: bc_mt_corresp_def comb_nil_def start_sttp_resp_cons_def) + apply (erule conjE)+ + apply (intro strip) + apply simp + + (* case start_sttp_resp_cons f2 *) + apply (simp add: bc_mt_corresp_def comb_def start_sttp_resp_cons_def del: all_simps) + apply (intro strip) + apply (erule conjE)+ + apply (drule mp, assumption) + apply (subgoal_tac "check_type cG (length (fst sttp1)) mxr (OK (Some sttp1))") + apply (erule conjE)+ + apply (drule mp, assumption) + apply (erule conjE)+ + + apply (rule conjI) + (* show wt_instr \ *) + + apply (drule_tac x=sttp1 in spec, simp) + apply (erule exE) + apply (intro strip) + apply (case_tac "pc < length mt1") + + (* case pc < length mt1 *) + apply (drule spec, drule mp, simp) + apply simp + apply (rule_tac mt="mt1 @ [Some sttp1]" in wt_instr_prefix) + apply assumption+ apply (rule HOL.refl) + apply (simp (no_asm_simp)) + apply (simp (no_asm_simp) add: max_ssize_def) + apply (simp add: max_of_list_def ac_simps) + apply arith + apply (simp (no_asm_simp))+ + + (* case pc \ length mt1 *) + apply (rule_tac bc=bc2 and mt=mt2 and bc_post="[]" and mt_post="[Some sttp2]" + and mxr=mxr + in wt_instr_offset) + apply simp + apply (simp (no_asm_simp))+ + apply simp + apply (simp add: max_ssize_def) apply (simp (no_asm_simp)) + + (* show check_type \ *) + apply (subgoal_tac "((mt2 @ [Some sttp2]) ! length bc2) = Some sttp2") + apply (simp only:) + apply (rule check_type_mono) apply assumption + apply (simp (no_asm_simp) add: max_ssize_def ac_simps) + apply (simp add: nth_append) + + apply (erule conjE)+ + apply (case_tac sttp1) + apply (simp add: check_type_def) + apply (rule states_lower, assumption) + apply (simp (no_asm_simp) add: max_ssize_def) + apply (simp (no_asm_simp) add: max_of_list_def ssize_sto_def) + apply (simp (no_asm_simp))+ + done + + +lemma bc_mt_corresp_zero [simp]: + "\ length (mt_of (f sttp)) = length bc; start_sttp_resp f\ \ bc_mt_corresp bc f sttp cG rT mxr 0" -apply (simp add: bc_mt_corresp_def start_sttp_resp_def split_beta) -apply (erule disjE) -apply (simp add: max_ssize_def max_of_list_def ssize_sto_def split: prod.split) -apply (intro strip) -apply (simp add: start_sttp_resp_cons_def split_beta) -apply (drule_tac x=sttp in spec, erule exE) -apply simp -apply (rule check_type_mono, assumption) -apply (simp add: max_ssize_def max_of_list_def ssize_sto_def split: prod.split) -done + apply (simp add: bc_mt_corresp_def start_sttp_resp_def split_beta) + apply (erule disjE) + apply (simp add: max_ssize_def max_of_list_def ssize_sto_def split: prod.split) + apply (intro strip) + apply (simp add: start_sttp_resp_cons_def split_beta) + apply (drule_tac x=sttp in spec, erule exE) + apply simp + apply (rule check_type_mono, assumption) + apply (simp add: max_ssize_def max_of_list_def ssize_sto_def split: prod.split) + done (* ********************************************************************** *) @@ -994,35 +978,35 @@ lemma mt_sttp_flatten_length [simp]: "n = (length (mt_of (f sttp))) - \ (mt_sttp_flatten (f sttp)) ! n = Some (sttp_of (f sttp))" -by (simp add: mt_sttp_flatten_def) + \ (mt_sttp_flatten (f sttp)) ! n = Some (sttp_of (f sttp))" + by (simp add: mt_sttp_flatten_def) lemma mt_sttp_flatten_comb: "(mt_sttp_flatten ((f1 \ f2) sttp)) = (mt_of (f1 sttp)) @ (mt_sttp_flatten (f2 (sttp_of (f1 sttp))))" -by (simp add: mt_sttp_flatten_def comb_def split_beta) + by (simp add: mt_sttp_flatten_def comb_def split_beta) lemma mt_sttp_flatten_comb_length [simp]: "\ n1 = length (mt_of (f1 sttp)); n1 \ n \ \ (mt_sttp_flatten ((f1 \ f2) sttp) ! n) = (mt_sttp_flatten (f2 (sttp_of (f1 sttp))) ! (n - n1))" -by (simp add: mt_sttp_flatten_comb nth_append) - -lemma mt_sttp_flatten_comb_zero [simp]: "start_sttp_resp f - \ (mt_sttp_flatten (f sttp)) ! 0 = Some sttp" -apply (simp only: start_sttp_resp_def) -apply (erule disjE) -apply (simp add: comb_nil_def mt_sttp_flatten_def) -apply (simp add: start_sttp_resp_cons_def mt_sttp_flatten_def split_beta) -apply (drule_tac x=sttp in spec) -apply (erule exE) -apply simp -done + by (simp add: mt_sttp_flatten_comb nth_append) + +lemma mt_sttp_flatten_comb_zero [simp]: + "start_sttp_resp f \ (mt_sttp_flatten (f sttp)) ! 0 = Some sttp" + apply (simp only: start_sttp_resp_def) + apply (erule disjE) + apply (simp add: comb_nil_def mt_sttp_flatten_def) + apply (simp add: start_sttp_resp_cons_def mt_sttp_flatten_def split_beta) + apply (drule_tac x=sttp in spec) + apply (erule exE) + apply simp + done (* move into prelude -- compare with nat_int_length *) lemma int_outside_right: "0 \ (m::int) \ m + (int n) = int ((nat m) + n)" -by simp + by simp lemma int_outside_left: "0 \ (m::int) \ (int n) + m = int (n + (nat m))" -by simp + by simp @@ -1040,33 +1024,33 @@ Opt.esl_def Listn.sup_def -lemma check_type_push: "\ - is_class cG cname; check_type cG (length ST) mxr (OK (Some (ST, LT))) \ +lemma check_type_push: + "\ is_class cG cname; check_type cG (length ST) mxr (OK (Some (ST, LT))) \ \ check_type cG (Suc (length ST)) mxr (OK (Some (Class cname # ST, LT)))" -apply (simp add: check_type_simps) -apply clarify -apply (rule_tac x="Suc (length ST)" in exI) -apply simp+ -done + apply (simp add: check_type_simps) + apply clarify + apply (rule_tac x="Suc (length ST)" in exI) + apply simp+ + done lemma bc_mt_corresp_New: "\is_class cG cname \ \ bc_mt_corresp [New cname] (pushST [Class cname]) (ST, LT) cG rT mxr (Suc 0)" -apply (simp add: bc_mt_corresp_def pushST_def wt_instr_altern_def - max_ssize_def max_of_list_def ssize_sto_def eff_def norm_eff_def max.absorb2) -apply (intro strip) -apply (rule conjI) -apply (rule check_type_mono, assumption, simp) -apply (simp add: check_type_push) -done + apply (simp add: bc_mt_corresp_def pushST_def wt_instr_altern_def + max_ssize_def max_of_list_def ssize_sto_def eff_def norm_eff_def max.absorb2) + apply (intro strip) + apply (rule conjI) + apply (rule check_type_mono, assumption, simp) + apply (simp add: check_type_push) + done lemma bc_mt_corresp_Pop: " bc_mt_corresp [Pop] (popST (Suc 0)) (T # ST, LT) cG rT mxr (Suc 0)" apply (simp add: bc_mt_corresp_def popST_def wt_instr_altern_def eff_def norm_eff_def) - apply (simp add: max_ssize_def ssize_sto_def max_of_list_def) + apply (simp add: max_ssize_def ssize_sto_def max_of_list_def) apply (simp add: check_type_simps max.absorb1) apply clarify apply (rule_tac x="(length ST)" in exI) - apply simp+ + apply simp done lemma bc_mt_corresp_Checkcast: "\ is_class cG cname; sttp = (ST, LT); @@ -1078,108 +1062,111 @@ apply (simp add: check_type_simps) apply clarify apply (rule_tac x="Suc (length STo)" in exI) - apply simp+ + apply simp done lemma bc_mt_corresp_LitPush: "\ typeof (\v. None) val = Some T \ \ bc_mt_corresp [LitPush val] (pushST [T]) sttp cG rT mxr (Suc 0)" -apply (subgoal_tac "\ ST LT. sttp= (ST, LT)", (erule exE)+) - apply (simp add: bc_mt_corresp_def pushST_def wt_instr_altern_def - max_ssize_def max_of_list_def ssize_sto_def eff_def norm_eff_def max.absorb2) - apply (intro strip) - apply (rule conjI) - apply (rule check_type_mono, assumption, simp) - apply (simp add: check_type_simps) -apply clarify -apply (rule_tac x="Suc (length ST)" in exI) -apply simp -apply (drule sym) -apply (case_tac val) -apply simp+ -done - - -lemma bc_mt_corresp_LitPush_CT: "\ typeof (\v. None) val = Some T \ cG \ T \ T'; - is_type cG T' \ + apply (subgoal_tac "\ST LT. sttp= (ST, LT)", (erule exE)+) + apply (simp add: bc_mt_corresp_def pushST_def wt_instr_altern_def + max_ssize_def max_of_list_def ssize_sto_def eff_def norm_eff_def max.absorb2) + apply (intro strip) + apply (rule conjI) + apply (rule check_type_mono, assumption, simp) + apply (simp add: check_type_simps) + apply clarify + apply (rule_tac x="Suc (length ST)" in exI) + apply simp + apply (drule sym) + apply (case_tac val) + apply simp+ + done + + +lemma bc_mt_corresp_LitPush_CT: + "\ typeof (\v. None) val = Some T \ cG \ T \ T'; is_type cG T' \ \ bc_mt_corresp [LitPush val] (pushST [T']) sttp cG rT mxr (Suc 0)" -apply (subgoal_tac "\ ST LT. sttp= (ST, LT)", (erule exE)+) - apply (simp add: bc_mt_corresp_def pushST_def wt_instr_altern_def - max_ssize_def max_of_list_def ssize_sto_def eff_def norm_eff_def max.absorb2) - apply (intro strip) - apply (rule conjI) - apply (rule check_type_mono, assumption, simp) - apply (simp add: check_type_simps) - apply (simp add: sup_state_Cons) -apply clarify -apply (rule_tac x="Suc (length ST)" in exI) -apply simp -apply simp+ -done + apply (subgoal_tac "\ST LT. sttp= (ST, LT)", (erule exE)+) + apply (simp add: bc_mt_corresp_def pushST_def wt_instr_altern_def max_ssize_def + max_of_list_def ssize_sto_def eff_def norm_eff_def max.absorb2) + apply (intro strip) + apply (rule conjI) + apply (rule check_type_mono, assumption, simp) + apply (simp add: check_type_simps) + apply (simp add: sup_state_Cons) + apply clarify + apply (rule_tac x="Suc (length ST)" in exI) + apply simp + apply simp + done declare not_Err_eq [iff del] lemma bc_mt_corresp_Load: "\ i < length LT; LT ! i \ Err; mxr = length LT \ \ bc_mt_corresp [Load i] (\(ST, LT). pushST [ok_val (LT ! i)] (ST, LT)) (ST, LT) cG rT mxr (Suc 0)" -apply (simp add: bc_mt_corresp_def pushST_def wt_instr_altern_def - max_ssize_def max_of_list_def ssize_sto_def eff_def norm_eff_def max.absorb2) + apply (simp add: bc_mt_corresp_def pushST_def wt_instr_altern_def max_ssize_def max_of_list_def + ssize_sto_def eff_def norm_eff_def max.absorb2) apply (intro strip) apply (rule conjI) - apply (rule check_type_mono, assumption, simp) -apply (simp add: check_type_simps) -apply clarify -apply (rule_tac x="Suc (length ST)" in exI) -apply (simp (no_asm_simp)) + apply (rule check_type_mono, assumption, simp) + apply (simp add: check_type_simps) + apply clarify + apply (rule_tac x="Suc (length ST)" in exI) + apply (simp (no_asm_simp)) apply (simp only: err_def) - apply (frule listE_nth_in) apply assumption -apply (subgoal_tac "LT ! i \ {x. \y\types cG. x = OK y}") -apply (drule CollectD) apply (erule bexE) -apply (simp (no_asm_simp)) -apply blast -apply blast -done - - -lemma bc_mt_corresp_Store_init: "\ i < length LT \ - \ bc_mt_corresp [Store i] (storeST i T) (T # ST, LT) cG rT mxr (Suc 0)" - apply (simp add: bc_mt_corresp_def storeST_def wt_instr_altern_def eff_def norm_eff_def) + apply (frule listE_nth_in) + apply assumption + apply (subgoal_tac "LT ! i \ {x. \y\types cG. x = OK y}") + apply (drule CollectD) + apply (erule bexE) + apply (simp (no_asm_simp)) + apply blast + apply blast + done + + +lemma bc_mt_corresp_Store_init: + "i < length LT \ bc_mt_corresp [Store i] (storeST i T) (T # ST, LT) cG rT mxr (Suc 0)" + apply (simp add: bc_mt_corresp_def storeST_def wt_instr_altern_def eff_def norm_eff_def) apply (simp add: max_ssize_def max_of_list_def) apply (simp add: ssize_sto_def) apply (intro strip) -apply (simp add: check_type_simps max.absorb1) -apply clarify -apply (rule conjI) -apply (rule_tac x="(length ST)" in exI) -apply simp+ -done - - -lemma bc_mt_corresp_Store: "\ i < length LT; cG \ LT[i := OK T] <=l LT \ + apply (simp add: check_type_simps max.absorb1) + apply clarify + apply (rule conjI) + apply (rule_tac x="(length ST)" in exI) + apply simp+ + done + + +lemma bc_mt_corresp_Store: + "\ i < length LT; cG \ LT[i := OK T] <=l LT \ \ bc_mt_corresp [Store i] (popST (Suc 0)) (T # ST, LT) cG rT mxr (Suc 0)" apply (simp add: bc_mt_corresp_def popST_def wt_instr_altern_def eff_def norm_eff_def) apply (simp add: sup_state_conv) apply (simp add: max_ssize_def max_of_list_def ssize_sto_def) - apply (intro strip) -apply (simp add: check_type_simps max.absorb1) -apply clarify -apply (rule_tac x="(length ST)" in exI) -apply simp+ -done + apply (intro strip) + apply (simp add: check_type_simps max.absorb1) + apply clarify + apply (rule_tac x="(length ST)" in exI) + apply simp + done lemma bc_mt_corresp_Dup: " bc_mt_corresp [Dup] dupST (T # ST, LT) cG rT mxr (Suc 0)" - apply (simp add: bc_mt_corresp_def dupST_def wt_instr_altern_def - max_ssize_def max_of_list_def ssize_sto_def eff_def norm_eff_def max.absorb2) + apply (simp add: bc_mt_corresp_def dupST_def wt_instr_altern_def + max_ssize_def max_of_list_def ssize_sto_def eff_def norm_eff_def max.absorb2) apply (intro strip) apply (rule conjI) - apply (rule check_type_mono, assumption, simp) -apply (simp add: check_type_simps) -apply clarify -apply (rule_tac x="Suc (Suc (length ST))" in exI) -apply simp+ -done + apply (rule check_type_mono, assumption, simp) + apply (simp add: check_type_simps) + apply clarify + apply (rule_tac x="Suc (Suc (length ST))" in exI) + apply simp + done lemma bc_mt_corresp_Dup_x1: " bc_mt_corresp [Dup_x1] dup_x1ST (T1 # T2 # ST, LT) cG rT mxr (Suc 0)" @@ -1187,12 +1174,12 @@ max_ssize_def max_of_list_def ssize_sto_def eff_def norm_eff_def max.absorb2) apply (intro strip) apply (rule conjI) - apply (rule check_type_mono, assumption, simp) -apply (simp add: check_type_simps) -apply clarify -apply (rule_tac x="Suc (Suc (Suc (length ST)))" in exI) -apply simp+ -done + apply (rule check_type_mono, assumption, simp) + apply (simp add: check_type_simps) + apply clarify + apply (rule_tac x="Suc (Suc (Suc (length ST)))" in exI) + apply simp+ + done @@ -1204,7 +1191,7 @@ apply (simp add: check_type_simps max.absorb1) apply clarify apply (rule_tac x="Suc (length ST)" in exI) - apply simp+ + apply simp done lemma bc_mt_corresp_Getfield: "\ wf_prog wf_mb G; @@ -1219,18 +1206,18 @@ apply (simp add: comp_field comp_subcls1 comp_widen comp_is_class) apply (simp add: max_ssize_def max_of_list_def ssize_sto_def) apply (intro strip) -apply (simp add: check_type_simps) -apply clarify -apply (rule_tac x="Suc (length ST)" in exI) -apply simp+ -apply (simp only: comp_is_type) -apply (rule_tac C=cname in fields_is_type) -apply (simp add: TypeRel.field_def) -apply (drule JBasis.table_of_remap_SomeD)+ -apply assumption+ -apply (erule wf_prog_ws_prog) -apply assumption -done + apply (simp add: check_type_simps) + apply clarify + apply (rule_tac x="Suc (length ST)" in exI) + apply simp+ + apply (simp only: comp_is_type) + apply (rule_tac C=cname in fields_is_type) + apply (simp add: TypeRel.field_def) + apply (drule JBasis.table_of_remap_SomeD)+ + apply assumption+ + apply (erule wf_prog_ws_prog) + apply assumption + done lemma bc_mt_corresp_Putfield: "\ wf_prog wf_mb G; field (G, C) vname = Some (cname, Ta); G \ T \ Ta; is_class G C \ @@ -1244,33 +1231,35 @@ apply (simp add: max_ssize_def max_of_list_def ssize_sto_def) apply (intro strip) -apply (simp add: check_type_simps max.absorb1) -apply clarify -apply (rule_tac x="Suc (length ST)" in exI) -apply simp+ -done - - - -lemma Call_app: "\ wf_prog wf_mb G; is_class G cname; -STs = rev pTsa @ Class cname # ST; -max_spec G cname (mname, pTsa) = {((md, T), pTs')} \ + apply (simp add: check_type_simps max.absorb1) + apply clarify + apply (rule_tac x="Suc (length ST)" in exI) + apply simp+ + done + + + +lemma Call_app: + "\ wf_prog wf_mb G; is_class G cname; + STs = rev pTsa @ Class cname # ST; + max_spec G cname (mname, pTsa) = {((md, T), pTs')} \ \ app (Invoke cname mname pTs') (comp G) (length (T # ST)) rT 0 empty_et (Some (STs, LTs))" apply (subgoal_tac "(\mD' rT' comp_b. - method (comp G, cname) (mname, pTs') = Some (mD', rT', comp_b))") - apply (simp add: comp_is_class) - apply (rule_tac x=pTsa in exI) - apply (rule_tac x="Class cname" in exI) - apply (simp add: max_spec_preserves_length comp_is_class) - apply (frule max_spec2mheads, (erule exE)+, (erule conjE)+) - apply (simp add: split_paired_all comp_widen list_all2_iff) + method (comp G, cname) (mname, pTs') = Some (mD', rT', comp_b))") + apply (simp add: comp_is_class) + apply (rule_tac x=pTsa in exI) + apply (rule_tac x="Class cname" in exI) + apply (simp add: max_spec_preserves_length comp_is_class) + apply (frule max_spec2mheads, (erule exE)+, (erule conjE)+) + apply (simp add: split_paired_all comp_widen list_all2_iff) apply (frule max_spec2mheads, (erule exE)+, (erule conjE)+) apply (rule exI)+ apply (simp add: wf_prog_ws_prog [THEN comp_method]) done -lemma bc_mt_corresp_Invoke: "\ wf_prog wf_mb G; +lemma bc_mt_corresp_Invoke: + "\ wf_prog wf_mb G; max_spec G cname (mname, pTsa) = {((md, T), fpTs)}; is_class G cname \ \ bc_mt_corresp [Invoke cname mname fpTs] (replST (Suc (length pTsa)) T) @@ -1280,10 +1269,12 @@ apply (intro strip) apply (rule conjI) - -- "app" - apply (rule Call_app [THEN app_mono_mxs]) apply assumption+ - apply (rule HOL.refl) apply assumption - apply (simp add: max_ssize_def max_of_list_elem ssize_sto_def) + -- "app" + apply (rule Call_app [THEN app_mono_mxs]) + apply assumption+ + apply (rule HOL.refl) + apply assumption + apply (simp add: max_ssize_def max_of_list_elem ssize_sto_def) -- {* @{text "<=s"} *} apply (frule max_spec2mheads, (erule exE)+, (erule conjE)+) @@ -1291,18 +1282,18 @@ apply (simp add: max_spec_preserves_length [symmetric]) -- "@{text check_type}" -apply (simp add: max_ssize_def ssize_sto_def) -apply (simp add: max_of_list_def) -apply (subgoal_tac "(max (length pTsa + length ST) (length ST)) = (length pTsa + length ST)") -apply simp -apply (simp add: check_type_simps) -apply clarify -apply (rule_tac x="Suc (length ST)" in exI) -apply simp+ -apply (simp only: comp_is_type) -apply (frule method_wf_mdecl) apply assumption apply assumption -apply (simp add: wf_mdecl_def wf_mhead_def) -apply (simp) + apply (simp add: max_ssize_def ssize_sto_def) + apply (simp add: max_of_list_def) + apply (subgoal_tac "(max (length pTsa + length ST) (length ST)) = (length pTsa + length ST)") + apply simp + apply (simp add: check_type_simps) + apply clarify + apply (rule_tac x="Suc (length ST)" in exI) + apply simp+ + apply (simp only: comp_is_type) + apply (frule method_wf_mdecl) apply assumption apply assumption + apply (simp add: wf_mdecl_def wf_mhead_def) + apply (simp) done @@ -1314,16 +1305,16 @@ mt_sttp_flatten f ! nat (int pc + i) = Some (ST,LT); check_type (TranslComp.comp G) mxs mxr (OK (Some (ts # ts' # ST, LT))) \ \ wt_instr_altern (Ifcmpeq i) (comp G) rT (mt_sttp_flatten f) mxs mxr max_pc empty_et pc" -by (simp add: wt_instr_altern_def eff_def norm_eff_def) + by (simp add: wt_instr_altern_def eff_def norm_eff_def) lemma wt_instr_Goto: "\ 0 \ (int pc + i); nat (int pc + i) < max_pc; mt_sttp_flatten f ! nat (int pc + i) = (mt_sttp_flatten f ! pc); check_type (TranslComp.comp G) mxs mxr (OK (mt_sttp_flatten f ! pc)) \ \ wt_instr_altern (Goto i) (comp G) rT (mt_sttp_flatten f) mxs mxr max_pc empty_et pc" -apply (case_tac "(mt_sttp_flatten f ! pc)") -apply (simp add: wt_instr_altern_def eff_def norm_eff_def app_def xcpt_app_def)+ -done + apply (case_tac "(mt_sttp_flatten f ! pc)") + apply (simp add: wt_instr_altern_def eff_def norm_eff_def app_def xcpt_app_def)+ + done @@ -1341,129 +1332,133 @@ length bc1 = length (mt_of (f1 sttp0)); start_sttp_resp f2; start_sttp_resp f3\ \ bc_mt_corresp bc' f' sttp0 cG rT mxr l12" -apply (subgoal_tac "\ mt1 sttp1. (f1 sttp0) = (mt1, sttp1)", (erule exE)+) -apply (subgoal_tac "\ mt2 sttp2. (f2 sttp1) = (mt2, sttp2)", (erule exE)+) -apply (subgoal_tac "\ mt3 sttp3. (f3 sttp2) = (mt3, sttp3)", (erule exE)+) - - (* unfold start_sttp_resp and make case distinction *) -apply (simp only: start_sttp_resp_def) -apply (erule_tac Q="start_sttp_resp_cons f2" in disjE) - (* case f2 = comb_nil *) -apply (simp add: bc_mt_corresp_def comb_nil_def start_sttp_resp_cons_def) - - (* case start_sttp_resp_cons f2 *) -apply (simp add: bc_mt_corresp_def comb_def start_sttp_resp_cons_def) -apply (drule_tac x=sttp1 in spec, simp, erule exE) -apply (intro strip, (erule conjE)+) - - - (* get rid of all check_type info *) -apply (subgoal_tac "check_type cG (length (fst sttp1)) mxr (OK (Some sttp1))") -apply (subgoal_tac "check_type cG (max_ssize (mt2 @ [Some sttp2])) mxr (OK (Some sttp2))") -apply (subgoal_tac "check_type cG (max_ssize (mt1 @ mt2 @ mt3 @ [Some sttp3])) mxr - (OK ((mt2 @ mt3 @ [Some sttp3]) ! length mt2))") -apply simp - - - -apply (intro strip, (erule conjE)+) -apply (case_tac "pc < length mt1") - - (* case pc < length mt1 *) -apply (drule spec, drule mp, assumption) -apply assumption - - (* case pc \ length mt1 *) - (* case distinction on start_sttp_resp f3 *) -apply (erule_tac P="f3 = comb_nil" in disjE) - - (* case f3 = comb_nil *) -apply (subgoal_tac "mt3 = [] \ sttp2 = sttp3") apply (erule conjE)+ -apply (subgoal_tac "bc3=[]") - -apply (rule_tac bc_pre=bc1 and bc=bc2 and bc_post=bc3 - and mt_pre=mt1 and mt=mt2 and mt_post="mt3@ [Some sttp3]" - and mxs="(max_ssize (mt2 @ [(Some sttp2)]))" - and max_pc="(Suc (length mt2))" - in wt_instr_offset) - apply simp - apply (rule HOL.refl)+ - apply (simp (no_asm_simp))+ - - apply (simp (no_asm_simp) add: max_ssize_def del: max_of_list_append) - apply (rule max_of_list_sublist) - apply (simp (no_asm_simp) only: set_append list.set list.map) apply blast - apply (simp (no_asm_simp)) - apply simp (* subgoal bc3 = [] *) - apply (simp add: comb_nil_def) (* subgoal mt3 = [] \ sttp2 = sttp3 *) - - (* case start_sttp_resp_cons f3 *) -apply (subgoal_tac "\mt3_rest. (mt3 = Some sttp2 # mt3_rest)", erule exE) -apply (rule_tac bc_pre=bc1 and bc=bc2 and bc_post=bc3 - and mt_pre=mt1 and mt=mt2 and mt_post="mt3@ [Some sttp3]" - and mxs="(max_ssize (mt2 @ [Some sttp2]))" - and max_pc="(Suc (length mt2))" - in wt_instr_offset) -apply (intro strip) -apply (rule_tac bc=bc2 and mt="(mt2 @ [Some sttp2])" - and mxs="(max_ssize (mt2 @ [Some sttp2]))" - and max_pc="(Suc (length mt2))" - in wt_instr_prefix) - - - (* preconditions of wt_instr_prefix *) - apply simp - apply (rule HOL.refl) - apply (simp (no_asm_simp))+ - apply simp+ - (* (some) preconditions of wt_instr_offset *) - apply (simp (no_asm_simp) add: max_ssize_def del: max_of_list_append) - apply (rule max_of_list_sublist) - apply (simp (no_asm_simp) only: set_append list.set list.map) apply blast - apply (simp (no_asm_simp)) - -apply (drule_tac x=sttp2 in spec, simp) (* subgoal \mt3_rest. \ *) - - (* subgoals check_type*) - (* \ ! length mt2 *) -apply simp - -apply (erule_tac P="f3 = comb_nil" in disjE) - - (* -- case f3 = comb_nil *) -apply (subgoal_tac "mt3 = [] \ sttp2 = sttp3") apply (erule conjE)+ -apply simp -apply (rule check_type_mono, assumption) -apply (simp only: max_ssize_def) apply (rule max_of_list_sublist) apply (simp (no_asm_simp)) -apply blast - apply simp (* subgoal bc3 = [] *) - apply (simp add: comb_nil_def) (* subgoal mt3 = [] \ sttp2 = sttp3 *) - - - (* -- case start_sttp_resp_cons f3 *) -apply (subgoal_tac "\mt3_rest. (mt3 = Some sttp2 # mt3_rest)", erule exE) -apply (simp (no_asm_simp) add: nth_append) -apply (erule conjE)+ -apply (rule check_type_mono, assumption) -apply (simp only: max_ssize_def) apply (rule max_of_list_sublist) apply (simp (no_asm_simp)) -apply blast -apply (drule_tac x=sttp2 in spec, simp) (* subgoal \mt3_rest. \ *) - - - (* subgoal check_type \ Some sttp2 *) -apply (simp add: nth_append) - - (* subgoal check_type \ Some sttp1 *) -apply (simp add: nth_append) -apply (erule conjE)+ -apply (case_tac "sttp1", simp) -apply (rule check_type_lower) apply assumption -apply (simp (no_asm_simp) add: max_ssize_def ssize_sto_def) -apply (simp (no_asm_simp) add: max_of_list_def) - - (* subgoals \ ... *) -apply (rule surj_pair)+ -done + apply (subgoal_tac "\ mt1 sttp1. (f1 sttp0) = (mt1, sttp1)", (erule exE)+) + apply (subgoal_tac "\ mt2 sttp2. (f2 sttp1) = (mt2, sttp2)", (erule exE)+) + apply (subgoal_tac "\ mt3 sttp3. (f3 sttp2) = (mt3, sttp3)", (erule exE)+) + + (* unfold start_sttp_resp and make case distinction *) + apply (simp only: start_sttp_resp_def) + apply (erule_tac Q="start_sttp_resp_cons f2" in disjE) + (* case f2 = comb_nil *) + apply (simp add: bc_mt_corresp_def comb_nil_def start_sttp_resp_cons_def) + + (* case start_sttp_resp_cons f2 *) + apply (simp add: bc_mt_corresp_def comb_def start_sttp_resp_cons_def) + apply (drule_tac x=sttp1 in spec, simp, erule exE) + apply (intro strip, (erule conjE)+) + + + (* get rid of all check_type info *) + apply (subgoal_tac "check_type cG (length (fst sttp1)) mxr (OK (Some sttp1))") + apply (subgoal_tac "check_type cG (max_ssize (mt2 @ [Some sttp2])) mxr (OK (Some sttp2))") + apply (subgoal_tac "check_type cG (max_ssize (mt1 @ mt2 @ mt3 @ [Some sttp3])) mxr + (OK ((mt2 @ mt3 @ [Some sttp3]) ! length mt2))") + apply simp + + + + apply (intro strip, (erule conjE)+) + apply (case_tac "pc < length mt1") + + (* case pc < length mt1 *) + apply (drule spec, drule mp, assumption) + apply assumption + + (* case pc \ length mt1 *) + (* case distinction on start_sttp_resp f3 *) + apply (erule_tac P="f3 = comb_nil" in disjE) + + (* case f3 = comb_nil *) + apply (subgoal_tac "mt3 = [] \ sttp2 = sttp3") apply (erule conjE)+ + apply (subgoal_tac "bc3=[]") + + apply (rule_tac bc_pre=bc1 and bc=bc2 and bc_post=bc3 + and mt_pre=mt1 and mt=mt2 and mt_post="mt3@ [Some sttp3]" + and mxs="(max_ssize (mt2 @ [(Some sttp2)]))" + and max_pc="(Suc (length mt2))" + in wt_instr_offset) + apply simp + apply (rule HOL.refl)+ + apply (simp (no_asm_simp))+ + + apply (simp (no_asm_simp) add: max_ssize_def del: max_of_list_append) + apply (rule max_of_list_sublist) + apply (simp (no_asm_simp) only: set_append list.set list.map) apply blast + apply (simp (no_asm_simp)) + apply simp (* subgoal bc3 = [] *) + apply (simp add: comb_nil_def) (* subgoal mt3 = [] \ sttp2 = sttp3 *) + + (* case start_sttp_resp_cons f3 *) + apply (subgoal_tac "\mt3_rest. (mt3 = Some sttp2 # mt3_rest)", erule exE) + apply (rule_tac bc_pre=bc1 and bc=bc2 and bc_post=bc3 + and mt_pre=mt1 and mt=mt2 and mt_post="mt3@ [Some sttp3]" + and mxs="(max_ssize (mt2 @ [Some sttp2]))" + and max_pc="(Suc (length mt2))" + in wt_instr_offset) + apply (intro strip) + apply (rule_tac bc=bc2 and mt="(mt2 @ [Some sttp2])" + and mxs="(max_ssize (mt2 @ [Some sttp2]))" + and max_pc="(Suc (length mt2))" + in wt_instr_prefix) + + (* preconditions of wt_instr_prefix *) + apply simp + apply (rule HOL.refl) + apply (simp (no_asm_simp))+ + apply simp+ + (* (some) preconditions of wt_instr_offset *) + apply (simp (no_asm_simp) add: max_ssize_def del: max_of_list_append) + apply (rule max_of_list_sublist) + apply (simp (no_asm_simp) only: set_append list.set list.map) + apply blast + apply (simp (no_asm_simp)) + + apply (drule_tac x=sttp2 in spec, simp) (* subgoal \mt3_rest. \ *) + + (* subgoals check_type*) + (* \ ! length mt2 *) + apply simp + + apply (erule_tac P="f3 = comb_nil" in disjE) + + (* -- case f3 = comb_nil *) + apply (subgoal_tac "mt3 = [] \ sttp2 = sttp3") apply (erule conjE)+ + apply simp + apply (rule check_type_mono, assumption) + apply (simp only: max_ssize_def) + apply (rule max_of_list_sublist) + apply (simp (no_asm_simp)) + apply blast + apply simp (* subgoal bc3 = [] *) + apply (simp add: comb_nil_def) (* subgoal mt3 = [] \ sttp2 = sttp3 *) + + + (* -- case start_sttp_resp_cons f3 *) + apply (subgoal_tac "\mt3_rest. (mt3 = Some sttp2 # mt3_rest)", erule exE) + apply (simp (no_asm_simp) add: nth_append) + apply (erule conjE)+ + apply (rule check_type_mono, assumption) + apply (simp only: max_ssize_def) + apply (rule max_of_list_sublist) + apply (simp (no_asm_simp)) + apply blast + apply (drule_tac x=sttp2 in spec, simp) (* subgoal \mt3_rest. \ *) + + + (* subgoal check_type \ Some sttp2 *) + apply (simp add: nth_append) + + (* subgoal check_type \ Some sttp1 *) + apply (simp add: nth_append) + apply (erule conjE)+ + apply (case_tac "sttp1", simp) + apply (rule check_type_lower, assumption) + apply (simp (no_asm_simp) add: max_ssize_def ssize_sto_def) + apply (simp (no_asm_simp) add: max_of_list_def) + + (* subgoals \ ... *) + apply (rule surj_pair)+ + done (* ******************** *) @@ -1475,62 +1470,61 @@ (* ### possibly move into HOL *) -lemma set_drop_Suc [rule_format]: "\ xs. set (drop (Suc n) xs) \ set (drop n xs)" -apply (induct n) -apply simp -apply (intro strip) -apply (rule list.induct) -apply simp -apply simp apply blast -apply (intro strip) -apply (rule_tac - P="\ xs. set (drop (Suc (Suc n)) xs) \ set (drop (Suc n) xs)" in list.induct) -apply simp+ -done - -lemma set_drop_le [rule_format,simp]: "\ n xs. n \ m \ set (drop m xs) \ set (drop n xs)" -apply (induct m) -apply simp -apply (intro strip) -apply (subgoal_tac "n \ m \ n = Suc m") -apply (erule disjE) -apply (frule_tac x=n in spec, drule_tac x=xs in spec, drule mp, assumption) -apply (rule set_drop_Suc [THEN subset_trans], assumption) -apply auto -done - -lemma set_drop [simp] : "set (drop m xs) \ set xs" -apply (rule_tac B="set (drop 0 xs)" in subset_trans) -apply (rule set_drop_le) -apply simp+ -done - - +lemma set_drop_Suc [rule_format]: "\xs. set (drop (Suc n) xs) \ set (drop n xs)" + apply (induct n) + apply simp + apply (intro strip) + apply (rule list.induct) + apply simp + apply simp + apply blast + apply (intro strip) + apply (rule_tac P="\ xs. set (drop (Suc (Suc n)) xs) \ set (drop (Suc n) xs)" in list.induct) + apply simp+ + done + +lemma set_drop_le [rule_format,simp]: "\n xs. n \ m \ set (drop m xs) \ set (drop n xs)" + apply (induct m) + apply simp + apply (intro strip) + apply (subgoal_tac "n \ m \ n = Suc m") + apply (erule disjE) + apply (frule_tac x=n in spec, drule_tac x=xs in spec, drule mp, assumption) + apply (rule set_drop_Suc [THEN subset_trans], assumption) + apply auto + done + +declare set_drop_subset [simp] lemma contracting_popST [simp]: "contracting (popST n)" -by (simp add: contracting_def popST_def) + by (simp add: contracting_def popST_def) lemma contracting_nochangeST [simp]: "contracting nochangeST" -by (simp add: contracting_def nochangeST_def) + by (simp add: contracting_def nochangeST_def) lemma check_type_contracting: "\ check_type cG mxs mxr (OK (Some sttp)); contracting f\ \ check_type cG mxs mxr (OK (Some (sttp_of (f sttp))))" -apply (subgoal_tac "\ ST LT. sttp = (ST, LT)", (erule exE)+) -apply (simp add: check_type_simps contracting_def) -apply clarify -apply (drule_tac x=ST in spec, drule_tac x=LT in spec) -apply (case_tac "(sttp_of (f (ST, LT)))") -apply simp -apply (erule conjE)+ - -apply (drule listE_set)+ -apply (rule conjI) -apply (rule_tac x="length a" in exI) apply simp -apply (rule listI) apply simp apply blast -apply (rule listI) apply simp apply blast -apply auto -done + apply (subgoal_tac "\ ST LT. sttp = (ST, LT)", (erule exE)+) + apply (simp add: check_type_simps contracting_def) + apply clarify + apply (drule_tac x=ST in spec, drule_tac x=LT in spec) + apply (case_tac "(sttp_of (f (ST, LT)))") + apply simp + apply (erule conjE)+ + + apply (drule listE_set)+ + apply (rule conjI) + apply (rule_tac x="length a" in exI) + apply simp + apply (rule listI) + apply simp + apply blast + apply (rule listI) + apply simp + apply blast + apply auto + done (* ******************** *) @@ -1556,78 +1550,76 @@ contracting f2 \ \ bc_mt_corresp bc' f' sttp0 cG rT mxr (length (bc1@[inst]))" -apply (subgoal_tac "\ mt1 sttp1. (f1 sttp0) = (mt1, sttp1)", (erule exE)+) -apply (subgoal_tac "\ mt2 sttp2. (f2 sttp1) = (mt2, sttp2)", (erule exE)+) -apply (subgoal_tac "\ mt3 sttp3. (f3 sttp2) = (mt3, sttp3)", (erule exE)+) - -apply (simp add: bc_mt_corresp_def comb_def start_sttp_resp_cons_def - mt_sttp_flatten_def) - -apply (intro strip, (erule conjE)+) -apply (drule mp, assumption)+ apply (erule conjE)+ -apply (drule mp, assumption) -apply (rule conjI) - - (* wt_instr \ *) -apply (intro strip) -apply (case_tac "pc < length mt1") - - (* case pc < length mt1 *) -apply (drule spec, drule mp, assumption) -apply assumption - - (* case pc \ length mt1 *) -apply (subgoal_tac "pc = length mt1") prefer 2 apply arith -apply (simp only:) -apply (simp add: nth_append mt_sttp_flatten_def) - - - (* check_type \ *) -apply (simp add: start_sttp_resp_def) -apply (drule_tac x="sttp0" in spec, simp, erule exE) -apply (drule_tac x="sttp1" in spec, simp, erule exE) - -apply (subgoal_tac "check_type cG (max_ssize (mt1 @ mt2 @ mt3 @ [Some sttp3])) mxr - (OK (Some (sttp_of (f2 sttp1))))") - -apply (simp only:) - -apply (erule disjE) - (* case f3 = comb_nil *) -apply (subgoal_tac "((mt1 @ mt2 @ mt3 @ [Some sttp3]) ! Suc (length mt1)) = (Some (snd (f2 sttp1)))")apply (subgoal_tac "mt3 = [] \ sttp2 = sttp3") apply (erule conjE)+ -apply (simp add: nth_append) -apply (simp add: comb_nil_def) (* subgoal mt3 = [] \ sttp2 = sttp3 *) -apply (simp add: nth_append comb_nil_def) (* subgoal \ ! Suc (length mt1) *) - - (* case start_sttp_resp_cons f3 *) -apply (simp add: start_sttp_resp_cons_def) -apply (drule_tac x="sttp2" in spec, simp, erule exE) -apply (simp add: nth_append) - - (* subgoal check_type *) -apply (rule check_type_contracting) -apply (subgoal_tac "((mt1 @ mt2 @ mt3 @ [Some sttp3]) ! length mt1) = (Some sttp1)") -apply (simp add: nth_append) -apply (simp add: nth_append) - -apply assumption - -(* subgoals *) -apply (rule surj_pair)+ -done - - -lemma compTpExpr_LT_ST_rewr [simp]: "\ - wf_java_prog G; - wf_java_mdecl G C ((mn, pTs), rT, (pns, lvars, blk, res)); - local_env G C (mn, pTs) pns lvars \ ex :: T; - is_inited_LT C pTs lvars LT\ -\ sttp_of (compTpExpr (pns, lvars, blk, res) G ex (ST, LT)) = (T # ST, LT)" -apply (rule compTpExpr_LT_ST) -apply auto -done - - + apply (subgoal_tac "\ mt1 sttp1. (f1 sttp0) = (mt1, sttp1)", (erule exE)+) + apply (subgoal_tac "\ mt2 sttp2. (f2 sttp1) = (mt2, sttp2)", (erule exE)+) + apply (subgoal_tac "\ mt3 sttp3. (f3 sttp2) = (mt3, sttp3)", (erule exE)+) + + apply (simp add: bc_mt_corresp_def comb_def start_sttp_resp_cons_def + mt_sttp_flatten_def) + + apply (intro strip, (erule conjE)+) + apply (drule mp, assumption)+ + apply (erule conjE)+ + apply (drule mp, assumption) + apply (rule conjI) + + (* wt_instr \ *) + apply (intro strip) + apply (case_tac "pc < length mt1") + + (* case pc < length mt1 *) + apply (drule spec, drule mp, assumption) + apply assumption + + (* case pc \ length mt1 *) + apply (subgoal_tac "pc = length mt1") prefer 2 apply arith + apply (simp only:) + apply (simp add: nth_append mt_sttp_flatten_def) + + + (* check_type \ *) + apply (simp add: start_sttp_resp_def) + apply (drule_tac x="sttp0" in spec, simp, erule exE) + apply (drule_tac x="sttp1" in spec, simp, erule exE) + + apply (subgoal_tac "check_type cG (max_ssize (mt1 @ mt2 @ mt3 @ [Some sttp3])) mxr + (OK (Some (sttp_of (f2 sttp1))))") + + apply (simp only:) + + apply (erule disjE) + (* case f3 = comb_nil *) + apply (subgoal_tac "((mt1 @ mt2 @ mt3 @ [Some sttp3]) ! Suc (length mt1)) = (Some (snd (f2 sttp1)))") + apply (subgoal_tac "mt3 = [] \ sttp2 = sttp3") + apply (erule conjE)+ + apply (simp add: nth_append) + apply (simp add: comb_nil_def) (* subgoal mt3 = [] \ sttp2 = sttp3 *) + apply (simp add: nth_append comb_nil_def) (* subgoal \ ! Suc (length mt1) *) + + (* case start_sttp_resp_cons f3 *) + apply (simp add: start_sttp_resp_cons_def) + apply (drule_tac x="sttp2" in spec, simp, erule exE) + apply (simp add: nth_append) + + (* subgoal check_type *) + apply (rule check_type_contracting) + apply (subgoal_tac "((mt1 @ mt2 @ mt3 @ [Some sttp3]) ! length mt1) = (Some sttp1)") + apply (simp add: nth_append) + apply (simp add: nth_append) + + apply assumption + + (* subgoals *) + apply (rule surj_pair)+ + done + + +lemma compTpExpr_LT_ST_rewr [simp]: + "\ wf_java_prog G; wf_java_mdecl G C ((mn, pTs), rT, (pns, lvars, blk, res)); + local_env G C (mn, pTs) pns lvars \ ex :: T; + is_inited_LT C pTs lvars LT\ + \ sttp_of (compTpExpr (pns, lvars, blk, res) G ex (ST, LT)) = (T # ST, LT)" + by (rule compTpExpr_LT_ST) auto lemma wt_method_compTpExpr_Exprs_corresp: " @@ -1647,265 +1639,291 @@ E \ exs [::] Ts \ (is_inited_LT C pTs lvars LT) \ bc_mt_corresp (compExprs jmb exs) (compTpExprs jmb G exs) (ST, LT) (comp G) rT (length LT) (length (compExprs jmb exs)))" - -apply (rule compat_expr_expr_list.induct) - - -(* expresssions *) - -(* NewC *) -apply (intro allI impI) -apply (simp only:) -apply (drule NewC_invers) -apply (simp (no_asm_use)) -apply (rule bc_mt_corresp_New) -apply (simp add: comp_is_class) - -(* Cast *) -apply (intro allI impI) -apply (simp only:) -apply (drule Cast_invers) -apply clarify -apply (simp (no_asm_use)) -apply (rule bc_mt_corresp_comb) apply (rule HOL.refl, simp (no_asm_simp), blast) -apply (simp (no_asm_simp), rule bc_mt_corresp_Checkcast) -apply (simp add: comp_is_class) -apply (simp only: compTpExpr_LT_ST) -apply (drule cast_RefT) -apply blast -apply (simp add: start_sttp_resp_def) - -(* Lit *) -apply (intro allI impI) -apply (simp only:) -apply (drule Lit_invers) -(* apply (simp (no_asm_use)) *) -apply simp -apply (rule bc_mt_corresp_LitPush) - apply assumption - - -(* BinOp *) - -apply (intro allI impI) -apply (simp (no_asm_simp) only:) -apply (drule BinOp_invers, erule exE, (erule conjE)+) -apply (rename_tac binop expr1 expr2 ST LT T bc' f' Ta, case_tac binop) -apply (simp (no_asm_simp)) - - (* case Eq *) -apply (subgoal_tac "bc_mt_corresp bc' f' (ST, LT) (comp G) rT (length LT) 0") -prefer 2 - apply (rule bc_mt_corresp_zero) apply (simp add: length_compTpExpr) - apply (simp (no_asm_simp)) - -apply (drule_tac ?bc1.0="[]" and ?bc2.0 = "compExpr jmb expr1" - and ?f1.0=comb_nil and ?f2.0 = "compTpExpr jmb G expr1" - in bc_mt_corresp_comb_inside) - apply (simp (no_asm_simp))+ - apply blast - apply (simp (no_asm_simp) add: length_compTpExpr)+ - -apply (drule_tac ?bc2.0 = "compExpr jmb expr2" and ?f2.0 = "compTpExpr jmb G expr2" - in bc_mt_corresp_comb_inside) - apply (simp (no_asm_simp))+ - apply (simp only: compTpExpr_LT_ST) - apply (simp (no_asm_simp) add: length_compTpExpr) - apply (simp (no_asm_simp)) - apply (simp (no_asm_simp)) - -apply (drule_tac ?bc1.0 = "compExpr jmb expr1 @ compExpr jmb expr2" - and inst = "Ifcmpeq 3" and ?bc3.0 = "[LitPush (Bool False),Goto 2, LitPush (Bool True)]" - and ?f1.0="compTpExpr jmb G expr1 \ compTpExpr jmb G expr2" - and ?f2.0="popST 2" and ?f3.0="pushST [PrimT Boolean] \ popST 1 \ pushST [PrimT Boolean]" - in bc_mt_corresp_comb_wt_instr) - apply (simp (no_asm_simp) add: length_compTpExpr)+ - - (* wt_instr *) - apply (intro strip) - apply (simp (no_asm_simp) add: wt_instr_altern_def length_compTpExpr eff_def) - apply (simp (no_asm_simp) add: norm_eff_def) - apply (simp (no_asm_simp) only: int_outside_left nat_int) - apply (simp (no_asm_simp) add: length_compTpExpr) - apply (simp only: compTpExpr_LT_ST)+ - apply (simp add: eff_def norm_eff_def popST_def pushST_def mt_sttp_flatten_def) - apply (case_tac Ta) apply (simp (no_asm_simp)) apply (simp (no_asm_simp)) - apply (rule contracting_popST) (* contracting (popST 2) *) - -apply (drule_tac ?bc1.0 = "compExpr jmb expr1 @ compExpr jmb expr2 @ [Ifcmpeq 3]" - and ?bc2.0 = "[LitPush (Bool False)]" - and ?bc3.0 = "[Goto 2, LitPush (Bool True)]" - and ?f1.0 = "compTpExpr jmb G expr1 \ compTpExpr jmb G expr2 \ popST 2" - and ?f2.0 = "pushST [PrimT Boolean]" - and ?f3.0 = "popST (Suc 0) \ pushST [PrimT Boolean]" - in bc_mt_corresp_comb_inside) - apply (simp (no_asm_simp))+ - apply simp - apply (rule_tac T="(PrimT Boolean)" in bc_mt_corresp_LitPush) apply (simp (no_asm_simp)) - apply (simp (no_asm_simp) add: length_compTpExpr) - apply (simp (no_asm_simp)) - apply (simp (no_asm_simp) add: start_sttp_resp_def) - - -apply (drule_tac ?bc1.0 = "compExpr jmb expr1 @ compExpr jmb expr2 @ [Ifcmpeq 3, LitPush (Bool False)]" - and inst = "Goto 2" and ?bc3.0 = "[LitPush (Bool True)]" - and ?f1.0="compTpExpr jmb G expr1 \ compTpExpr jmb G expr2 \ popST 2 \ pushST [PrimT Boolean]" - and ?f2.0="popST 1" and ?f3.0="pushST [PrimT Boolean]" - in bc_mt_corresp_comb_wt_instr) - apply (simp (no_asm_simp) add: length_compTpExpr)+ - - (* wt_instr *) - apply (simp (no_asm_simp) add: wt_instr_altern_def length_compTpExpr) - apply (simp (no_asm_simp) add: eff_def norm_eff_def) - apply (simp (no_asm_simp) only: int_outside_right nat_int) - apply (simp (no_asm_simp) add: length_compTpExpr) - apply (simp only: compTpExpr_LT_ST)+ - apply (simp add: eff_def norm_eff_def popST_def pushST_def) - apply (rule contracting_popST) (* contracting (popST 1) *) - -apply (drule_tac - ?bc1.0 = "compExpr jmb expr1 @ compExpr jmb expr2 @ [Ifcmpeq 3, LitPush (Bool False), Goto 2]" - and ?bc2.0 = "[LitPush (Bool True)]" - and ?bc3.0 = "[]" - and ?f1.0 = "compTpExpr jmb G expr1 \ compTpExpr jmb G expr2 \ popST 2 \ - pushST [PrimT Boolean] \ popST (Suc 0)" - and ?f2.0 = "pushST [PrimT Boolean]" - and ?f3.0 = "comb_nil" - in bc_mt_corresp_comb_inside) - apply (simp (no_asm_simp))+ + apply (rule compat_expr_expr_list.induct) + + (* expresssions *) + + (* NewC *) + apply (intro allI impI) + apply (simp only:) + apply (drule NewC_invers) + apply (simp (no_asm_use)) + apply (rule bc_mt_corresp_New) + apply (simp add: comp_is_class) + + (* Cast *) + apply (intro allI impI) + apply (simp only:) + apply (drule Cast_invers) + apply clarify + apply (simp (no_asm_use)) + apply (rule bc_mt_corresp_comb) + apply (rule HOL.refl, simp (no_asm_simp), blast) + apply (simp (no_asm_simp), rule bc_mt_corresp_Checkcast) + apply (simp add: comp_is_class) + apply (simp only: compTpExpr_LT_ST) + apply (drule cast_RefT) + apply blast + apply (simp add: start_sttp_resp_def) + + (* Lit *) + apply (intro allI impI) + apply (simp only:) + apply (drule Lit_invers) + apply simp + apply (rule bc_mt_corresp_LitPush) + apply assumption + + + (* BinOp *) + + apply (intro allI impI) + apply (simp (no_asm_simp) only:) + apply (drule BinOp_invers, erule exE, (erule conjE)+) + apply (rename_tac binop expr1 expr2 ST LT T bc' f' Ta, case_tac binop) + apply (simp (no_asm_simp)) + + (* case Eq *) + apply (subgoal_tac "bc_mt_corresp bc' f' (ST, LT) (comp G) rT (length LT) 0") + prefer 2 + apply (rule bc_mt_corresp_zero) + apply (simp add: length_compTpExpr) + apply (simp (no_asm_simp)) + + apply (drule_tac ?bc1.0="[]" and ?bc2.0 = "compExpr jmb expr1" + and ?f1.0=comb_nil and ?f2.0 = "compTpExpr jmb G expr1" + in bc_mt_corresp_comb_inside) + apply (simp (no_asm_simp))+ + apply blast + apply (simp (no_asm_simp) add: length_compTpExpr)+ + + apply (drule_tac ?bc2.0 = "compExpr jmb expr2" and ?f2.0 = "compTpExpr jmb G expr2" + in bc_mt_corresp_comb_inside) + apply (simp (no_asm_simp))+ + apply (simp only: compTpExpr_LT_ST) + apply (simp (no_asm_simp) add: length_compTpExpr) + apply (simp (no_asm_simp)) + apply (simp (no_asm_simp)) + + apply (drule_tac ?bc1.0 = "compExpr jmb expr1 @ compExpr jmb expr2" + and inst = "Ifcmpeq 3" and ?bc3.0 = "[LitPush (Bool False),Goto 2, LitPush (Bool True)]" + and ?f1.0="compTpExpr jmb G expr1 \ compTpExpr jmb G expr2" + and ?f2.0="popST 2" and ?f3.0="pushST [PrimT Boolean] \ popST 1 \ pushST [PrimT Boolean]" + in bc_mt_corresp_comb_wt_instr) + apply (simp (no_asm_simp) add: length_compTpExpr)+ + + (* wt_instr *) + apply (intro strip) + apply (simp (no_asm_simp) add: wt_instr_altern_def length_compTpExpr eff_def) + apply (simp (no_asm_simp) add: norm_eff_def) + apply (simp (no_asm_simp) only: int_outside_left nat_int) + apply (simp (no_asm_simp) add: length_compTpExpr) + apply (simp only: compTpExpr_LT_ST)+ + apply (simp add: eff_def norm_eff_def popST_def pushST_def mt_sttp_flatten_def) + apply (case_tac Ta) + apply (simp (no_asm_simp)) + apply (simp (no_asm_simp)) + apply (rule contracting_popST) (* contracting (popST 2) *) + + apply (drule_tac ?bc1.0 = "compExpr jmb expr1 @ compExpr jmb expr2 @ [Ifcmpeq 3]" + and ?bc2.0 = "[LitPush (Bool False)]" + and ?bc3.0 = "[Goto 2, LitPush (Bool True)]" + and ?f1.0 = "compTpExpr jmb G expr1 \ compTpExpr jmb G expr2 \ popST 2" + and ?f2.0 = "pushST [PrimT Boolean]" + and ?f3.0 = "popST (Suc 0) \ pushST [PrimT Boolean]" + in bc_mt_corresp_comb_inside) + apply (simp (no_asm_simp))+ + apply simp + apply (rule_tac T="(PrimT Boolean)" in bc_mt_corresp_LitPush) apply (simp (no_asm_simp)) + apply (simp (no_asm_simp) add: length_compTpExpr) + apply (simp (no_asm_simp)) + apply (simp (no_asm_simp) add: start_sttp_resp_def) + + + apply (drule_tac ?bc1.0 = "compExpr jmb expr1 @ compExpr jmb expr2 @ [Ifcmpeq 3, LitPush (Bool False)]" + and inst = "Goto 2" and ?bc3.0 = "[LitPush (Bool True)]" + and ?f1.0="compTpExpr jmb G expr1 \ compTpExpr jmb G expr2 \ popST 2 \ pushST [PrimT Boolean]" + and ?f2.0="popST 1" and ?f3.0="pushST [PrimT Boolean]" + in bc_mt_corresp_comb_wt_instr) + apply (simp (no_asm_simp) add: length_compTpExpr)+ + + (* wt_instr *) + apply (simp (no_asm_simp) add: wt_instr_altern_def length_compTpExpr) + apply (simp (no_asm_simp) add: eff_def norm_eff_def) + apply (simp (no_asm_simp) only: int_outside_right nat_int) + apply (simp (no_asm_simp) add: length_compTpExpr) + apply (simp only: compTpExpr_LT_ST)+ + apply (simp add: eff_def norm_eff_def popST_def pushST_def) + apply (rule contracting_popST) (* contracting (popST 1) *) + + apply (drule_tac ?bc1.0 = "compExpr jmb expr1 @ compExpr jmb expr2 @ [Ifcmpeq 3, LitPush (Bool False), Goto 2]" + and ?bc2.0 = "[LitPush (Bool True)]" + and ?bc3.0 = "[]" + and ?f1.0 = "compTpExpr jmb G expr1 \ compTpExpr jmb G expr2 \ popST 2 \ + pushST [PrimT Boolean] \ popST (Suc 0)" + and ?f2.0 = "pushST [PrimT Boolean]" + and ?f3.0 = "comb_nil" + in bc_mt_corresp_comb_inside) + apply (simp (no_asm_simp))+ + apply simp + apply (rule_tac T="(PrimT Boolean)" in bc_mt_corresp_LitPush) + apply (simp (no_asm_simp)) + apply (simp (no_asm_simp) add: length_compTpExpr) + apply (simp (no_asm_simp) add: start_sttp_resp_def) + apply (simp (no_asm_simp)) + + apply simp + + (* case Add *) + apply simp + apply (rule bc_mt_corresp_comb) + apply (rule HOL.refl) + apply simp + apply blast + apply (rule bc_mt_corresp_comb, rule HOL.refl) + apply (simp only: compTpExpr_LT_ST) + apply (simp only: compTpExpr_LT_ST) + apply blast + + apply (simp only: compTpExpr_LT_ST) + apply simp + apply (rule bc_mt_corresp_IAdd) + apply (simp (no_asm_simp) add: start_sttp_resp_def) + apply (simp (no_asm_simp) add: start_sttp_resp_def) + + + (* LAcc *) + apply (intro allI impI) + apply (simp only:) + apply (drule LAcc_invers) + apply (frule wf_java_mdecl_length_pTs_pns) + apply clarify + apply (simp add: is_inited_LT_def) + apply (rule bc_mt_corresp_Load) + apply (rule index_in_bounds) + apply simp + apply assumption + apply (rule inited_LT_at_index_no_err) + apply (rule index_in_bounds) + apply simp + apply assumption + apply (rule HOL.refl) + + + (* LAss *) + apply (intro allI impI) + apply (simp only:) + apply (drule LAss_invers, erule exE, (erule conjE)+) + apply (drule LAcc_invers) + apply (frule wf_java_mdecl_disjoint_varnames, simp add: disjoint_varnames_def) + apply (frule wf_java_mdecl_length_pTs_pns) + apply clarify + apply (simp (no_asm_use)) + apply (rule bc_mt_corresp_comb) + apply (rule HOL.refl, simp (no_asm_simp), blast) + apply (rename_tac vname x2 ST LT T Ta) + apply (rule_tac ?bc1.0="[Dup]" and ?bc2.0="[Store (index (pns, lvars, blk, res) vname)]" + and ?f1.0="dupST" and ?f2.0="popST (Suc 0)" + in bc_mt_corresp_comb) + apply (simp (no_asm_simp))+ + apply (rule bc_mt_corresp_Dup) + apply (simp only: compTpExpr_LT_ST) + apply (simp add: dupST_def is_inited_LT_def) + apply (rule bc_mt_corresp_Store) + apply (rule index_in_bounds) + apply simp + apply assumption + apply (rule sup_loc_update_index, assumption+) + apply simp + apply assumption+ + apply (simp add: start_sttp_resp_def) + apply (simp add: start_sttp_resp_def) + + (* FAcc *) + apply (intro allI impI) + apply (simp only:) + apply (drule FAcc_invers) + apply clarify + apply (simp (no_asm_use)) + apply (rule bc_mt_corresp_comb) + apply (rule HOL.refl, simp (no_asm_simp), blast) + apply (simp (no_asm_simp)) + apply (rule bc_mt_corresp_Getfield) + apply assumption+ + apply (fast intro: wt_class_expr_is_class) + apply (simp (no_asm_simp) add: start_sttp_resp_def) + + + (* FAss *) + apply (intro allI impI) + apply (simp only:) + apply (drule FAss_invers, erule exE, (erule conjE)+) + apply (drule FAcc_invers) + apply clarify + apply (simp (no_asm_use)) + apply (rule bc_mt_corresp_comb) + apply (rule HOL.refl) + apply simp + apply blast + apply (simp only: compTpExpr_LT_ST) + apply (rule bc_mt_corresp_comb, (rule HOL.refl)+) + apply blast + apply (simp only: compTpExpr_LT_ST) + apply (rename_tac cname x2 vname x4 ST LT T Ta Ca) + apply (rule_tac ?bc1.0="[Dup_x1]" and ?bc2.0="[Putfield vname cname]" in bc_mt_corresp_comb) + apply (simp (no_asm_simp))+ + apply (rule bc_mt_corresp_Dup_x1) + apply (simp (no_asm_simp) add: dup_x1ST_def) + apply (rule bc_mt_corresp_Putfield, assumption+) + apply (fast intro: wt_class_expr_is_class) + apply (simp (no_asm_simp) add: start_sttp_resp_def) + apply (simp (no_asm_simp) add: start_sttp_resp_def) + apply (simp (no_asm_simp) add: start_sttp_resp_def) + + (* Call *) + apply (intro allI impI) + apply (simp only:) + apply (drule Call_invers) + apply clarify + apply (simp (no_asm_use)) + apply (rule bc_mt_corresp_comb) + apply (rule HOL.refl) + apply simp + apply blast + apply (simp only: compTpExpr_LT_ST) + apply (rule bc_mt_corresp_comb, (rule HOL.refl)+) + apply blast + apply (simp only: compTpExprs_LT_ST) + apply (simp (no_asm_simp)) + apply (rule bc_mt_corresp_Invoke) + apply assumption+ + apply (fast intro: wt_class_expr_is_class) + apply (simp (no_asm_simp) add: start_sttp_resp_def) + apply (rule start_sttp_resp_comb) + apply (simp (no_asm_simp)) + apply (simp (no_asm_simp) add: start_sttp_resp_def) + + + (* expression lists *) + (* nil *) + + apply (intro allI impI) + apply (drule Nil_invers) + apply simp + + (* cons *) + + apply (intro allI impI) + apply (drule Cons_invers, (erule exE)+, (erule conjE)+) + apply clarify + apply (simp (no_asm_use)) + apply (rule bc_mt_corresp_comb) + apply (rule HOL.refl) + apply simp + apply blast + apply (simp only: compTpExpr_LT_ST) + apply blast apply simp - apply (rule_tac T="(PrimT Boolean)" in bc_mt_corresp_LitPush) apply (simp (no_asm_simp)) - apply (simp (no_asm_simp) add: length_compTpExpr) - apply (simp (no_asm_simp) add: start_sttp_resp_def) - apply (simp (no_asm_simp)) - -apply simp - - (* case Add *) -apply simp -apply (rule bc_mt_corresp_comb) apply (rule HOL.refl) apply simp apply blast -apply (rule bc_mt_corresp_comb, rule HOL.refl) - apply (simp only: compTpExpr_LT_ST) -apply (simp only: compTpExpr_LT_ST) apply blast - -apply (simp only: compTpExpr_LT_ST) -apply simp -apply (rule bc_mt_corresp_IAdd) -apply (simp (no_asm_simp) add: start_sttp_resp_def) -apply (simp (no_asm_simp) add: start_sttp_resp_def) - - - (* LAcc *) -apply (intro allI impI) -apply (simp only:) -apply (drule LAcc_invers) -apply (frule wf_java_mdecl_length_pTs_pns) -apply clarify -apply (simp add: is_inited_LT_def) -apply (rule bc_mt_corresp_Load) - apply (rule index_in_bounds) apply simp apply assumption - apply (rule inited_LT_at_index_no_err) - apply (rule index_in_bounds) apply simp apply assumption -apply (rule HOL.refl) - - - (* LAss *) -apply (intro allI impI) -apply (simp only:) -apply (drule LAss_invers, erule exE, (erule conjE)+) -apply (drule LAcc_invers) -apply (frule wf_java_mdecl_disjoint_varnames, simp add: disjoint_varnames_def) -apply (frule wf_java_mdecl_length_pTs_pns) -apply clarify -apply (simp (no_asm_use)) -apply (rule bc_mt_corresp_comb) apply (rule HOL.refl, simp (no_asm_simp), blast) -apply (rename_tac vname x2 ST LT T Ta) -apply (rule_tac ?bc1.0="[Dup]" and ?bc2.0="[Store (index (pns, lvars, blk, res) vname)]" - and ?f1.0="dupST" and ?f2.0="popST (Suc 0)" - in bc_mt_corresp_comb) - apply (simp (no_asm_simp))+ - apply (rule bc_mt_corresp_Dup) - apply (simp only: compTpExpr_LT_ST) - apply (simp add: dupST_def is_inited_LT_def) - apply (rule bc_mt_corresp_Store) - apply (rule index_in_bounds) - apply simp apply assumption - apply (rule sup_loc_update_index, assumption+) - apply simp apply assumption+ -apply (simp add: start_sttp_resp_def) -apply (simp add: start_sttp_resp_def) - - (* FAcc *) -apply (intro allI impI) -apply (simp only:) -apply (drule FAcc_invers) -apply clarify -apply (simp (no_asm_use)) -apply (rule bc_mt_corresp_comb) apply (rule HOL.refl, simp (no_asm_simp), blast) - apply (simp (no_asm_simp)) - apply (rule bc_mt_corresp_Getfield) apply assumption+ - apply (fast intro: wt_class_expr_is_class) -apply (simp (no_asm_simp) add: start_sttp_resp_def) - - - (* FAss *) -apply (intro allI impI) -apply (simp only:) -apply (drule FAss_invers, erule exE, (erule conjE)+) -apply (drule FAcc_invers) -apply clarify -apply (simp (no_asm_use)) -apply (rule bc_mt_corresp_comb) apply (rule HOL.refl) apply simp apply blast - apply (simp only: compTpExpr_LT_ST) -apply (rule bc_mt_corresp_comb, (rule HOL.refl)+) apply blast - apply (simp only: compTpExpr_LT_ST) -apply (rename_tac cname x2 vname x4 ST LT T Ta Ca) -apply (rule_tac ?bc1.0="[Dup_x1]" and ?bc2.0="[Putfield vname cname]" in bc_mt_corresp_comb) - apply (simp (no_asm_simp))+ -apply (rule bc_mt_corresp_Dup_x1) - apply (simp (no_asm_simp) add: dup_x1ST_def) -apply (rule bc_mt_corresp_Putfield) apply assumption+ - apply (fast intro: wt_class_expr_is_class) -apply (simp (no_asm_simp) add: start_sttp_resp_def) -apply (simp (no_asm_simp) add: start_sttp_resp_def) -apply (simp (no_asm_simp) add: start_sttp_resp_def) - -(* Call *) -apply (intro allI impI) -apply (simp only:) -apply (drule Call_invers) -apply clarify -apply (simp (no_asm_use)) -apply (rule bc_mt_corresp_comb) apply (rule HOL.refl) apply simp apply blast - apply (simp only: compTpExpr_LT_ST) -apply (rule bc_mt_corresp_comb, (rule HOL.refl)+) apply blast - apply (simp only: compTpExprs_LT_ST) - apply (simp (no_asm_simp)) -apply (rule bc_mt_corresp_Invoke) apply assumption+ - apply (fast intro: wt_class_expr_is_class) -apply (simp (no_asm_simp) add: start_sttp_resp_def) -apply (rule start_sttp_resp_comb) - apply (simp (no_asm_simp)) - apply (simp (no_asm_simp) add: start_sttp_resp_def) - - -(* expression lists *) -(* nil *) - -apply (intro allI impI) -apply (drule Nil_invers) -apply simp - -(* cons *) - -apply (intro allI impI) -apply (drule Cons_invers, (erule exE)+, (erule conjE)+) -apply clarify -apply (simp (no_asm_use)) -apply (rule bc_mt_corresp_comb) apply (rule HOL.refl) apply simp apply blast - apply (simp only: compTpExpr_LT_ST) -apply blast -apply simp done @@ -1932,301 +1950,295 @@ f' = (compTpStmt jmb G s) \ bc_mt_corresp bc' f' (ST, LT) (comp G) rT (length LT) (length bc'))" -apply (rule stmt.induct) - -(* Skip *) -apply (intro allI impI) -apply simp - - -(* Expr *) -apply (intro allI impI) -apply (drule Expr_invers, erule exE) -apply (simp (no_asm_simp)) -apply (rule bc_mt_corresp_comb) apply (rule HOL.refl, simp (no_asm_simp)) -apply (rule wt_method_compTpExpr_corresp) apply assumption+ -apply (simp add: compTpExpr_LT_ST [of _ pns lvars blk res])+ -apply (rule bc_mt_corresp_Pop) -apply (simp add: start_sttp_resp_def) - - -(* Comp *) -apply (intro allI impI) -apply (drule Comp_invers) -apply clarify -apply (simp (no_asm_use)) -apply (rule bc_mt_corresp_comb) apply (rule HOL.refl) - apply (simp (no_asm_simp)) apply blast -apply (simp only: compTpStmt_LT_ST) -apply (simp (no_asm_simp)) - -(* Cond *) -apply (intro allI impI) -apply (simp (no_asm_simp) only:) -apply (drule Cond_invers, (erule conjE)+) -apply (simp (no_asm_simp)) - -apply (subgoal_tac "bc_mt_corresp bc' f' (ST, LT) (comp G) rT (length LT) 0") -prefer 2 -apply (rule bc_mt_corresp_zero) - apply (simp (no_asm_simp) add: length_compTpStmt length_compTpExpr) + apply (rule stmt.induct) + + (* Skip *) + apply (intro allI impI) + apply simp + + + (* Expr *) + apply (intro allI impI) + apply (drule Expr_invers, erule exE) + apply (simp (no_asm_simp)) + apply (rule bc_mt_corresp_comb) apply (rule HOL.refl, simp (no_asm_simp)) + apply (rule wt_method_compTpExpr_corresp) apply assumption+ + apply (simp add: compTpExpr_LT_ST [of _ pns lvars blk res])+ + apply (rule bc_mt_corresp_Pop) + apply (simp add: start_sttp_resp_def) + + + (* Comp *) + apply (intro allI impI) + apply (drule Comp_invers) + apply clarify + apply (simp (no_asm_use)) + apply (rule bc_mt_corresp_comb) apply (rule HOL.refl) + apply (simp (no_asm_simp)) apply blast + apply (simp only: compTpStmt_LT_ST) + apply (simp (no_asm_simp)) + + (* Cond *) + apply (intro allI impI) + apply (simp (no_asm_simp) only:) + apply (drule Cond_invers, (erule conjE)+) + apply (simp (no_asm_simp)) + + apply (subgoal_tac "bc_mt_corresp bc' f' (ST, LT) (comp G) rT (length LT) 0") + prefer 2 + apply (rule bc_mt_corresp_zero) + apply (simp (no_asm_simp) add: length_compTpStmt length_compTpExpr) + apply (simp (no_asm_simp)) + + apply (rename_tac expr stmt1 stmt2 ST LT bc' f') + apply (drule_tac ?bc1.0="[]" and ?bc2.0 = "[LitPush (Bool False)]" + and ?bc3.0="compExpr jmb expr @ Ifcmpeq (2 + int (length (compStmt jmb stmt1))) # + compStmt jmb stmt1 @ Goto (1 + int (length (compStmt jmb stmt2))) # + compStmt jmb stmt2" + and ?f1.0=comb_nil and ?f2.0 = "pushST [PrimT Boolean]" + and ?f3.0="compTpExpr jmb G expr \ popST 2 \ compTpStmt jmb G stmt1 \ + nochangeST \ compTpStmt jmb G stmt2" + in bc_mt_corresp_comb_inside) + apply (simp (no_asm_simp))+ + apply (rule_tac T="(PrimT Boolean)" in bc_mt_corresp_LitPush) + apply (simp (no_asm_simp) add: start_sttp_resp_def)+ + + apply (drule_tac ?bc1.0="[LitPush (Bool False)]" and ?bc2.0 = "compExpr jmb expr" + and ?bc3.0="Ifcmpeq (2 + int (length (compStmt jmb stmt1))) # + compStmt jmb stmt1 @ Goto (1 + int (length (compStmt jmb stmt2))) # + compStmt jmb stmt2" + and ?f1.0="pushST [PrimT Boolean]" and ?f2.0 = "compTpExpr jmb G expr" + and ?f3.0="popST 2 \ compTpStmt jmb G stmt1 \ nochangeST \ compTpStmt jmb G stmt2" + in bc_mt_corresp_comb_inside) + apply (simp (no_asm_simp))+ + apply (simp (no_asm_simp) add: pushST_def) + apply (rule wt_method_compTpExpr_corresp, assumption+) + apply (simp (no_asm_simp))+ + + + apply (drule_tac ?bc1.0 = "[LitPush (Bool False)] @ compExpr jmb expr" + and inst = "Ifcmpeq (2 + int (length (compStmt jmb stmt1)))" + and ?bc3.0 = "compStmt jmb stmt1 @ Goto (1 + int (length (compStmt jmb stmt2))) # + compStmt jmb stmt2" + and ?f1.0="pushST [PrimT Boolean] \ compTpExpr jmb G expr" and ?f2.0 = "popST 2" + and ?f3.0="compTpStmt jmb G stmt1 \ nochangeST \ compTpStmt jmb G stmt2" + in bc_mt_corresp_comb_wt_instr) + apply (simp (no_asm_simp) add: length_compTpExpr)+ + apply (simp (no_asm_simp) add: start_sttp_resp_comb) + + (* wt_instr *) + apply (intro strip) + apply (rule_tac ts="PrimT Boolean" and ts'="PrimT Boolean" and ST=ST and LT=LT + in wt_instr_Ifcmpeq) + apply (simp (no_asm_simp)) + apply (simp (no_asm_simp) only: int_outside_right nat_int, simp (no_asm_simp)) + apply (simp (no_asm_simp) only: int_outside_right nat_int, simp (no_asm_simp)) + (* current pc *) + apply (simp add: length_compTpExpr pushST_def) + apply (simp only: compTpExpr_LT_ST) + (* Suc pc *) + apply (simp add: length_compTpExpr pushST_def) + apply (simp add: popST_def start_sttp_resp_comb) + (* jump goal *) + apply (simp (no_asm_simp) only: int_outside_right nat_int, simp (no_asm_simp)) + apply (simp add: length_compTpExpr pushST_def) + apply (simp add: popST_def start_sttp_resp_comb length_compTpStmt) + apply (simp only: compTpStmt_LT_ST) + apply (simp add: nochangeST_def) + (* check_type *) + apply (subgoal_tac " + (mt_sttp_flatten (f' (ST, LT)) ! length ([LitPush (Bool False)] @ compExpr jmb expr)) = + (Some (PrimT Boolean # PrimT Boolean # ST, LT))") + apply (simp only:) + apply (simp (no_asm_simp)) apply (rule trans, rule mt_sttp_flatten_comb_length) + apply (rule HOL.refl) apply (simp (no_asm_simp) add: length_compTpExpr) + apply (simp (no_asm_simp) add: length_compTpExpr pushST_def) + apply (simp only: compTpExpr_LT_ST_rewr) + (* contracting\ *) + apply (rule contracting_popST) + + apply (drule_tac ?bc1.0="[LitPush (Bool False)] @ compExpr jmb expr @ + [Ifcmpeq (2 + int (length (compStmt jmb stmt1)))] " + and ?bc2.0 = "compStmt jmb stmt1" + and ?bc3.0="Goto (1 + int (length (compStmt jmb stmt2))) # compStmt jmb stmt2" + and ?f1.0="pushST [PrimT Boolean] \ compTpExpr jmb G expr \ popST 2" + and ?f2.0 = "compTpStmt jmb G stmt1" + and ?f3.0="nochangeST \ compTpStmt jmb G stmt2" + in bc_mt_corresp_comb_inside) + apply (simp (no_asm_simp))+ + apply (simp (no_asm_simp) add: pushST_def popST_def compTpExpr_LT_ST) + apply (simp only: compTpExpr_LT_ST) + apply (simp (no_asm_simp)) + apply (simp (no_asm_simp) add: length_compTpExpr)+ + + + apply (drule_tac ?bc1.0 = "[LitPush (Bool False)] @ compExpr jmb expr @ [Ifcmpeq (2 + int (length (compStmt jmb stmt1)))] @ compStmt jmb stmt1" + and inst = "Goto (1 + int (length (compStmt jmb stmt2)))" + and ?bc3.0 = "compStmt jmb stmt2" + and ?f1.0="pushST [PrimT Boolean] \ compTpExpr jmb G expr \ popST 2 \ compTpStmt jmb G stmt1" + and ?f2.0 = "nochangeST" + and ?f3.0="compTpStmt jmb G stmt2" + in bc_mt_corresp_comb_wt_instr) + apply (simp (no_asm_simp) add: length_compTpExpr length_compTpStmt)+ + apply (intro strip) + apply (rule wt_instr_Goto) + apply (simp (no_asm_simp) only: int_outside_right nat_int, simp (no_asm_simp)) + apply (simp (no_asm_simp) only: int_outside_right nat_int, simp (no_asm_simp)) + (* \ ! nat (int pc + i) = \ ! pc *) + apply (simp (no_asm_simp) add: length_compTpExpr length_compTpStmt) + apply (simp (no_asm_simp) add: pushST_def popST_def nochangeST_def) + apply (simp only: compTpExpr_LT_ST compTpStmt_LT_ST) + apply (simp (no_asm_simp) add: pushST_def popST_def nochangeST_def) + apply (simp only: compTpExpr_LT_ST compTpStmt_LT_ST) + apply (simp only:) + apply (simp add: length_compTpExpr length_compTpStmt) + apply (rule contracting_nochangeST) + + + apply (drule_tac + ?bc1.0= "[LitPush (Bool False)] @ compExpr jmb expr @ + [Ifcmpeq (2 + int (length (compStmt jmb stmt1)))] @ + compStmt jmb stmt1 @ [Goto (1 + int (length (compStmt jmb stmt2)))]" + and ?bc2.0 = "compStmt jmb stmt2" + and ?bc3.0="[]" + and ?f1.0="pushST [PrimT Boolean] \ compTpExpr jmb G expr \ popST 2 \ compTpStmt jmb G stmt1 \ nochangeST" + and ?f2.0 = "compTpStmt jmb G stmt2" + and ?f3.0="comb_nil" + in bc_mt_corresp_comb_inside) + apply (simp (no_asm_simp))+ + apply (simp (no_asm_simp) add: pushST_def popST_def nochangeST_def compTpExpr_LT_ST) + apply (simp only: compTpExpr_LT_ST) + apply (simp (no_asm_simp)) + apply (simp only: compTpStmt_LT_ST) + apply (simp (no_asm_simp) add: length_compTpExpr length_compTpStmt)+ + + apply simp + + + (* Loop *) + apply (intro allI impI) + apply (simp (no_asm_simp) only:) + apply (drule Loop_invers, (erule conjE)+) apply (simp (no_asm_simp)) -apply (rename_tac expr stmt1 stmt2 ST LT bc' f') -apply (drule_tac ?bc1.0="[]" and ?bc2.0 = "[LitPush (Bool False)]" - and ?bc3.0="compExpr jmb expr @ Ifcmpeq (2 + int (length (compStmt jmb stmt1))) # - compStmt jmb stmt1 @ Goto (1 + int (length (compStmt jmb stmt2))) # - compStmt jmb stmt2" - and ?f1.0=comb_nil and ?f2.0 = "pushST [PrimT Boolean]" - and ?f3.0="compTpExpr jmb G expr \ popST 2 \ compTpStmt jmb G stmt1 \ - nochangeST \ compTpStmt jmb G stmt2" - in bc_mt_corresp_comb_inside) - apply (simp (no_asm_simp))+ - apply (rule_tac T="(PrimT Boolean)" in bc_mt_corresp_LitPush) - apply (simp (no_asm_simp) add: start_sttp_resp_def)+ - -apply (drule_tac ?bc1.0="[LitPush (Bool False)]" and ?bc2.0 = "compExpr jmb expr" - and ?bc3.0="Ifcmpeq (2 + int (length (compStmt jmb stmt1))) # - compStmt jmb stmt1 @ Goto (1 + int (length (compStmt jmb stmt2))) # - compStmt jmb stmt2" - and ?f1.0="pushST [PrimT Boolean]" and ?f2.0 = "compTpExpr jmb G expr" - and ?f3.0="popST 2 \ compTpStmt jmb G stmt1 \ - nochangeST \ compTpStmt jmb G stmt2" - in bc_mt_corresp_comb_inside) - apply (simp (no_asm_simp))+ - apply (simp (no_asm_simp) add: pushST_def) - apply (rule wt_method_compTpExpr_corresp) apply assumption+ - apply (simp (no_asm_simp))+ - - -apply (drule_tac ?bc1.0 = "[LitPush (Bool False)] @ compExpr jmb expr" - and inst = "Ifcmpeq (2 + int (length (compStmt jmb stmt1)))" - and ?bc3.0 = "compStmt jmb stmt1 @ Goto (1 + int (length (compStmt jmb stmt2))) # - compStmt jmb stmt2" - and ?f1.0="pushST [PrimT Boolean] \ compTpExpr jmb G expr" and ?f2.0 = "popST 2" - and ?f3.0="compTpStmt jmb G stmt1 \ nochangeST \ compTpStmt jmb G stmt2" - in bc_mt_corresp_comb_wt_instr) - apply (simp (no_asm_simp) add: length_compTpExpr)+ - apply (simp (no_asm_simp) add: start_sttp_resp_comb) + apply (subgoal_tac "bc_mt_corresp bc' f' (ST, LT) (comp G) rT (length LT) 0") + prefer 2 + apply (rule bc_mt_corresp_zero) + apply (simp (no_asm_simp) add: length_compTpStmt length_compTpExpr) + apply (simp (no_asm_simp)) + + apply (rename_tac expr stmt ST LT bc' f') + apply (drule_tac ?bc1.0="[]" and ?bc2.0 = "[LitPush (Bool False)]" + and ?bc3.0="compExpr jmb expr @ Ifcmpeq (2 + int (length (compStmt jmb stmt))) # + compStmt jmb stmt @ + [Goto (-2 + (- int (length (compStmt jmb stmt)) - int (length (compExpr jmb expr))))]" + and ?f1.0=comb_nil and ?f2.0 = "pushST [PrimT Boolean]" + and ?f3.0="compTpExpr jmb G expr \ popST 2 \ compTpStmt jmb G stmt \ nochangeST" + in bc_mt_corresp_comb_inside) + apply (simp (no_asm_simp))+ + apply (rule_tac T="(PrimT Boolean)" in bc_mt_corresp_LitPush) + apply (simp (no_asm_simp) add: start_sttp_resp_def)+ + + apply (drule_tac ?bc1.0="[LitPush (Bool False)]" and ?bc2.0 = "compExpr jmb expr" + and ?bc3.0="Ifcmpeq (2 + int (length (compStmt jmb stmt))) # + compStmt jmb stmt @ + [Goto (-2 + (- int (length (compStmt jmb stmt)) - int (length (compExpr jmb expr))))]" + and ?f1.0="pushST [PrimT Boolean]" and ?f2.0 = "compTpExpr jmb G expr" + and ?f3.0="popST 2 \ compTpStmt jmb G stmt \ nochangeST" + in bc_mt_corresp_comb_inside) + apply (simp (no_asm_simp))+ + apply (simp (no_asm_simp) add: pushST_def) + apply (rule wt_method_compTpExpr_corresp, assumption+) + apply (simp (no_asm_simp))+ + + + apply (drule_tac ?bc1.0 = "[LitPush (Bool False)] @ compExpr jmb expr" + and inst = "Ifcmpeq (2 + int (length (compStmt jmb stmt)))" + and ?bc3.0 = "compStmt jmb stmt @ + [Goto (-2 + (- int (length (compStmt jmb stmt)) - int (length (compExpr jmb expr))))]" + and ?f1.0="pushST [PrimT Boolean] \ compTpExpr jmb G expr" and ?f2.0 = "popST 2" + and ?f3.0="compTpStmt jmb G stmt \ nochangeST" + in bc_mt_corresp_comb_wt_instr) + apply (simp (no_asm_simp) add: length_compTpExpr)+ + apply (simp (no_asm_simp) add: start_sttp_resp_comb) (* wt_instr *) - apply (intro strip) - apply (rule_tac ts="PrimT Boolean" and ts'="PrimT Boolean" - and ST=ST and LT=LT - in wt_instr_Ifcmpeq) - apply (simp (no_asm_simp)) - apply (simp (no_asm_simp) only: int_outside_right nat_int, simp (no_asm_simp)) - apply (simp (no_asm_simp) only: int_outside_right nat_int, simp (no_asm_simp)) - (* current pc *) - apply (simp add: length_compTpExpr pushST_def) - apply (simp only: compTpExpr_LT_ST) - (* Suc pc *) - apply (simp add: length_compTpExpr pushST_def) - apply (simp add: popST_def start_sttp_resp_comb) - (* jump goal *) - apply (simp (no_asm_simp) only: int_outside_right nat_int, simp (no_asm_simp)) - apply (simp add: length_compTpExpr pushST_def) - apply (simp add: popST_def start_sttp_resp_comb length_compTpStmt) - apply (simp only: compTpStmt_LT_ST) - apply (simp add: nochangeST_def) + apply (intro strip) + apply (rule_tac ts="PrimT Boolean" and ts'="PrimT Boolean" + and ST=ST and LT=LT + in wt_instr_Ifcmpeq) + apply (simp (no_asm_simp)) + apply (simp (no_asm_simp) only: int_outside_right nat_int, simp (no_asm_simp)) + apply (simp (no_asm_simp) only: int_outside_right nat_int, simp (no_asm_simp)) + (* current pc *) + apply (simp add: length_compTpExpr pushST_def) + apply (simp only: compTpExpr_LT_ST) + (* Suc pc *) + apply (simp add: length_compTpExpr pushST_def) + apply (simp add: popST_def start_sttp_resp_comb) + (* jump goal *) + apply (simp (no_asm_simp) only: int_outside_right nat_int, simp (no_asm_simp)) + apply (simp add: length_compTpExpr pushST_def) + apply (simp add: popST_def start_sttp_resp_comb length_compTpStmt) + apply (simp only: compTpStmt_LT_ST) + apply (simp add: nochangeST_def) (* check_type *) - apply (subgoal_tac " - (mt_sttp_flatten (f' (ST, LT)) ! length ([LitPush (Bool False)] @ compExpr jmb expr)) = - (Some (PrimT Boolean # PrimT Boolean # ST, LT))") - apply (simp only:) - apply (simp (no_asm_simp)) apply (rule trans, rule mt_sttp_flatten_comb_length) - apply (rule HOL.refl) apply (simp (no_asm_simp) add: length_compTpExpr) + apply (subgoal_tac " + (mt_sttp_flatten (f' (ST, LT)) ! length ([LitPush (Bool False)] @ compExpr jmb expr)) = + (Some (PrimT Boolean # PrimT Boolean # ST, LT))") + apply (simp only:) + apply (simp (no_asm_simp)) apply (rule trans, rule mt_sttp_flatten_comb_length) + apply (rule HOL.refl) apply (simp (no_asm_simp) add: length_compTpExpr) apply (simp (no_asm_simp) add: length_compTpExpr pushST_def) - apply (simp only: compTpExpr_LT_ST_rewr) - (* contracting\ *) - apply (rule contracting_popST) - -apply (drule_tac - ?bc1.0="[LitPush (Bool False)] @ compExpr jmb expr @ - [Ifcmpeq (2 + int (length (compStmt jmb stmt1)))] " - and ?bc2.0 = "compStmt jmb stmt1" - and ?bc3.0="Goto (1 + int (length (compStmt jmb stmt2))) # compStmt jmb stmt2" - and ?f1.0="pushST [PrimT Boolean] \ compTpExpr jmb G expr \ popST 2" - and ?f2.0 = "compTpStmt jmb G stmt1" - and ?f3.0="nochangeST \ compTpStmt jmb G stmt2" - in bc_mt_corresp_comb_inside) - apply (simp (no_asm_simp))+ - apply (simp (no_asm_simp) add: pushST_def popST_def compTpExpr_LT_ST) - apply (simp only: compTpExpr_LT_ST) - apply (simp (no_asm_simp)) - apply (simp (no_asm_simp) add: length_compTpExpr)+ - - -apply (drule_tac ?bc1.0 = "[LitPush (Bool False)] @ compExpr jmb expr @ [Ifcmpeq (2 + int (length (compStmt jmb stmt1)))] @ compStmt jmb stmt1" - and inst = "Goto (1 + int (length (compStmt jmb stmt2)))" - and ?bc3.0 = "compStmt jmb stmt2" - and ?f1.0="pushST [PrimT Boolean] \ compTpExpr jmb G expr \ popST 2 \ - compTpStmt jmb G stmt1" - and ?f2.0 = "nochangeST" - and ?f3.0="compTpStmt jmb G stmt2" - in bc_mt_corresp_comb_wt_instr) - apply (simp (no_asm_simp) add: length_compTpExpr length_compTpStmt)+ - apply (intro strip) - apply (rule wt_instr_Goto) - apply (simp (no_asm_simp) only: int_outside_right nat_int, simp (no_asm_simp)) - apply (simp (no_asm_simp) only: int_outside_right nat_int, simp (no_asm_simp)) - (* \ ! nat (int pc + i) = \ ! pc *) - apply (simp (no_asm_simp) add: length_compTpExpr length_compTpStmt) - apply (simp (no_asm_simp) add: pushST_def popST_def nochangeST_def) - apply (simp only: compTpExpr_LT_ST compTpStmt_LT_ST) - apply (simp (no_asm_simp) add: pushST_def popST_def nochangeST_def) - apply (simp only: compTpExpr_LT_ST compTpStmt_LT_ST) - apply (simp only:) - apply (simp add: length_compTpExpr length_compTpStmt) - apply (rule contracting_nochangeST) - - -apply (drule_tac - ?bc1.0= "[LitPush (Bool False)] @ compExpr jmb expr @ - [Ifcmpeq (2 + int (length (compStmt jmb stmt1)))] @ - compStmt jmb stmt1 @ [Goto (1 + int (length (compStmt jmb stmt2)))]" - and ?bc2.0 = "compStmt jmb stmt2" - and ?bc3.0="[]" - and ?f1.0="pushST [PrimT Boolean] \ compTpExpr jmb G expr \ popST 2 \ - compTpStmt jmb G stmt1 \ nochangeST" - and ?f2.0 = "compTpStmt jmb G stmt2" - and ?f3.0="comb_nil" - in bc_mt_corresp_comb_inside) - apply (simp (no_asm_simp))+ - apply (simp (no_asm_simp) add: pushST_def popST_def nochangeST_def compTpExpr_LT_ST) - apply (simp only: compTpExpr_LT_ST) - apply (simp (no_asm_simp)) - apply (simp only: compTpStmt_LT_ST) - apply (simp (no_asm_simp) add: length_compTpExpr length_compTpStmt)+ - -apply simp - - -(* Loop *) -apply (intro allI impI) -apply (simp (no_asm_simp) only:) -apply (drule Loop_invers, (erule conjE)+) -apply (simp (no_asm_simp)) - -apply (subgoal_tac "bc_mt_corresp bc' f' (ST, LT) (comp G) rT (length LT) 0") -prefer 2 -apply (rule bc_mt_corresp_zero) - apply (simp (no_asm_simp) add: length_compTpStmt length_compTpExpr) - apply (simp (no_asm_simp)) - -apply (rename_tac expr stmt ST LT bc' f') -apply (drule_tac ?bc1.0="[]" and ?bc2.0 = "[LitPush (Bool False)]" - and ?bc3.0="compExpr jmb expr @ Ifcmpeq (2 + int (length (compStmt jmb stmt))) # - compStmt jmb stmt @ - [Goto (-2 + (- int (length (compStmt jmb stmt)) - int (length (compExpr jmb expr))))]" - and ?f1.0=comb_nil and ?f2.0 = "pushST [PrimT Boolean]" - and ?f3.0="compTpExpr jmb G expr \ popST 2 \ compTpStmt jmb G stmt \ nochangeST" - in bc_mt_corresp_comb_inside) - apply (simp (no_asm_simp))+ - apply (rule_tac T="(PrimT Boolean)" in bc_mt_corresp_LitPush) - apply (simp (no_asm_simp) add: start_sttp_resp_def)+ - -apply (drule_tac ?bc1.0="[LitPush (Bool False)]" and ?bc2.0 = "compExpr jmb expr" - and ?bc3.0="Ifcmpeq (2 + int (length (compStmt jmb stmt))) # - compStmt jmb stmt @ - [Goto (-2 + (- int (length (compStmt jmb stmt)) - int (length (compExpr jmb expr))))]" - and ?f1.0="pushST [PrimT Boolean]" and ?f2.0 = "compTpExpr jmb G expr" - and ?f3.0="popST 2 \ compTpStmt jmb G stmt \ nochangeST" - in bc_mt_corresp_comb_inside) - apply (simp (no_asm_simp))+ - apply (simp (no_asm_simp) add: pushST_def) - apply (rule wt_method_compTpExpr_corresp) apply assumption+ - apply (simp (no_asm_simp))+ - - -apply (drule_tac ?bc1.0 = "[LitPush (Bool False)] @ compExpr jmb expr" - and inst = "Ifcmpeq (2 + int (length (compStmt jmb stmt)))" - and ?bc3.0 = "compStmt jmb stmt @ - [Goto (-2 + (- int (length (compStmt jmb stmt)) - int (length (compExpr jmb expr))))]" - and ?f1.0="pushST [PrimT Boolean] \ compTpExpr jmb G expr" and ?f2.0 = "popST 2" - and ?f3.0="compTpStmt jmb G stmt \ nochangeST" - in bc_mt_corresp_comb_wt_instr) - apply (simp (no_asm_simp) add: length_compTpExpr)+ - apply (simp (no_asm_simp) add: start_sttp_resp_comb) - - (* wt_instr *) - apply (intro strip) - apply (rule_tac ts="PrimT Boolean" and ts'="PrimT Boolean" - and ST=ST and LT=LT - in wt_instr_Ifcmpeq) - apply (simp (no_asm_simp)) - apply (simp (no_asm_simp) only: int_outside_right nat_int, simp (no_asm_simp)) - apply (simp (no_asm_simp) only: int_outside_right nat_int, simp (no_asm_simp)) - (* current pc *) - apply (simp add: length_compTpExpr pushST_def) - apply (simp only: compTpExpr_LT_ST) - (* Suc pc *) - apply (simp add: length_compTpExpr pushST_def) - apply (simp add: popST_def start_sttp_resp_comb) - (* jump goal *) - apply (simp (no_asm_simp) only: int_outside_right nat_int, simp (no_asm_simp)) - apply (simp add: length_compTpExpr pushST_def) - apply (simp add: popST_def start_sttp_resp_comb length_compTpStmt) - apply (simp only: compTpStmt_LT_ST) - apply (simp add: nochangeST_def) - (* check_type *) - apply (subgoal_tac " - (mt_sttp_flatten (f' (ST, LT)) ! length ([LitPush (Bool False)] @ compExpr jmb expr)) = - (Some (PrimT Boolean # PrimT Boolean # ST, LT))") - apply (simp only:) - apply (simp (no_asm_simp)) apply (rule trans, rule mt_sttp_flatten_comb_length) - apply (rule HOL.refl) apply (simp (no_asm_simp) add: length_compTpExpr) - apply (simp (no_asm_simp) add: length_compTpExpr pushST_def) - apply (simp only: compTpExpr_LT_ST_rewr) - (* contracting\ *) - apply (rule contracting_popST) - -apply (drule_tac - ?bc1.0="[LitPush (Bool False)] @ compExpr jmb expr @ - [Ifcmpeq (2 + int (length (compStmt jmb stmt)))] " - and ?bc2.0 = "compStmt jmb stmt" - and ?bc3.0="[Goto (-2 + (- int (length (compStmt jmb stmt)) - int (length (compExpr jmb expr))))]" - and ?f1.0="pushST [PrimT Boolean] \ compTpExpr jmb G expr \ popST 2" - and ?f2.0 = "compTpStmt jmb G stmt" - and ?f3.0="nochangeST" - in bc_mt_corresp_comb_inside) - apply (simp (no_asm_simp))+ - apply (simp (no_asm_simp) add: pushST_def popST_def compTpExpr_LT_ST) - apply (simp only: compTpExpr_LT_ST) - apply (simp (no_asm_simp)) - apply (simp (no_asm_simp) add: length_compTpExpr)+ - - -apply (drule_tac ?bc1.0 = "[LitPush (Bool False)] @ compExpr jmb expr @ [Ifcmpeq (2 + int (length (compStmt jmb stmt)))] @ compStmt jmb stmt" - and inst = "Goto (-2 + (- int (length (compStmt jmb stmt)) - int (length (compExpr jmb expr))))" - and ?bc3.0 = "[]" - and ?f1.0="pushST [PrimT Boolean] \ compTpExpr jmb G expr \ popST 2 \ - compTpStmt jmb G stmt" - and ?f2.0 = "nochangeST" - and ?f3.0="comb_nil" - in bc_mt_corresp_comb_wt_instr) - apply (simp (no_asm_simp) add: length_compTpExpr length_compTpStmt)+ - apply (intro strip) - apply (rule wt_instr_Goto) - apply arith - apply arith - (* \ ! nat (int pc + i) = \ ! pc *) - apply (simp (no_asm_simp)) - apply (simp (no_asm_simp) add: length_compTpExpr length_compTpStmt) - apply (simp (no_asm_simp) add: pushST_def popST_def nochangeST_def) - apply (simp only: compTpExpr_LT_ST compTpStmt_LT_ST) - apply (simp (no_asm_simp) add: length_compTpExpr length_compTpStmt) - apply (simp only: compTpExpr_LT_ST compTpStmt_LT_ST) - apply (simp (no_asm_simp) add: pushST_def popST_def nochangeST_def) - apply (simp (no_asm_simp) add: length_compTpExpr length_compTpStmt) - apply (simp only: compTpExpr_LT_ST compTpStmt_LT_ST) - - apply (simp add: length_compTpExpr length_compTpStmt) (* check_type *) + apply (simp only: compTpExpr_LT_ST_rewr) + (* contracting\ *) + apply (rule contracting_popST) + + apply (drule_tac + ?bc1.0="[LitPush (Bool False)] @ compExpr jmb expr @ + [Ifcmpeq (2 + int (length (compStmt jmb stmt)))] " + and ?bc2.0 = "compStmt jmb stmt" + and ?bc3.0="[Goto (-2 + (- int (length (compStmt jmb stmt)) - int (length (compExpr jmb expr))))]" + and ?f1.0="pushST [PrimT Boolean] \ compTpExpr jmb G expr \ popST 2" + and ?f2.0 = "compTpStmt jmb G stmt" + and ?f3.0="nochangeST" + in bc_mt_corresp_comb_inside) + apply (simp (no_asm_simp))+ + apply (simp (no_asm_simp) add: pushST_def popST_def compTpExpr_LT_ST) + apply (simp only: compTpExpr_LT_ST) + apply (simp (no_asm_simp)) + apply (simp (no_asm_simp) add: length_compTpExpr)+ + + + apply (drule_tac ?bc1.0 = "[LitPush (Bool False)] @ compExpr jmb expr @ [Ifcmpeq (2 + int (length (compStmt jmb stmt)))] @ compStmt jmb stmt" + and inst = "Goto (-2 + (- int (length (compStmt jmb stmt)) - int (length (compExpr jmb expr))))" + and ?bc3.0 = "[]" + and ?f1.0="pushST [PrimT Boolean] \ compTpExpr jmb G expr \ popST 2 \ compTpStmt jmb G stmt" + and ?f2.0 = "nochangeST" + and ?f3.0="comb_nil" + in bc_mt_corresp_comb_wt_instr) + apply (simp (no_asm_simp) add: length_compTpExpr length_compTpStmt)+ + apply (intro strip) + apply (rule wt_instr_Goto) + apply arith + apply arith + (* \ ! nat (int pc + i) = \ ! pc *) + apply (simp (no_asm_simp)) + apply (simp (no_asm_simp) add: length_compTpExpr length_compTpStmt) + apply (simp (no_asm_simp) add: pushST_def popST_def nochangeST_def) + apply (simp only: compTpExpr_LT_ST compTpStmt_LT_ST) + apply (simp (no_asm_simp) add: length_compTpExpr length_compTpStmt) + apply (simp only: compTpExpr_LT_ST compTpStmt_LT_ST) + apply (simp (no_asm_simp) add: pushST_def popST_def nochangeST_def) + apply (simp (no_asm_simp) add: length_compTpExpr length_compTpStmt) + apply (simp only: compTpExpr_LT_ST compTpStmt_LT_ST) + + apply (simp add: length_compTpExpr length_compTpStmt) (* check_type *) apply (simp add: pushST_def popST_def compTpExpr_LT_ST compTpStmt_LT_ST) - apply (rule contracting_nochangeST) -apply simp - -done + apply (rule contracting_nochangeST) + apply simp + + done (**********************************************************************************) @@ -2239,21 +2251,21 @@ bc = (compInit jmb (vn,ty)); f = (compTpInit jmb (vn,ty)); is_type G ty \ \ bc_mt_corresp bc f (ST, LT) (comp G) rT mxr (length bc)" -apply (simp add: compInit_def compTpInit_def split_beta) -apply (rule_tac ?bc1.0="[load_default_val ty]" and ?bc2.0="[Store (index jmb vn)]" - in bc_mt_corresp_comb) - apply simp+ - apply (simp add: load_default_val_def) - apply (rule typeof_default_val [THEN exE]) - - apply (rule bc_mt_corresp_LitPush_CT) apply assumption + apply (simp add: compInit_def compTpInit_def split_beta) + apply (rule_tac ?bc1.0="[load_default_val ty]" and ?bc2.0="[Store (index jmb vn)]" + in bc_mt_corresp_comb) + apply simp+ + apply (simp add: load_default_val_def) + apply (rule typeof_default_val [THEN exE]) + + apply (rule bc_mt_corresp_LitPush_CT, assumption) apply (simp add: comp_is_type) -apply (simp add: pushST_def) - apply (rule bc_mt_corresp_Store_init) - apply simp - apply (rule index_length_lvars [THEN conjunct2]) -apply auto -done + apply (simp add: pushST_def) + apply (rule bc_mt_corresp_Store_init) + apply simp + apply (rule index_length_lvars [THEN conjunct2]) + apply auto + done lemma wt_method_compTpInitLvars_corresp_aux [rule_format (no_asm)]: " @@ -2264,27 +2276,28 @@ wf_java_mdecl G C ((mn, pTs), rT, jmb) \ bc_mt_corresp (compInitLvars jmb lvars) (compTpInitLvars jmb lvars) (ST, LT) (comp G) rT (length LT) (length (compInitLvars jmb lvars))" -apply (induct lvars) -apply (simp add: compInitLvars_def) - -apply (intro strip, (erule conjE)+) -apply (subgoal_tac "\ vn ty. a = (vn, ty)") - prefer 2 apply (simp (no_asm_simp)) + apply (induct lvars) + apply (simp add: compInitLvars_def) + + apply (intro strip, (erule conjE)+) + apply (subgoal_tac "\ vn ty. a = (vn, ty)") + prefer 2 + apply (simp (no_asm_simp)) apply ((erule exE)+, simp (no_asm_simp)) -apply (drule_tac x="lvars_pre @ [a]" in spec) -apply (drule_tac x="lvars0" in spec) -apply (simp (no_asm_simp) add: compInitLvars_def) -apply (rule_tac ?bc1.0="compInit jmb a" and ?bc2.0="compInitLvars jmb lvars" - in bc_mt_corresp_comb) -apply (simp (no_asm_simp) add: compInitLvars_def)+ - -apply (rule_tac vn=vn and ty=ty in wt_method_compTpInit_corresp) -apply assumption+ -apply (simp (no_asm_simp))+ -apply (simp add: wf_java_mdecl_def) (* is_type G ty *) -apply (simp add: compTpInit_def storeST_def pushST_def) -apply simp -done + apply (drule_tac x="lvars_pre @ [a]" in spec) + apply (drule_tac x="lvars0" in spec) + apply (simp (no_asm_simp) add: compInitLvars_def) + apply (rule_tac ?bc1.0="compInit jmb a" and ?bc2.0="compInitLvars jmb lvars" + in bc_mt_corresp_comb) + apply (simp (no_asm_simp) add: compInitLvars_def)+ + + apply (rule_tac vn=vn and ty=ty in wt_method_compTpInit_corresp) + apply assumption+ + apply (simp (no_asm_simp))+ + apply (simp add: wf_java_mdecl_def) (* is_type G ty *) + apply (simp add: compTpInit_def storeST_def pushST_def) + apply simp + done lemma wt_method_compTpInitLvars_corresp: "\ jmb = (pns,lvars,blk,res); @@ -2292,14 +2305,14 @@ length LT = (length pns) + (length lvars) + 1; mxr = (length LT); bc = (compInitLvars jmb lvars); f= (compTpInitLvars jmb lvars) \ \ bc_mt_corresp bc f (ST, LT) (comp G) rT mxr (length bc)" -apply (simp only:) -apply (subgoal_tac "bc_mt_corresp (compInitLvars (pns, lvars, blk, res) lvars) - (compTpInitLvars (pns, lvars, blk, res) lvars) (ST, LT) (TranslComp.comp G) rT - (length LT) (length (compInitLvars (pns, lvars, blk, res) lvars))") -apply simp -apply (rule_tac lvars_pre="[]" in wt_method_compTpInitLvars_corresp_aux) -apply auto -done + apply (simp only:) + apply (subgoal_tac "bc_mt_corresp (compInitLvars (pns, lvars, blk, res) lvars) + (compTpInitLvars (pns, lvars, blk, res) lvars) (ST, LT) (TranslComp.comp G) rT + (length LT) (length (compInitLvars (pns, lvars, blk, res) lvars))") + apply simp + apply (rule_tac lvars_pre="[]" in wt_method_compTpInitLvars_corresp_aux) + apply auto + done (**********************************************************************************) @@ -2315,69 +2328,71 @@ li = (length (inited_LT C pTs lvars)) \ \ bc_mt_corresp bc f sttp (comp G) rT li (length bc)" -apply (subgoal_tac "\ E. (E = (local_env G C (mn, pTs) pns lvars) \ E \ blk \ \ - (\T. E\res::T \ G\T\rT))") + apply (subgoal_tac "\E. (E = (local_env G C (mn, pTs) pns lvars) \ E \ blk \ \ + (\T. E\res::T \ G\T\rT))") apply (erule exE, (erule conjE)+)+ -apply (simp only:) -apply (rule bc_mt_corresp_comb) apply (rule HOL.refl)+ - - (* InitLvars *) -apply (rule wt_method_compTpInitLvars_corresp) - apply assumption+ apply (simp only:) - apply (simp (no_asm_simp) add: start_LT_def) - apply (rule wf_java_mdecl_length_pTs_pns, assumption) - apply (simp (no_asm_simp) only: start_LT_def) - apply (simp (no_asm_simp) add: inited_LT_def)+ - -apply (rule bc_mt_corresp_comb) apply (rule HOL.refl)+ - apply (simp (no_asm_simp) add: compTpInitLvars_LT_ST) - - (* stmt *) -apply (simp only: compTpInitLvars_LT_ST) -apply (subgoal_tac "(Suc (length pTs + length lvars)) = (length (inited_LT C pTs lvars))") - prefer 2 apply (simp (no_asm_simp) add: inited_LT_def) -apply (simp only:) -apply (rule_tac s=blk in wt_method_compTpStmt_corresp) - apply assumption+ - apply (simp only:)+ - apply (simp (no_asm_simp) add: is_inited_LT_def) - apply (simp only:)+ + apply (rule bc_mt_corresp_comb, (rule HOL.refl)+) + + (* InitLvars *) + apply (rule wt_method_compTpInitLvars_corresp) + apply assumption+ + apply (simp only:) + apply (simp (no_asm_simp) add: start_LT_def) + apply (rule wf_java_mdecl_length_pTs_pns, assumption) + apply (simp (no_asm_simp) only: start_LT_def) + apply (simp (no_asm_simp) add: inited_LT_def)+ + + apply (rule bc_mt_corresp_comb, (rule HOL.refl)+) + apply (simp (no_asm_simp) add: compTpInitLvars_LT_ST) + + (* stmt *) + apply (simp only: compTpInitLvars_LT_ST) + apply (subgoal_tac "(Suc (length pTs + length lvars)) = (length (inited_LT C pTs lvars))") + prefer 2 apply (simp (no_asm_simp) add: inited_LT_def) + apply (simp only:) + apply (rule_tac s=blk in wt_method_compTpStmt_corresp) + apply assumption+ + apply (simp only:)+ + apply (simp (no_asm_simp) add: is_inited_LT_def) + apply (simp only:)+ (* expr *) -apply (simp only: compTpInitLvars_LT_ST compTpStmt_LT_ST is_inited_LT_def) -apply (subgoal_tac "(Suc (length pTs + length lvars)) = (length (inited_LT C pTs lvars))") - prefer 2 apply (simp (no_asm_simp) add: inited_LT_def) -apply (simp only:) -apply (rule_tac ex=res in wt_method_compTpExpr_corresp) - apply assumption+ - apply (simp only:)+ - apply (simp (no_asm_simp) add: is_inited_LT_def) - apply (simp only:)+ - - (* start_sttp_resp *) -apply (simp add: start_sttp_resp_comb)+ + apply (simp only: compTpInitLvars_LT_ST compTpStmt_LT_ST is_inited_LT_def) + apply (subgoal_tac "(Suc (length pTs + length lvars)) = (length (inited_LT C pTs lvars))") + prefer 2 apply (simp (no_asm_simp) add: inited_LT_def) + apply (simp only:) + apply (rule_tac ex=res in wt_method_compTpExpr_corresp) + apply assumption+ + apply (simp only:)+ + apply (simp (no_asm_simp) add: is_inited_LT_def) + apply (simp only:)+ + + (* start_sttp_resp *) + apply (simp add: start_sttp_resp_comb)+ (* subgoal *) -apply (simp add: wf_java_mdecl_def local_env_def) -done + apply (simp add: wf_java_mdecl_def local_env_def) + done (**********************************************************************************) -lemma check_type_start: "\ wf_mhead cG (mn, pTs) rT; is_class cG C\ +lemma check_type_start: + "\ wf_mhead cG (mn, pTs) rT; is_class cG C\ \ check_type cG (length start_ST) (Suc (length pTs + mxl)) (OK (Some (start_ST, start_LT C pTs mxl)))" -apply (simp add: check_type_def wf_mhead_def start_ST_def start_LT_def) -apply (simp add: check_type_simps) -apply (simp only: list_def) -apply (auto simp: err_def) -done - - -lemma wt_method_comp_aux: "\ bc' = bc @ [Return]; f' = (f \ nochangeST); + apply (simp add: check_type_def wf_mhead_def start_ST_def start_LT_def) + apply (simp add: check_type_simps) + apply (simp only: list_def) + apply (auto simp: err_def) + done + + +lemma wt_method_comp_aux: + "\ bc' = bc @ [Return]; f' = (f \ nochangeST); bc_mt_corresp bc f sttp0 cG rT (1+length pTs+mxl) (length bc); start_sttp_resp_cons f'; sttp0 = (start_ST, start_LT C pTs mxl); @@ -2390,74 +2405,74 @@ (Suc (length bc)) empty_et (length bc) \ \ wt_method_altern cG C pTs rT mxs mxl bc' empty_et (mt_of (f' sttp0))" -apply (subgoal_tac "check_type cG (length start_ST) (Suc (length pTs + mxl)) - (OK (Some (start_ST, start_LT C pTs mxl)))") -apply (subgoal_tac "check_type cG mxs (1+length pTs+mxl) (OK (Some (T # ST, LT)))") - -apply (simp add: wt_method_altern_def) - - (* length (.. f ..) = length bc *) -apply (rule conjI) -apply (simp add: bc_mt_corresp_def split_beta) - - (* check_bounded *) -apply (rule conjI) -apply (simp add: bc_mt_corresp_def split_beta check_bounded_def) -apply (erule conjE)+ -apply (intro strip) -apply (subgoal_tac "pc < (length bc) \ pc = length bc") - apply (erule disjE) - (* case pc < length bc *) - apply (subgoal_tac "(bc' ! pc) = (bc ! pc)") - apply (simp add: wt_instr_altern_def eff_def) + apply (subgoal_tac "check_type cG (length start_ST) (Suc (length pTs + mxl)) + (OK (Some (start_ST, start_LT C pTs mxl)))") + apply (subgoal_tac "check_type cG mxs (1+length pTs+mxl) (OK (Some (T # ST, LT)))") + + apply (simp add: wt_method_altern_def) + + (* length (.. f ..) = length bc *) + apply (rule conjI) + apply (simp add: bc_mt_corresp_def split_beta) + + (* check_bounded *) + apply (rule conjI) + apply (simp add: bc_mt_corresp_def split_beta check_bounded_def) + apply (erule conjE)+ + apply (intro strip) + apply (subgoal_tac "pc < (length bc) \ pc = length bc") + apply (erule disjE) + (* case pc < length bc *) + apply (subgoal_tac "(bc' ! pc) = (bc ! pc)") + apply (simp add: wt_instr_altern_def eff_def) + (* subgoal *) + apply (simp add: nth_append) + (* case pc = length bc *) + apply (subgoal_tac "(bc' ! pc) = Return") + apply (simp add: wt_instr_altern_def) (* subgoal *) - apply (simp add: nth_append) - (* case pc = length bc *) - apply (subgoal_tac "(bc' ! pc) = Return") - apply (simp add: wt_instr_altern_def) - (* subgoal *) + apply (simp add: nth_append) + + (* subgoal pc < length bc \ pc = length bc *) + apply arith + + (* wt_start *) + apply (rule conjI) + apply (simp add: wt_start_def start_sttp_resp_cons_def split_beta) + apply (drule_tac x=sttp0 in spec) apply (erule exE) + apply (simp add: mt_sttp_flatten_def start_ST_def start_LT_def) + + (* wt_instr *) + apply (intro strip) + apply (subgoal_tac "pc < (length bc) \ pc = length bc") + apply (erule disjE) + + (* pc < (length bc) *) + apply (simp (no_asm_use) add: bc_mt_corresp_def mt_sttp_flatten_def split_beta) + apply (erule conjE)+ + apply (drule mp, assumption)+ + apply (erule conjE)+ + apply (drule spec, drule mp, assumption) + apply (simp add: nth_append) + apply (simp (no_asm_simp) add: comb_def split_beta nochangeST_def) + + (* pc = length bc *) + apply (simp add: nth_append) + + (* subgoal pc < (length bc) \ pc = length bc *) + apply arith + + (* subgoals *) + apply (simp (no_asm_use) add: bc_mt_corresp_def split_beta) + apply (subgoal_tac "check_type cG (length (fst sttp0)) (Suc (length pTs + mxl)) + (OK (Some sttp0))") + apply ((erule conjE)+, drule mp, assumption) apply (simp add: nth_append) - - (* subgoal pc < length bc \ pc = length bc *) -apply arith - - (* wt_start *) -apply (rule conjI) -apply (simp add: wt_start_def start_sttp_resp_cons_def split_beta) - apply (drule_tac x=sttp0 in spec) apply (erule exE) - apply (simp add: mt_sttp_flatten_def start_ST_def start_LT_def) - - (* wt_instr *) -apply (intro strip) -apply (subgoal_tac "pc < (length bc) \ pc = length bc") -apply (erule disjE) - - (* pc < (length bc) *) -apply (simp (no_asm_use) add: bc_mt_corresp_def mt_sttp_flatten_def split_beta) -apply (erule conjE)+ -apply (drule mp, assumption)+ -apply (erule conjE)+ -apply (drule spec, drule mp, assumption) -apply (simp add: nth_append) -apply (simp (no_asm_simp) add: comb_def split_beta nochangeST_def) - - (* pc = length bc *) -apply (simp add: nth_append) - - (* subgoal pc < (length bc) \ pc = length bc *) -apply arith - - (* subgoals *) -apply (simp (no_asm_use) add: bc_mt_corresp_def split_beta) -apply (subgoal_tac "check_type cG (length (fst sttp0)) (Suc (length pTs + mxl)) - (OK (Some sttp0))") -apply ((erule conjE)+, drule mp, assumption) -apply (simp add: nth_append) -apply (simp (no_asm_simp) add: comb_def nochangeST_def split_beta) -apply (simp (no_asm_simp)) - -apply (rule check_type_start, assumption+) -done + apply (simp (no_asm_simp) add: comb_def nochangeST_def split_beta) + apply (simp (no_asm_simp)) + + apply (rule check_type_start, assumption+) + done lemma wt_instr_Return: "\fst f ! pc = Some (T # ST, LT); (G \ T \ rT); pc < max_pc; @@ -2465,7 +2480,7 @@ \ \ wt_instr_altern Return (comp G) rT (mt_of f) mxs mxr max_pc empty_et pc" apply (case_tac "(mt_of f ! pc)") - apply (simp add: wt_instr_altern_def eff_def norm_eff_def app_def)+ + apply (simp add: wt_instr_altern_def eff_def norm_eff_def app_def)+ apply (drule sym) apply (simp add: comp_widen xcpt_app_def) done @@ -2479,82 +2494,88 @@ \ wt_method (comp G) C pTs rT mxs mxl bc et mt" (* show statement for wt_method_altern *) -apply (rule wt_method_altern_wt_method) - -apply (subgoal_tac "wf_java_mdecl G C jmdcl") -apply (subgoal_tac "wf_mhead G (mn, pTs) rT") -apply (subgoal_tac "is_class G C") -apply (subgoal_tac "\ jmb. \ pns lvars blk res. jmb = (pns, lvars, blk, res)") - apply (drule_tac x=jmb in spec, (erule exE)+) -apply (subgoal_tac "\ E. (E = (local_env G C (mn, pTs) pns lvars) \ E \ blk \ \ - (\T. E\res::T \ G\T\rT))") - apply (erule exE, (erule conjE)+)+ -apply (simp add: compMethod_def compTpMethod_def split_beta) -apply (rule_tac T=T and LT="inited_LT C pTs lvars" and ST=start_ST in wt_method_comp_aux) - - (* bc *) -apply (simp only: append_assoc [symmetric]) - - (* comb *) -apply (simp only: comb_assoc [symmetric]) - - (* bc_corresp *) -apply (rule wt_method_comp_wo_return) - apply assumption+ - apply (simp (no_asm_use) only: append_assoc) - apply (rule HOL.refl) - apply (simp (no_asm_simp))+ - apply (simp (no_asm_simp) add: inited_LT_def) - - (* start_sttp_resp *) -apply (simp add: start_sttp_resp_cons_comb_cons_r)+ - - (* wf_mhead / is_class *) -apply (simp add: wf_mhead_def comp_is_type) -apply (simp add: comp_is_class) - - (* sttp_of .. = (T # ST, LT) *) -apply (simp (no_asm_simp) add: compTpInitLvars_LT_ST compTpExpr_LT_ST compTpStmt_LT_ST is_inited_LT_def) -apply (subgoal_tac "(snd (compTpInitLvars (pns, lvars, blk, res) lvars - (start_ST, start_LT C pTs (length lvars)))) - = (start_ST, inited_LT C pTs lvars)") - prefer 2 apply (rule compTpInitLvars_LT_ST) apply (rule HOL.refl) apply assumption -apply (subgoal_tac "(snd (compTpStmt (pns, lvars, blk, res) G blk - (start_ST, inited_LT C pTs lvars))) - = (start_ST, inited_LT C pTs lvars)") - prefer 2 apply (erule conjE)+ - apply (rule compTpStmt_LT_ST) apply (rule HOL.refl) apply assumption+ - apply (simp only:)+ apply (simp (no_asm_simp) add: is_inited_LT_def) - apply (simp (no_asm_simp) add: is_inited_LT_def) - - - (* Return *) -apply (intro strip) -apply (rule_tac T=T and ST=start_ST and LT="inited_LT C pTs lvars" in wt_instr_Return) -apply (simp (no_asm_simp) add: nth_append - length_compTpInitLvars length_compTpStmt length_compTpExpr) -apply (simp only: compTpInitLvars_LT_ST compTpStmt_LT_ST compTpExpr_LT_ST - nochangeST_def) -apply (simp only: is_inited_LT_def compTpStmt_LT_ST compTpExpr_LT_ST) -apply (simp (no_asm_simp))+ -apply simp - - (* subgoal \ E. \ *) -apply (simp add: wf_java_mdecl_def local_env_def) - - (* subgoal jmb = (\) *) -apply (simp only: split_paired_All, simp) - - (* subgoal is_class / wf_mhead / wf_java_mdecl *) -apply (blast intro: methd [THEN conjunct2]) -apply (frule wf_prog_wf_mdecl, assumption+) apply (simp only:) apply (simp add: wf_mdecl_def) -apply (rule wf_java_prog_wf_java_mdecl, assumption+) -done + apply (rule wt_method_altern_wt_method) + + apply (subgoal_tac "wf_java_mdecl G C jmdcl") + apply (subgoal_tac "wf_mhead G (mn, pTs) rT") + apply (subgoal_tac "is_class G C") + apply (subgoal_tac "\jmb. \ pns lvars blk res. jmb = (pns, lvars, blk, res)") + apply (drule_tac x=jmb in spec, (erule exE)+) + apply (subgoal_tac "\E. (E = (local_env G C (mn, pTs) pns lvars) \ E \ blk \ \ + (\T. E\res::T \ G\T\rT))") + apply (erule exE, (erule conjE)+)+ + apply (simp add: compMethod_def compTpMethod_def split_beta) + apply (rule_tac T=T and LT="inited_LT C pTs lvars" and ST=start_ST in wt_method_comp_aux) + + (* bc *) + apply (simp only: append_assoc [symmetric]) + + (* comb *) + apply (simp only: comb_assoc [symmetric]) + + (* bc_corresp *) + apply (rule wt_method_comp_wo_return) + apply assumption+ + apply (simp (no_asm_use) only: append_assoc) + apply (rule HOL.refl) + apply (simp (no_asm_simp))+ + apply (simp (no_asm_simp) add: inited_LT_def) + + (* start_sttp_resp *) + apply (simp add: start_sttp_resp_cons_comb_cons_r)+ + + (* wf_mhead / is_class *) + apply (simp add: wf_mhead_def comp_is_type) + apply (simp add: comp_is_class) + + (* sttp_of .. = (T # ST, LT) *) + apply (simp (no_asm_simp) add: compTpInitLvars_LT_ST compTpExpr_LT_ST compTpStmt_LT_ST is_inited_LT_def) + apply (subgoal_tac "(snd (compTpInitLvars (pns, lvars, blk, res) lvars + (start_ST, start_LT C pTs (length lvars)))) + = (start_ST, inited_LT C pTs lvars)") + prefer 2 + apply (rule compTpInitLvars_LT_ST) + apply (rule HOL.refl) + apply assumption + apply (subgoal_tac "(snd (compTpStmt (pns, lvars, blk, res) G blk + (start_ST, inited_LT C pTs lvars))) + = (start_ST, inited_LT C pTs lvars)") + prefer 2 apply (erule conjE)+ + apply (rule compTpStmt_LT_ST) + apply (rule HOL.refl) + apply assumption+ + apply (simp only:)+ + apply (simp (no_asm_simp) add: is_inited_LT_def) + apply (simp (no_asm_simp) add: is_inited_LT_def) + + + (* Return *) + apply (intro strip) + apply (rule_tac T=T and ST=start_ST and LT="inited_LT C pTs lvars" in wt_instr_Return) + apply (simp (no_asm_simp) add: nth_append length_compTpInitLvars length_compTpStmt length_compTpExpr) + apply (simp only: compTpInitLvars_LT_ST compTpStmt_LT_ST compTpExpr_LT_ST nochangeST_def) + apply (simp only: is_inited_LT_def compTpStmt_LT_ST compTpExpr_LT_ST) + apply (simp (no_asm_simp))+ + apply simp + + (* subgoal \ E. \ *) + apply (simp add: wf_java_mdecl_def local_env_def) + + (* subgoal jmb = (\) *) + apply (simp only: split_paired_All, simp) + + (* subgoal is_class / wf_mhead / wf_java_mdecl *) + apply (blast intro: methd [THEN conjunct2]) + apply (frule wf_prog_wf_mdecl, assumption+) + apply (simp only:) + apply (simp add: wf_mdecl_def) + apply (rule wf_java_prog_wf_java_mdecl, assumption+) + done lemma comp_set_ms: "(C, D, fs, cms)\set (comp G) \ \ ms. (C, D, fs, ms) \set G \ cms = map (compMethod G C) ms" -by (auto simp: comp_def compClass_def) + by (auto simp: comp_def compClass_def) (* ---------------------------------------------------------------------- *) @@ -2563,50 +2584,53 @@ (* MAIN THEOREM: Methodtype computed by comp is correct for bytecode generated by compTp *) theorem wt_prog_comp: "wf_java_prog G \ wt_jvm_prog (comp G) (compTp G)" -apply (simp add: wf_prog_def) - -apply (subgoal_tac "wf_java_prog G") prefer 2 apply (simp add: wf_prog_def) -apply (simp (no_asm_simp) add: wf_prog_def wt_jvm_prog_def) -apply (simp add: comp_ws_prog) - -apply (intro strip) -apply (subgoal_tac "\ C D fs cms. c = (C, D, fs, cms)") -apply clarify -apply (frule comp_set_ms) -apply clarify -apply (drule bspec, assumption) -apply (rule conjI) - - (* wf_mrT *) -apply (case_tac "C = Object") -apply (simp add: wf_mrT_def) -apply (subgoal_tac "is_class G D") -apply (simp add: comp_wf_mrT) -apply (simp add: wf_prog_def ws_prog_def ws_cdecl_def) -apply blast - - (* wf_cdecl_mdecl *) -apply (simp add: wf_cdecl_mdecl_def) -apply (simp add: split_beta) -apply (intro strip) - - (* show wt_method \ *) -apply (subgoal_tac "\ sig rT mb. x = (sig, rT, mb)") -apply (erule exE)+ -apply (simp (no_asm_simp) add: compMethod_def split_beta) -apply (erule conjE)+ -apply (drule_tac x="(sig, rT, mb)" in bspec) apply simp -apply (rule_tac mn="fst sig" and pTs="snd sig" in wt_method_comp) - apply assumption+ - apply simp -apply (simp (no_asm_simp) add: compTp_def) -apply (simp (no_asm_simp) add: compMethod_def split_beta) -apply (frule WellForm.methd) apply assumption+ -apply simp -apply simp -apply (simp add: compMethod_def split_beta) -apply auto -done + apply (simp add: wf_prog_def) + + apply (subgoal_tac "wf_java_prog G") + prefer 2 + apply (simp add: wf_prog_def) + apply (simp (no_asm_simp) add: wf_prog_def wt_jvm_prog_def) + apply (simp add: comp_ws_prog) + + apply (intro strip) + apply (subgoal_tac "\C D fs cms. c = (C, D, fs, cms)") + apply clarify + apply (frule comp_set_ms) + apply clarify + apply (drule bspec, assumption) + apply (rule conjI) + + (* wf_mrT *) + apply (case_tac "C = Object") + apply (simp add: wf_mrT_def) + apply (subgoal_tac "is_class G D") + apply (simp add: comp_wf_mrT) + apply (simp add: wf_prog_def ws_prog_def ws_cdecl_def) + apply blast + + (* wf_cdecl_mdecl *) + apply (simp add: wf_cdecl_mdecl_def) + apply (simp add: split_beta) + apply (intro strip) + + (* show wt_method \ *) + apply (subgoal_tac "\sig rT mb. x = (sig, rT, mb)") + apply (erule exE)+ + apply (simp (no_asm_simp) add: compMethod_def split_beta) + apply (erule conjE)+ + apply (drule_tac x="(sig, rT, mb)" in bspec) + apply simp + apply (rule_tac mn="fst sig" and pTs="snd sig" in wt_method_comp) + apply assumption+ + apply simp + apply (simp (no_asm_simp) add: compTp_def) + apply (simp (no_asm_simp) add: compMethod_def split_beta) + apply (frule WellForm.methd) apply assumption+ + apply simp + apply simp + apply (simp add: compMethod_def split_beta) + apply auto + done diff -r 00c06f1315d0 -r 3f429b7d8eb5 src/HOL/MicroJava/Comp/Index.thy --- a/src/HOL/MicroJava/Comp/Index.thy Tue May 26 21:58:04 2015 +0100 +++ b/src/HOL/MicroJava/Comp/Index.thy Thu May 28 17:25:57 2015 +1000 @@ -18,82 +18,82 @@ lemma index_length_pns: " \ i = index (pns,lvars,blk,res) vn; - wf_java_mdecl G C ((mn,pTs),rT, (pns,lvars,blk,res)); - vn \ set pns\ + wf_java_mdecl G C ((mn,pTs),rT, (pns,lvars,blk,res)); + vn \ set pns\ \ 0 < i \ i < Suc (length pns)" -apply (simp add: wf_java_mdecl_def index_def) -apply (subgoal_tac "vn \ This") -apply (auto intro: length_takeWhile) -done + apply (simp add: wf_java_mdecl_def index_def) + apply (subgoal_tac "vn \ This") + apply (auto intro: length_takeWhile) + done lemma index_length_lvars: " \ i = index (pns,lvars,blk,res) vn; - wf_java_mdecl G C ((mn,pTs),rT, (pns,lvars,blk,res)); - vn \ set (map fst lvars)\ + wf_java_mdecl G C ((mn,pTs),rT, (pns,lvars,blk,res)); + vn \ set (map fst lvars)\ \ (length pns) < i \ i < Suc((length pns) + (length lvars))" -apply (simp add: wf_java_mdecl_def index_def) -apply (subgoal_tac "vn \ This") -apply simp -apply (subgoal_tac "\ x \ set pns. (\z. z \ vn) x") -apply simp -apply (subgoal_tac "length (takeWhile (\z. z \ vn) (map fst lvars)) < length (map fst lvars)") -apply simp -apply (rule length_takeWhile) -apply simp -apply (simp add: map_of_in_set) -apply (intro strip notI) apply simp apply blast -done + apply (simp add: wf_java_mdecl_def index_def) + apply (subgoal_tac "vn \ This") + apply simp + apply (subgoal_tac "\ x \ set pns. (\z. z \ vn) x") + apply simp + apply (subgoal_tac "length (takeWhile (\z. z \ vn) (map fst lvars)) < length (map fst lvars)") + apply simp + apply (rule length_takeWhile) + apply simp + apply (simp add: map_of_in_set) + apply (intro strip notI) + apply simp + apply blast + done (*** index ***) lemma select_at_index : - "x \ set (gjmb_plns (gmb G C S)) \ x = This - \ (the (loc This) # glvs (gmb G C S) loc) ! (index (gmb G C S) x) = - the (loc x)" -apply (simp only: index_def gjmb_plns_def) -apply (case_tac "gmb G C S" rule: prod.exhaust) -apply (simp add: galldefs del: set_append map_append) -apply (rename_tac a b) -apply (case_tac b, simp add: gmb_def gjmb_lvs_def del: set_append map_append) -apply (intro strip) -apply (simp del: set_append map_append) -apply (frule length_takeWhile) -apply (frule_tac f = "(the \ loc)" in nth_map) -apply simp -done + "x \ set (gjmb_plns (gmb G C S)) \ x = This + \ (the (loc This) # glvs (gmb G C S) loc) ! (index (gmb G C S) x) = the (loc x)" + apply (simp only: index_def gjmb_plns_def) + apply (case_tac "gmb G C S" rule: prod.exhaust) + apply (simp add: galldefs del: set_append map_append) + apply (rename_tac a b) + apply (case_tac b, simp add: gmb_def gjmb_lvs_def del: set_append map_append) + apply (intro strip) + apply (simp del: set_append map_append) + apply (frule length_takeWhile) + apply (frule_tac f = "(the \ loc)" in nth_map) + apply simp + done lemma lift_if: "(f (if b then t else e)) = (if b then (f t) else (f e))" -apply auto -done + by auto lemma update_at_index: " - \ distinct (gjmb_plns (gmb G C S)); - x \ set (gjmb_plns (gmb G C S)); x \ This \ \ + \ distinct (gjmb_plns (gmb G C S)); + x \ set (gjmb_plns (gmb G C S)); x \ This \ \ locvars_xstate G C S (Norm (h, l))[index (gmb G C S) x := val] = - locvars_xstate G C S (Norm (h, l(x\val)))" -apply (simp only: locvars_xstate_def locvars_locals_def index_def) -apply (case_tac "gmb G C S" rule: prod.exhaust, simp) -apply (rename_tac a b) -apply (case_tac b, simp) -apply (rule conjI) -apply (simp add: gl_def) -apply (simp add: galldefs del: set_append map_append) -done + locvars_xstate G C S (Norm (h, l(x\val)))" + apply (simp only: locvars_xstate_def locvars_locals_def index_def) + apply (case_tac "gmb G C S" rule: prod.exhaust, simp) + apply (rename_tac a b) + apply (case_tac b, simp) + apply (rule conjI) + apply (simp add: gl_def) + apply (simp add: galldefs del: set_append map_append) + done (* !!!! incomprehensible: why can't List.takeWhile_append2 be applied the same way in the second case as in the first case ? *) lemma index_of_var: "\ xvar \ set pns; xvar \ set (map fst zs); xvar \ This \ \ index (pns, zs @ ((xvar, xval) # xys), blk, res) xvar = Suc (length pns + length zs)" -apply (simp add: index_def) -apply (subgoal_tac "(\x. ((x \ (set pns)) \ ((\z. (z \ xvar))x)))") -apply simp -apply (subgoal_tac "(takeWhile (\z. z \ xvar) (map fst zs @ xvar # map fst xys)) = map fst zs @ (takeWhile (\z. z \ xvar) (xvar # map fst xys))") -apply simp -apply (rule List.takeWhile_append2) -apply auto -done + apply (simp add: index_def) + apply (subgoal_tac "(\x. ((x \ (set pns)) \ ((\z. (z \ xvar))x)))") + apply simp + apply (subgoal_tac "(takeWhile (\z. z \ xvar) (map fst zs @ xvar # map fst xys)) = map fst zs @ (takeWhile (\z. z \ xvar) (xvar # map fst xys))") + apply simp + apply (rule List.takeWhile_append2) + apply auto + done @@ -101,12 +101,6 @@ (* The following def should replace the conditions in WellType.thy / wf_java_mdecl *) definition disjoint_varnames :: "[vname list, (vname \ ty) list] \ bool" where -(* This corresponds to the original def in wf_java_mdecl: - "disjoint_varnames pns lvars \ - nodups pns \ unique lvars \ This \ set pns \ This \ set (map fst lvars) \ - (\pn\set pns. map_of lvars pn = None)" -*) - "disjoint_varnames pns lvars \ distinct pns \ unique lvars \ This \ set pns \ This \ set (map fst lvars) \ (\pn\set pns. pn \ set (map fst lvars))" @@ -116,26 +110,26 @@ disjoint_varnames pns (lvars_pre @ (vn, ty) # lvars_post) \ index (pns, lvars_pre @ (vn, ty) # lvars_post, blk, res) vn = Suc (length pns + length lvars_pre)" -apply (simp add: disjoint_varnames_def index_def unique_def) -apply (subgoal_tac "vn \ This", simp) -apply (subgoal_tac - "takeWhile (\z. z \ vn) (map fst lvars_pre @ vn # map fst lvars_post) = - map fst lvars_pre @ takeWhile (\z. z \ vn) (vn # map fst lvars_post)") -apply simp -apply (rule List.takeWhile_append2) -apply auto -done + apply (simp add: disjoint_varnames_def index_def unique_def) + apply (subgoal_tac "vn \ This", simp) + apply (subgoal_tac + "takeWhile (\z. z \ vn) (map fst lvars_pre @ vn # map fst lvars_post) = + map fst lvars_pre @ takeWhile (\z. z \ vn) (vn # map fst lvars_post)") + apply simp + apply (rule List.takeWhile_append2) + apply auto + done lemma wf_java_mdecl_disjoint_varnames: - "wf_java_mdecl G C (S,rT,(pns,lvars,blk,res)) + "wf_java_mdecl G C (S,rT,(pns,lvars,blk,res)) \ disjoint_varnames pns lvars" -apply (cases S) -apply (simp add: wf_java_mdecl_def disjoint_varnames_def map_of_in_set) -done + apply (cases S) + apply (simp add: wf_java_mdecl_def disjoint_varnames_def map_of_in_set) + done lemma wf_java_mdecl_length_pTs_pns: "wf_java_mdecl G C ((mn, pTs), rT, pns, lvars, blk, res) \ length pTs = length pns" -by (simp add: wf_java_mdecl_def) + by (simp add: wf_java_mdecl_def) end diff -r 00c06f1315d0 -r 3f429b7d8eb5 src/HOL/MicroJava/Comp/LemmasComp.thy --- a/src/HOL/MicroJava/Comp/LemmasComp.thy Tue May 26 21:58:04 2015 +0100 +++ b/src/HOL/MicroJava/Comp/LemmasComp.thy Thu May 28 17:25:57 2015 +1000 @@ -9,6 +9,9 @@ begin +context +begin + declare split_paired_All [simp del] declare split_paired_Ex [simp del] @@ -16,29 +19,19 @@ (**********************************************************************) (* misc lemmas *) - - -lemma split_pairs: "(\(a,b). (F a b)) (ab) = F (fst ab) (snd ab)" -proof - - have "(\(a,b). (F a b)) (fst ab,snd ab) = F (fst ab) (snd ab)" - by (simp add: split_def) - then show ?thesis by simp -qed - - lemma c_hupd_conv: "c_hupd h' (xo, (h,l)) = (xo, (if xo = None then h' else h),l)" -by (simp add: c_hupd_def) + by (simp add: c_hupd_def) lemma gl_c_hupd [simp]: "(gl (c_hupd h xs)) = (gl xs)" -by (simp add: gl_def c_hupd_def split_pairs) + by (simp add: gl_def c_hupd_def split_beta) lemma c_hupd_xcpt_invariant [simp]: "gx (c_hupd h' (xo, st)) = xo" -by (cases st) (simp only: c_hupd_conv gx_conv) + by (cases st) (simp only: c_hupd_conv gx_conv) (* not added to simpset because of interference with c_hupd_conv *) lemma c_hupd_hp_invariant: "gh (c_hupd hp (None, st)) = hp" -by (cases st) (simp add: c_hupd_conv gh_def) + by (cases st) (simp add: c_hupd_conv gh_def) lemma unique_map_fst [rule_format]: "(\ x \ set xs. (fst x = fst (f x) )) \ @@ -73,10 +66,10 @@ qed lemma comp_unique: "unique (comp G) = unique G" -apply (simp add: comp_def) -apply (rule unique_map_fst) -apply (simp add: compClass_def split_beta) -done + apply (simp add: comp_def) + apply (rule unique_map_fst) + apply (simp add: compClass_def split_beta) + done (**********************************************************************) @@ -85,58 +78,59 @@ lemma comp_class_imp: "(class G C = Some(D, fs, ms)) \ (class (comp G) C = Some(D, fs, map (compMethod G C) ms))" -apply (simp add: class_def comp_def compClass_def) -apply (rule HOL.trans) -apply (rule map_of_map2) -apply auto -done + apply (simp add: class_def comp_def compClass_def) + apply (rule trans) + apply (rule map_of_map2) + apply auto + done lemma comp_class_None: "(class G C = None) = (class (comp G) C = None)" -apply (simp add: class_def comp_def compClass_def) -apply (simp add: map_of_in_set) -apply (simp add: image_comp [symmetric] o_def split_beta) -done + apply (simp add: class_def comp_def compClass_def) + apply (simp add: map_of_in_set) + apply (simp add: image_comp [symmetric] o_def split_beta) + done lemma comp_is_class: "is_class (comp G) C = is_class G C" -by (cases "class G C") (auto simp: is_class_def comp_class_None dest: comp_class_imp) + by (cases "class G C") (auto simp: is_class_def comp_class_None dest: comp_class_imp) lemma comp_is_type: "is_type (comp G) T = is_type G T" - apply (cases T) - apply simp + apply (cases T, simp) apply (induct G) - apply simp - apply (simp only: comp_is_class) + apply simp + apply (simp only: comp_is_class) apply (simp add: comp_is_class) apply (simp only: comp_is_class) done -lemma comp_classname: "is_class G C - \ fst (the (class G C)) = fst (the (class (comp G) C))" -by (cases "class G C") (auto simp: is_class_def dest: comp_class_imp) +lemma comp_classname: + "is_class G C \ fst (the (class G C)) = fst (the (class (comp G) C))" + by (cases "class G C") (auto simp: is_class_def dest: comp_class_imp) lemma comp_subcls1: "subcls1 (comp G) = subcls1 G" -by (auto simp add: subcls1_def2 comp_classname comp_is_class) + by (auto simp add: subcls1_def2 comp_classname comp_is_class) lemma comp_widen: "widen (comp G) = widen G" apply (simp add: fun_eq_iff) apply (intro allI iffI) - apply (erule widen.cases) - apply (simp_all add: comp_subcls1 widen.null) - apply (erule widen.cases) - apply (simp_all add: comp_subcls1 widen.null) + apply (erule widen.cases) + apply (simp_all add: comp_subcls1 widen.null) + apply (erule widen.cases) + apply (simp_all add: comp_subcls1 widen.null) done lemma comp_cast: "cast (comp G) = cast G" apply (simp add: fun_eq_iff) apply (intro allI iffI) - apply (erule cast.cases) - apply (simp_all add: comp_subcls1 cast.widen cast.subcls) - apply (rule cast.widen) apply (simp add: comp_widen) + apply (erule cast.cases) + apply (simp_all add: comp_subcls1 cast.widen cast.subcls) + apply (rule cast.widen) + apply (simp add: comp_widen) apply (erule cast.cases) - apply (simp_all add: comp_subcls1 cast.widen cast.subcls) - apply (rule cast.widen) apply (simp add: comp_widen) + apply (simp_all add: comp_subcls1 cast.widen cast.subcls) + apply (rule cast.widen) + apply (simp add: comp_widen) done lemma comp_cast_ok: "cast_ok (comp G) = cast_ok G" @@ -144,129 +138,127 @@ lemma compClass_fst [simp]: "(fst (compClass G C)) = (fst C)" -by (simp add: compClass_def split_beta) + by (simp add: compClass_def split_beta) lemma compClass_fst_snd [simp]: "(fst (snd (compClass G C))) = (fst (snd C))" -by (simp add: compClass_def split_beta) + by (simp add: compClass_def split_beta) lemma compClass_fst_snd_snd [simp]: "(fst (snd (snd (compClass G C)))) = (fst (snd (snd C)))" -by (simp add: compClass_def split_beta) + by (simp add: compClass_def split_beta) lemma comp_wf_fdecl [simp]: "wf_fdecl (comp G) fd = wf_fdecl G fd" -by (cases fd) (simp add: wf_fdecl_def comp_is_type) + by (cases fd) (simp add: wf_fdecl_def comp_is_type) -lemma compClass_forall [simp]: " - (\x\set (snd (snd (snd (compClass G C)))). P (fst x) (fst (snd x))) = +lemma compClass_forall [simp]: + "(\x\set (snd (snd (snd (compClass G C)))). P (fst x) (fst (snd x))) = (\x\set (snd (snd (snd C))). P (fst x) (fst (snd x)))" -by (simp add: compClass_def compMethod_def split_beta) + by (simp add: compClass_def compMethod_def split_beta) lemma comp_wf_mhead: "wf_mhead (comp G) S rT = wf_mhead G S rT" -by (simp add: wf_mhead_def split_beta comp_is_type) + by (simp add: wf_mhead_def split_beta comp_is_type) -lemma comp_ws_cdecl: " - ws_cdecl (TranslComp.comp G) (compClass G C) = ws_cdecl G C" -apply (simp add: ws_cdecl_def split_beta comp_is_class comp_subcls1) -apply (simp (no_asm_simp) add: comp_wf_mhead) -apply (simp add: compClass_def compMethod_def split_beta unique_map_fst) -done +lemma comp_ws_cdecl: + "ws_cdecl (TranslComp.comp G) (compClass G C) = ws_cdecl G C" + apply (simp add: ws_cdecl_def split_beta comp_is_class comp_subcls1) + apply (simp (no_asm_simp) add: comp_wf_mhead) + apply (simp add: compClass_def compMethod_def split_beta unique_map_fst) + done lemma comp_wf_syscls: "wf_syscls (comp G) = wf_syscls G" -apply (simp add: wf_syscls_def comp_def compClass_def split_beta) -apply (simp add: image_comp) -apply (subgoal_tac "(Fun.comp fst (\(C, cno::cname, fdls::fdecl list, jmdls). - (C, cno, fdls, map (compMethod G C) jmdls))) = fst") -apply simp -apply (simp add: fun_eq_iff split_beta) -done + apply (simp add: wf_syscls_def comp_def compClass_def split_beta) + apply (simp add: image_comp) + apply (subgoal_tac "(Fun.comp fst (\(C, cno::cname, fdls::fdecl list, jmdls). + (C, cno, fdls, map (compMethod G C) jmdls))) = fst") + apply simp + apply (simp add: fun_eq_iff split_beta) + done lemma comp_ws_prog: "ws_prog (comp G) = ws_prog G" -apply (rule sym) -apply (simp add: ws_prog_def comp_ws_cdecl comp_unique) -apply (simp add: comp_wf_syscls) -apply (auto simp add: comp_ws_cdecl [symmetric] TranslComp.comp_def) -done + apply (rule sym) + apply (simp add: ws_prog_def comp_ws_cdecl comp_unique) + apply (simp add: comp_wf_syscls) + apply (auto simp add: comp_ws_cdecl [symmetric] TranslComp.comp_def) + done -lemma comp_class_rec: " wf ((subcls1 G)^-1) \ -class_rec (comp G) C t f = +lemma comp_class_rec: + "wf ((subcls1 G)^-1) \ + class_rec (comp G) C t f = class_rec G C t (\ C' fs' ms' r'. f C' fs' (map (compMethod G C') ms') r')" -apply (rule_tac a = C in wf_induct) apply assumption -apply (subgoal_tac "wf ((subcls1 (comp G))^-1)") -apply (subgoal_tac "(class G x = None) \ (\ D fs ms. (class G x = Some (D, fs, ms)))") -apply (erule disjE) - - (* case class G x = None *) -apply (simp (no_asm_simp) add: class_rec_def comp_subcls1 - wfrec [where R="(subcls1 G)^-1", simplified]) -apply (simp add: comp_class_None) + apply (rule_tac a = C in wf_induct) + apply assumption + apply (subgoal_tac "wf ((subcls1 (comp G))^-1)") + apply (subgoal_tac "(class G x = None) \ (\ D fs ms. (class G x = Some (D, fs, ms)))") + apply (erule disjE) - (* case \ D fs ms. (class G x = Some (D, fs, ms)) *) -apply (erule exE)+ -apply (frule comp_class_imp) -apply (frule_tac G="comp G" and C=x and t=t and f=f in class_rec_lemma) - apply assumption -apply (frule_tac G=G and C=x and t=t - and f="(\C' fs' ms'. f C' fs' (map (compMethod G C') ms'))" in class_rec_lemma) - apply assumption -apply (simp only:) + (* case class G x = None *) + apply (simp (no_asm_simp) add: class_rec_def comp_subcls1 + wfrec [where R="(subcls1 G)^-1", simplified]) + apply (simp add: comp_class_None) -apply (case_tac "x = Object") - apply simp - apply (frule subcls1I, assumption) + (* case \ D fs ms. (class G x = Some (D, fs, ms)) *) + apply (erule exE)+ + apply (frule comp_class_imp) + apply (frule_tac G="comp G" and C=x and t=t and f=f in class_rec_lemma) + apply assumption + apply (frule_tac G=G and C=x and t=t + and f="(\C' fs' ms'. f C' fs' (map (compMethod G C') ms'))" in class_rec_lemma) + apply assumption + apply (simp only:) + apply (case_tac "x = Object") + apply simp + apply (frule subcls1I, assumption) apply (drule_tac x=D in spec, drule mp, simp) apply simp - (* subgoals *) -apply (case_tac "class G x") -apply auto -apply (simp add: comp_subcls1) -done + (* subgoals *) + apply (case_tac "class G x") + apply auto + apply (simp add: comp_subcls1) + done lemma comp_fields: "wf ((subcls1 G)^-1) \ fields (comp G,C) = fields (G,C)" -by (simp add: fields_def comp_class_rec) + by (simp add: fields_def comp_class_rec) lemma comp_field: "wf ((subcls1 G)^-1) \ field (comp G,C) = field (G,C)" -by (simp add: TypeRel.field_def comp_fields) - + by (simp add: TypeRel.field_def comp_fields) lemma class_rec_relation [rule_format (no_asm)]: "\ ws_prog G; - \ fs ms. R (f1 Object fs ms t1) (f2 Object fs ms t2); - \ C fs ms r1 r2. (R r1 r2) \ (R (f1 C fs ms r1) (f2 C fs ms r2)) \ - \ ((class G C) \ None) \ - R (class_rec G C t1 f1) (class_rec G C t2 f2)" -apply (frule wf_subcls1) (* establish wf ((subcls1 G)^-1) *) -apply (rule_tac a = C in wf_induct) apply assumption -apply (intro strip) -apply (subgoal_tac "(\D rT mb. class G x = Some (D, rT, mb))") - apply (erule exE)+ -apply (frule_tac C=x and t=t1 and f=f1 in class_rec_lemma) - apply assumption -apply (frule_tac C=x and t=t2 and f=f2 in class_rec_lemma) - apply assumption -apply (simp only:) + \fs ms. R (f1 Object fs ms t1) (f2 Object fs ms t2); + \C fs ms r1 r2. (R r1 r2) \ (R (f1 C fs ms r1) (f2 C fs ms r2)) \ + \ ((class G C) \ None) \ R (class_rec G C t1 f1) (class_rec G C t2 f2)" + apply (frule wf_subcls1) (* establish wf ((subcls1 G)^-1) *) + apply (rule_tac a = C in wf_induct) + apply assumption + apply (intro strip) + apply (subgoal_tac "(\D rT mb. class G x = Some (D, rT, mb))") + apply (erule exE)+ + apply (frule_tac C=x and t=t1 and f=f1 in class_rec_lemma) + apply assumption + apply (frule_tac C=x and t=t2 and f=f2 in class_rec_lemma) + apply assumption + apply (simp only:) -apply (case_tac "x = Object") - apply simp - apply (frule subcls1I, assumption) - apply (drule_tac x=D in spec, drule mp, simp) + apply (case_tac "x = Object") apply simp - apply (subgoal_tac "(\D' rT' mb'. class G D = Some (D', rT', mb'))") + apply (frule subcls1I, assumption) + apply (drule_tac x=D in spec, drule mp, simp) + apply simp + apply (subgoal_tac "(\D' rT' mb'. class G D = Some (D', rT', mb'))") apply blast - (* subgoals *) - -apply (frule class_wf_struct) apply assumption -apply (simp add: ws_cdecl_def is_class_def) - -apply (simp add: subcls1_def2 is_class_def) -apply auto -done + (* subgoals *) + apply (frule class_wf_struct, assumption) + apply (simp add: ws_cdecl_def is_class_def) + apply (simp add: subcls1_def2 is_class_def) + apply auto + done abbreviation (input) @@ -276,111 +268,111 @@ "map_of (map (\(k, v). (k, f v)) xs) k = map_option f (map_of xs k)" by (simp add: map_of_map) -lemma map_of_map_fst: "\ inj f; - \x\set xs. fst (f x) = fst x; \x\set xs. fst (g x) = fst x \ - \ map_of (map g xs) k - = map_option (\ e. (snd (g ((inv f) (k, e))))) (map_of (map f xs) k)" -apply (induct xs) -apply simp -apply simp -apply (case_tac "k = fst a") -apply simp -apply (subgoal_tac "(inv f (fst a, snd (f a))) = a", simp) -apply (subgoal_tac "(fst a, snd (f a)) = f a", simp) -apply (erule conjE)+ -apply (drule_tac s ="fst (f a)" and t="fst a" in sym) -apply simp -apply (simp add: surjective_pairing) -done +lemma map_of_map_fst: + "\ inj f; \x\set xs. fst (f x) = fst x; \x\set xs. fst (g x) = fst x \ + \ map_of (map g xs) k = map_option (\ e. (snd (g ((inv f) (k, e))))) (map_of (map f xs) k)" + apply (induct xs) + apply simp + apply simp + apply (case_tac "k = fst a") + apply simp + apply (subgoal_tac "(inv f (fst a, snd (f a))) = a", simp) + apply (subgoal_tac "(fst a, snd (f a)) = f a", simp) + apply (erule conjE)+ + apply (drule_tac s ="fst (f a)" and t="fst a" in sym) + apply simp + apply (simp add: surjective_pairing) + done -lemma comp_method [rule_format (no_asm)]: "\ ws_prog G; is_class G C\ \ +lemma comp_method [rule_format (no_asm)]: + "\ ws_prog G; is_class G C\ \ ((method (comp G, C) S) = - map_option (\ (D,rT,b). (D, rT, mtd_mb (compMethod G D (S, rT, b)))) - (method (G, C) S))" -apply (simp add: method_def) -apply (frule wf_subcls1) -apply (simp add: comp_class_rec) -apply (simp add: split_iter split_compose map_map [symmetric] del: map_map) -apply (rule_tac R="%x y. ((x S) = (map_option (\(D, rT, b). - (D, rT, snd (snd (compMethod G D (S, rT, b))))) (y S)))" - in class_rec_relation) -apply assumption + map_option (\ (D,rT,b). (D, rT, mtd_mb (compMethod G D (S, rT, b)))) + (method (G, C) S))" + apply (simp add: method_def) + apply (frule wf_subcls1) + apply (simp add: comp_class_rec) + apply (simp add: split_iter split_compose map_map [symmetric] del: map_map) + apply (rule_tac R="\x y. ((x S) = (map_option (\(D, rT, b). + (D, rT, snd (snd (compMethod G D (S, rT, b))))) (y S)))" + in class_rec_relation) + apply assumption -apply (intro strip) -apply simp - -apply (rule trans) + apply (intro strip) + apply simp + apply (rule trans) -apply (rule_tac f="(\(s, m). (s, Object, m))" and - g="(Fun.comp (\(s, m). (s, Object, m)) (compMethod G Object))" - in map_of_map_fst) -defer (* inj \ *) -apply (simp add: inj_on_def split_beta) -apply (simp add: split_beta compMethod_def) -apply (simp add: map_of_map [symmetric]) -apply (simp add: split_beta) -apply (simp add: Fun.comp_def split_beta) -apply (subgoal_tac "(\x\(vname list \ fdecl list \ stmt \ expr) mdecl. - (fst x, Object, - snd (compMethod G Object - (inv (\(s\sig, m\ty \ vname list \ fdecl list \ stmt \ expr). - (s, Object, m)) - (S, Object, snd x))))) - = (\x. (fst x, Object, fst (snd x), - snd (snd (compMethod G Object (S, snd x)))))") -apply (simp only:) -apply (simp add: fun_eq_iff) -apply (intro strip) -apply (subgoal_tac "(inv (\(s, m). (s, Object, m)) (S, Object, snd x)) = (S, snd x)") -apply (simp only:) -apply (simp add: compMethod_def split_beta) -apply (rule inv_f_eq) -defer -defer + apply (rule_tac f="(\(s, m). (s, Object, m))" and + g="(Fun.comp (\(s, m). (s, Object, m)) (compMethod G Object))" + in map_of_map_fst) + defer (* inj \ *) + apply (simp add: inj_on_def split_beta) + apply (simp add: split_beta compMethod_def) + apply (simp add: map_of_map [symmetric]) + apply (simp add: split_beta) + apply (simp add: Fun.comp_def split_beta) + apply (subgoal_tac "(\x\(vname list \ fdecl list \ stmt \ expr) mdecl. + (fst x, Object, + snd (compMethod G Object + (inv (\(s\sig, m\ty \ vname list \ fdecl list \ stmt \ expr). + (s, Object, m)) + (S, Object, snd x))))) + = (\x. (fst x, Object, fst (snd x), + snd (snd (compMethod G Object (S, snd x)))))") + apply (simp only:) + apply (simp add: fun_eq_iff) + apply (intro strip) + apply (subgoal_tac "(inv (\(s, m). (s, Object, m)) (S, Object, snd x)) = (S, snd x)") + apply (simp only:) + apply (simp add: compMethod_def split_beta) + apply (rule inv_f_eq) + defer + defer -apply (intro strip) -apply (simp add: map_add_Some_iff map_of_map) -apply (simp add: map_add_def) -apply (subgoal_tac "inj (\(s, m). (s, Ca, m))") -apply (drule_tac g="(Fun.comp (\(s, m). (s, Ca, m)) (compMethod G Ca))" and xs=ms - and k=S in map_of_map_fst) -apply (simp add: split_beta) -apply (simp add: compMethod_def split_beta) -apply (case_tac "(map_of (map (\(s, m). (s, Ca, m)) ms) S)") -apply simp -apply simp apply (simp add: split_beta map_of_map) apply (erule exE) apply (erule conjE)+ -apply (drule_tac t=a in sym) -apply (subgoal_tac "(inv (\(s, m). (s, Ca, m)) (S, a)) = (S, snd a)") -apply simp -apply (subgoal_tac "\x\set ms. fst ((Fun.comp (\(s, m). (s, Ca, m)) (compMethod G Ca)) x) = fst x") - prefer 2 apply (simp (no_asm_simp) add: compMethod_def split_beta) -apply (simp add: map_of_map2) -apply (simp (no_asm_simp) add: compMethod_def split_beta) + apply (intro strip) + apply (simp add: map_add_Some_iff map_of_map) + apply (simp add: map_add_def) + apply (subgoal_tac "inj (\(s, m). (s, Ca, m))") + apply (drule_tac g="(Fun.comp (\(s, m). (s, Ca, m)) (compMethod G Ca))" and xs=ms + and k=S in map_of_map_fst) + apply (simp add: split_beta) + apply (simp add: compMethod_def split_beta) + apply (case_tac "(map_of (map (\(s, m). (s, Ca, m)) ms) S)") + apply simp + apply (simp add: split_beta map_of_map) + apply (elim exE conjE) + apply (drule_tac t=a in sym) + apply (subgoal_tac "(inv (\(s, m). (s, Ca, m)) (S, a)) = (S, snd a)") + apply simp + apply (subgoal_tac "\x\set ms. fst ((Fun.comp (\(s, m). (s, Ca, m)) (compMethod G Ca)) x) = fst x") + prefer 2 apply (simp (no_asm_simp) add: compMethod_def split_beta) + apply (simp add: map_of_map2) + apply (simp (no_asm_simp) add: compMethod_def split_beta) - -- "remaining subgoals" -apply (auto intro: inv_f_eq simp add: inj_on_def is_class_def) -done + -- "remaining subgoals" + apply (auto intro: inv_f_eq simp add: inj_on_def is_class_def) + done lemma comp_wf_mrT: "\ ws_prog G; is_class G D\ \ wf_mrT (TranslComp.comp G) (C, D, fs, map (compMethod G a) ms) = wf_mrT G (C, D, fs, ms)" -apply (simp add: wf_mrT_def compMethod_def split_beta) -apply (simp add: comp_widen) -apply (rule iffI) -apply (intro strip) -apply simp -apply (drule bspec) apply assumption -apply (drule_tac x=D' in spec) apply (drule_tac x=rT' in spec) apply (drule mp) -prefer 2 apply assumption -apply (simp add: comp_method [of G D]) -apply (erule exE)+ -apply (simp add: split_paired_all) -apply (intro strip) -apply (simp add: comp_method) -apply auto -done + apply (simp add: wf_mrT_def compMethod_def split_beta) + apply (simp add: comp_widen) + apply (rule iffI) + apply (intro strip) + apply simp + apply (drule (1) bspec) + apply (drule_tac x=D' in spec) + apply (drule_tac x=rT' in spec) + apply (drule mp) + prefer 2 apply assumption + apply (simp add: comp_method [of G D]) + apply (erule exE)+ + apply (simp add: split_paired_all) + apply (auto simp: comp_method) + done (**********************************************************************) @@ -388,37 +380,32 @@ (**********************************************************************) lemma max_spec_preserves_length: - "max_spec G C (mn, pTs) = {((md,rT),pTs')} - \ length pTs = length pTs'" -apply (frule max_spec2mheads) -apply (erule exE)+ -apply (simp add: list_all2_iff) -done + "max_spec G C (mn, pTs) = {((md,rT),pTs')} \ length pTs = length pTs'" + apply (frule max_spec2mheads) + apply (clarsimp simp: list_all2_iff) + done lemma ty_exprs_length [simp]: "(E\es[::]Ts \ length es = length Ts)" -apply (subgoal_tac "(E\e::T \ True) \ (E\es[::]Ts \ length es = length Ts) \ (E\s\ \ True)") -apply blast -apply (rule ty_expr_ty_exprs_wt_stmt.induct) -apply auto -done + apply (subgoal_tac "(E\e::T \ True) \ (E\es[::]Ts \ length es = length Ts) \ (E\s\ \ True)") + apply blast + apply (rule ty_expr_ty_exprs_wt_stmt.induct, auto) + done lemma max_spec_preserves_method_rT [simp]: "max_spec G C (mn, pTs) = {((md,rT),pTs')} \ method_rT (the (method (G, C) (mn, pTs'))) = rT" -apply (frule max_spec2mheads) -apply (erule exE)+ -apply (simp add: method_rT_def) -done + apply (frule max_spec2mheads) + apply (clarsimp simp: method_rT_def) + done (**********************************************************************************) +end (* context *) + declare compClass_fst [simp del] declare compClass_fst_snd [simp del] declare compClass_fst_snd_snd [simp del] -declare split_paired_All [simp add] -declare split_paired_Ex [simp add] - end diff -r 00c06f1315d0 -r 3f429b7d8eb5 src/HOL/MicroJava/Comp/TranslCompTp.thy --- a/src/HOL/MicroJava/Comp/TranslCompTp.thy Tue May 26 21:58:04 2015 +0100 +++ b/src/HOL/MicroJava/Comp/TranslCompTp.thy Thu May 28 17:25:57 2015 +1000 @@ -20,20 +20,18 @@ comb (infixr "\" 55) lemma comb_nil_left [simp]: "comb_nil \ f = f" -by (simp add: comb_def comb_nil_def split_beta) + by (simp add: comb_def comb_nil_def split_beta) lemma comb_nil_right [simp]: "f \ comb_nil = f" -by (simp add: comb_def comb_nil_def split_beta) + by (simp add: comb_def comb_nil_def split_beta) lemma comb_assoc [simp]: "(fa \ fb) \ fc = fa \ (fb \ fc)" -by (simp add: comb_def split_beta) + by (simp add: comb_def split_beta) -lemma comb_inv: "(xs', x') = (f1 \ f2) x0 \ - \ xs1 x1 xs2 x2. (xs1, x1) = (f1 x0) \ (xs2, x2) = f2 x1 \ xs'= xs1 @ xs2 \ x'=x2" -apply (case_tac "f1 x0") -apply (case_tac "f2 x1") -apply (simp add: comb_def split_beta) -done +lemma comb_inv: + "(xs', x') = (f1 \ f2) x0 \ + \xs1 x1 xs2 x2. (xs1, x1) = (f1 x0) \ (xs2, x2) = f2 x1 \ xs'= xs1 @ xs2 \ x'=x2" + by (case_tac "f1 x0") (simp add: comb_def split_beta) (**********************************************************************) diff -r 00c06f1315d0 -r 3f429b7d8eb5 src/HOL/MicroJava/Comp/TypeInf.thy --- a/src/HOL/MicroJava/Comp/TypeInf.thy Tue May 26 21:58:04 2015 +0100 +++ b/src/HOL/MicroJava/Comp/TypeInf.thy Thu May 28 17:25:57 2015 +1000 @@ -4,7 +4,7 @@ (* Exact position in theory hierarchy still to be determined *) theory TypeInf -imports "../J/WellType" +imports "../J/WellType" "~~/src/HOL/Eisbach/Eisbach" begin @@ -102,6 +102,9 @@ declare split_paired_All [simp del] declare split_paired_Ex [simp del] +method ty_case_simp = ((erule ty_exprs.cases ty_expr.cases; simp)+)[] +method strip_case_simp = (intro strip, ty_case_simp) + (* Uniqueness of types property *) lemma uniqueness_of_types: " @@ -109,83 +112,63 @@ E\e :: T1 \ E\e :: T2 \ T1 = T2) \ (\ (E\'a prog \ (vname \ ty option)) Ts1 Ts2. E\es [::] Ts1 \ E\es [::] Ts2 \ Ts1 = Ts2)" -apply (rule compat_expr_expr_list.induct) + apply (rule compat_expr_expr_list.induct) + (* NewC *) + apply strip_case_simp -(* NewC *) -apply (intro strip) -apply (erule ty_expr.cases) apply simp+ -apply (erule ty_expr.cases) apply simp+ + (* Cast *) + apply strip_case_simp -(* Cast *) -apply (intro strip) -apply (erule ty_expr.cases) apply simp+ -apply (erule ty_expr.cases) apply simp+ + (* Lit *) + apply strip_case_simp -(* Lit *) -apply (intro strip) -apply (erule ty_expr.cases) apply simp+ -apply (erule ty_expr.cases) apply simp+ + (* BinOp *) + apply (intro strip) + apply (rename_tac binop x2 x3 E T1 T2, case_tac binop) + (* Eq *) + apply ty_case_simp + (* Add *) + apply ty_case_simp -(* BinOp *) -apply (intro strip) -apply (rename_tac binop x2 x3 E T1 T2, case_tac binop) -(* Eq *) -apply (erule ty_expr.cases) apply simp+ -apply (erule ty_expr.cases) apply simp+ -(* Add *) -apply (erule ty_expr.cases) apply simp+ -apply (erule ty_expr.cases) apply simp+ - -(* LAcc *) -apply (intro strip) -apply (erule ty_expr.cases) apply simp+ -apply (erule ty_expr.cases) apply simp+ + (* LAcc *) + apply (strip_case_simp) -(* LAss *) -apply (intro strip) -apply (erule ty_expr.cases) apply simp+ -apply (erule ty_expr.cases) apply simp+ - + (* LAss *) + apply (strip_case_simp) -(* FAcc *) -apply (intro strip) -apply (drule FAcc_invers)+ apply (erule exE)+ - apply (subgoal_tac "C = Ca", simp) apply blast - + (* FAcc *) + apply (intro strip) + apply (drule FAcc_invers)+ + apply fastforce -(* FAss *) -apply (intro strip) -apply (drule FAss_invers)+ apply (erule exE)+ apply (erule conjE)+ -apply (drule FAcc_invers)+ apply (erule exE)+ apply blast - + (* FAss *) + apply (intro strip) + apply (drule FAss_invers)+ + apply (elim conjE exE) + apply (drule FAcc_invers)+ + apply fastforce -(* Call *) -apply (intro strip) -apply (drule Call_invers)+ apply (erule exE)+ apply (erule conjE)+ -apply (subgoal_tac "pTs = pTsa", simp) apply blast + (* Call *) + apply (intro strip) + apply (drule Call_invers)+ + apply fastforce -(* expression lists *) -apply (intro strip) -apply (erule ty_exprs.cases)+ apply simp+ + (* expression lists *) + apply (strip_case_simp) -apply (intro strip) -apply (erule ty_exprs.cases, simp) -apply (erule ty_exprs.cases, simp) -apply (subgoal_tac "e = ea", simp) apply simp -done + apply (strip_case_simp) + done lemma uniqueness_of_types_expr [rule_format (no_asm)]: " - (\ E T1 T2. E\e :: T1 \ E\e :: T2 \ T1 = T2)" -by (rule uniqueness_of_types [THEN conjunct1]) + (\E T1 T2. E\e :: T1 \ E\e :: T2 \ T1 = T2)" + by (rule uniqueness_of_types [THEN conjunct1]) lemma uniqueness_of_types_exprs [rule_format (no_asm)]: " - (\ E Ts1 Ts2. E\es [::] Ts1 \ E\es [::] Ts2 \ Ts1 = Ts2)" -by (rule uniqueness_of_types [THEN conjunct2]) + (\E Ts1 Ts2. E\es [::] Ts1 \ E\es [::] Ts2 \ Ts1 = Ts2)" + by (rule uniqueness_of_types [THEN conjunct2]) - - definition inferred_tp :: "[java_mb env, expr] \ ty" where "inferred_tp E e == (SOME T. E\e :: T)" @@ -194,11 +177,9 @@ (* get inferred type(s) for well-typed term *) lemma inferred_tp_wt: "E\e :: T \ (inferred_tp E e) = T" -by (auto simp: inferred_tp_def intro: uniqueness_of_types_expr) + by (auto simp: inferred_tp_def intro: uniqueness_of_types_expr) lemma inferred_tps_wt: "E\es [::] Ts \ (inferred_tps E es) = Ts" -by (auto simp: inferred_tps_def intro: uniqueness_of_types_exprs) - + by (auto simp: inferred_tps_def intro: uniqueness_of_types_exprs) end -