Up to index of Isabelle/HOL/Jinja
theory Progress = Equivalence + WellTypeRT + DefAss + Conform:(* Title: Jinja/J/SmallProgress.thy ID: $Id: Progress.html 1910 2004-05-19 04:46:04Z kleing $ Author: Tobias Nipkow Copyright 2003 Technische Universitaet Muenchen *) header {* \isaheader{Progress of Small Step Semantics} *} theory Progress = Equivalence + WellTypeRT + DefAss + Conform: lemma final_addrE: "[| P,E,h \<turnstile> e : Class C; final e; !!a. e = addr a ==> R; !!a. e = Throw a ==> R |] ==> R" (*<*)by(auto simp:final_def)(*>*) lemma finalRefE: "[| P,E,h \<turnstile> e : T; is_refT T; final e; e = null ==> R; !!a C. [| e = addr a; T = Class C |] ==> R; !!a. e = Throw a ==> R |] ==> R" (*<*)by(auto simp:final_def is_refT_def)(*>*) text{* Derivation of new induction scheme for well typing: *} consts WTrt' :: "J_prog => (env × heap × expr × ty )set" WTrts':: "J_prog => (env × heap × expr list × ty list)set" (*<*) syntax (xsymbols) WTrt' :: "[J_prog,env,heap,expr,ty] => bool" ("_,_,_ \<turnstile> _ :' _" [51,51,51]50) WTrts':: "[J_prog,env,heap,expr list, ty list] => bool" ("_,_,_ \<turnstile> _ [:''] _" [51,51,51]50) (*>*) translations "P,E,h \<turnstile> e :' T" == "(E,h,e,T) ∈ WTrt' P" "P,E,h \<turnstile> es [:'] Ts" == "(E,h,es,Ts) ∈ WTrts' P" inductive "WTrt' P" "WTrts' P" intros "is_class P C ==> P,E,h \<turnstile> new C :' Class C" "[| P,E,h \<turnstile> e :' T; is_refT T; is_class P C |] ==> P,E,h \<turnstile> Cast C e :' Class C" "typeofh v = Some T ==> P,E,h \<turnstile> Val v :' T" "E v = Some T ==> P,E,h \<turnstile> Var v :' T" "[| P,E,h \<turnstile> e1 :' T1; P,E,h \<turnstile> e2 :' T2; case bop of Eq => T' = Boolean | Add => T1 = Integer ∧ T2 = Integer ∧ T' = Integer |] ==> P,E,h \<turnstile> e1 «bop» e2 :' T'" "[| P,E,h \<turnstile> Var V :' T; P,E,h \<turnstile> e :' T'; P \<turnstile> T' ≤ T (* V ≠ This*) |] ==> P,E,h \<turnstile> V:=e :' Void" "[| P,E,h \<turnstile> e :' Class C; P \<turnstile> C has F:T in D |] ==> P,E,h \<turnstile> e\<bullet>F{D} :' T" "P,E,h \<turnstile> e :' NT ==> P,E,h \<turnstile> e\<bullet>F{D} :' T" "[| P,E,h \<turnstile> e1 :' Class C; P \<turnstile> C has F:T in D; P,E,h \<turnstile> e2 :' T2; P \<turnstile> T2 ≤ T |] ==> P,E,h \<turnstile> e1\<bullet>F{D}:=e2 :' Void" "[| P,E,h \<turnstile> e1:'NT; P,E,h \<turnstile> e2 :' T2 |] ==> P,E,h \<turnstile> e1\<bullet>F{D}:=e2 :' Void" "[| P,E,h \<turnstile> e :' Class C; P \<turnstile> C sees M:Ts -> T = (pns,body) in D; P,E,h \<turnstile> es [:'] Ts'; P \<turnstile> Ts' [≤] Ts |] ==> P,E,h \<turnstile> e\<bullet>M(es) :' T" "[| P,E,h \<turnstile> e :' NT; P,E,h \<turnstile> es [:'] Ts |] ==> P,E,h \<turnstile> e\<bullet>M(es) :' T" "P,E,h \<turnstile> [] [:'] []" "[| P,E,h \<turnstile> e :' T; P,E,h \<turnstile> es [:'] Ts |] ==> P,E,h \<turnstile> e#es [:'] T#Ts" "[| typeofh v = Some T1; P \<turnstile> T1 ≤ T; P,E(V\<mapsto>T),h \<turnstile> e2 :' T2 |] ==> P,E,h \<turnstile> {V:T := Val v; e2} :' T2" "[| P,E(V\<mapsto>T),h \<turnstile> e :' T'; ¬ assigned V e |] ==> P,E,h \<turnstile> {V:T; e} :' T'" "[| P,E,h \<turnstile> e1:' T1; P,E,h \<turnstile> e2:'T2 |] ==> P,E,h \<turnstile> e1;;e2 :' T2" "[| P,E,h \<turnstile> e :' Boolean; P,E,h \<turnstile> e1:' T1; P,E,h \<turnstile> e2:' T2; P \<turnstile> T1 ≤ T2 ∨ P \<turnstile> T2 ≤ T1; P \<turnstile> T1 ≤ T2 --> T = T2; P \<turnstile> T2 ≤ T1 --> T = T1 |] ==> P,E,h \<turnstile> if (e) e1 else e2 :' T" "[| P,E,h \<turnstile> e :' Boolean; P,E,h \<turnstile> c:' T |] ==> P,E,h \<turnstile> while(e) c :' Void" "[| P,E,h \<turnstile> e :' Tr; is_refT Tr |] ==> P,E,h \<turnstile> throw e :' T" "[| P,E,h \<turnstile> e1 :' T1; P,E(V \<mapsto> Class C),h \<turnstile> e2 :' T2; P \<turnstile> T1 ≤ T2 |] ==> P,E,h \<turnstile> try e1 catch(C V) e2 :' T2" (*<*) lemmas WTrt'_induct = WTrt'_WTrts'.induct[split_format (complete)] inductive_cases WTrt'_elim_cases[elim!]: "P,E,h \<turnstile> V :=e :' T" (*>*) lemma [iff]: "P,E,h \<turnstile> e1;;e2 :' T2 = (∃T1. P,E,h \<turnstile> e1:' T1 ∧ P,E,h \<turnstile> e2:' T2)" (*<*) apply(rule iffI) apply (auto elim: WTrt'_WTrts'.elims intro!:WTrt'_WTrts'.intros) done (*>*) lemma [iff]: "P,E,h \<turnstile> Val v :' T = (typeofh v = Some T)" (*<*) apply(rule iffI) apply (auto elim: WTrt'_WTrts'.elims intro!:WTrt'_WTrts'.intros) done (*>*) lemma [iff]: "P,E,h \<turnstile> Var v :' T = (E v = Some T)" (*<*) apply(rule iffI) apply (auto elim: WTrt'_WTrts'.elims intro!:WTrt'_WTrts'.intros) done (*>*) lemma wt_wt': "P,E,h \<turnstile> e : T ==> P,E,h \<turnstile> e :' T" and wts_wts': "P,E,h \<turnstile> es [:] Ts ==> P,E,h \<turnstile> es [:'] Ts" (*<*) apply (induct rule:WTrt_induct) prefer 13 apply(case_tac "assigned V e") apply(clarsimp simp add:fun_upd_same assigned_def simp del:fun_upd_apply) apply(erule (2) WTrt'_WTrts'.intros) apply(erule (1) WTrt'_WTrts'.intros) apply(blast intro:WTrt'_WTrts'.intros)+ done (*>*) lemma wt'_wt: "P,E,h \<turnstile> e :' T ==> P,E,h \<turnstile> e : T" and wts'_wts: "P,E,h \<turnstile> es [:'] Ts ==> P,E,h \<turnstile> es [:] Ts" (*<*) apply (induct rule:WTrt'_induct) prefer 15 apply(rule WTrt_WTrts.intros) apply(rule WTrt_WTrts.intros) apply(rule WTrt_WTrts.intros) apply simp apply(erule (2) WTrt_WTrts.intros) apply(blast intro:WTrt_WTrts.intros)+ done (*>*) corollary wt'_iff_wt: "(P,E,h \<turnstile> e :' T) = (P,E,h \<turnstile> e : T)" (*<*)by(blast intro:wt_wt' wt'_wt)(*>*) corollary wts'_iff_wts: "(P,E,h \<turnstile> es [:'] Ts) = (P,E,h \<turnstile> es [:] Ts)" (*<*)by(blast intro:wts_wts' wts'_wts)(*>*) (*<*) lemmas WTrt_induct2 = WTrt'_induct[simplified wt'_iff_wt wts'_iff_wts, case_names WTrtNew WTrtCast WTrtVal WTrtVar WTrtBinOp WTrtLAss WTrtFAcc WTrtFAccNT WTrtFAss WTrtFAssNT WTrtCall WTrtCallNT WTrtNil WTrtCons WTrtInitBlock WTrtBlock WTrtSeq WTrtCond WTrtWhile WTrtThrow WTrtTry] (*>*) theorem assumes wf: "wwf_J_prog P" shows progress: "P,E,h \<turnstile> e : T ==> (!!l. [| P \<turnstile> h \<surd>; \<D> e ⌊dom l⌋; ¬ final e |] ==> ∃e' s'. P \<turnstile> 〈e,(h,l)〉 -> 〈e',s'〉)" and "P,E,h \<turnstile> es [:] Ts ==> (!!l. [| P \<turnstile> h \<surd>; \<D>s es ⌊dom l⌋; ¬ finals es |] ==> ∃es' s'. P \<turnstile> 〈es,(h,l)〉 [->] 〈es',s'〉)" (*<*) proof (induct rule:WTrt_induct2) case (WTrtNew C E h) show ?case proof cases assume "∃a. h a = None" from prems show ?thesis by (fastsimp del:exE intro!:RedNew simp add:new_Addr_def elim!:wf_Fields_Ex[THEN exE]) next assume "¬(∃a. h a = None)" from prems show ?thesis by(fastsimp intro:RedNewFail simp add:new_Addr_def) qed next case (WTrtCast C E T e h) have wte: "P,E,h \<turnstile> e : T" and ref: "is_refT T" and IH: "!!l. [|P \<turnstile> h \<surd>; \<D> e ⌊dom l⌋; ¬ final e|] ==> ∃e' s'. P \<turnstile> 〈e,(h,l)〉 -> 〈e',s'〉" and D: "\<D> (Cast C e) ⌊dom l⌋" and hconf: "P \<turnstile> h \<surd>" . from D have De: "\<D> e ⌊dom l⌋" by auto show ?case proof cases assume "final e" with wte ref show ?thesis proof (rule finalRefE) assume "e = null" thus ?case by(fastsimp intro:RedCastNull) next fix D a assume A: "T = Class D" "e = addr a" show ?thesis proof cases assume "P \<turnstile> D \<preceq>* C" thus ?thesis using A wte by(fastsimp intro:RedCast) next assume "¬ P \<turnstile> D \<preceq>* C" thus ?thesis using A wte by(force intro!:RedCastFail) qed next fix a assume "e = Throw a" thus ?thesis by(blast intro!:red_reds.CastThrow) qed next assume nf: "¬ final e" from IH[OF hconf De nf] show ?thesis by (blast intro:CastRed) qed next case WTrtVal thus ?case by(simp add:final_def) next case WTrtVar thus ?case by(fastsimp intro:RedVar simp:hyper_isin_def) next case (WTrtBinOp E T' T1 T2 bop e1 e2 h) show ?case proof cases assume "final e1" thus ?thesis proof (rule finalE) fix v1 assume [simp]: "e1 = Val v1" show ?thesis proof cases assume "final e2" thus ?thesis proof (rule finalE) fix v2 assume [simp]: "e2 = Val v2" show ?thesis proof (cases bop) assume "bop = Eq" thus ?thesis using WTrtBinOp by(fastsimp intro:RedBinOp) next assume "bop = Add" thus ?thesis using WTrtBinOp by(fastsimp intro:RedBinOp) qed next fix a assume "e2 = Throw a" thus ?thesis by(auto intro:red_reds.BinOpThrow2) qed next assume "¬ final e2" from prems show ?thesis by simp (fast intro!:BinOpRed2) qed next fix a assume "e1 = Throw a" thus ?thesis by simp (fast intro:red_reds.BinOpThrow1) qed next assume "¬ final e1" from prems show ?thesis by simp (fast intro:BinOpRed1) qed next case (WTrtLAss E T T' V e) show ?case proof cases assume "final e" from prems show ?thesis by(auto simp:final_def intro!:RedLAss red_reds.LAssThrow) next assume "¬ final e" from prems show ?thesis by simp (fast intro:LAssRed) qed next case (WTrtFAcc C D E F T e h) have wte: "P,E,h \<turnstile> e : Class C" and field: "P \<turnstile> C has F:T in D" and hconf: "P \<turnstile> h \<surd>" . show ?case proof cases assume "final e" with wte show ?thesis proof (rule final_addrE) fix a assume e: "e = addr a" with wte obtain fs where hp: "h a = Some(C,fs)" by auto with hconf have "P,h \<turnstile> (C,fs) \<surd>" using hconf_def by fastsimp then obtain v where "fs(F,D) = Some v" using field by(fastsimp dest:has_fields_fun simp:oconf_def fconf_def has_field_def) with hp e show ?thesis by(fastsimp intro:RedFAcc) next fix a assume "e = Throw a" thus ?thesis by(fastsimp intro:red_reds.FAccThrow) qed next assume "¬ final e" from prems show ?thesis by(fastsimp intro!:FAccRed) qed next case (WTrtFAccNT D E F T e) show ?case proof cases assume "final e" --"@{term e} is @{term null} or @{term throw}" from prems show ?thesis by(fastsimp simp:final_def intro: RedFAccNull red_reds.FAccThrow) next assume "¬ final e" --"@{term e} reduces by IH" from prems show ?thesis by simp (fast intro:FAccRed) qed next case (WTrtFAss C D E F T T2 e1 e2 h) have wte1: "P,E,h \<turnstile> e1 : Class C" . show ?case proof cases assume "final e1" with wte1 show ?thesis proof (rule final_addrE) fix a assume e1: "e1 = addr a" show ?thesis proof cases assume "final e2" thus ?thesis proof (rule finalE) fix v assume "e2 = Val v" thus ?thesis using e1 wte1 by(fastsimp intro:RedFAss) next fix a assume "e2 = Throw a" thus ?thesis using e1 by(fastsimp intro:red_reds.FAssThrow2) qed next assume "¬ final e2" from prems show ?thesis by simp (fast intro!:FAssRed2) qed next fix a assume "e1 = Throw a" thus ?thesis by(fastsimp intro:red_reds.FAssThrow1) qed next assume "¬ final e1" from prems show ?thesis by simp (blast intro!:FAssRed1) qed next case (WTrtFAssNT D E F T2 e1 e2) show ?case proof cases assume "final e1" --"@{term e1} is @{term null} or @{term throw}" show ?thesis proof cases assume "final e2" --"@{term e2} is @{term Val} or @{term throw}" from prems show ?thesis by(fastsimp simp:final_def intro: RedFAssNull red_reds.FAssThrow1 red_reds.FAssThrow2) next assume "¬ final e2" --"@{term e2} reduces by IH" from prems show ?thesis by (fastsimp simp:final_def intro!:red_reds.FAssRed2 red_reds.FAssThrow1) qed next assume "¬ final e1" --"@{term e1} reduces by IH" from prems show ?thesis by (fastsimp intro:FAssRed1) qed next case (WTrtCall C D E M T Ts Ts' body e es h pns) have wte: "P,E,h \<turnstile> e : Class C" and method: "P \<turnstile> C sees M:Ts->T = (pns,body) in D" and wtes: "P,E,h \<turnstile> es [:] Ts'"and sub: "P \<turnstile> Ts' [≤] Ts" and IHes: "!!l. [|P \<turnstile> h \<surd>; \<D>s es ⌊dom l⌋; ¬ finals es|] ==> ∃es' s'. P \<turnstile> 〈es,(h,l)〉 [->] 〈es',s'〉" and hconf: "P \<turnstile> h \<surd>" and D: "\<D> (e\<bullet>M(es)) ⌊dom l⌋" . show ?case proof cases assume "final e" with wte show ?thesis proof (rule final_addrE) fix a assume e_addr: "e = addr a" show ?thesis proof cases assume es: "∃vs. es = map Val vs" from wte e_addr obtain fs where ha: "h a = Some(C,fs)" by auto show ?thesis using e_addr ha method WTrts_same_length[OF wtes] sub es sees_wf_mdecl[OF wf method] by(fastsimp intro: RedCall simp:list_all2_def wf_mdecl_def) next assume "¬(∃vs. es = map Val vs)" hence not_all_Val: "¬(∀e ∈ set es. ∃v. e = Val v)" by(simp add:ex_map_conv) let ?ves = "takeWhile (λe. ∃v. e = Val v) es" let ?rest = "dropWhile (λe. ∃v. e = Val v) es" let ?ex = "hd ?rest" let ?rst = "tl ?rest" from not_all_Val have nonempty: "?rest ≠ []" by auto hence es: "es = ?ves @ ?ex # ?rst" by simp have "∀e ∈ set ?ves. ∃v. e = Val v" by(fastsimp dest:set_take_whileD) then obtain vs where ves: "?ves = map Val vs" using ex_map_conv by blast show ?thesis proof cases assume "final ?ex" moreover from nonempty have "¬(∃v. ?ex = Val v)" by(auto simp:neq_Nil_conv simp del:dropWhile_eq_Nil_conv) (simp add:dropWhile_eq_Cons_conv) ultimately obtain b where ex_Throw: "?ex = Throw b" by(fast elim!:finalE) show ?thesis using e_addr es ex_Throw ves by(fastsimp intro:CallThrowParams) next assume not_fin: "¬ final ?ex" have "finals es = finals(?ves @ ?ex # ?rst)" using es by(rule arg_cong) also have "… = finals(?ex # ?rst)" using ves by simp finally have "finals es = finals(?ex # ?rst)" . hence "¬ finals es" using not_finals_ConsI[OF not_fin] by blast thus ?thesis using e_addr D IHes[OF hconf] by(fastsimp intro!:CallParams) qed qed next fix a assume "e = Throw a" with WTrtCall.prems show ?thesis by(fast intro!:CallThrowObj) qed next assume "¬ final e" with prems show ?thesis by simp (blast intro!:CallObj) qed next case (WTrtCallNT E M T Ts e es) show ?case proof cases assume "final e" moreover { fix v assume "e = Val v" hence "e = null" using prems by simp have ?case proof cases assume "finals es" moreover { fix vs assume "es = map Val vs" from prems have ?thesis by(fastsimp intro: RedCallNull) } moreover { fix vs a es' assume "es = map Val vs @ Throw a # es'" from prems have ?thesis by(fastsimp intro: CallThrowParams) } ultimately show ?thesis by(fastsimp simp:finals_def) next assume "¬ finals es" --"@{term es} reduces by IH" from prems show ?thesis by(fastsimp intro: CallParams) qed } moreover { fix a assume "e = Throw a" from prems have ?case by(fastsimp intro: CallThrowObj) } ultimately show ?thesis by(fastsimp simp:final_def) next assume "¬ final e" --"@{term e} reduces by IH" from prems show ?thesis by (fastsimp intro:CallObj) qed next case WTrtNil thus ?case by simp next case (WTrtCons E T Ts e es h) have IHe: "!!l. [|P \<turnstile> h \<surd>; \<D> e ⌊dom l⌋; ¬ final e|] ==> ∃e' s'. P \<turnstile> 〈e,(h,l)〉 -> 〈e',s'〉" and IHes: "!!l. [|P \<turnstile> h \<surd>; \<D>s es ⌊dom l⌋; ¬ finals es|] ==> ∃es' s'. P \<turnstile> 〈es,(h,l)〉 [->] 〈es',s'〉" and hconf: "P \<turnstile> h \<surd>" and D: "\<D>s (e#es) ⌊dom l⌋" and not_fins: "¬ finals(e # es)" . have De: "\<D> e ⌊dom l⌋" and Des: "\<D>s es (⌊dom l⌋ \<squnion> \<A> e)" using D by auto show ?case proof cases assume "final e" thus ?thesis proof (rule finalE) fix v assume e: "e = Val v" hence Des': "\<D>s es ⌊dom l⌋" using De Des by auto have not_fins_tl: "¬ finals es" using not_fins e by simp show ?thesis using e IHes[OF hconf Des' not_fins_tl] by (blast intro!:ListRed2) next fix a assume "e = Throw a" hence False using not_fins by simp thus ?thesis .. qed next assume "¬ final e" with IHe[OF hconf De] show ?thesis by(fast intro!:ListRed1) qed next case (WTrtInitBlock E T T1 T2 V e2 h v) have IH2: "!!l. [|P \<turnstile> h \<surd>; \<D> e2 ⌊dom l⌋; ¬ final e2|] ==> ∃e' s'. P \<turnstile> 〈e2,(h,l)〉 -> 〈e',s'〉" and hconf: "P \<turnstile> h \<surd>" and D: "\<D> {V:T := Val v; e2} ⌊dom l⌋" . show ?case proof cases assume "final e2" show ?thesis proof (rule finalE) fix v2 assume "e2 = Val v2" thus ?thesis by(fast intro:RedInitBlock) next fix a assume "e2 = Throw a" thus ?thesis by(fast intro:red_reds.InitBlockThrow) qed next assume not_fin2: "¬ final e2" from D have D2: "\<D> e2 ⌊dom(l(V\<mapsto>v))⌋" by (auto simp:hyperset_defs) from IH2[OF hconf D2 not_fin2] obtain h' l' e' where red2: "P \<turnstile> 〈e2,(h, l(V\<mapsto>v))〉 -> 〈e',(h', l')〉" by auto from red_lcl_incr[OF red2] have "V ∈ dom l'" by auto with red2 show ?thesis by(fastsimp intro:InitBlockRed) qed next case (WTrtBlock E T T' V e h) have IH: "!!l. [|P \<turnstile> h \<surd>; \<D> e ⌊dom l⌋; ¬ final e|] ==> ∃e' s'. P \<turnstile> 〈e,(h,l)〉 -> 〈e',s'〉" and unass: "¬ assigned V e" and hconf: "P \<turnstile> h \<surd>" and D: "\<D> {V:T; e} ⌊dom l⌋" . show ?case proof cases assume "final e" show ?thesis proof (rule finalE) fix v assume "e = Val v" thus ?thesis by(fast intro:RedBlock) next fix a assume "e = Throw a" thus ?thesis by(fast intro:red_reds.BlockThrow) qed next assume not_fin: "¬ final e" from D have De: "\<D> e ⌊dom(l(V:=None))⌋" by(simp add:hyperset_defs) from IH[OF hconf De not_fin] obtain h' l' e' where red: "P \<turnstile> 〈e,(h,l(V:=None))〉 -> 〈e',(h',l')〉" by auto show ?thesis proof (cases "l' V") assume "l' V = None" with red unass show ?thesis by(blast intro: BlockRedNone) next fix v assume "l' V = Some v" with red unass show ?thesis by(blast intro: BlockRedSome) qed qed next case (WTrtSeq E T1 T2 e1 e2) show ?case proof cases assume "final e1" thus ?thesis by(fast elim:finalE intro:intro:RedSeq red_reds.SeqThrow) next assume "¬ final e1" from prems show ?thesis by simp (blast intro:SeqRed) qed next case (WTrtCond E T T1 T2 e e1 e2 h) have wt: "P,E,h \<turnstile> e : Boolean" . show ?case proof cases assume "final e" thus ?thesis proof (rule finalE) fix v assume val: "e = Val v" then obtain b where v: "v = Bool b" using wt by auto show ?thesis proof (cases b) case True with val v show ?thesis by(auto intro:RedCondT) next case False with val v show ?thesis by(auto intro:RedCondF) qed next fix a assume "e = Throw a" thus ?thesis by(fast intro:red_reds.CondThrow) qed next assume "¬ final e" from prems show ?thesis by simp (fast intro:CondRed) qed next case WTrtWhile show ?case by(fast intro:RedWhile) next case (WTrtThrow C E T e) show ?case proof cases assume "final e" -- {*Then @{term e} must be @{term throw} or @{term null}*} from prems show ?thesis by(fastsimp simp:final_def is_refT_def intro:red_reds.ThrowThrow red_reds.RedThrowNull) next assume "¬ final e" -- {*Then @{term e} must reduce*} from prems show ?thesis by simp (blast intro:ThrowRed) qed next case (WTrtTry C E T1 T2 V e1 e2 h) have wt1: "P,E,h \<turnstile> e1 : T1" . show ?case proof cases assume "final e1" thus ?thesis proof (rule finalE) fix v assume "e1 = Val v" thus ?thesis by(fast intro:RedTry) next fix a assume e1_Throw: "e1 = Throw a" with wt1 obtain D fs where ha: "h a = Some(D,fs)" by fastsimp show ?thesis proof cases assume "P \<turnstile> D \<preceq>* C" with e1_Throw ha show ?thesis by(fastsimp intro!:RedTryCatch) next assume "¬ P \<turnstile> D \<preceq>* C" with e1_Throw ha show ?thesis by(force intro!:RedTryFail) qed qed next assume "¬ final e1" show ?thesis using prems by simp (fast intro:TryRed) qed qed (*>*) end
lemma final_addrE:
[| WTrt P E h e (Class C); final e; !!a. e = addr a ==> R; !!a. e = Throw a ==> R |] ==> R
lemma finalRefE:
[| WTrt P E h e T; is_refT T; final e; e = null ==> R; !!a C. [| e = addr a; T = Class C |] ==> R; !!a. e = Throw a ==> R |] ==> R
lemmas WTrt'_induct:
[| !!C E h. is_class P C ==> P1 E h (new C) (Class C); !!C E T e h. [| WTrt' P E h e T; P1 E h e T; is_refT T; is_class P C |] ==> P1 E h (Cast C e) (Class C); !!E T h v. typeofh v = ⌊T⌋ ==> P1 E h (Val v) T; !!E T h v. E v = ⌊T⌋ ==> P1 E h (Var v) T; !!E T' T1 T2 bop e1 e2 h. [| WTrt' P E h e1 T1; P1 E h e1 T1; WTrt' P E h e2 T2; P1 E h e2 T2; case bop of Eq => T' = Boolean | Add => T1 = Integer & T2 = Integer & T' = Integer |] ==> P1 E h (e1 «bop» e2) T'; !!E T T' V e h. [| WTrt' P E h (Var V) T; P1 E h (Var V) T; WTrt' P E h e T'; P1 E h e T'; widen P T' T |] ==> P1 E h (V:=e) Void; !!C D E F T e h. [| WTrt' P E h e (Class C); P1 E h e (Class C); P \<turnstile> C has F:T in D |] ==> P1 E h (e\<bullet>F{D}) T; !!D E F T e h. [| WTrt' P E h e NT; P1 E h e NT |] ==> P1 E h (e\<bullet>F{D}) T; !!C D E F T T2 e1 e2 h. [| WTrt' P E h e1 (Class C); P1 E h e1 (Class C); P \<turnstile> C has F:T in D; WTrt' P E h e2 T2; P1 E h e2 T2; widen P T2 T |] ==> P1 E h (e1\<bullet>F{D} := e2) Void; !!D E F T2 e1 e2 h. [| WTrt' P E h e1 NT; P1 E h e1 NT; WTrt' P E h e2 T2; P1 E h e2 T2 |] ==> P1 E h (e1\<bullet>F{D} := e2) Void; !!C D E M T Ts Ts' body e es h pns. [| WTrt' P E h e (Class C); P1 E h e (Class C); P \<turnstile> C sees M: Ts->T = (pns, body) in D; WTrts' P E h es Ts'; P2 E h es Ts'; widens P Ts' Ts |] ==> P1 E h (e\<bullet>M(es)) T; !!E M T Ts e es h. [| WTrt' P E h e NT; P1 E h e NT; WTrts' P E h es Ts; P2 E h es Ts |] ==> P1 E h (e\<bullet>M(es)) T; !!E h. P2 E h [] []; !!E T Ts e es h. [| WTrt' P E h e T; P1 E h e T; WTrts' P E h es Ts; P2 E h es Ts |] ==> P2 E h (e # es) (T # Ts); !!E T T1 T2 V e2 h v. [| typeofh v = ⌊T1⌋; widen P T1 T; WTrt' P (E(V |-> T)) h e2 T2; P1 (E(V |-> T)) h e2 T2 |] ==> P1 E h {V:T; V:=Val v;; e2} T2; !!E T T' V e h. [| WTrt' P (E(V |-> T)) h e T'; P1 (E(V |-> T)) h e T'; ¬ assigned V e |] ==> P1 E h {V:T; e} T'; !!E T1 T2 e1 e2 h. [| WTrt' P E h e1 T1; P1 E h e1 T1; WTrt' P E h e2 T2; P1 E h e2 T2 |] ==> P1 E h (e1;; e2) T2; !!E T T1 T2 e e1 e2 h. [| WTrt' P E h e Boolean; P1 E h e Boolean; WTrt' P E h e1 T1; P1 E h e1 T1; WTrt' P E h e2 T2; P1 E h e2 T2; widen P T1 T2 | widen P T2 T1; widen P T1 T2 --> T = T2; widen P T2 T1 --> T = T1 |] ==> P1 E h (if (e) e1 else e2) T; !!E T c e h. [| WTrt' P E h e Boolean; P1 E h e Boolean; WTrt' P E h c T; P1 E h c T |] ==> P1 E h (while (e) c) Void; !!E T Tr e h. [| WTrt' P E h e Tr; P1 E h e Tr; is_refT Tr |] ==> P1 E h (throw e) T; !!C E T1 T2 V e1 e2 h. [| WTrt' P E h e1 T1; P1 E h e1 T1; WTrt' P (E(V |-> Class C)) h e2 T2; P1 (E(V |-> Class C)) h e2 T2; widen P T1 T2 |] ==> P1 E h (try e1 catch(C V) e2) T2 |] ==> (WTrt' P xha xga xfa xea --> P1 xha xga xfa xea) & (WTrts' P xda xca xba xaa --> P2 xda xca xba xaa)
lemmas WTrt'_elim_cases:
[| WTrt' P E h (V:=e) T; !!Ta T'. [| WTrt' P E h (Var V) Ta; WTrt' P E h e T'; widen P T' Ta; T = Void |] ==> Pa |] ==> Pa
lemma
WTrt' P E h (e1;; e2) T2 = (EX T1. WTrt' P E h e1 T1 & WTrt' P E h e2 T2)
lemma
WTrt' P E h (Val v) T = (typeofh v = ⌊T⌋)
lemma
WTrt' P E h (Var v) T = (E v = ⌊T⌋)
lemma wt_wt':
WTrt P E h e T ==> WTrt' P E h e T
and wts_wts':
WTrts P E h es Ts ==> WTrts' P E h es Ts
lemma wt'_wt:
WTrt' P E h e T ==> WTrt P E h e T
and wts'_wts:
WTrts' P E h es Ts ==> WTrts P E h es Ts
corollary wt'_iff_wt:
WTrt' P E h e T = WTrt P E h e T
corollary wts'_iff_wts:
WTrts' P E h es Ts = WTrts P E h es Ts
lemmas WTrt_induct2:
[| !!C E h. is_class P C ==> P1 E h (new C) (Class C); !!C E T e h. [| WTrt P E h e T; P1 E h e T; is_refT T; is_class P C |] ==> P1 E h (Cast C e) (Class C); !!E T h v. typeofh v = ⌊T⌋ ==> P1 E h (Val v) T; !!E T h v. E v = ⌊T⌋ ==> P1 E h (Var v) T; !!E T' T1 T2 bop e1 e2 h. [| WTrt P E h e1 T1; P1 E h e1 T1; WTrt P E h e2 T2; P1 E h e2 T2; case bop of Eq => T' = Boolean | Add => T1 = Integer & T2 = Integer & T' = Integer |] ==> P1 E h (e1 «bop» e2) T'; !!E T T' V e h. [| WTrt P E h (Var V) T; P1 E h (Var V) T; WTrt P E h e T'; P1 E h e T'; widen P T' T |] ==> P1 E h (V:=e) Void; !!C D E F T e h. [| WTrt P E h e (Class C); P1 E h e (Class C); P \<turnstile> C has F:T in D |] ==> P1 E h (e\<bullet>F{D}) T; !!D E F T e h. [| WTrt P E h e NT; P1 E h e NT |] ==> P1 E h (e\<bullet>F{D}) T; !!C D E F T T2 e1 e2 h. [| WTrt P E h e1 (Class C); P1 E h e1 (Class C); P \<turnstile> C has F:T in D; WTrt P E h e2 T2; P1 E h e2 T2; widen P T2 T |] ==> P1 E h (e1\<bullet>F{D} := e2) Void; !!D E F T2 e1 e2 h. [| WTrt P E h e1 NT; P1 E h e1 NT; WTrt P E h e2 T2; P1 E h e2 T2 |] ==> P1 E h (e1\<bullet>F{D} := e2) Void; !!C D E M T Ts Ts' body e es h pns. [| WTrt P E h e (Class C); P1 E h e (Class C); P \<turnstile> C sees M: Ts->T = (pns, body) in D; WTrts P E h es Ts'; P2 E h es Ts'; widens P Ts' Ts |] ==> P1 E h (e\<bullet>M(es)) T; !!E M T Ts e es h. [| WTrt P E h e NT; P1 E h e NT; WTrts P E h es Ts; P2 E h es Ts |] ==> P1 E h (e\<bullet>M(es)) T; !!E h. P2 E h [] []; !!E T Ts e es h. [| WTrt P E h e T; P1 E h e T; WTrts P E h es Ts; P2 E h es Ts |] ==> P2 E h (e # es) (T # Ts); !!E T T1 T2 V e2 h v. [| typeofh v = ⌊T1⌋; widen P T1 T; WTrt P (E(V |-> T)) h e2 T2; P1 (E(V |-> T)) h e2 T2 |] ==> P1 E h {V:T; V:=Val v;; e2} T2; !!E T T' V e h. [| WTrt P (E(V |-> T)) h e T'; P1 (E(V |-> T)) h e T'; ¬ assigned V e |] ==> P1 E h {V:T; e} T'; !!E T1 T2 e1 e2 h. [| WTrt P E h e1 T1; P1 E h e1 T1; WTrt P E h e2 T2; P1 E h e2 T2 |] ==> P1 E h (e1;; e2) T2; !!E T T1 T2 e e1 e2 h. [| WTrt P E h e Boolean; P1 E h e Boolean; WTrt P E h e1 T1; P1 E h e1 T1; WTrt P E h e2 T2; P1 E h e2 T2; widen P T1 T2 | widen P T2 T1; widen P T1 T2 --> T = T2; widen P T2 T1 --> T = T1 |] ==> P1 E h (if (e) e1 else e2) T; !!E T c e h. [| WTrt P E h e Boolean; P1 E h e Boolean; WTrt P E h c T; P1 E h c T |] ==> P1 E h (while (e) c) Void; !!E T Tr e h. [| WTrt P E h e Tr; P1 E h e Tr; is_refT Tr |] ==> P1 E h (throw e) T; !!C E T1 T2 V e1 e2 h. [| WTrt P E h e1 T1; P1 E h e1 T1; WTrt P (E(V |-> Class C)) h e2 T2; P1 (E(V |-> Class C)) h e2 T2; widen P T1 T2 |] ==> P1 E h (try e1 catch(C V) e2) T2 |] ==> (WTrt P xha xga xfa xea --> P1 xha xga xfa xea) & (WTrts P xda xca xba xaa --> P2 xda xca xba xaa)
theorem progress:
[| wwf_J_prog P; WTrt P E h e T; P \<turnstile> h \<surd>; \<D> e ⌊dom l⌋; ¬ final e |] ==> EX e' s'. red P e (h, l) e' s'
and
[| wwf_J_prog P; WTrts P E h es Ts; P \<turnstile> h \<surd>; \<D>s es ⌊dom l⌋; ¬ finals es |] ==> EX es' s'. reds P es (h, l) es' s'