diff --git a/thys/Dominance_CHK/Dom_Kildall.thy b/thys/Dominance_CHK/Dom_Kildall.thy --- a/thys/Dominance_CHK/Dom_Kildall.thy +++ b/thys/Dominance_CHK/Dom_Kildall.thy @@ -1,132 +1,93 @@ section \A kildall's algorithm for computing dominators\ theory Dom_Kildall -imports Dom_Semi_List Listn "HOL-Library.While_Combinator" +imports Dom_Semi_List "HOL-Library.While_Combinator" "Jinja.SemilatAlg" begin text \A kildall's algorithm for computing dominators. It uses the ideas and the framework of kildall's algorithm implemented in Jinja \cite{Kildall-AFP}, and modifications are needed to make it work for a fast algorithm for computing dominators\ -type_synonym - 's step_type = "nat \ 's \ (nat \ 's) list" - type_synonym state_dom = "nat list " primrec propa :: "'s binop \ (nat \ 's) list \ 's list \ nat list \ 's list * nat list" where "propa f [] \s wl = (\s,wl)" | "propa f (q'# qs) \s wl = (let (q,\) = q'; u = (\ \\<^bsub>f\<^esub> \s!q); wl' = (if u = \s!q then wl else (insort q (remove1 q wl))) in propa f qs (\s[q := u]) wl')" definition iter :: "'s binop \ 's step_type \ 's list \ nat list \ 's list \ nat list" where "iter f step \s w = while (\(\s,w). w \ []) (\(\s,w). let p = hd w in propa f (step p (\s!p)) \s (tl w)) (\s,w)" -definition stable :: "'s ord \ 's step_type \ 's list \ nat \ bool" -where - "stable r step \s p \ (\(q,\) \ set (step p (\s!p)). \ \\<^sub>r \s!q)" - definition unstables :: "state_dom ord \ state_dom step_type \ state_dom list \ nat list" where "unstables r step \s = sorted_list_of_set {p. p < size \s \ \ stable r step \s p}" +definition kildall :: "state_dom ord \state_dom binop \ state_dom step_type \ state_dom list \ state_dom list" where + "kildall r f step \s = fst(iter f step \s (unstables r step \s))" + lemma init_worklist_is_sorted: "sorted (unstables r step \s)" by (simp add:sorted_less_sorted_list_of_set unstables_def) -definition kildall :: "state_dom ord \state_dom binop \ state_dom step_type \ state_dom list \ state_dom list" where - "kildall r f step \s = fst(iter f step \s (unstables r step \s))" context cfg_doms begin definition transf :: "node \ state_dom \ state_dom " where "transf n input \ (n # input)" definition exec :: "node \ state_dom \ (node \ state_dom) list" where "exec n xs = map (\pc. (pc, (transf n xs))) (rev (sorted_list_of_set(succs n)))" lemma transf_res_is_rev: "sorted (rev ns) \ n > hd ns \ sorted (rev ((transf n ( ns))))" by (induct ns) (auto simp add:transf_def sorted_wrt_append) abbreviation "start \ [] # (replicate (length (g_V G) - 1) ( (rev[0.. state_dom list" where "dom_kildall = kildall (fst (snd nodes_semi)) (snd (snd nodes_semi)) exec" definition dom:: "nat \ nat \ bool" where "dom i j \(let res = (dom_kildall start) !j in i \ (set res) \ i = j )" lemma dom_refl: "dom i i" by (unfold dom_def) simp definition strict_dom :: "nat \ nat \ bool" where "strict_dom i j \ (let res = (dom_kildall start) !j in i \ set res)" lemma strict_domI1: "(dom_kildall ([] # (replicate (length (g_V G) - 1) ( (rev[0.. i \ set res \ strict_dom i j" by (auto simp add:strict_dom_def ) lemma strict_domD: "strict_dom i j \ dom_kildall (( [] # (replicate (length (g_V G) - 1) ( (rev[0.. i \ set a" by (auto simp add:strict_dom_def ) lemma sdom_dom: "strict_dom i j \ dom i j" by (unfold strict_dom_def) (auto simp add:dom_def) lemma not_sdom_not_dom: "\strict_dom i j \ i \ j \ \dom i j" by (unfold strict_dom_def) (auto simp add:dom_def) lemma dom_sdom: "dom i j \ i \ j \ strict_dom i j" by (unfold dom_def) (auto simp add:dom_def strict_dom_def) end -definition stables :: "'s ord \ 's step_type \ 's list \ bool" -where - "stables r step \s \ (\p < size \s. stable r step \s p)" - -definition lesubstep_type :: "(nat \ 's) set \ 's ord \ (nat \ 's) set \ bool" - ("(_ /{\\<^bsub>_\<^esub>} _)" [50, 0, 51] 50) - where "A {\\<^bsub>r\<^esub>} B \ \(p,\) \ A. \\'. (p,\') \ B \ \ \\<^sub>r \'" - -notation (ASCII) - lesubstep_type ("(_ /{<='__} _)" [50, 0, 51] 50) - -primrec pluslussub :: "'a list \ ('a \ 'a \ 'a) \ 'a \ 'a" ("(_ /\\<^bsub>_\<^esub> _)" [65, 0, 66] 65) -where - "pluslussub [] f y = y" -| "pluslussub (x#xs) f y = pluslussub xs f (x \\<^sub>f y)" -(*<*) -notation (ASCII) - pluslussub ("(_ /++'__ _)" [65, 1000, 66] 65) -(*>*) - -definition bounded :: "'s step_type \ nat \ bool" -where - "bounded step n \ (\p\. \(q,\') \ set (step p \). q nat \ 's set \ bool" -where - "pres_type step n A \ (\\\A. \p(q,\') \ set (step p \). \' \ A)" - -definition mono :: "'s ord \ 's step_type \ nat \ 's set \ bool" -where - "mono r step n A \ - (\\ p \'. \ \ A \ p \ \\<^sub>r \' \ set (step p \) {\\<^bsub>r\<^esub>} set (step p \'))" - end \ No newline at end of file diff --git a/thys/Dominance_CHK/Dom_Kildall_Property.thy b/thys/Dominance_CHK/Dom_Kildall_Property.thy --- a/thys/Dominance_CHK/Dom_Kildall_Property.thy +++ b/thys/Dominance_CHK/Dom_Kildall_Property.thy @@ -1,2241 +1,2160 @@ (* Author: Nan Jiang *) section \Properties of the kildall's algorithm on the semilattice \ theory Dom_Kildall_Property -imports Dom_Kildall +imports Dom_Kildall "Jinja.Listn" "Jinja.Kildall_1" begin lemma sorted_list_len_lt: "x \ y \ finite y \ length (sorted_list_of_set x) < length (sorted_list_of_set y)" proof- let ?x = "sorted_list_of_set x" let ?y = "sorted_list_of_set y" assume x_y: "x \ y" and fin_y: "finite y" then have card_x_y: "card x < card y" and fin_x: "finite x" by (auto simp add:psubset_card_mono finite_subset) with fin_y have "length ?x = card x" and "length ?y = card y" by auto with card_x_y show ?thesis by auto qed lemma wf_sorted_list: "wf ((\(x,y). (sorted_list_of_set x, sorted_list_of_set y)) ` finite_psubset)" apply (unfold finite_psubset_def) apply (rule wf_measure [THEN wf_subset]) apply (simp add: measure_def inv_image_def image_def) by (auto intro: sorted_list_len_lt) lemma sorted_list_psub: "sorted w \ w \ [] \ (sorted_list_of_set (set (tl w)), w) \ (\(x, y). (sorted_list_of_set x, sorted_list_of_set y)) ` {(A, B). A \ B \ finite B}" proof(intro strip, simp add:image_iff) assume sorted_w: "sorted w" and w_n_nil: "w \ []" let ?a = "set (tl w)" let ?b = "set w" from sorted_w have sorted_tl_w: "sorted (tl w)" and dist: "distinct w" by (induct w) (auto simp add: sorted_wrt_append ) with w_n_nil have a_psub_b: "?a \ ?b" by (induct w)auto from sorted_w sorted_tl_w have "w = sorted_list_of_set ?b" and "tl w = sorted_list_of_set (set (tl w))" by (auto simp add: sorted_less_set_eq) with a_psub_b show "\a b. a \ b \ finite b \ sorted_list_of_set (set (tl w)) = sorted_list_of_set a \ w = sorted_list_of_set b" by auto qed -primrec merges :: "'s binop \ (nat \ 's) list \ 's list \ 's list" -where - "merges f [] \s = \s" -| "merges f (p'#ps) \s = (let (p,\) = p' in merges f ps (\s[p := \ \\<^bsub>f\<^esub> \s!p]))" - locale dom_sl = cfg_doms + fixes A and r and f and step and start and n defines "A \ ((rev \ sorted_list_of_set) ` (Pow (set (nodes))))" defines "r \ nodes_le" defines "f \ nodes_sup" defines "n \ (size nodes)" defines "step \ exec" defines "start \ ([] # (replicate (length (g_V G) - 1) ( (rev[0..\used by acc\_le\_listI\ lemma Cons_less_Conss [simp]: "x#xs [\\<^sub>r] y#ys = (x \\<^sub>r y \ xs [\\<^bsub>r\<^esub>] ys \ x = y \ xs [\\<^sub>r] ys)" -proof(unfold lesssub_def, rule iffI) - assume "x # xs [\\<^bsub>r\<^esub>] y # ys \ x # xs \ y # ys " - then have ass1: "x # xs [\\<^bsub>r\<^esub>] y # ys" and ass2:"x # xs \ y # ys " by auto - from ass1 have g1: "x \\<^bsub>r\<^esub> y" and g2: "xs [\\<^bsub>r\<^esub>] ys" by auto - - from ass2 have "x \ y \ x = y \ xs \ ys" by auto - then show "(x \\<^bsub>r\<^esub> y \ x \ y) \ xs [\\<^bsub>r\<^esub>] ys \ x = y \ xs [\\<^bsub>r\<^esub>] ys \ xs \ ys" - proof - assume "x \ y" - with g1 g2 have "(x \\<^bsub>r\<^esub> y \ x \ y) \ xs [\\<^bsub>r\<^esub>] ys " by auto - then show ?thesis by auto - next - assume "x = y \ xs \ ys " - then have "x = y \ xs [\\<^bsub>r\<^esub>] ys \ xs \ ys" using g2 by auto - then show ?thesis by auto - qed -next - assume ass:"((x \\<^bsub>r\<^esub> y \ x \ y) \ xs [\\<^bsub>r\<^esub>] ys) \ (x = y \ xs [\\<^bsub>r\<^esub>] ys \ xs \ ys)" - from ass show "x # xs [\\<^bsub>r\<^esub>] y # ys \ x # xs \ y # ys" - proof - assume "(x \\<^bsub>r\<^esub> y \ x \ y) \ xs [\\<^bsub>r\<^esub>] ys " - then have "x \\<^bsub>r\<^esub> y \ x \ y" and ass1: "xs [\\<^bsub>r\<^esub>] ys " by auto - then have ass2: "x \\<^bsub>r\<^esub> y" and ass3: " x \ y" by auto - from ass3 have g2: "x # xs \ y # ys" by auto - from ass1 ass2 have g1: "x # xs [\\<^bsub>r\<^esub>] y # ys" by auto - with g2 show "x # xs [\\<^bsub>r\<^esub>] y # ys \ x # xs \ y # ys" by auto - next - assume "x = y \ xs [\\<^bsub>r\<^esub>] ys \ xs \ ys " - then have "x = y" and "xs [\\<^bsub>r\<^esub>] ys " and " xs \ ys " by auto - then have g2: "x # xs \ y # ys" by auto - - from `x = y` have "x \\<^bsub>r\<^esub> y" by(simp add: lesssub_def lesub_def r_def nodes_le_def ) - with `xs [\\<^bsub>r\<^esub>] ys` have g1: "x # xs [\\<^bsub>r\<^esub>] y # ys" by auto - with g2 show " x # xs [\\<^bsub>r\<^esub>] y # ys \ x # xs \ y # ys" by auto - qed -qed + apply (unfold lesssub_def) + apply auto + apply (unfold lesssub_def lesub_def r_def) + apply (simp only: nodes_le_refl) + done lemma acc_le_listI [intro!]: "acc r \ acc (Listn.le r) " apply (unfold acc_def) apply (subgoal_tac "Wellfounded.wf(UN n. {(ys,xs). size xs = n \ size ys = n \ xs <_(Listn.le r) ys})") apply (erule wf_subset) apply (blast intro: lesssub_lengthD) apply (rule wf_UN) prefer 2 apply (rename_tac m n) apply (case_tac "m=n") apply simp apply (fast intro!: equals0I dest: not_sym) apply (rename_tac n) apply (induct_tac n) apply (simp add: lesssub_def cong: conj_cong) apply (rename_tac k) apply (simp add: wf_eq_minimal) apply (simp (no_asm) add: length_Suc_conv cong: conj_cong) apply clarify apply (rename_tac M m) apply (case_tac "\x xs. size xs = k \ x#xs \ M") prefer 2 apply (erule thin_rl) apply (erule thin_rl) apply blast apply (erule_tac x = "{a. \xs. size xs = k \ a#xs:M}" in allE) apply (erule impE) apply blast apply (thin_tac "\x xs. P x xs" for P) apply clarify apply (rename_tac maxA xs) apply (erule_tac x = "{ys. size ys = size xs \ maxA#ys \ M}" in allE) apply (erule impE) apply blast apply clarify apply (thin_tac "m \ M") apply (thin_tac "maxA#xs \ M") apply (rule bexI) prefer 2 apply assumption apply clarify apply simp apply blast done lemma wf_listn: "wf {(y,x). x \\<^bsub>Listn.le r\<^esub> y}" by(insert acc_nodes_le acc_le_listI r_def) (simp add:acc_def) lemma wf_listn': "wf {(y,x). x [\\<^sub>r] y}" by (rule wf_listn) lemma wf_listn_termination_rel: "wf ({(y,x). x \\<^bsub>Listn.le r\<^esub> y} <*lex*> (\(x, y). (sorted_list_of_set x, sorted_list_of_set y)) ` finite_psubset)" by (insert wf_listn wf_sorted_list) (fastforce dest:wf_lex_prod) lemma inA_is_sorted: "xs \ A \ sorted (rev xs)" by (auto simp add:A_def sorted_less_sorted_list_of_set) lemma list_nA_lt_refl: "xs \ list n A \ xs [\\<^bsub>r\<^esub>] xs" proof assume "xs \ list n A" then have "set xs \ A" by (rule listE_set) then have "\i A" by auto then have "\i\<^bsub>r\<^esub>] xs" by(unfold Listn.le_def lesub_def) (auto simp add:list_all2_conv_all_nth Listn.le_def r_def nodes_le_def) qed lemma nil_inA: "[] \ A" apply (unfold A_def) apply (subgoal_tac "{} \ Pow (set nodes)") apply (subgoal_tac "[] = (\x. rev (sorted_list_of_set x)) {}") apply (fastforce intro:rev_image_eqI) by auto lemma upt_n_in_pow_nodes: "{0.. Pow (set nodes)" by(auto simp add:n_def nodes_def verts_set) lemma rev_all_inA: "rev [0.. A" proof(unfold A_def,simp) let ?f = "\x. rev (sorted_list_of_set x)" have "rev [0.. ?f ` Pow (set nodes)" by (fastforce intro: image_eqI) qed lemma len_start_is_n: "length start = n" by (insert len_verts_gt0) (auto simp add:start_def n_def nodes_def dest:Suc_pred) lemma len_start_is_len_verts: "length start = length (g_V G)" using len_verts_gt0 by (simp add:start_def) lemma start_len_gt_0: "length start > 0" by (insert len_verts_gt0) (simp add:start_def) lemma start_subset_A: "set start \ A" by(auto simp add:nil_inA rev_all_inA start_def) lemma start_in_A : "start \ (list n A)" by (insert start_subset_A len_start_is_n)(fastforce intro:listI) lemma sorted_start_nth: "i < n \ sorted (rev (start!i))" apply(subgoal_tac "start!i \ A") apply (fastforce dest:inA_is_sorted) by (auto simp add:start_subset_A len_start_is_n) lemma start_nth0_empty: "start!0 = []" by (simp add:start_def) lemma start_nth_lt0_all: "\p\{1..< length start}. start!p = (rev [0.. set (g_V G) \ x < n" by (simp add:n_def nodes_def verts_set) lemma start_nth0_unstable_auxi: "\ [0] \\<^bsub>r\<^esub> (rev [0.. stable r step start 0" proof(rule notI,auto simp add: start_nth0_empty stable_def step_def exec_def transf_def) assume ass: "\x\set (sorted_list_of_set (succs 0)). [0] \\<^bsub>r\<^esub> start ! x" from succ_of_entry0 obtain s where "s \ succs 0" and "s \ 0 \ s \ set (g_V G)" using head_is_vert by (auto simp add:succs_def) then have "s \ set (sorted_list_of_set (succs 0))" and "start!s = (rev [0.. {1 ..< length (g_V G)} " and "succs p \ {}" shows "\ stable r step start p" proof (rule notI, unfold stable_def) let ?step_p = "step p (start ! p)" let ?rev_all = "rev[0..(q, \)\set ?step_p. \ \\<^bsub>r\<^esub> start ! q" from assms(1) have n_sorted: "\ sorted (rev (p # ?rev_all))" and p:"p \ set (g_V G)" and "start!p = ?rev_all" using verts_set by (auto simp add:n_def nodes_def start_def sorted_wrt_append) with sta have step_p: "\(q, \)\set ?step_p. sorted (rev (p # ?rev_all)) \ (p # ?rev_all = start!q)" by (auto simp add:step_def exec_def transf_def lesssub_def lesub_def r_def nodes_le_def) from assms(2) fin_succs p obtain a b where a_b: "(a, b) \ set ?step_p" by (auto simp add:step_def exec_def transf_def) with step_p have "sorted (rev (p # ?rev_all)) \ (p # ?rev_all = start!a)" by auto with n_sorted have eq_p_cons: "(p # ?rev_all = start!a)" by auto from p have "\(q, \)\set ?step_p. q < n" using succ_in_G fin_succs verts_set n_def nodes_def by (auto simp add:step_def exec_def) with a_b have "a < n" using len_start_is_n by auto then have "sorted (rev (start!a))" using sorted_start_nth by auto with eq_p_cons n_sorted show False by auto qed lemma start_unstable_cond: assumes "succs p \ {} " and "p < length (g_V G)" shows "\ stable r step start p" using assms start_nth0_unstable start_nth_unstable by(cases "p = 0") auto lemma unstable_start: "unstables r step start = sorted_list_of_set ({p. succs p \ {} \ p < length start})" using len_start_is_len_verts start_unstable_cond by (subgoal_tac "{p. p < length start \ \ stable r step start p} = {p. succs p \ {} \ p < length start}") (auto simp add: unstables_def stable_def step_def exec_def) -lemma nth_merges: - "\ss. \p < length ss; ss \ list n A; \(p,t)\set ps. p t\A \ \ - (merges f ps ss)!p = map snd [(p',t') \ ps. p'=p] \\<^bsub>f\<^esub> ss!p" - (is "\ss. \_; _; ?steptype ps\ \ ?P ss ps") -(*<*) -proof (induct ps) - show "\ss. ?P ss []" by simp - - fix ss p' ps' - assume ss: "ss \ list n A" - assume l: "p < length ss" - assume "?steptype (p'#ps')" - then obtain a b where - p': "p'=(a,b)" and ab: "aA" and ps': "?steptype ps'" - by (cases p') auto - assume "\ss. p< length ss \ ss \ list n A \ ?steptype ps' \ ?P ss ps'" - hence IH: "\ss. ss \ list n A \ p < length ss \ ?P ss ps'" using ps' by iprover - - from is_semi have "closed A f" by (simp add:semilat_def) - with ss ab - have "ss[a := b \\<^bsub>f\<^esub> ss!a] \ list n A" by (simp add:closedD) - moreover - with l have "p < length (ss[a := b \\<^bsub>f\<^esub> ss!a])" by simp - ultimately - have "?P (ss[a := b \\<^bsub>f\<^esub> ss!a]) ps'" by (rule IH) - with p' l - show "?P ss (p'#ps')" by simp -qed - -lemma length_merges [simp]: - "\ss. size(merges f ps ss) = size ss" -(*<*) by (induct ps, auto) (*>*) - -lemma merges_preserves_type_lemma: -shows "\xs. xs \ list n A \ (\(p,x) \ set ps. p x\A) \ merges f ps xs \ list n A" - apply(insert is_semi) - apply (unfold semilat_def) - apply (unfold closed_def) - apply (induct ps) - apply simp - apply clarsimp - done - -lemma merges_preserves_type [simp]: - "\ xs \ list n A; \(p,x) \ set ps. p x\A \ \ merges f ps xs \ list n A" -by (simp only: merges_preserves_type_lemma) - -declare sorted_list_of_set_insert[simp del] - -lemmas [simp] = Let_def Semilat.le_iff_plus_unchanged [OF Semilat.intro, symmetric] +end declare sorted_list_of_set_insert_remove[simp del] -lemma decomp_propa: "\ss w. +context dom_sl +begin + +lemma (in dom_sl) decomp_propa: "\ss w. (\(q,t)\set qs. q < size ss \ t \ A ) \ sorted w \ set ss \ A \ propa f qs ss w = (merges f qs ss, (sorted_list_of_set ({q. \t.(q,t)\set qs \ t \\<^bsub>f\<^esub> (ss!q) \ ss!q} \ set w)))" (*<*) apply (induct qs) apply (fastforce intro:sorted_less_set_eq) apply (simp (no_asm)) apply clarify apply (subst sorted_insort_remove1) apply simp apply (simp add: sorted_less_sorted_list_of_set Semilat.closed_f[OF Semilat.intro, OF is_semi]) apply (rule conjI) apply (blast intro: arg_cong) apply(simp add: nth_list_update) by (auto intro: arg_cong) lemma distinct_pair_filter_n: "distinct (map fst xs) \ a \ set (map fst xs) \ (map snd (filter (\(x,y). x = a) xs)) = []" by (induct xs) (auto simp add: distinct_map_filter image_def) lemma map_fst: "map fst (map (\pc. (pc, x)) xs) = xs" by (induct xs) auto lemma distinct_p: "p < n \ distinct (map fst (step p (ss!p)))" proof- let ?qs = "step p (ss!p)" have "map fst ?qs = (rev (sorted_list_of_set(succs p)))" using map_fst[of _ "(rev (sorted_list_of_set(succs p)))"] by (auto simp add:step_def exec_def) then show ?thesis by auto qed lemma distinct_pair_filter: "distinct (map fst xs) \ (a,b)\ set xs \ map snd (filter (\x. fst x = a) xs) = [b]" apply (induct xs) apply simp apply (auto simp add: distinct_map_filter image_def) proof- { fix xs assume "\x\set xs. a \ fst x " and " distinct (map fst xs)" then show "filter (\x. fst x = a) xs = []" by (induct xs) auto } qed lemma split_comp_eq_pair: "(\x. fst x = a) = (\(x,y). x = a)" by (rule split_comp_eq) lemma distinct_pair_filter': "distinct (map fst xs) \ (a,b)\ set xs \ (map snd (filter (\(x,y). x = a) xs)) = [b]" using distinct_pair_filter by (simp only: split_comp_eq_pair) lemma merges_property1: fixes ss w qs assumes step_bounded_pres: "\(q, \) \ set qs. q < size ss \ \ \ A" and subset_ss_A: "set ss \ A " and len_ss_n: "length ss = n " and dist: "distinct (map fst qs) " shows "\(q, \) \ set qs. merges f qs ss!q = \ \\<^bsub>f\<^esub>ss!q" proof fix x assume "x \ set qs" from dist have \: "\(q, \) \ set qs. (map snd (filter (\(x,y). x = q) qs)) = [\]" using distinct_pair_filter' by fastforce from len_ss_n subset_ss_A have "ss \ list n A" by (rule listI) with step_bounded_pres have merge_nth: "\(q, \) \ set qs. (merges f qs ss)!q = map snd [(p',t') \ qs. p'=q] \\<^bsub>f\<^esub> ss!q" - by (fastforce intro:nth_merges) \ \use lemma: listE_length\ + by (fastforce intro:Semilat.nth_merges[OF Semilat.intro, OF is_semi]) \ \use lemma: listE_length\ with \ have "\(q, \) \ set qs. (merges f qs ss)!q = [\]\\<^bsub>f\<^esub> ss!q" by fastforce then show "case x of (q, \) \ merges f qs ss ! q = \ \\<^bsub>f\<^esub> ss ! q" using `x \ set qs` by auto qed lemma propa_property1: fixes ss w qs assumes step_bounded_pres: "\(q, \) \ set qs. q < size ss \ \ \ A " and sorted_w: "sorted (w ::nat list)" and subset_ss_A: "set ss \ A " and len_ss_n: "length ss = n " and dist: "distinct (map fst qs) " shows "\(q, \) \ set qs. (fst(propa f qs ss w))!q = \ \\<^bsub>f\<^esub>ss!q" using assms apply (subgoal_tac "fst (propa f qs ss w) = merges f qs ss") apply(simp add: merges_property1) by (auto dest:decomp_propa) lemma merges_property2: fixes ss w qs q assumes step_bounded_pres: "\(q, \) \ set qs. q < size ss \ \ \ A" and subset_ss_A: "set ss \ A" and len_ss_n: "length ss = n " and dist: "distinct (map fst qs) " and q_nin: "q \ set(map fst qs) " and q_lt_len_ss: "q < length ss " shows "(merges f qs ss)!q = ss!q" proof- from len_ss_n subset_ss_A have "ss \ list n A" by (rule listI) - with step_bounded_pres q_lt_len_ss have merge_nth: "(merges f qs ss)!q = map snd [(p',t') \ qs. p'=q] \\<^bsub>f\<^esub> ss!q" by (fastforce intro:nth_merges) + with step_bounded_pres q_lt_len_ss have merge_nth: "(merges f qs ss)!q = map snd [(p',t') \ qs. p'=q] \\<^bsub>f\<^esub> ss!q" + by (fastforce intro:Semilat.nth_merges[OF Semilat.intro, OF is_semi]) from dist have "q \ set(map fst qs) \ (map snd (filter (\(x,y). x = q) qs)) = []" using distinct_pair_filter_n by fastforce with merge_nth q_nin have "(merges f qs ss)!q = []\\<^bsub>f\<^esub> ss!q" by fastforce then show "(merges f qs ss)!q = ss ! q" by auto qed lemma propa_property2: fixes ss w qs q assumes step_bounded_pres: "\(q, \) \ set qs. q < size ss \ \ \ A" and sorted_w: "sorted (w ::nat list)" and subset_ss_A: "set ss \ A" and len_ss_n: "length ss = n " and dist: "distinct (map fst qs) " and q_nin: "q \ set(map fst qs) " and q_lt_len_ss: "q < length ss " shows "(fst(propa f qs ss w))!q = ss!q" using assms apply (subgoal_tac "fst (propa f qs ss w) = merges f qs ss") apply(simp add: merges_property2) by (auto dest:decomp_propa) -lemma merges_incr_lemma: +lemma merges_incr_lemma_dom: "\xs. xs \ list n A \ distinct (map fst ps) \ (\(p,x)\set ps. p x \ A) \ xs [\\<^bsub>r\<^esub>] merges f ps xs" proof(intro strip) fix xs assume xs_inA: "xs \ list n A " and step_bounded_pres: "\(p, x)\set ps. p < length xs \ x \ A" and dist: "distinct (map fst ps)" then have len_xs_n: "length xs = n" and subset_xs_inA: "set xs \ A" by (auto simp add:listE_length) with step_bounded_pres dist have merge1: "\(q, \) \ set ps. (merges f ps xs)!q = \ \\<^bsub>f\<^esub>xs!q" using merges_property1 by auto from step_bounded_pres dist len_xs_n subset_xs_inA have merge2: "\q. q \ set(map fst ps) \ q < length xs \ (merges f ps xs)!q = xs!q" using merges_property2 by auto have len_eq: "length xs = length (merges f ps xs)" by simp have "\i\<^bsub>r\<^esub> (merges f ps xs)!i" proof(intro strip) fix i assume i_lt_len_xs: "i < length xs" with xs_inA have xs_i_inA: "xs!i \ A" by auto show " xs ! i \\<^bsub>r\<^esub> (merges f ps xs) ! i " proof(cases "i \ set (map fst ps)") case True then obtain s' where s': "(i, s') \ set ps" by auto with merge1 have merge1': "(merges f ps xs)!i = s' \\<^bsub>f\<^esub> xs ! i" by auto from s' step_bounded_pres have "s' \ A" by auto with xs_i_inA merge1' show ?thesis by (auto intro: Semilat.ub2[OF Semilat.intro, OF is_semi]) next case False note i_nin = this with i_lt_len_xs merge2 have "(merges f ps xs)!i = xs!i" by auto with xs_i_inA show ?thesis by (auto simp add:Semilat.refl_r[OF Semilat.intro, OF is_semi]) qed qed then have "\i\<^bsub>r\<^esub> (merges f ps xs)!i" by auto then have "\iix y. lesub x r y) (xs!i)((merges f ps xs)!i)" by simp then have nth_lt: "\i. i (\x y. lesub x r y) (xs!i)((merges f ps xs)!i)" by simp with len_eq show "xs [\\<^bsub>r\<^esub>] merges f ps xs" by (auto simp only:list_all2_conv_all_nth Listn.le_def lesssub_def lesub_def) qed -lemma merges_incr: +lemma merges_incr_dom: "\ xs \ list n A; distinct (map fst ps); \(p,x)\set ps. p x \ A \ \ xs [\\<^bsub>r\<^esub>] merges f ps xs" - by (simp add: merges_incr_lemma) + by (simp add: merges_incr_lemma_dom) -lemma merges_same_conv [rule_format]: +lemma merges_same_conv_dom [rule_format]: "(\xs. xs \ list n A \ distinct (map fst ps) \ (\(p,x)\set ps. p x\A) \ (merges f ps xs = xs) = (\(p,x)\set ps. x \\<^bsub>r\<^esub> xs!p))" apply (induct_tac ps) apply (simp) apply clarsimp apply (rename_tac p x ps xs) apply (rule iffI) apply (rule context_conjI) apply (subgoal_tac "xs[p := x \\<^bsub>f\<^esub> xs!p] [\\<^bsub>r\<^esub>] xs") apply (fastforce dest!: le_listD simp add: nth_list_update Semilat.plus_le_conv[OF Semilat.intro,OF is_semi] Semilat.ub1[OF Semilat.intro,OF is_semi] Semilat.ub2[OF Semilat.intro,OF is_semi] Semilat.lub[OF Semilat.intro,OF is_semi] ) - apply (erule subst, rule merges_incr) + apply (erule subst, rule merges_incr_dom) apply (blast intro!: listE_set intro: closedD listE_length [THEN nth_in] Semilat.closed_f[OF Semilat.intro, OF is_semi]) apply clarify apply(intro strip) apply auto apply (erule allE) apply (erule impE) apply assumption apply (drule bspec) apply assumption apply (simp add: Semilat.le_iff_plus_unchanged [OF Semilat.intro, OF is_semi, THEN iffD1] list_update_same_conv [THEN iffD2]) apply blast apply (simp add: Semilat.le_iff_plus_unchanged[OF Semilat.intro, OF is_semi, THEN iffD1] list_update_same_conv [THEN iffD2]) done (*>*) -lemma list_update_le_listI [rule_format]: +lemma (in Semilat)list_update_le_listI [rule_format]: "set xs \ A \ set ys \ A \ xs [\\<^sub>r] ys \ p < size xs \ - x \\<^sub>r ys!p \ semilat(A,r,f) \ x\A \ + x \\<^sub>r ys!p \ x\A \ xs[p := x \\<^sub>f xs!p] [\\<^sub>r] ys" (*<*) + apply (insert semilat) apply (simp only: Listn.le_def lesub_def semilat_def) apply (simp add: list_all2_conv_all_nth nth_list_update) done -lemma merges_pres_le_ub: +lemma (in Semilat) merges_pres_le_ub: assumes "set ts \ A" "set ss \ A" "\(p,t)\set ps. t \\<^bsub>r\<^esub> ts!p \ t \ A \ p < size ts" "ss [\\<^bsub>r\<^esub>] ts" shows "merges f ps ss [\\<^bsub>r\<^esub>] ts" (*<*) proof - { fix t ts ps have "\qs. \set ts \ A; \(p,t)\set ps. t \\<^bsub>r\<^esub> ts!p \ t \ A \ p< size ts \ \ set qs \ set ps \ (\ss. set ss \ A \ ss [\\<^bsub>r\<^esub>] ts \ merges f qs ss [\\<^bsub>r\<^esub>] ts)" apply (induct_tac qs) apply simp apply (simp (no_asm_simp)) apply clarify apply simp apply (erule allE, erule impE, erule_tac [2] mp) apply (drule bspec, assumption) - apply (simp add: closedD Semilat.closed_f[OF Semilat.intro, OF is_semi]) + apply (simp add: closedD Semilat.closed_f) apply (drule bspec, assumption) apply (simp ) apply(rule list_update_le_listI) apply auto - apply(insert is_semi) - apply auto done } note this [dest] from assms show ?thesis by blast qed (*>*) lemma termination_lemma: shows "\ss \ list n A; distinct (map fst qs); \(q,t)\set qs. q t\A; sorted w; w \ [] \ \ ss [\\<^sub>r] merges f qs ss \ merges f qs ss = ss \ (sorted_list_of_set ({q. \t. (q,t)\set qs \ t \\<^bsub>f\<^esub> ss!q \ ss!q} \ set (tl w)),w) \ (\x. case x of (x, y) \ (sorted_list_of_set x, sorted_list_of_set y)) ` {(A, B). A \ B \ finite B}" apply(insert is_semi) apply (unfold lesssub_def) - apply (simp (no_asm_simp) add: merges_incr) + apply (simp (no_asm_simp) add: merges_incr_dom) apply (rule impI) - apply (rule merges_same_conv [THEN iffD1, elim_format]) + apply (rule merges_same_conv_dom [THEN iffD1, elim_format]) apply assumption+ defer apply (rule sym, assumption) defer apply simp apply (subgoal_tac "\q t. \((q, t) \ set qs \ t \\<^bsub>f\<^esub> ss ! q \ ss ! q)") defer apply clarsimp apply (drule bspec, assumption) apply (drule bspec, assumption) apply clarsimp apply(subgoal_tac "{q. \t. (q, t) \ set qs \ t \\<^bsub>f\<^esub> ss ! q \ ss ! q} \ set (tl w) = set (tl w)") apply (auto simp add: sorted_list_psub) done lemma bounded_exec: "bounded step n" apply (insert is_semi,unfold semilat_def bounded_def step_def exec_def transf_def ) by (auto simp add: n_def nodes_def verts_set fin_succs' succ_range) lemma bounded_exec': fixes ss w a b assumes w_n_nil: " w \ [] " and step_hdw: " (a, b) \ set (step (hd w) (ss ! hd w)) " and w_lt_n:"\p\set w. p < n " shows" a < n" using assms proof- from w_lt_n have "hd w < n" using w_n_nil by auto with bounded_exec have "( \\. \(q,\') \ set (step (hd w) \). q (\\\set ss. \ \ A) \ (\p < n. sorted (rev ( (ss!p))) \ (ss!p \ rev [0..< n] \ (\x\set ( (ss!p)). x < p)) \ (ss!p = rev [0.. (\x\ set w. (x,p)\ g_E G \ x < p)) \ (p \ set w \ stable r step ss p)) \ sorted w \ length ss = n \ (\x\set w. x < n) " lemma wf_dom_w: assumes "wf_dom ss w" shows "sorted w" using assms by (auto simp add:wf_dom_def) lemma wf_dom_w2: assumes "wf_dom ss w" shows "(\x\set w. x < n)" using assms by (auto simp add:wf_dom_def) lemma wf_dom_len_eq: assumes "wf_dom ss w" shows "(length ss = n)" using assms by (auto simp add:wf_dom_def) lemma wf_dom_inA: assumes "wf_dom ss w" shows "ss \ list n A" using assms by (auto simp add:wf_dom_def intro: listI) lemma wf_dom_le: assumes wf_ss_w: "wf_dom ss w" shows "ss [\\<^bsub>r\<^esub>] ss" using assms proof- from wf_ss_w have "\i A" by (auto simp add:wf_dom_def) then have "\iss\A. p < n \ (\x ss!x) \ transf p ss \ A" proof(intro strip, unfold transf_def) fix ss assume p_lt_n: "p < n" and p_gt: "\x_in_A: "ss \ A" then have "set ss \ set nodes" by (unfold A_def) (rule inpow_subset_nodes) with p_lt_n have p_\_in: "set (p # ss) \ set nodes" by (auto simp add:n_def nodes_def verts_set) from \_in_A have sorted_\: "sorted (rev ss)" by (simp add:inA_is_sorted) with p_gt have "sorted (rev (p # ss))" using Cons_sorted_less_nth by auto with p_\_in show "(p # ss) \ A" by (unfold A_def) (fastforce intro: subset_nodes_inpow) qed lemma pres_exec: assumes "(q,\) \ set (step p (ss!p))" and "\n \ set (ss!p). p > n" and "ss!p \ A" and "p < n" shows "\ \ A " using assms by (auto simp add:step_def exec_def pres_transf) lemma wf_hd_neq_all: assumes wf_ss_w: "wf_dom ss w " and w_n_nil: "w \ [] " shows " ss!(hd w) \ rev [0.. ss ! hd w \ (rev [0..x\ set w. x < n" and "sorted w" by (auto simp add:wf_dom_def) then have "hd w < n" and y:"\x \ set w. x \ hd w" using w_n_nil by (induct w) (auto simp add:sorted_wrt_append) with wf_ss_w x have "(\x\ set w. (x,hd w)\ g_E G \ x < hd w)" by (auto simp add:wf_dom_def) with y show False by auto qed lemma pres_wf_exec: fixes ss w a b assumes wf_ss_w: "wf_dom ss w " and w_n_nil: "w \ [] " shows "\(q,\) \ set (step (hd w) (ss!(hd w))). \ \ A " using assms proof(intro strip, auto) fix a b assume a_b: "(a, b) \ set (step (hd w) (ss ! hd w))" from wf_ss_w have sorted_w: "sorted w" and hd_w_lt_n: "hd w < n" and ss_hdw_inA: "ss!hd w \ A" using w_n_nil by (auto simp add:wf_dom_def) from assms have ss_hd_w_neq_all: "ss!hd w \ (rev [0..x\set (ss!hd w). hd w > x" by (auto simp add:wf_dom_def) with ss_hdw_inA hd_w_lt_n a_b show "b \ A" using pres_exec by auto qed lemma propa_dom_invariant_auxi: assumes wf_a_b: "wf_dom a b" and b_n_nil: "b \ [] " shows "a!hd b \ rev [0.. sorted (rev (hd b # (a!hd b))) \ set (hd b # (a!hd b)) \ set nodes \ hd b # (a!hd b) \ A \ (\(q, \)\set (step (hd b) (a!hd b)). q < length a \ \ \ A)" using assms proof- from assms have "a!hd b \ (rev \ sorted_list_of_set) ` (Pow (set nodes))" and sorted_hdb: "sorted (rev (a!hd b))" and hd_b_lt_n: "hd b < n" and sorted_b: "sorted b" and len_eq: "length a = n" by (auto simp add:A_def wf_dom_def) then have as_subset_nodes: "set (a!hd b) \ set nodes" by (auto simp add:inpow_subset_nodes) with n_def verts_set assms nodes_def have hdb_cons_subset_nodes: "set (hd b # (a!hd b)) \ set nodes" by (auto simp add:wf_dom_def) then have hdb_subset_n: "set (hd b # (a!hd b)) \ {0.. (rev [0.. Pow (set (g_V G))" by (auto simp add:nodes_def) then have pow1: "set ((hd b # (a!hd b))) \ set nodes" by (auto simp add:nodes_def) from hdb_cons_subset_nodes sorted_hd_b_cons have "hd b # (a!hd b) \ (rev \ sorted_list_of_set) ` (Pow (set nodes))" by (fastforce intro: subset_nodes_inpow) then have hd_b_cons_in_A: "(hd b # (a!hd b)) \ A" by (unfold A_def ) (auto simp add:nodes_def) have "(\p\. \(q,\') \ set (step p \). q(q,\') \ set (step (hd b) (a!hd b)). q(q, \)\set(step (hd b) (a!hd b)). \ \ A" by (rule pres_wf_exec) with bounded pres have step_pres_bounded: "\(q, \)\set (step (hd b) (a!hd b)). q < length a \ \ \ A " using len_eq by auto with a_hd_b_neq_all sorted_hd_b_cons hdb_cons_subset_nodes hd_b_cons_in_A show ?thesis by auto qed lemma propa_dom_invariant_aux: fixes a b ss w assumes propa: "propa f (step (hd b) (a ! hd b)) a (tl b) = (ss, w) " and b_n_nil: "b \ [] " and a_in_A: "\\\set a. \ \ A" and ass: "\p (a!p \ rev [0.. (\x\set (a!p). x < p)) \ (a!p = rev [0.. (\x\set b. (x,p)\ g_E G \ x < p)) \ (p \ set b \ stable r step a p)" and sorted_b: "sorted b " and len_eq: "length a = n " and b_lt_n: "\x\set b. x < n " shows "(\\\set ss. \ \ A) \ (\p (ss!p \ rev [0.. (\x\set (ss!p). x < p)) \ (ss!p = rev [0.. (\x\ set w. (x,p)\ g_E G \ x < p)) \ (p \ set w \ stable r step ss p)) \ sorted w \ length ss = n \ (\x\set w. x list n A" by (auto intro: listI) from a_in_A have "\p< length a. a!p \ A" by (auto simp add:A_def) then have a_p_subset: "\p< length a. set (a!p) \ set nodes" by (auto simp add:inpow_subset_nodes A_def) from a_in_A ass sorted_b len_eq b_lt_n b_n_nil have wf_a_b: "wf_dom a b" by (auto simp add:wf_dom_def) with b_n_nil have a_hd_b_neq_all: "?a_hdb \ rev [0.. set nodes" and hd_b_cons_in_A: "hd b # ?a_hdb \ A" and step_pres_bounded: "(\(q, \)\set (step (hd b) ?a_hdb). q < length a \ \ \ A)" using propa_dom_invariant_auxi by auto then have hdb_subset_n: "set (hd b # ?a_hdb) \ {0..x\ set (tl b). x < n" by (induct b)(auto simp add:wf_dom_def) then have dist: "distinct (map fst ?qs_a)" using distinct_p by auto from b_lt_n len_eq sorted_b have sorted_tl_b: "sorted (tl b)" by (induct b) auto from b_lt_n verts n_def nodes_def verts_set have b_in_verts: "\x\set b. x \ set (g_V G)" by auto then have hd_b_in_verts: "hd b \ set (g_V G)" using b_n_nil by auto then have fin_succ_hd_b: "finite (succs (hd b))" using fin_succs by auto have fin_q1: "finite {q. \t.(q,t)\set ?qs_a}" and fin_q2: "finite ?q_a_wl" by (auto simp add:step_def exec_def image_def) then have fin: "finite (?q_a_wl \ set (tl b))" by auto from a_in_A have set_a: "set a \ A" by auto with step_pres_bounded sorted_tl_b len_eq dist have "\(q, \) \ set ?qs_a. (fst(propa f ?qs_a a (tl b)))!q = \ \\<^bsub>f\<^esub>a!q" by (auto dest:propa_property1) with propa have propa_ss1: "\(q, \) \ set ?qs_a. ss!q = \ \\<^bsub>f\<^esub>a!q" by simp then have propa_ss1': "\(q, \) \ set ?qs_a. ss!q = (hd b # ?a_hdb) \\<^bsub>f\<^esub>a!q" by (auto simp add:step_def exec_def transf_def) then have succ_self_eq: "\(q, \) \ set ?qs_a. q = hd b \ ss!q = a!q" by(auto simp add: f_def nodes_sup_def plussub_def inter_sorted_cons[OF sorted_hd_b_cons]) have step_hd_b: "\(q,\)\set ?qs_a. \ = (hd b # ?a_hdb)" by(auto simp add:step_def exec_def transf_def) from step_pres_bounded sorted_tl_b set_a len_eq dist have "\q. q \ set(map fst ?qs_a) \ q < length a \ (fst(propa f ?qs_a a (tl b)))!q = a!q" by (auto intro:propa_property2) with propa have exec2: "\q. q \ set(map fst ?qs_a) \ q < length a \ ss!q = a!q" by auto then have ss_hd_b_eq_a: "ss!hd b = a!hd b" using hd_b_lt_n len_eq fin_succ_hd_b succ_self_eq by (auto simp add:step_def exec_def) then have hdb_nin_w: "hd b \ ?q_a_wl" using fin_succ_hd_b propa_ss1 by (auto simp add:step_def exec_def) from step_pres_bounded sorted_tl_b set_a have "snd (propa f ?qs_a a (tl b)) = (sorted_list_of_set (?q_a_wl \ set (tl b)))" by (fastforce dest:decomp_propa) with propa have ww: "w = sorted_list_of_set (?q_a_wl \ set (tl b))" by simp then have sorted_w:"sorted w" by (simp add:sorted_less_sorted_list_of_set) from ww have set_ww: "set w =?q_a_wl \ set (tl b)" using fin by auto with step_pres_bounded tl_b_lt_n ww fin len_eq have w_lt_n: "(\x\set w. x < n)" using len_eq by auto then have w_set': "\x\set w. x \ {0.. {0.. set (tl b) " by (induct b) (auto simp add:sorted_wrt_append) with hdb_nin_w ww have "hd b \ set w" using fin by auto from step_pres_bounded sorted_tl_b set_a propa have ss_merge: "ss = merges f ?qs_a a" by (auto dest: decomp_propa) from step_pres_bounded a_in_list_nA have "\(q, \)\set (step (hd b) (a ! hd b)). q < n \ \ \ A" using len_eq by simp - with a_in_list_nA have "merges f ?qs_a a \ list n A" by (rule merges_preserves_type) + with a_in_list_nA have "merges f ?qs_a a \ list n A" + by (rule Semilat.merges_preserves_type[OF Semilat.intro, OF is_semi]) with ss_merge have ss_in_A: "ss \ list n A" by simp then have len_ss_n: "length ss = n" using len_eq by simp with len_eq have len_ss_n: "length ss = n" by auto from ss_in_A have set_ss: "set ss \ A " by (rule listE_set) then have ss_inA: "\\\set ss. \ \ A" by auto then have ss_p_subset: "\p< length ss. set (ss!p) \ set nodes" by (auto simp add:inpow_subset_nodes A_def) from w_lt_n len_ss_n have bounded_a: "\a b. w \ [] \ (a, b) \ set ?qs_ss \ a < length ss" by (auto simp add:bounded_exec') have sorted_a_hdw_n: "w \ [] \ sorted (rev ?a_hdw)" using wf_a_b len_eq w_set' by (auto simp add:wf_dom_def) have sorted_ss_hdw_n: "w \ [] \ sorted (rev ?ss_hdw)" using ss_in_A inA_is_sorted w_lt_n by auto have "w \ [] \ (hd w # ?ss_hdw) \ A" proof assume w_n_nil: "w \ []" with w_lt_n have hd_w_lt_n: "hd w < n" by auto with a_in_A have a_hdw_inA: "?a_hdw \ A" using len_eq by auto then have a_hdw_in_nodes: "set ?a_hdw \ set nodes" by (unfold A_def )(rule inpow_subset_nodes) from hd_w_lt_n have hdw_in_nodes: "hd w \ set nodes" using verts_set by (simp add:n_def nodes_def) from sorted_a_hdw_n w_n_nil have sorted_a_hdw: "sorted (rev ?a_hdw)" by auto show "(hd w # ?ss_hdw) \ A" proof(cases "hd w \ succs (hd b)") case True then obtain s where s: "(hd w, s) \ set ?qs_a" using fin_succ_hd_b by (auto simp add:step_def exec_def) with step_hd_b have "s = (hd b # ?a_hdb)" by auto with s propa_ss1 have ss_hd_w: "?ss_hdw = (hd b # ?a_hdb) \\<^bsub>f\<^esub> ?a_hdw" by auto with hd_b_cons_in_A a_hdw_inA have "?ss_hdw \ A" by (simp add: Semilat.closed_f[OF Semilat.intro, OF is_semi]) then have ss_hdw_in_nodes: "set ?ss_hdw \ set nodes" by (auto simp add:inpow_subset_nodes A_def) with hdw_in_nodes have hdw_cons_ss_in_nodes: "set (hd w # ?ss_hdw) \ set nodes" by auto then have in_pow: "set (hd w # ?ss_hdw) \ Pow (set (g_V G))" by (auto simp add:nodes_def) from sorted_a_hdw ss_hd_w sorted_hd_b_cons have "sorted (rev ?ss_hdw)" and "set ?ss_hdw = set (hd b # ?a_hdb) \ set ?a_hdw" by (auto simp add:f_def plussub_def nodes_sup_def inter_sorted_correct) then have sorted_ss_hdw: "sorted (rev ?ss_hdw)" and ss_hdw_subset_a_hdb_cons: "set ?ss_hdw \ set (hd b # ?a_hdb)" and ss_hdw_subset_a_hdw: "set ?ss_hdw \ set ?a_hdw" by auto from sorted_hd_b_cons have "\x\ set (hd b # ?a_hdb). x \ hd b" using b_n_nil by (induct b) (auto simp add:sorted_wrt_append) with ss_hdw_subset_a_hdb_cons have ss_hdw_lt_hdb: "\x\set ?ss_hdw. x \ hd b" using sorted_hd_b_cons by auto have sorted_rev_hdw: "sorted (rev (hd w # ?ss_hdw))" proof(cases "?a_hdw = rev [0..x\ set b. (x,hd w)\ g_E G \ x < hd w" using hd_w_lt_n by (auto simp add:wf_dom_def) then obtain prev_hd_w where prev_hd_w_in_b: "prev_hd_w \ set b" and prev_hd_w: "(prev_hd_w, hd w)\ g_E G" and prev_hd_w_lt: "prev_hd_w < hd w" by auto show ?thesis proof(cases "prev_hd_w = hd b") case True with prev_hd_w_lt have "hd b < hd w" by auto with ss_hdw_lt_hdb have "\x\set ?ss_hdw. hd w > x" by auto with sorted_ss_hdw show ?thesis by (auto simp add:sorted_wrt_append) next case False with prev_hd_w_in_b have "prev_hd_w \ set (tl b)" by (induct b) auto with ww have prev_hd_w_nin_w: "prev_hd_w \ set w" using fin by auto from sorted_w have "\x\ set w. x \ hd w" by(induct w) (auto simp add:sorted_wrt_append sorted_less_sorted_list_of_set) with prev_hd_w_lt have "prev_hd_w \ set w" by auto with prev_hd_w_nin_w show ?thesis by auto qed next case False with wf_a_b have "(\x\set ?a_hdw. x < hd w)" using hd_w_lt_n by (auto simp add:wf_dom_def) with ss_hdw_subset_a_hdw have "(\x\set ?ss_hdw. x < hd w)" by auto with sorted_ss_hdw show ?thesis by (auto simp add:sorted_wrt_append) qed with hdw_cons_ss_in_nodes show "hd w # ?ss_hdw \ A" using A_def by (fastforce dest:subset_nodes_inpow) next case False note hd_w_nin_succ_hdb = this then have "hd w \ set (map fst ?qs_a)" using fin_succ_hd_b by (auto simp add:step_def exec_def) with exec2 have ss_hd_w_eq_a: "?ss_hdw = ?a_hdw" using hd_w_lt_n len_ss_n len_eq by auto with a_hdw_in_nodes sorted_a_hdw have ss_hdw_in_nodes: "set ?ss_hdw \ set nodes" and sorted_ss_hdw: "sorted (rev ?ss_hdw)" by auto with hdw_in_nodes have hdw_cons_in_nodes: "set (hd w # ?ss_hdw) \ set nodes" by (auto) from hd_w_nin_succ_hdb ww have hd_w_non: "hd w \{q. \t. (q, t) \ set ?qs_a \ t \\<^bsub>f\<^esub> a ! q \ a ! q}" using fin_succ_hd_b by (auto simp add:step_def exec_def ) from set_ww hd_w_non have hd_w_in_tl_b: "hd w \ set (tl b)" using sorted_tl_b `w \ []` by auto have sorted_rev_hdw: "sorted (rev (hd w # ?ss_hdw))" proof(cases "?a_hdw = rev [0..x\ set b. (x,hd w)\ g_E G \ x < hd w" using hd_w_lt_n by (auto simp add:wf_dom_def) then obtain prev_hd_w where prev_hd_w_in_b: "prev_hd_w \ set b" and prev_hd_w: "(prev_hd_w, hd w)\ g_E G" and prev_hd_w_lt: "prev_hd_w < hd w" by auto from ww have "sorted w" by(simp add:sorted_less_sorted_list_of_set) then have "\x\ set w. x \ hd w" using w_n_nil by (induct w) (auto simp add:sorted_wrt_append) with prev_hd_w_lt have prev_hd_w_nin_w: "prev_hd_w \ set w" using w_n_nil by auto with set_ww have "prev_hd_w \ set (tl b)" by auto with prev_hd_w_in_b have "prev_hd_w = hd b" using sorted_b by (induct b) auto with prev_hd_w have "(hd b, hd w) \ g_E G" by auto with hd_w_nin_succ_hdb show ?thesis by (auto simp add:succs_def) next case False with wf_a_b have a_hdw_lt: "(\x\set ?a_hdw. x < hd w)" and "sorted (rev ?a_hdw)" using hd_w_lt_n by (auto simp add:wf_dom_def) with ss_hd_w_eq_a have "sorted (rev ?ss_hdw)" by simp from ss_hd_w_eq_a have "?a_hdw = ?ss_hdw" by auto with a_hdw_lt have "\x\set ?ss_hdw. x < hd w" by auto with sorted_ss_hdw show ?thesis by (auto simp add:sorted_wrt_append) qed with hdw_cons_in_nodes show " (hd w # ss ! hd w) \ A" by (unfold A_def) (rule subset_nodes_inpow) qed qed then have pres_ss: "\q \. w \ [] \ (q, \) \ set (step (hd w) ?ss_hdw) \ \ \ A" by (auto simp add:step_def exec_def transf_def) have hd_b_ss_sta: "stable r step ss (hd b)" proof(unfold stable_def, clarsimp) fix q \ assume "(q, \) \ set ?qs_ss_hdb " then have "q \ succs (hd b)" and "\ = transf (hd b) ?ss_hdb" using hd_b_lt_n fin_succ_hd_b by (auto simp add:step_def exec_def) then have \:"\ = (hd b # ?a_hdb)" using ss_hd_b_eq_a by (auto simp add:transf_def) from `q \ succs (hd b)` hd_b_lt_n have "q\ set (g_V G)" using succ_in_G by auto then have "q < n" using verts_set by (auto simp add:n_def nodes_def) with wf_a_b have a_q_inA: "a!q \ A" by (auto simp add:wf_dom_def) from wf_a_b a_hd_b_neq_all hd_b_lt_n have "\x\ set ( (a!hd b)). x < hd b" by (auto simp add:wf_dom_def) with sorted_hd_b_cons have "sorted (rev (hd b # ?a_hdb))" by (auto simp add:sorted_wrt_append) from propa_ss1 `(q, \) \ set (step (hd b) (ss ! hd b))` have "ss!q = \ \\<^bsub>f\<^esub> a ! q" using `ss!hd b = a!hd b` by auto with \ have "ss!q = (hd b # ?a_hdb) \\<^bsub>f\<^esub> a ! q" by simp with hd_b_cons_in_A a_q_inA have " (hd b # ?a_hdb)\\<^bsub>r\<^esub> ss!q " by (auto simp add: Semilat.ub1[OF Semilat.intro, OF is_semi]) with \ show "\ \\<^bsub>r\<^esub> ss ! q" by auto qed have wf_dom_2: "\p (ss ! p \ rev [0.. (\x\set ( (ss ! p)). x < p)) \ (ss ! p = rev [0.. (\x\ set w. (x,p)\ g_E G \ x < p)) \ (p \ set w \ stable r step ss p)" proof (intro strip) fix p let ?a_p = "a!p" let ?ss_p = "ss!p" assume p_lt_n: "p < n" with wf_a_b have a_p_inA: "?a_p \ A" and sorted_a_p: "sorted (rev ?a_p)" and sta_a_p: "p \ set b \ stable r step a p" by (auto simp add:wf_dom_def) from p_lt_n have "p \ set (g_V G)" using verts_set n_def nodes_def by auto then have fin_succ_p: "finite (succs p)" using fin_succs by auto from set_a p_lt_n have a_p_inA: "?a_p \ A" using `length a = n` by (auto simp add:A_def) then have "set ?a_p \ set nodes" using inpow_subset_nodes by (auto simp add:A_def) with p_lt_n have set_p_a_p: "set (p#?a_p) \ set nodes" using n_def nodes_def verts_set by auto from p_lt_n len_eq have p_lt_len_a: "p < length a" by auto with ss_inA have ss_p_in_A: "ss!p \ A" using len_eq len_ss_n by auto with p_lt_n have sorted_ss_p: "sorted (rev ?ss_p)" by (auto simp add:A_def sorted_less_sorted_list_of_set) have len_ss_eq_a: "length ss = length a" using len_eq len_ss_n by auto have p_nin_w_eq: "p \ set w \ (p \ set b \ p = hd b) \ (\s. (p, s) \ set ?qs_a \ s \\<^bsub>f\<^esub> ?a_p = ?a_p)" proof assume "p \ set w" with set_ww have p_nin1: "p \ {q. \t. (q, t) \ set ?qs_a\ t \\<^bsub>f\<^esub> a ! q \ a ! q}" and p_nin2: "p \ set (tl b)" by auto from p_nin1 have p_nin: "(\s. (p, s) \ set ?qs_a \ s \\<^bsub>f\<^esub> ?a_p = ?a_p)" by auto from p_nin2 have "p \ set b \ p = hd b" by (induct b) auto with p_nin show "(p \ set b \ p = hd b) \ (\s. (p, s) \ set ?qs_a \ s \\<^bsub>f\<^esub> ?a_p = ?a_p)" by auto next assume "(p \ set b \ p = hd b) \ (\s. (p, s) \ set ?qs_a \ s \\<^bsub>f\<^esub> ?a_p = ?a_p)" then have p1: "p \ set b \ p = hd b" and p2: "\s. (p, s) \ set ?qs_a \ s \\<^bsub>f\<^esub> a ! p = ?a_p" by auto from p1 have "p \ set (tl b)" proof assume "p \ set b" then show ?thesis by(induct b) auto next assume "p = hd b" with sorted_b show ?thesis by (induct b) (auto simp add:sorted_wrt_append) qed with p2 set_ww show "p \ set w" using set_ww by auto qed have stable_ss_p: "p \ set w \ w \ Nil \ stable r step ss p" proof (clarsimp simp add: stable_def split_paired_all) fix q \ assume p_nin_w: "p \ set w" and w_n_nil: "w \ Nil" and step_ss_p: "(q, \) \ set (step p ?ss_p) " from p_nin_w have p_cond: "(p \ set b \ p = hd b) \ (\s. (p, s) \ set ?qs_a \ (hd b # ?a_hdb) \\<^bsub>f\<^esub> ?a_p = ?a_p)" using p_nin_w_eq by (auto simp add:transf_def step_def exec_def) then have p_cond1: "(p \ set b \ p = hd b)" and p_cond2: "(\s. (p, s) \ set ?qs_a \ (hd b # ?a_hdb) \\<^bsub>f\<^esub> ?a_p = ?a_p)"by auto from bounded_a pres_ss w_n_nil have " \(q, t)\set ?qs_ss. q < length ss \ t \ A" by auto then have "\(q, t)\set (step (hd w) (ss ! hd w)). q < length ss \ (transf (hd w) (ss!hd w)) \ A" by (auto simp add:step_def exec_def) have ss_a_p_eq: "?ss_p = ?a_p" proof(cases "p \ succs (hd b)") case True note p_in_succ_hd_b = this from `p \ succs (hd b)` propa_ss1' have ss_p: "?ss_p = (hd b # ?a_hdb) \\<^bsub>f\<^esub> ?a_p" using fin_succ_hd_b by (auto simp add:step_def exec_def) from p_in_succ_hd_b p_cond2 have " (hd b # ?a_hdb) \\<^bsub>f\<^esub> ?a_p= ?a_p" using fin_succ_hd_b by (auto simp add:step_def exec_def) with ss_p show ?thesis by auto next case False then have "p \ {q. \t. (q, t) \ set ?qs_a}" using fin fin_succ_hd_b by (auto simp add:step_def exec_def) then have "p \ set ( map fst ?qs_a)" by auto with p_lt_n show ?thesis using exec2 len_eq by auto qed have "(\(q, \)\set (step p ?ss_p). transf p ?ss_p \\<^bsub>r\<^esub> ss ! q) " proof(intro strip, clarsimp) fix succ_p z let ?a_succ_p = "a!succ_p" let ?ss_succ_p = "ss!succ_p" assume "(succ_p, z) \ set (step p ?ss_p)" then have succ_p: "succ_p \ succs p" using p_lt_n fin_succ_p by (auto simp add:step_def exec_def) with p_lt_n have succ_p_lt_n: "succ_p < n" using succ_in_verts n_def nodes_def verts_set by auto with wf_a_b have a_succ_p_inA: "?a_succ_p \ A" by (auto simp add:wf_dom_def) from succ_p_lt_n ss_in_A have ss_succ_p_inA: "?ss_succ_p \ A" by auto have p_nin_b_ssp: "p \ set b \ transf p ?ss_p \ A \ transf p ?ss_p \\<^sub>r ?a_succ_p" proof- assume "p \ set b" with sta_a_p have "(\(q,\) \ set (step p ?a_p). \ \\<^sub>r a!q)" by (simp add:stable_def) with succ_p have transf_p_succp':"transf p ?a_p \\<^sub>r ?a_succ_p" using fin_succ_p by (auto simp add:stable_def step_def exec_def) with ss_a_p_eq have transf_lt_p: "transf p ?ss_p \\<^sub>r ?a_succ_p" by auto from transf_p_succp' have "(p# ?a_p) \\<^sub>r ?a_succ_p" by (auto simp add:transf_def) then have "sorted (rev (p#?a_p)) \ (p#?a_p = ?a_succ_p)" by (auto simp add:step_def exec_def transf_def r_def lesssub_def lesub_def nodes_le_def) with set_p_a_p a_succ_p_inA have "(p#?a_p) \ (rev \ sorted_list_of_set) ` (Pow (set nodes))" using subset_nodes_inpow by (auto simp add:A_def) then have "transf p ?a_p \ A" using transf_def by (auto simp add:A_def transf_def) then show ?thesis using ss_a_p_eq transf_lt_p by auto qed show "transf p ?ss_p \\<^bsub>r\<^esub> ?ss_succ_p" proof(cases "p \ succs (hd b)") case True note p_in_succ_hd_b = this with p_cond have " (hd b # ?a_hdb) \\<^bsub>f\<^esub> ?a_p =?a_p" using fin_succ_hd_b by (auto simp add:step_def exec_def) from p_in_succ_hd_b propa_ss1' have ss_p: "?ss_p = (hd b # ?a_hdb) \\<^bsub>f\<^esub> ?a_p" using fin_succ_hd_b by (auto simp add:step_def exec_def) then have transf_p2: "transf p ?ss_p = (p # (inter_sorted_rev (hd b # ?a_hdb) ?a_p))" by (auto simp add:f_def plussub_def nodes_sup_def transf_def ) then show ?thesis proof(cases "succ_p \ succs (hd b)") case True note succ_p_is_succ_hdb = this with propa_ss1 have ss_succ_p: "?ss_succ_p = (hd b # ?a_hdb) \\<^bsub>f\<^esub> ?a_succ_p" using fin_succ_hd_b by (auto simp add:step_def exec_def transf_def) with a_succ_p_inA hd_b_cons_in_A have succ_p1: "(hd b # ?a_hdb) \\<^bsub>r\<^esub>?ss_succ_p" and succ_p2: "?a_succ_p\\<^bsub>r\<^esub> ?ss_succ_p " by (auto simp add:Semilat.ub1[OF Semilat.intro, OF is_semi] Semilat.ub2[OF Semilat.intro, OF is_semi]) show ?thesis proof(cases "p \ set b") case True note p_in_set_b = this then have p_eq_hdb: "p = hd b" using p_cond by auto with succ_p have succ_p_is_succ_hdb: "succ_p \ succs (hd b)" by auto from ss_a_p_eq p_eq_hdb have "(hd b # ?ss_hdb) = hd b # ?a_hdb" by auto with succ_p1 have "hd b # ss ! hd b \\<^bsub>r\<^esub> ss ! succ_p" by auto with p_eq_hdb show ?thesis by (auto simp add:transf_def) next case False then have "transf p ?ss_p \ A" and "transf p ?ss_p \\<^sub>r ?a_succ_p" using p_nin_b_ssp by auto with succ_p2 a_succ_p_inA ss_succ_p_inA show ?thesis by (auto intro: order_trans Semilat.orderI[OF Semilat.intro, OF is_semi]) qed next case False note succ_p_n_succ_hdb = this with exec2 have ss_succ_p_eq_a: "?ss_succ_p = ?a_succ_p" using fin_succ_hd_b succ_p_lt_n len_eq by (auto simp add:step_def exec_def) show ?thesis proof(cases "p \ set b") case True with p_cond have p_eq_hdb: "p = hd b" by auto with succ_p have succ_p_is_succ_hdb: "succ_p \ succs (hd b)" by auto with succ_p_n_succ_hdb show ?thesis by auto next case False with sta_a_p have "\(q,\) \ set (step p ?a_p). \ \\<^sub>r a!q" by (simp add:stable_def) with succ_p fin_succ_p have "transf p ?a_p \\<^sub>r ?a_succ_p" by (auto simp add:step_def exec_def) with ss_succ_p_eq_a ss_a_p_eq show ?thesis by auto qed qed next case False note p_nin_succ_hd_b = this from p_nin_succ_hd_b p_cond have "p \ set b \ p = hd b" by auto then show ?thesis proof assume "p \ set b" then have transf_ss_p_inA: "transf p ?ss_p \ A" and transf_lt_p: "transf p ?ss_p \\<^sub>r ?a_succ_p" using p_nin_b_ssp by auto show "transf p ?ss_p \\<^bsub>r\<^esub> ?ss_succ_p" proof(cases "succ_p \ succs (hd b)") case True with succ_p_lt_n propa_ss1' len_eq fin_succ_hd_b have "?ss_succ_p = (hd b # ?a_hdb) \\<^bsub>f\<^esub> ?a_succ_p" by (auto simp add:step_def exec_def) with a_succ_p_inA hd_b_cons_in_A have "?a_succ_p \\<^sub>r ?ss_succ_p" by (auto simp add:Semilat.ub2[OF Semilat.intro, OF is_semi]) with transf_lt_p transf_ss_p_inA a_succ_p_inA ss_succ_p_inA show ?thesis by (auto intro: order_trans Semilat.orderI[OF Semilat.intro, OF is_semi]) next case False with succ_p_lt_n exec2 len_eq fin_succ_hd_b have "?ss_succ_p = ?a_succ_p" by (auto simp add:step_def exec_def) with transf_lt_p show ?thesis by auto qed next assume "p = hd b" with hd_b_ss_sta have "(\(q,\) \ set (step p ?ss_p). \ \\<^sub>r ss!q)" by (simp add:stable_def) with succ_p `p = hd b` show "transf p ?ss_p \\<^bsub>r\<^esub> ?ss_succ_p" using fin_succ_hd_b by (auto simp add:stable_def step_def exec_def transf_def) qed qed qed with step_ss_p show "\ \\<^bsub>r\<^esub> ss ! q" by (auto simp add:step_def exec_def) qed show "sorted (rev ?ss_p) \ (?ss_p \ rev [0.. (\x\set?ss_p. x < p)) \ (?ss_p = rev [0.. (\x\ set w. (x,p)\ g_E G \ x < p)) \ (p \ set w \ stable r step ss p)" proof(cases "w \ []") case True note w_n_nil = this show ?thesis proof (cases "p \ set( map fst (?qs_a))") case True with propa_ss1 have ss_p: "?ss_p = (hd b # ?a_hdb) \\<^bsub>f\<^esub> ?a_p" by (simp add:step_def exec_def transf_def) with sorted_hd_b_cons sorted_a_p have ss_p_lt_hdb: "set ?ss_p \ set (hd b # ?a_hdb)" and ss_p_lt: "set ?ss_p \ set ?a_p" and ss_p_inter: "set ?ss_p = set (hd b # ?a_hdb) \ set ?a_p" by (auto simp add:f_def plussub_def nodes_sup_def inter_sorted_correct) from sorted_hd_b_cons sorted_a_p have "inter_sorted_rev (hd b # ?a_hdb) ?a_p = rev (sorted_list_of_set (set (hd b # ?a_hdb) \ set ?a_p))" by (rule inter_sorted_correct_col) with ss_p have ss_p2: "?ss_p = (rev (sorted_list_of_set (set (hd b # ?a_hdb) \ set ?a_p)))" by (auto simp add:f_def plussub_def nodes_sup_def) show ?thesis proof(cases "?a_p \ (rev [0.. {0..x\set ?a_p. x < p)" using n_def len_eq by (auto simp add:wf_dom_def) with ss_p_lt have "\x\set ?ss_p. x < p" by auto then have ss_lt_p: "?ss_p \ rev [0.. (\x\set ?ss_p. x < p)" by auto have "set ?ss_p \ {0.. set ?ss_p \ {0.. set ?a_p" by auto from a_in_A p_lt_len_a have " ?a_p \ ((rev \ sorted_list_of_set) ` (Pow (set nodes)))" by (auto simp add: A_def) then have ass_set_in_nodes: "set ?a_p \ set nodes" by (rule inpow_subset_nodes) then have "set ?a_p \ {0.. rev [0.. (\x\set w. (x,p)\ g_E G \ x < p)" by auto with sorted_ss_p ss_lt_p stable_ss_p w_n_nil show ?thesis by fastforce next case False then have a_p_all: "?a_p = (rev [0..x\ set b. (x,p)\ g_E G \ x < p)" by (auto simp add:wf_dom_def) have "hd b \ set b" using b_n_nil by auto from a_p_all have "set ?a_p = {0.. {0.. set ?a_p) = set (hd b # ?a_hdb)" by auto with ss_p2 have ss_p3: "?ss_p = (rev (sorted_list_of_set (set (hd b # ?a_hdb))))" by auto from sorted_hd_b_cons have "sorted_list_of_set (set (hd b # ?a_hdb)) = rev (hd b # ?a_hdb)" by (fastforce dest: sorted_less_rev_set_eq) with ss_p3 have ss_p_4: "?ss_p = (hd b # ?a_hdb)" by auto have "(?ss_p \ rev [0.. (\x\set ?ss_p. x < p)) \ (?ss_p = rev [0.. (\x\ set w. (x,p)\ g_E G \ x < p)) \ (p \ set w \ stable r step ss p)" proof(cases "?ss_p \ (rev [0..x\set ( (ss ! p)). x < p" by (auto simp add:sorted_wrt_append) with ss_p_4 ss_p_n_all stable_ss_p `w \ []` show ?thesis by auto next case False note hd_b_ge_p = this from ex_lt_p obtain x where "x\ set b " and " (x,p)\ g_E G " and "x < p"by auto from `x \ set b` `x < p` hd_b_ge_p have "tl b \ []" by (induct b) auto with hd_b_ge_p sorted_b have temp_t: "\x\ set (tl b). x \ p" by (induct b) auto with `\hd b < p` have "x \ set (tl b)" using `(x,p)\ g_E G ` `x \ set b` `x < p` by (induct b) auto with `x < p` temp_t have False by auto then show ?thesis by auto qed next case False then have "?ss_p = (rev [0.. 0 "using n_def nodes_def len_verts_gt0 by auto then have last_hd: "last [0.. 0` by auto then have "hd (rev [0..set b \ (x, p) \ g_E G \ x < p" by auto then have "x\set b " and " (x, p) \ g_E G " and " x < p"by auto from x `\hd b hd b" by auto with `x\set b ` have "tl b \ []" by (induct b) auto with `x\set b ` x_n_hd_b have "x \ set (tl b)" by (induct b) auto with ww have "x \ set w" using fin by auto then show ?thesis using `ss!p = (rev [0.. g_E G` `x []` by auto qed qed then show ?thesis using sorted_ss_p by auto qed next case False note p_not_in_succ = this with exec2 p_lt_n len_eq have ss_a_p_eq: "ss!p = a!p" by auto with ass p_lt_n wf_a_b have cond1: "(?ss_p \ rev [0.. (\x\set ?ss_p. x < p))" by (auto simp add:wf_dom_def) show ?thesis proof(cases "?a_p \ (rev [0..x\set ?a_p. x < p)" using p_lt_n len_eq by (auto simp add:wf_dom_def) with ss_a_p_eq have "(\x\set ?ss_p. x < p)" by auto with cond1 sorted_ss_p show ?thesis using p_lt_n len_eq stable_ss_p w_n_nil by auto next case False then have "?a_p = (rev [0.. (\x\ set b. (x,p)\ g_E G \ x < p)" by (auto simp add:wf_dom_def) with ss_a_p_eq `a!p = (rev [0..x\ set b. (x,p)\ g_E G \ x < p)" using len_eq by auto then obtain x where "x\ set b " and " (x,p) \ g_E G" and " x < p" by auto from fin_succ_hd_b `(x,p) \ g_E G` p_not_in_succ have "x \ hd b " by (auto simp add:step_def exec_def succs_def) with `x\ set b` have "x \ set (tl b)" using b_n_nil by (induct b) auto with ww have "x \ set w" using fin by auto with `(x,p) \ g_E G` and ` x < p` have "(\x\ set w. (x,p)\ g_E G \ x < p)" by auto with cond1 sorted_ss_p show ?thesis using stable_ss_p w_n_nil by auto qed qed next case False then have w_n_nil: "w =[]" by auto with set_ww have "tl b = []" and succ_hd_b_eq: "\(q, t) \ set ?qs_a. t \\<^bsub>f\<^esub> a ! q = a ! q" by auto from succ_hd_b_eq propa have "\p(q, t)\set ?qs_a. t \\<^bsub>f\<^esub> a ! q = a ! q " and " propa f ?qs_a a (tl b) = (ss, w) " and p_lt_n: " p < n " show "ss ! p = a ! p" proof(cases "p \ succs (hd b)") case True with fin_succ_hd_b propa_ss1 have ss_p_eq: "ss!p = transf (hd b) (a!hd b) \\<^bsub>f\<^esub> a ! p" by (auto simp add:step_def exec_def) with ass_eq `p \ succs (hd b)` fin_succ_hd_b have "transf (hd b) (a!hd b) \\<^bsub>f\<^esub> a ! p = a ! p" by (auto simp add:step_def exec_def) with ss_p_eq show ?thesis by auto next case False with fin_succ_hd_b exec2 p_lt_n n_def len_eq nodes_def show ?thesis by (auto simp add:step_def exec_def) qed qed then have "\p< length ss. ss!p = a!p" using `length ss = n` by auto then have ss_eq_a: "ss = a" using n_def len_eq nodes_def `length ss = n` by (auto simp add:list_eq_iff_nth_eq) with wf_a_b p_lt_n have t3: "(?ss_p \ rev [0.. (\x\set ?ss_p. x < p))" and t4: "(?ss_p = rev [0.. (\x\set b. (x, p) \ g_E G \ x < p))" and sta_temp: "(p \ set b \ stable r step ss p)" by (auto simp add:wf_dom_def) from `tl b = []` `b \ []` have "p \ set b \ p \ hd b " by (induct b) auto with sta_temp have "p \ hd b \ stable r step ss p" by auto with hd_b_ss_sta have "stable r step ss p" by auto then have sta_temp': "p \ set w \ stable r step ss p" using w_n_nil by auto have "?ss_p \ (rev [0..?ss_p \ (rev [0..x\set b. (x, p) \ g_E G \ x < p)" by auto then obtain x where x: "x \ set b \ (x, p) \ g_E G \ x < p" by auto with `tl b = []` `b \ []` have "x =hd b" by (induct b) auto with x have " (hd b, p) \ g_E G" and hdb_lt_p: "hd b < p" by auto then have "p \ succs (hd b)" by (simp add:succs_def) with succ_hd_b_eq have transf_hd_b_ap: "transf (hd b) (a!hd b) \\<^bsub>f\<^esub> a ! p = a ! p" using fin_succ_hd_b by (auto simp add:step_def exec_def) have a_p_neq_all: "?a_p \ (rev [0.. ?a_p \ (rev [0..\<^bsub>f\<^esub> ?a_p = (hd b # ?a_hdb) \\<^bsub>f\<^esub> (rev [0..\<^bsub>f\<^esub>?a_p = ( (inter_sorted_rev (hd b # ?a_hdb) (rev [0.. set (hd b # ?a_hdb)" by (auto simp add: inter_sorted_correct) with tx have "set (rev [0.. set (hd b # ?a_hdb)" by auto then have ty: "{0.. set (hd b # ?a_hdb)" by auto from hdb_subset_n ty have "{0.. 0" by auto with tz have tzz: "hd b = n - 1" by (induct n) auto from p_lt_n tzz hdb_lt_p show False by auto qed with ss_p_all ss_eq_a show False by auto qed with sta_temp' sorted_ss_p t3 show ?thesis by auto qed qed from ss_inA wf_dom_2 sorted_w len_ss_n w_lt_n show ?thesis by auto qed lemma propa_dom_invariant_aux': fixes a b ss w assumes propa: "propa f (step (hd b) (a ! hd b)) a (tl b) = (ss, w) " and b_n_nil: "b \ [] " and a_in_A: "\\\set a. \ \ A " and ass: "\p (a ! p \ rev [0.. (\x\set ( (a ! p)). x < p)) \ (a ! p = rev [0.. (\x\ set b. (x,p)\ g_E G \ x < p)) \ (p \ set b \ stable r step a p)" and sorted_b: "sorted b " and n_len: "n = length ss0 " and len_eq: "length a = length ss0 " and b_lt_n: "\x\set b. x < length ss0 " shows "(\\\set ss. \ \ A) \ (\p (ss ! p \ rev [0.. (\x\set (ss ! p). x < p)) \ (ss ! p = rev [0.. (\x\ set w. (x,p)\ g_E G \ x < p)) \ (p \ set w \ stable r step ss p)) \ sorted w \ length ss = length ss0 \ (\x\set w. x < length ss0) " using assms by (auto dest: propa_dom_invariant_aux) lemma propa_dom_invariant: assumes wf_ss_w: "wf_dom ss w " and w_n_nil: "w \ []" and propa: "propa f (step (hd w) (ss ! hd w)) ss (tl w) = (ss', w') " shows "wf_dom ss' w'" using assms proof- from wf_ss_w have ass: "(\p< n. sorted (rev (ss!p)) \ (ss!p \ rev [0..< n] \ (\x\set (ss!p). x < p)) \ (ss!p = rev [0..< n] \ (\x\ set w. (x,p)\ g_E G \ x < p)) \ (p \ set w \ stable r step ss p)) " and sorted_b: "sorted w" and a_in_A: "\\\set ss. \ \ A" and len_eq: "length ss = n" and b_lt_n: "(\x\set w. x < n)" by (auto simp add:wf_dom_def) from propa w_n_nil a_in_A ass sorted_b len_eq b_lt_n show ?thesis by (unfold wf_dom_def) (rule propa_dom_invariant_aux) qed lemma step_dom_mono_aux: fixes \ p \' a b assumes sorted: "sorted (rev (transf p \)) " and a_b_step: "(a, b) \ set (step p \) " and "\ \ A " and " p < n " and " \ \\<^bsub>r\<^esub> \'" shows "\\''. (a, \'') \ set (step p \') \ b \\<^bsub>r\<^esub> \''" proof- have step1: "step p \ = map (\pc. (pc, (transf p \))) (rev (sorted_list_of_set(succs p)))" by (simp add:step_def exec_def) from a_b_step have "set (step p \) \ {}" by auto with step1 have succ_p_n_nil: "(rev (sorted_list_of_set(succs p))) \ []" by auto from `p set (g_V G)" using n_def nodes_def verts_set len_verts_gt0 by auto then have fin: "finite (succs p)" using fin_succs by auto with step1 have "\(x,y)\ set (step p \). x \ succs p" and step2: "\(x,y)\ set (step p \). y = transf p \" by (auto simp add:step_def exec_def) with a_b_step have "a \ succs p" by auto then have "succs p \ {}" by auto from step2 a_b_step have b: "b = transf p \" by auto have step2: "step p \' = map (\pc. (pc, (transf p \'))) (rev (sorted_list_of_set(succs p)))" by (simp add:step_def exec_def) with fin have g1: "\(x,y)\ set (step p \'). x \ succs p" and g2: "\(x,y)\ set (step p \'). y = transf p \'" by (auto simp add:step_def exec_def) with `a \ succs p` have "\t. (a,t)\ set (step p \')" using fin by (auto simp add:step_def exec_def) then obtain t where ex: "(a,t)\ set (step p \')" by auto with g2 have t: "t = transf p \'" by auto from` \ \\<^bsub>r\<^esub> \'` have g: "sorted (rev \) \ sorted (rev \')\ set \' \ set \ \ \ = \'" by (auto simp add:r_def lesssub_def lesub_def nodes_le_def) then have subset_p: "set (p#\') \ set (p# \)" and "set \' \ set \" by auto from sorted have "\x\ set \. x < p" and "sorted (rev \')" using g by (auto simp add:sorted_wrt_append transf_def) with `set \' \ set \` have "\x\ set \'. x < p" by auto with `sorted (rev \')` have "sorted (rev (p#\'))" by (auto simp add:sorted_wrt_append) with sorted b t subset_p have "b \\<^bsub>r\<^esub> t" by (auto simp add:r_def lesssub_def lesub_def nodes_le_def transf_def) with ex show ?thesis by auto qed lemma step_dom_mono: "(\\ p \'. \ \ A \ p \ \\<^sub>r \' \ sorted (rev (transf p \)) \ set (step p \) {\\<^bsub>r\<^esub>} set (step p \'))" apply(unfold mono_def lesubstep_type_def) by(auto simp add:step_dom_mono_aux n_def nodes_def transf_def) lemma propa_termination: fixes a b assumes wf_a_b: "wf_dom a b" and b_n_nil: "b \ [] " shows "(propa f (step (hd b) (a ! hd b)) a (tl b), a, b) \ {(ss', ss). ss [\\<^bsub>r\<^esub>] ss'} <*lex*> (\x. case x of (x, y) \ (sorted_list_of_set x, sorted_list_of_set y)) ` finite_psubset" proof- let ?qs = "(step (hd b) (a ! hd b))" from wf_a_b have "\x\set b. x < n" and n_len: "length a = n" and sorted_b: "sorted b" and set_a: "set a \ A" by (auto simp add:wf_dom_def) then have sorted_tl_b: "sorted (tl b)" and hd_b_lt_n: "hd b < n" using b_n_nil by (induct b) (auto simp add:sorted_wrt_append) from set_a have a_inA: "a \ list n A" using n_len by (auto intro: listI) from hd_b_lt_n have dist: "distinct (map fst ?qs)" using distinct_p by auto from wf_a_b b_n_nil have step_pres_bounded: "\(q, \)\set ?qs. q < n \ \ \ A" using propa_dom_invariant_auxi n_len by fastforce with sorted_tl_b set_a n_len have propa: "propa f ?qs a (tl b) = (merges f ?qs a, (sorted_list_of_set ({q. \t.(q,t)\set ?qs \ t \\<^bsub>f\<^esub> (a!q) \ a!q} \ set(tl b))))" by (auto dest: decomp_propa) from a_inA step_pres_bounded sorted_b b_n_nil dist have "((merges f ?qs a, (sorted_list_of_set ({q. \t.(q,t)\set ?qs \ t \\<^bsub>f\<^esub> (a!q) \ a!q} \ set(tl b)))),(a,b)) \ {(ss', ss). ss [\\<^bsub>r\<^esub>] ss'} <*lex*> (\x. case x of (x, y) \ (sorted_list_of_set x, sorted_list_of_set y)) ` finite_psubset" by (auto simp add: finite_psubset_def dest: termination_lemma) with propa show ?thesis by auto qed lemma iter_dom_invariant: assumes "wf_dom ss0 w0" shows "iter f step ss0 w0 = (ss',w') \ wf_dom ss' w'" using assms apply (unfold iter_def) apply(rule_tac P = "(\(ss, w). wf_dom ss w )" and r = "{(ss',ss). ss [\\<^sub>r] ss'} <*lex*> (\(x,y). (sorted_list_of_set x, sorted_list_of_set y)) ` finite_psubset" in while_rule) \ \Invariant holds initially:\ apply (clarsimp) \ \Invariant is preserved:\ apply(simp add: wf_dom_def) apply clarsimp apply(rule propa_dom_invariant_aux') apply assumption+ \ \Postcondition holds upon termination:\ apply clarsimp \ \Well-foundedness of the termination relation:\ apply (simp add: wf_listn_termination_rel) \ \Loop decreases along termination relation:\ apply clarsimp apply (fastforce intro:propa_termination) done lemma propa_dom_invariant_aux2: fixes ss w ssa wa assumes wf_dom_ss0_w0: "wf_dom ss0 w0" and w_n_nil: "w \ [] " and propa: "propa f (step (hd w) (ss ! hd w)) ss (tl w) = (ssa, wa) " and wf_ss_w: "wf_dom ss w " and ss0_lt_ss: "ss0 [\\<^bsub>r\<^esub>] ss" and sta: " \ts\list n A. ss0 [\\<^bsub>r\<^esub>] ts \ (\p ss [\\<^bsub>r\<^esub>] ts" shows "ss0 [\\<^bsub>r\<^esub>] ssa \ (\ts\list n A. ss0 [\\<^bsub>r\<^esub>] ts \ (\p ssa [\\<^bsub>r\<^esub>] ts)" using assms proof- let ?ss_hdw = "ss!hd w" from wf_dom_ss0_w0 have len_ss0: "length ss0 = n" and "\x\ set ss0. x \ A" by (auto simp add:wf_dom_def) then have ss0_inA: "ss0 \ list n A" and set_ss0: "set ss0 \ A"by (auto intro:listI) then have ss0_nth_inA: "\i A" by auto then have ss0_p_subset: "\p< length ss0. set (ss0!p) \ set nodes" by (auto simp add:inpow_subset_nodes A_def) from len_ss0 n_def nodes_def len_verts_gt0 have "0 < length ss0" by auto from ss0_lt_ss have "list_all2 (\x y. x \\<^sub>r y) ss0 ss" by (auto simp add:lesssub_def lesub_def list_all2_def Listn.le_def) then have lt1: "\p\<^bsub>r\<^esub> ss!p" by (auto simp add: list_all2_conv_all_nth) from wf_ss_w have "\\\set ss. \ \ A" and ass1: "\p (ss!p \ rev [0.. (\x\set (ss ! p). x < p)) \ (ss!p = rev [0.. (\x\ set w. (x,p)\ g_E G \ x < p)) \ (p \ set w \ stable r step ss p)" and sorted_w: "sorted w" and len_ss: "length ss = n" and w_lt_n: "\x\set w. x < n "by (auto simp add:wf_dom_def) then have ss_inA: "ss \ list n A" and set_ss: "set ss \ A" by (auto intro:listI) then have ss_nth_inA: "\i A" by auto then have ss_p_subset: "\p< length ss. set (ss!p) \ set nodes" by (auto simp add:inpow_subset_nodes A_def) then have ss_hdw_nodes: "set ?ss_hdw \ set nodes" using w_lt_n w_n_nil len_ss by auto let ?qs = "step (hd w) ?ss_hdw" from w_lt_n have hd_w_lt_n: "hd w < n" using w_n_nil by auto then have dist: "distinct (map fst ?qs)" using distinct_p by auto from ss_nth_inA have ss_hdw_inA: "?ss_hdw \ A" using w_lt_n w_n_nil len_ss by auto from ass1 have sorted_ss_hdw: "sorted (rev ?ss_hdw)" using w_lt_n w_n_nil by auto then have "\q\succs (hd w). q \ set (g_V G)" by (auto simp add:succ_in_G) then have hd_w_suc_lt_n: "\q\succs (hd w). q < n" using n_def verts_set nodes_def by auto have hdw_in_nodes:"hd w \ set (g_V G)" using verts_set n_def nodes_def w_lt_n w_n_nil by auto then have fin_succ_hd_w: "finite (succs (hd w))" using fin_succs by auto from sorted_w have sorted_tl_w': "sorted (tl w)" using w_n_nil by (induct w) auto from wf_ss_w w_n_nil have ss_hd_w_neq_all: "?ss_hdw \ (rev [0.. set nodes" and hd_w_ss_in_A: " (hd w # ?ss_hdw) \ A" and step_pres_bounded: " \(q, \)\set (step (hd w) ?ss_hdw). q < length ss \ \ \ A" using propa_dom_invariant_auxi by auto from ss_hd_w_neq_all have ss_lt_hd_w: "\x\set ?ss_hdw. x < hd w" using hd_w_lt_n wf_ss_w by (auto simp add:wf_dom_def) from wf_ss_w w_n_nil propa have wf_ssa_wa: "wf_dom ssa wa" using propa_dom_invariant by auto then have sorted_ssa_p: "\px\ set ssa. x \ A" and sorted_wa: "sorted wa" and len_ssa: "length ssa = n" and wa_lt_n: "\x\ set wa. x < n" by (auto simp add:wf_dom_def) then have ssa_inA: "ssa\ list n A" and set_ssa: "set ssa \ A"by (auto intro:listI) then have ssa_nth_inA: "\i A" by auto then have ssa_p_subset: "\p< length ssa. set (ssa!p) \ set nodes" by (auto simp add:inpow_subset_nodes A_def) from len_ss0 len_ssa have len_ss0_ssa: "length ss0 = length ssa" by simp from len_ss0 len_ss have len_ss0_ss: "length ss0 = length ss" by simp have "\(q, \)\set ?qs. \ = transf (hd w) (ss!(hd w))" by (simp add:step_def exec_def) then have transf_ss_hd_w: "\(q, \)\set ?qs. \ = (hd w # ?ss_hdw)" by (simp add:transf_def) from set_ss step_pres_bounded sorted_tl_w' len_ss dist have "\(q, \) \ set ?qs. (fst(propa f ?qs ss (tl w)))!q = \ \\<^bsub>f\<^esub>ss!q" by (auto dest:propa_property1) with propa have propa_ss: "\(q, \) \ set ?qs. ssa!q = \ \\<^bsub>f\<^esub>ss!q" by simp with transf_ss_hd_w have propa_ss1: "\(q, \) \ set ?qs. ssa!q = (hd w # ?ss_hdw) \\<^bsub>f\<^esub>ss!q" by auto from ss_nth_inA step_pres_bounded have "\(q, \) \ set ?qs. ss!q \ A" using hd_w_suc_lt_n fin_succ_hd_w by (auto simp add:step_def exec_def) from hd_w_ss_in_A this propa_ss1 have ss_lt_ssa_q: "\(q, \) \ set ?qs. ss!q \\<^bsub>r\<^esub> ssa!q" by (fastforce dest:Semilat.ub2[OF Semilat.intro, OF is_semi]) from step_pres_bounded sorted_tl_w' set_ss len_ss dist have "\q. q \ set(map fst ?qs) \ q < length ss \ (fst(propa f ?qs ss (tl w)))!q = ss!q" by (auto intro:propa_property2) with propa have exec2: "\q. q \ set(map fst ?qs) \ q < length ss \ ssa!q = ss!q" by auto have tran_ss_ssa: "\p\<^bsub>r\<^esub> ssa!p" proof(intro strip) fix p assume "p < length ss" with ssa_nth_inA have ssa_p_inA: "ssa!p \ A" using `length ssa = n` `length ss = n` by auto from ss_nth_inA have ss_p_inA: "ss!p \ A" using `length ss = n` `p < length ss` by auto show " ss ! p \\<^bsub>r\<^esub> ssa ! p" proof(cases "p \ succs (hd w)") case True then show "ss!p \\<^bsub>r\<^esub> ssa!p" using ss_lt_ssa_q using fin_succ_hd_w by (auto simp add:step_def exec_def) next case False then have "p \ set (map fst ?qs)" using fin_succ_hd_w by (auto simp add:step_def exec_def) then show ?thesis using exec2 `p < length ss` using ssa_p_inA ss_p_inA by(auto simp add:step_def exec_def intro: Semilat.orderI[OF Semilat.intro, OF is_semi]) qed qed then have "\p\<^bsub>r\<^esub> ssa ! p" using len_ss0_ss by auto with lt1 have "\p\<^bsub>r\<^esub> ssa!p" using ss0_nth_inA ssa_nth_inA ss_nth_inA using len_ss0_ss len_ss0_ssa by (auto intro: order_trans Semilat.orderI[OF Semilat.intro, OF is_semi]) with len_ss0_ssa have g1: "ss0 [\\<^bsub>r\<^esub>] ssa" by (auto simp only:Listn.le_def lesssub_def lesub_def intro:list_all2_all_nthI) have "(\ts\list n A. ss0 [\\<^bsub>r\<^esub>] ts \ (\p ssa [\\<^bsub>r\<^esub>] ts)" proof(intro strip) fix ts assume ts_inA: "ts \ list n A" and ass: "ss0 [\\<^bsub>r\<^esub>] ts \ (\p\<^bsub>r\<^esub>] ts" and sta_ts: "(\p\<^bsub>r\<^esub> ?ts_hdw "using le_listD len_ss hd_w_lt_n by auto then have "sorted (rev (?ts_hdw)) \ set ?ts_hdw \ set ?ss_hdw \ ?ss_hdw =?ts_hdw " by (auto simp add:r_def lesssub_def lesub_def nodes_le_def) then have sorted_ts_hdw: "sorted (rev (?ts_hdw))" and ts_ss_subset: "set ?ts_hdw \ set ?ss_hdw" using sorted_ss_hdw by (auto simp add:r_def lesssub_def lesub_def nodes_le_def) then have "\x\ set ?ts_hdw. x < hd w" using ss_lt_hd_w by auto with sorted_ts_hdw have sorted_hd_w_ts: "sorted (rev (hd w # ?ts_hdw))" by (auto simp add:sorted_wrt_append) with sorted_hd_w_ss ts_ss_subset have "(hd w # ?ss_hdw) \\<^bsub>r\<^esub> (hd w # ?ts_hdw)" by (auto simp add:transf_def lesssub_def lesub_def r_def nodes_le_def) then have transf_ss_ts: "transf (hd w) ?ss_hdw \\<^bsub>r\<^esub> transf (hd w) ?ts_hdw" by (auto simp add:transf_def) from sta_ts hd_w_lt_n have sta_ts_hdw: "stable r step ts (hd w)" by auto from ss_hdw_nodes ts_ss_subset have "set ?ts_hdw \ set nodes" by auto with hd_w_lt_n have hdw_ts_subset_nodes: "set (hd w # ?ts_hdw) \ set nodes" using nodes_def n_def verts_set by auto with sorted_hd_w_ts have "(hd w # ?ts_hdw) \ ( (rev \ sorted_list_of_set) ` (Pow (set nodes)))" by (fastforce intro: subset_nodes_inpow) then have transf_ts_inA: "(hd w #?ts_hdw) \ A" by (simp add:A_def) then have sorted_hdw_ts_hdw: "sorted (rev (hd w # ?ts_hdw))" by (rule inA_is_sorted) have "\p\<^bsub>r\<^esub> ts!p" proof(intro strip) fix p assume p_lt_len_ssa: "p < length ssa" then have "p < length ss" and "p < n" using len_ssa len_ss by auto with ss_ts have ss_ts_p: "ss!p \\<^bsub>r\<^esub> ts!p " using le_listD by auto show "ssa ! p \\<^bsub>r\<^esub> ts ! p" proof(cases "p \ succs (hd w)") case True note p_in_succs_hdw = this then have "ss!p \\<^bsub>r\<^esub> ssa!p" using ss_lt_ssa_q using fin_succ_hd_w by (auto simp add:step_def exec_def) from p_in_succs_hdw have ssa_p: "ssa!p = transf (hd w) (ss!hd w) \\<^bsub>f\<^esub>ss!p" using propa_ss fin_succ_hd_w by (auto simp add:step_def exec_def) from sta_ts_hdw have transf_hdw_ts_hdw: "transf (hd w) (ts!hd w) \\<^bsub>r\<^esub> ts!p" using p_in_succs_hdw fin_succ_hd_w by (auto simp add:step_def exec_def stable_def) then have ts_p_subset: "(hd w # ?ts_hdw) \\<^bsub>r\<^esub> ts!p" by (auto simp add:transf_def) then have "(sorted (rev (ts!p)) \ set (ts!p)\ set (hd w # ?ts_hdw)\ sorted (rev (hd w # ?ts_hdw))) \ hd w # ?ts_hdw = ts!p" by (auto simp add:r_def lesssub_def lesub_def nodes_le_def) then have "sorted (rev (ts!p)) \ set (ts!p)\ set (hd w # ?ts_hdw)" proof assume "sorted (rev (ts ! p)) \ set (ts ! p) \ set (hd w # ts ! hd w) \ sorted (rev (hd w # ts ! hd w))" then show ?thesis by auto next assume " hd w # ts ! hd w = ts ! p" with sorted_hdw_ts_hdw show ?thesis by auto qed then have sorted_ts_p: "sorted (rev (ts!p))" and ts_as_subset: "set (ts!p)\ set (hd w # ?ts_hdw)" by auto with hdw_ts_subset_nodes have "set (ts!p) \ set nodes" by auto with sorted_ts_p have "(ts!p) \ ( (rev \ sorted_list_of_set) ` (Pow (set nodes)))" by (fastforce intro: subset_nodes_inpow) then have ts_p_inA: "ts!p \ A" by (simp add:A_def) from sorted_hdw_ts_hdw have "\x\ set ?ts_hdw. x < hd w" by (auto simp add:sorted_wrt_append) with `hd w < n` have "\x\ set ?ts_hdw. x < n" by auto then have "set(hd w # ?ts_hdw) \ set nodes"using `hd w < n` n_def verts_set nodes_def by auto with sorted_hdw_ts_hdw have "hd w # ?ts_hdw \ ( (rev \ sorted_list_of_set) ` (Pow (set nodes)))" by (fastforce intro: subset_nodes_inpow) then have "(hd w # ?ts_hdw) \ A" by (auto simp add:A_def) then have trans_hdw_ts_inA: "transf (hd w) (ts!hd w) \ A" by (auto simp add:transf_def) have transf_hdw_ss_inA: "transf (hd w) ?ss_hdw \ A" using hd_w_ss_in_A by (auto simp add:transf_def) have ss_p_inA: "ss!p \ A" using `p\<^bsub>r\<^esub> ts ! p" by (auto intro: order_trans Semilat.orderI[OF Semilat.intro, OF is_semi]) with `ss!p \\<^bsub>r\<^esub> ts ! p` trans_hdw_ts_inA ss_p_inA transf_hdw_ss_inA ssa_p ts_p_inA show "ssa ! p \\<^bsub>r\<^esub> ts ! p" by (auto intro: Semilat.lub[OF Semilat.intro, OF is_semi]) next case False then have "p \ set (map fst ?qs)" using fin_succ_hd_w by (auto simp add:step_def exec_def) then have "ssa!p = ss!p"using exec2 p_lt_len_ssa len_ss len_ssa by(auto simp add:step_def exec_def) with ss_ts_p show ?thesis by auto qed qed with `length ss = length ts` len_ss len_ssa show "ssa [\\<^bsub>r\<^esub>] ts" by (auto simp only:Listn.le_def lesssub_def lesub_def intro:list_all2_all_nthI) qed with g1 show ?thesis by auto qed lemma in_list_nA_refl: "ss \ list n A \ ss [\\<^bsub>r\<^esub>] ss" apply (unfold Listn.le_def lesssub_def lesub_def) proof- assume "ss \ list n A" then have "set ss \ A" and "length ss = n" by auto then have "\i A" by auto then have "\i\<^bsub>r\<^esub> ss!i" by (auto simp add:r_def lesssub_def lesub_def nodes_le_def ) then show " list_all2 r ss ss" by (auto simp add:lesssub_def lesub_def intro: list_all2_all_nthI) qed lemma iter_dom: assumes "wf_dom ss0 w0" shows "iter f step ss0 w0 = (ss',w') \ wf_dom ss' w' \ stables r step ss' \ ss0 [\\<^sub>r] ss' \ (\ts\list n A. ss0 [\\<^sub>r] ts \ stables r step ts \ ss' [\\<^sub>r] ts)" using assms apply (unfold iter_def stables_def) apply (rule_tac P = "\(ss,w). wf_dom ss w \ ss0 [\\<^sub>r] ss \ (\ts\list n A. ss0 [\\<^sub>r] ts \ stables r step ts \ ss [\\<^sub>r] ts)" and r = "{(ss',ss). ss [\\<^sub>r] ss'} <*lex*> (\(x,y). (sorted_list_of_set x, sorted_list_of_set y)) ` finite_psubset" in while_rule) \ \Invariant holds initially:\ apply(fastforce simp add: wf_dom_def intro:wf_dom_le) \ \Invariant is preserved:\ apply clarsimp apply (rule conjI) apply(fastforce dest:propa_dom_invariant) apply(simp add:stables_def) apply (rule propa_dom_invariant_aux2) apply assumption+ \ \Postcondition holds upon termination:\ apply(clarsimp simp add: stables_def split_paired_all) apply(subgoal_tac "length ss0 = n") apply(simp add:wf_dom_def)+ \ \Well-foundedness of the termination relation:\ apply clarsimp apply(simp add:wf_listn_termination_rel) \ \Loop decreases along termination relation:\ apply clarsimp apply (fastforce intro:propa_termination) done lemma wf_start: "wf_dom start (unstables r step start)" proof- let ?w0 = "unstables r step start" have sorted_w: "sorted ?w0" using unstables_def by (simp add:sorted_less_sorted_list_of_set) have w0_lt_n: "\x\set ?w0. x < n" using unstables_def len_start_is_n by auto have neq_all: "\p < n. start!p \ rev [0..< n] \ (\x\set (start!p). x < p) " proof(intro strip) fix p x assume p_lt_n: "p < n" and p_neq_all: "start ! p \ rev [0..< n]" and x_in: "x \ set (start ! p)" then have "p = 0" using start_nth_lt0_all len_start_is_n by auto with start_nth0_empty show "x < p" using x_in by auto qed have eq_all:"(\p < n. start!p = rev [0..< n] \ (\x\ set ?w0. (x,p)\ g_E G \ x < p))" proof(intro strip) fix p assume p_lt_n: "p < n" and p_eq_all: "start ! p = rev [0..< n]" from `p < n` have "p = 0 \ p > 0 \ p < length start" using len_start_is_n by auto with p_eq_all have "p > 0" and p_lt_len_start: "p < length start" using start_nth0_empty n_def nodes_def len_verts_gt0 by auto then have "p \ set (g_V G) - {0}" using len_start_is_n n_def nodes_def verts_set by auto with dfst obtain prev where "(prev, p) \ g_E G" and "prev < p" by auto then have "succs prev \ {}" and "prev < length start" using p_lt_len_start by (auto simp add:succs_def) with unstable_start have "prev \ set ?w0" by auto with `(prev, p) \ g_E G` `prev < p` show "\x\set (unstables r step start). (x, p) \ g_E G \ x < p" by auto qed have "\p set (unstables r step start)\ stable r step start p)" by (unfold unstables_def) (simp add:n_def start_def nodes_def) with sorted_start_nth neq_all eq_all start_subset_A sorted_w len_start_is_n w0_lt_n show ?thesis by (auto simp only:wf_dom_def) qed lemma iter_dom_properties: "iter f step start (unstables r step start) = (ss',w') \ wf_dom ss' w' \ stables r step ss' \ start [\\<^sub>r] ss' \ (\ts\list n A. start [\\<^sub>r] ts \ stables r step ts \ ss' [\\<^sub>r] ts)" using iter_dom[OF wf_start] by auto lemma iter_dom_properties2: "iter f step start (unstables r step start) = (ss',w') \ ss' \ list n A" using iter_dom_properties by (auto simp only:wf_dom_def list_def) lemma iter_dom_termination: "iter f step start (unstables r step start) = (ss,w) \ w \ [] \ propa f (step (hd w) (ss!(hd w))) ss (tl w) = (ss, tl w)" proof (intro strip) assume iter: "iter f step start (unstables r step start) = (ss, w)" and w_n_nil: "w \ []" with iter_dom_properties have stas: "stables r step ss" and wf_ss_w: "wf_dom ss w" and start_le_ss: "start [\\<^sub>r] ss" by auto from stas have sta_p: "(\p < size ss. stable r step ss p)"by (auto simp add:stables_def) from wf_ss_w have n_w_sta: "\pset w \ stable r step ss p" and len_eq: "length ss = n" and w_lt_n: "\x\ set w. x < n" and ss_inA: "\x\set ss. x \ A" and sorted_w: "sorted w " by (auto simp add:wf_dom_def) from w_lt_n have hd_w_lt_n: "hd w < n" using w_n_nil by auto then have hd_w_in_verts: "hd w \ set (g_V G)" using n_def nodes_def verts_set by auto then have fin_succ_hd_w: "finite (succs (hd w))" using fin_succs hd_w_in_verts by auto let ?ss_hdw = "ss!hd w" let ?qs = "step (hd w) ?ss_hdw" from hd_w_lt_n have dist: "distinct (map fst ?qs)" using distinct_p by auto from wf_ss_w w_n_nil have hdw_ss_subset_nodes: "set (hd w # ?ss_hdw) \ set nodes" and sorted_hd_w_ss: "sorted (rev (hd w # ?ss_hdw))" and hd_w_ss_in_A: " ((hd w # ?ss_hdw)) \ A" and step_pres_bounded: "\(q, \)\set ?qs. q < length ss \ \ \ A" using propa_dom_invariant_auxi by auto let ?res = "propa f (step (hd w) (ss!(hd w))) ss (tl w) " have "propa f (step (hd w) (ss!(hd w))) ss (tl w) = ?res" by simp then obtain ss' w' where propa: "propa f (step (hd w) (ss!(hd w))) ss (tl w) = (ss', w')" by (cases ?res) auto from sorted_w have sorted_tl_b: "sorted (tl w)" by (induct w) auto from ss_inA have set_a: "set ss \ A" by auto with step_pres_bounded sorted_tl_b len_eq dist have "\(q, \) \ set ?qs. (fst(propa f ?qs ss (tl w)))!q = \ \\<^bsub>f\<^esub>ss!q" by (auto dest:propa_property1) with propa have propa_ss1: "\(q, \) \ set ?qs. ss'!q = (hd w # ?ss_hdw) \\<^bsub>f\<^esub>ss!q" by (simp add:step_def exec_def transf_def) from step_pres_bounded sorted_tl_b set_a len_eq dist have "\q. q \ set(map fst ?qs) \ q < length ss \ (fst(propa f ?qs ss (tl w)))!q = ss!q" by (auto intro:propa_property2) with propa have g2: "\q. q \ set(map fst ?qs) \ q < length ss \ ss'!q = ss!q" by auto from sorted_w have sorted_tl_w: "sorted (tl w)" by (induct w) auto with step_pres_bounded set_a have fst_propa: "fst (propa f ?qs ss (tl w)) = (merges f ?qs ss)" and snd_propa: "snd (propa f ?qs ss (tl w)) = (sorted_list_of_set ({q. \t.(q,t)\set ?qs \ t \\<^bsub>f\<^esub> (ss!q) \ ss!q} \ set (tl w)))" using decomp_propa by auto with propa have len_ss_eq_ss': "length ss' = length ss" using length_merges by auto have ss_ss'_eq: "\i set(map fst ?qs)") case True assume ass1: "i \ set (map fst ?qs)" with propa_ss1 have ss': "ss'!i = (hd w # ?ss_hdw) \\<^bsub>f\<^esub>ss!i" by (auto simp add: step_def exec_def) from ass1 have "i \ set (g_V G)" using succ_in_G fin_succ_hd_w by (auto simp add:step_def exec_def) then have q_lt_len_ss: "i < length ss" using len_eq by (auto simp add:n_def nodes_def verts_set) from hd_w_lt_n len_eq stas have "stable r step ss (hd w)" by (auto simp add:stables_def) with ass1 have "(hd w # ?ss_hdw) \\<^sub>r ss!i" by (auto simp add:stables_def stable_def step_def exec_def transf_def) then have "set (ss!i) \ set (hd w # ?ss_hdw) \ (hd w # ?ss_hdw) = ss!i" by (auto simp add:lesssub_def lesub_def r_def nodes_le_def) then have set_ss_q_subst_hdw_ss: "set (ss!i) \ set (hd w # ?ss_hdw)" by (rule disjE)(auto simp add:lesssub_def lesub_def r_def nodes_le_def) then have ss_q: "set (ss!i) \ set (hd w # ?ss_hdw) = set (ss!i)" by auto from wf_ss_w q_lt_len_ss have sorted_ss_q: "sorted (rev (ss!i))" by (simp add:wf_dom_def) with sorted_hd_w_ss have ss_q': "set (ss'!i) = set (ss!i) \ set (hd w # ?ss_hdw)" and sorrted_ss_q': "sorted (rev (ss'!i))" using ss' by (auto simp add:plussub_def f_def nodes_sup_def inter_sorted_correct) with ss_q sorted_ss_q sorrted_ss_q' show ?thesis using sorted_less_rev_set_unique propa by auto next case False note i_nin = this from step_pres_bounded sorted_tl_b set_a len_eq dist propa i_lt_len_ss i_nin show ?thesis by (fastforce dest:propa_property2) qed qed with len_ss_eq_ss' have eq_ss: "ss' = ss" using len_eq propa by (auto simp add:list_eq_iff_nth_eq) then have qs_empty: "(({q. \t.(q,t)\set ?qs \ t \\<^bsub>f\<^esub> (ss!q) \ ss!q} \ set (tl w))) = (set (tl w))" using propa_ss1 propa transf_def step_def exec_def by fastforce with snd_propa have "snd (propa f ?qs ss (tl w)) = sorted_list_of_set (set (tl w))" using sorted_w by auto with sorted_tl_w have "snd (propa f ?qs ss (tl w)) = tl w" by (fastforce dest:sorted_less_set_eq) with propa have "w' = tl w" by simp with eq_ss show "propa f (step (hd w) (ss ! hd w)) ss (tl w) = (ss, tl w)" using propa by auto qed lemma dom_iter_properties: "iter f step start (unstables r step start) = (ss, w) \ ss!0 = [] \ (\p 0 \ ss!p \ (rev [0..\<^sub>r] ss" and ss_inA: "ss \ list n A" by auto - then have len_ss_n: "length ss = n" by auto + then have len_ss_n: "length ss = n" by (auto simp only:list_def) from start_le_ss have "start!0 \\<^sub>r ss!0" using start_len_gt_0 by (unfold Listn.le_def lesssub_def lesub_def) (auto simp add:lesssub_def lesub_def intro:list_all2_nthD) then have ss_0th: "ss!0 = []" by (auto simp add:r_def nodes_le_def lesssub_def lesub_def start_def) have "(\p 0 \ ss ! p \ (rev [0..0" and ss_p_eq_all: "\ ss ! p \ rev [0..(q,\) \ set (step p (ss!p)). \ \\<^sub>r ss!q" by (simp add: stables_def stable_def) from p_lt_n len_start_is_n verts_set have p_is_vert: "p \ set (g_V G)" by (auto simp add:n_def nodes_def) then have step_p: "\(q,\) \ set (step p (ss!p)). q \ succs p" using fin_succs by (auto simp add:step_def exec_def) from p_is_vert p_neq_0 dfst obtain prev where "(prev, p) \ g_E G" and prev_lt_p: "prev < p" by auto then have prev_lt_n: "prev < n" and prev_p: "p \ succs prev" and "prev \ set (g_V G)" using p_lt_n tail_is_vert by (auto simp add:succs_def) then have fin_suc_prev : "finite (succs prev)" using fin_succs by auto let ?prev_\ = "ss!prev" from prev_lt_n stas `length ss = n` have "stable r step ss prev" by (auto simp add:stables_def) then have "\(q,\) \ set (step prev ?prev_\). (prev # ?prev_\) \\<^sub>r ss!q" by (auto simp add: stable_def transf_def step_def exec_def) with prev_p have "(prev # ?prev_\) \\<^bsub>r\<^esub> ss ! p" using fin_suc_prev by (auto simp add: stable_def transf_def step_def exec_def) with ss_p_eq_all have "sorted (rev (prev # ?prev_\)) \ {0.. set (prev # ?prev_\) \ (prev # ss ! prev) = (rev [0..)) \ {0.. set (prev # ?prev_\)" by(rule disjE) auto then have sorted_rev: "sorted (rev (prev # ?prev_\))" and pres_subset: "{0.. set (prev # ?prev_\)" by auto then have prev_set: "\x\ set ?prev_\. x < prev" by (induct ?prev_\) (auto simp add:sorted_wrt_append) from p_lt_n prev_lt_p have "prev < n - 1" using n_def nodes_def len_verts_gt0 by auto with prev_set have prev_lt_n: "\x\set(prev # ?prev_\). x < n - 1" by auto from pres_subset have "n - 1 \ set (prev # ?prev_\)" using n_def nodes_def len_verts_gt0 by fastforce with prev_lt_n show False by auto qed with ss_0th show " ss ! 0 = [] \ (\p 0 \ ss ! p \ rev [0.. (\p (rev [0.. (rev [0.. list n A \ stables r step (kildall r f step start) \ start [\\<^sub>r] (kildall r f step start) \ (\ts\list n A. start [\\<^sub>r] ts \ stables r step ts \ (kildall r f step start) [\\<^sub>r] ts)" (is "PROP ?P") by (case_tac "iter f step start (unstables r step start)")(simp add: kildall_def iter_dom_properties iter_dom_properties2) lemma dom_kildall_stables: "stables r step (dom_kildall start)" using kildall_dom_properties by(unfold dom_kildall_def nodes_semi_def) (simp add: r_def f_def step_def) lemma dom_kildall_entry: "dom_kildall start !0 = []" by (case_tac "iter f step start (unstables r step start)") (auto simp add:dom_kildall_def nodes_semi_def dom_iter_properties r_def f_def step_def kildall_def) lemma zero_dom_zero: "dom i 0 \ i = 0" using start_def n_def nodes_def dom_kildall_entry by (simp add:dom_def) (*lemma sadom_succs: "strict_dom i j \ j \ succs k \ strict_dom i k"*) lemma sdom_dom_succs: "strict_dom i j \ j \ succs k \ dom i k" proof - assume sdom_i_j: "strict_dom i j" and k_j: "j \ succs k" then have j: "j \ set (g_V G)" and k: "k \ set (g_V G)" using verts_set succ_in_verts by auto then have j_lt_n: "j < n" and k_lt_n: "k < n" using n_def nodes_def verts_set by auto have fin_succs_k: "finite (succs k)" using fin_succs k by auto with k_j have k_j2: "j \ set (rev (sorted_list_of_set(succs k)))" by auto let ?ss0 = "start" let ?w0 = "unstables r step start" let ?res = "dom_kildall start" have dom_kil: "?res = kildall r f step ?ss0" by (auto simp add:dom_kildall_def r_def f_def step_def nodes_semi_def) have sorted_unstables: "sorted ?w0" by (auto simp add:unstables_def sorted_less_sorted_list_of_set) have ss: "?res = fst (iter f step ?ss0 ?w0)" by (auto simp add:dom_kildall_def kildall_def f_def step_def nodes_semi_def start_def r_def) then obtain ww where dom_iter: "iter f step ?ss0 ?w0 = (?res, ww)" by (cases "iter f step ?ss0 ?w0") auto with iter_dom_properties have wf_res: "wf_dom (dom_kildall start) ww" by auto with sdom_i_j have i_dom_j: "i \ set (?res!j)" by (auto simp add:strict_dom_def start_def n_def nodes_def) from wf_res have len_res: "length ?res = n" by (auto simp add:wf_dom_def) show "dom i k" proof(rule ccontr) assume ass: "\ dom i k" then have i_neq_k: "i \ k" by (auto simp add:dom_refl) with ass have "(\res. ?res!k = res \ i \ set res)" using ass by (auto simp add:dom_def start_def nodes_def n_def) then show False proof - assume "\res. ?res ! k = res \ i \ set res" then obtain rs where k_dom: "?res ! k = rs" and i_notin_rs: "i \ set rs" by auto from iter_dom_properties dom_iter have "(\p < length ?res. stable r step ?res p)" by (auto simp add:stables_def) with k_lt_n have "stable r step ?res k" using len_res by auto with k_dom have aux: "\(q,\) \ set (map (\pc. (pc, (k # rs))) (rev (sorted_list_of_set(succs k)))). \ \\<^sub>r ?res!q" by (simp add:stable_def r_def step_def exec_def transf_def) with k_j2 have "(k # rs) \\<^sub>r ?res!j" by auto then have "set (?res!j) \ set (k # rs) \ ?res!j = k#rs" by (auto simp add:lesssub_def lesub_def nodes_semi_def nodes_le_def r_def f_def) then have "set (?res!j) \ set (k # rs)" by auto with i_dom_j i_neq_k have " i \ set rs" by auto with i_notin_rs show False by auto qed qed qed lemma adom_succs: "dom i j \ i \ j \ j \ succs k \ dom i k" by (auto intro: dom_sdom sdom_dom_succs) lemma dom_kildall_is_strict: "j < length start \ dom_kildall start ! j = res \ j \ set res" proof - assume j_dom: "dom_kildall start ! j = res" and j_lt: "j < length start" from j_dom have iter_fst: "(fst (iter f step start (unstables r step start))) ! j = res" by (auto simp add:dom_kildall_def r_def f_def step_def start_def nodes_semi_def kildall_def) then obtain ss w where iter: "iter f step start (unstables r step start) = (ss, w)" by fastforce with iter_fst have res: "ss!j = res" by auto with dom_iter_properties2 iter have res_neq_all: "res \ rev [0..p < n. (ss!p) \ rev [0..< n] \ (\x\set ( (ss!p)). x < p)" by (auto simp add:wf_dom_def) with j_lt len_start_is_n res res_neq_all have "(\x\set res. x < j)" by (auto simp add:wf_dom_def) then show "j \ set res" by auto qed lemma sdom_asyc: "strict_dom i j \ j \ set (g_V G) \ i \ j" proof- assume sdom_i_j: "strict_dom i j" and "j \ set (g_V G)" then have j_lt: "j < length start" using start_def n_def nodes_def verts_set by auto let ?start = " [] # (replicate (length (g_V G) - 1) ( (rev[0.. set (dom_kildall ?start !j)" by (auto simp add:strict_dom_def) from j_lt have j_nin: "j \ set (dom_kildall ?start !j)" using eq_start by (simp add: dom_kildall_is_strict) with i_in show "i \ j" by auto qed lemma propa_dom_invariant_complete: fixes i a b ss w assumes b_n_nil: "b \ [] " and propa: "propa f (step (hd b) (a ! hd b)) a (tl b) = (ss, w) " and wf_a_b: "wf_dom a b" and non_strict: "\i. i < n \ k \ set (a!i) \ non_strict_dominate k i " shows "wf_dom ss w \ (\i. i < n \ k \ set ( ss ! i) \ non_strict_dominate k i)" (is "?PROP ?P") using assms proof- let ?a_hdb = "a!hd b" let ?qs = "step (hd b) (a!hd b)" from wf_a_b b_n_nil propa have wf_ss_w: "wf_dom ss w" using propa_dom_invariant by auto from wf_a_b have "\\\set a. \ \ A" and sorted_b: "sorted b" and len_a: "length a = n" and b_lt_n: "\x\set b. x < n "by (auto simp add:wf_dom_def) then have set_a: "set a \ A" by (auto intro:listI) from sorted_b have sorted_tl_b: "sorted (tl b)" using b_n_nil by (induct b) auto from b_lt_n b_n_nil have hd_b_lt_n: "hd b < n" by auto with n_def nodes_def verts_set have "hd b \ set (g_V G)" using b_n_nil by auto then have fin_succ_hd_b: "finite (succs (hd b))" using fin_succs by auto from wf_a_b b_n_nil have sorted_hd_b_a: "sorted (rev (hd b # ?a_hdb))" and step_pres_bounded: " \(q, \)\set (step (hd b) ?a_hdb). q < length a \ \ \ A" using propa_dom_invariant_auxi by auto from hd_b_lt_n have dist: "distinct (map fst ?qs)" using distinct_p by auto with set_a step_pres_bounded sorted_tl_b len_a have "\(q, \) \ set ?qs. (fst(propa f ?qs a (tl b)))!q = \ \\<^bsub>f\<^esub>a!q" by (auto dest:propa_property1) with propa have propa_ss1: "\(q, \) \ set ?qs. ss!q = (hd b # ?a_hdb) \\<^bsub>f\<^esub>a!q" by (auto simp add:step_def exec_def transf_def) from step_pres_bounded sorted_tl_b set_a len_a dist have "\q. q \ set(map fst ?qs) \ q < length a \ (fst(propa f ?qs a (tl b)))!q = a!q" by (auto intro:propa_property2) with propa have exec2: "\q. q \ set(map fst ?qs) \ q < length a \ ss!q = a!q" by auto have "(\i. i < n \ k \ set ( ss ! i) \ non_strict_dominate k i)" proof(intro strip) fix i let ?a_i = "a!i" assume "i < n \ k \ set ( ss ! i) " then have i_lt_n: "i < n" and k_nin_ss: "k \ set (ss ! i) " by auto from non_strict have g: "i < n \ a ! i = ?a_i \ k \ set ?a_i \ non_strict_dominate k i" by auto have sorted_a_i: "sorted (rev ?a_i)" using wf_a_b i_lt_n by (auto simp add:_wf_dom_def) show "non_strict_dominate k i" proof(cases "i \ (succs (hd b))") case True note i_succ_hdb = this with propa_ss1 have ss_i: "ss!i = (hd b # ?a_hdb) \\<^bsub>f\<^esub> a!i" using fin_succ_hd_b by (auto simp add:step_def exec_def) then have "ss!i = (hd b # ?a_hdb ) \\<^bsub>f\<^esub> ?a_i" by auto with sorted_hd_b_a sorted_a_i have "set (ss!i) = set (hd b # ?a_hdb) \ set ?a_i" and ss_i_merge: "ss!i = ((inter_sorted_rev (hd b # ?a_hdb) ?a_i))" by (auto simp add:inter_sorted_correct f_def nodes_sup_def plussub_def) with k_nin_ss have k_nin': "k \ set (hd b # ?a_hdb) \ set ?a_i" by auto show ?thesis proof(cases "k \ set ?a_i") case True then show ?thesis using i_lt_n g by auto next case False with k_nin' have "k \ set ?a_hdb" and k_neq_hdb: "k \ hd b" by auto with hd_b_lt_n non_strict have n_k_hdb: "non_strict_dominate k (hd b)" by auto from i_succ_hdb have "(hd b, i)\ g_E G" by (auto simp add:succs_def) with n_k_hdb k_neq_hdb show ?thesis using non_sdominate_succ by auto qed next case False with exec2 have "ss!i = a!i" using fin_succ_hd_b len_a i_lt_n by (auto simp add:step_def exec_def) with k_nin_ss have "k \ set (a!i)" by auto with non_strict show ?thesis using i_lt_n by fastforce qed qed with wf_ss_w show "?PROP ?P" by auto qed lemma start_non_sdom: " i < n \ start!i = res \ k \ set res \ non_strict_dominate k i" proof(auto) assume i_lt_n: "i < n" and k_nin: "k \ set (start ! i)" then have reach_i: "reachable i" using n_def nodes_def verts_set len_verts_gt0 reachable by (simp add:reachable_def) then obtain pa where pa_i: "path_entry (g_E G) pa i" using reachable_path_entry by auto show "non_strict_dominate k i" proof(cases "k \ set (g_V G)") case True then have "k < n" using verts_set by (auto simp add:n_def nodes_def) then have k_in_verts: "k \ set (g_V G)" and k_in_verts': "k \ {0..pa. path_entry (g_E G) pa i \ set pa \ set (g_V G)" using path_in_verts nodes_def by auto with k_nin_verts have k_nin: "\pa. path_entry (g_E G) pa i \ k \ set pa" by (fastforce dest: contra_subsetD) with pa_i have "k \ set pa" by auto with pa_i show ?thesis by (auto simp add: non_strict_dominate_def) qed qed lemma iter_dom_invariant_complete: shows "\res. iter f step start (unstables r step start) = (ss',w') \ i < n \ ss'!i = res \ k \ set res \ non_strict_dominate k i" apply (unfold iter_def) apply(rule_tac P = "(\(ss, w). wf_dom ss w \ (\i. (\rs. i < n \ ss!i = rs \ k \ set rs) \ non_strict_dominate k i))" and r = "{(ss',ss). ss [\\<^sub>r] ss'} <*lex*> (\(x,y). (sorted_list_of_set x, sorted_list_of_set y)) ` finite_psubset" in while_rule) \ \Invariant holds initially:\ apply clarsimp apply (fastforce simp add:start_non_sdom wf_start) \ \Invariant is preserved:\ apply (fastforce dest:propa_dom_invariant_complete) \ \Postcondition holds upon termination:\ apply clarsimp \ \Well-foundedness of the termination relation:\ apply (simp add:wf_listn_termination_rel) \ \Loop decreases along termination relation:\ apply clarsimp apply (fastforce dest:propa_termination) done end end diff --git a/thys/Dominance_CHK/Dom_Semi_List.thy b/thys/Dominance_CHK/Dom_Semi_List.thy --- a/thys/Dominance_CHK/Dom_Semi_List.thy +++ b/thys/Dominance_CHK/Dom_Semi_List.thy @@ -1,220 +1,222 @@ (* Author: Nan Jiang *) section \A semilattice of reversed-ordered list\ theory Dom_Semi_List -imports Main Semilat Sorted_List_Operations2 Sorted_Less2 Cfg +imports Main "Jinja.Semilat" Sorted_List_Operations2 Sorted_Less2 Cfg begin type_synonym node = nat context cfg_doms begin definition nodes :: "nat list" where "nodes \ (g_V G)" definition nodes_le :: "node list \ node list \ bool" where "nodes_le xs ys \ (sorted (rev ys) \ sorted (rev xs) \ (set ys) \ (set xs)) \ xs = ys" definition nodes_sup ::"node list \ node list \node list " where "nodes_sup = (\x y. inter_sorted_rev x y)" definition nodes_semi :: "node list sl" where "nodes_semi \ ((rev \ sorted_list_of_set) ` (Pow (set (nodes))), nodes_le, nodes_sup )" lemma subset_nodes_inpow: assumes "sorted (rev xs)" and "set xs \ set nodes" shows "xs \ (rev \ sorted_list_of_set) ` (Pow (set nodes))" proof - from assms(1) have "(sorted_list_of_set (set xs)) = rev xs" by (auto intro:sorted_less_rev_set_eq) then have "rev (rev xs) = rev (sorted_list_of_set (set xs))" by simp with assms(2) show ?thesis by auto qed lemma nil_in_A: "[] \ (rev \ sorted_list_of_set) ` (Pow (set nodes))" proof(simp add: Pow_def image_def) have "sorted_list_of_set {} = []" by auto then show "\x\set nodes. sorted_list_of_set x = []" by blast qed lemma single_n_in_A: "p < length nodes \ [p] \ (rev \ sorted_list_of_set) ` (Pow (set nodes))" proof (unfold nodes_def) let ?S = "(rev \ sorted_list_of_set) ` (Pow (set (g_V G)))" assume "p < length (g_V G)" then have p: "{p} \ Pow (set (g_V G))" by (auto simp add:Pow_def verts_set) then have "[p] \?S" by (unfold image_def) force then show "[p] \ ?S" by auto qed lemma inpow_subset_nodes: assumes "xs \ (rev \ sorted_list_of_set) ` (Pow (set nodes))" shows "set xs \ set nodes" proof - from assms obtain x where x: "x \ Pow (set nodes)" and "xs = (rev \ sorted_list_of_set) x" by auto then have eq: "set xs = set (sorted_list_of_set x)" by auto have "\x \ Pow (set nodes). finite x" by (auto intro: rev_finite_subset) with x eq show "set xs \ set nodes" by auto qed lemma inter_in_pow_nodes: assumes "xs \ (rev \ sorted_list_of_set) ` (Pow (set nodes))" shows "(rev \ sorted_list_of_set)(set xs \ set ys) \ (rev \ (sorted_list_of_set)) ` (Pow (set nodes))" using assms proof - let ?res = "set xs \ set ys" from assms have "set xs \ set nodes" using inpow_subset_nodes by auto then have "?res \ set nodes" by auto then show ?thesis using subset_nodes_inpow by auto qed lemma nodes_le_order: "order nodes_le ((rev \ sorted_list_of_set) ` (Pow (set nodes)))" proof - let ?A = "(rev \ sorted_list_of_set) ` (Pow (set nodes))" have "\x \ ?A. sorted (rev x)" by (auto intro: sorted_less_sorted_list_of_set) then have "\x\?A. nodes_le x x" by (auto simp add:nodes_le_def) moreover have "\x\?A. \y\?A. (nodes_le x y \ nodes_le y x \ x = y)" proof (intro strip) fix x y assume "x \ ?A" and "y \ ?A" and "nodes_le x y \ nodes_le y x" then have "sorted (rev x) \ sorted (rev (y::nat list)) \ set x = set y \ x = y" by (auto simp add: nodes_le_def intro:subset_antisym sorted_less_sorted_list_of_set) then show "x = y" by (auto dest: sorted_less_rev_set_unique) qed moreover have "\x\ ?A. \y\ ?A. \z\ ?A . nodes_le x y \ nodes_le y z \ nodes_le x z" by (auto simp add: nodes_le_def) ultimately show ?thesis by (unfold order_def lesub_def lesssub_def) fastforce qed lemma nodes_semi_auxi: "let A = (rev \ sorted_list_of_set) ` (Pow (set (nodes))); r = nodes_le; f = (\x y. (inter_sorted_rev x y)) in semilat(A, r, f)" proof - let ?A = "(rev \ sorted_list_of_set) ` (Pow (set (nodes)))" let ?r = "nodes_le" let ?f = "(\x y. (inter_sorted_rev x y))" have "order ?r ?A" by (rule nodes_le_order) moreover have "closed ?A ?f" proof (unfold closed_def, intro strip) fix xs ys assume xs_in: "xs \ ?A" and ys_in: "ys \ ?A" then have sorted_xs: "sorted (rev xs)" and sorted_ys: "sorted (rev ys)" by (auto intro: sorted_less_sorted_list_of_set) then have inter_xs_ys: "set (?f xs ys) = set xs \ set ys" and sorted_res: "sorted (rev (?f xs ys))" using inter_sorted_correct by auto from xs_in have "set xs \ set nodes" using inpow_subset_nodes by auto with inter_xs_ys have "set (?f xs ys) \ set nodes" by auto with sorted_res show "xs \\<^bsub>?f\<^esub> ys\ ?A" using subset_nodes_inpow by (auto simp add:plussub_def) qed moreover have "(\x\?A. \y\?A. x \\<^bsub>?r\<^esub> x \\<^bsub>?f\<^esub> y) \ (\x\?A. \y\?A. y \\<^bsub>?r\<^esub> x \\<^bsub>?f\<^esub> y)" proof(rule conjI, intro strip) fix xs ys assume xs_in: "xs \ ?A" and ys_in: "ys \ ?A" then have sorted_xs: "sorted (rev xs)" and sorted_ys: "sorted (rev ys)" by (auto intro: sorted_less_sorted_list_of_set) then have "set (?f xs ys) \ set xs" and sorted_f_xs_ys: "sorted (rev (?f xs ys))" by (auto simp add: inter_sorted_correct) then show "xs \\<^bsub>?r\<^esub> xs \\<^bsub>?f\<^esub> ys" by (simp add: lesub_def sorted_xs sorted_ys sorted_f_xs_ys nodes_le_def plussub_def) next show "\x\?A. \y\?A. y \\<^bsub>?r\<^esub> x \\<^bsub>?f\<^esub> y" proof (intro strip) fix xs ys assume xs_in: "xs \ ?A" and ys_in: "ys \ ?A" then have sorted_xs: "sorted (rev xs)" and sorted_ys: "sorted (rev ys)" by (auto intro: sorted_less_sorted_list_of_set) then have "set (?f xs ys) \ set ys" and sorted_f_xs_ys: "sorted (rev (?f xs ys))" by (auto simp add: inter_sorted_correct) then show "ys \\<^bsub>?r\<^esub> xs \\<^bsub>?f\<^esub> ys" by (simp add: lesub_def sorted_ys sorted_xs sorted_f_xs_ys nodes_le_def plussub_def) qed qed moreover have "\x\?A. \y\?A. \z\?A. x \\<^bsub>?r\<^esub> z \ y \\<^bsub>?r\<^esub> z \ x \\<^bsub>?f\<^esub> y \\<^bsub>?r\<^esub>z" proof (intro strip) fix xs ys zs assume xin: "xs \ ?A" and yin: "ys \ ?A" and zin: "zs \ ?A" and "xs \\<^bsub>?r\<^esub> zs \ ys \\<^bsub>?r\<^esub> zs" then have xs_zs: "xs \\<^bsub>?r\<^esub> zs" and ys_zs: "ys \\<^bsub>?r\<^esub> zs" and sorted_xs: "sorted (rev xs)" and sorted_ys: "sorted (rev ys)" by (auto simp add: sorted_less_sorted_list_of_set) then have inter_xs_ys: "set (?f xs ys) = (set xs \ set ys)" and sorted_f_xs_ys: "sorted (rev (?f xs ys))" by (auto simp add: inter_sorted_correct) from xs_zs ys_zs sorted_xs have sorted_zs: "sorted (rev zs)" and "set zs \ set xs" and "set zs \ set ys" by (auto simp add: lesub_def nodes_le_def) then have zs: "set zs \ set xs \ set ys" by auto with inter_xs_ys sorted_zs sorted_f_xs_ys show "xs \\<^bsub>?f\<^esub> ys \\<^bsub>?r\<^esub> zs" by (auto simp add:plussub_def lesub_def sorted_xs sorted_ys sorted_f_xs_ys sorted_zs nodes_le_def) qed ultimately show ?thesis by (unfold semilat_def) simp qed lemma nodes_semi_is_semilat: "semilat (nodes_semi)" using nodes_semi_auxi by (auto simp add: nodes_sup_def nodes_semi_def) lemma sorted_rev_subset_len_lt: assumes "sorted (rev a)" and "sorted (rev b)" and "set a \ set b" shows "length a < length b" using assms proof- from assms(1) assms(2) have dist_a: "distinct a" and dist_b: "distinct b" by (auto dest: distinct_sorted_rev) from assms(3) have "card (set a) < card (set b)" by (auto intro: psubset_card_mono) with dist_a dist_b show ?thesis by (auto simp add: distinct_card) qed lemma wf_nodes_le_auxi: "wf {(y, x). (sorted (rev y) \ sorted (rev x) \ set y \ set x) \ x \ y}" apply(rule wf_measure [THEN wf_subset]) apply(simp only: measure_def inv_image_def) apply clarify apply(frule sorted_rev_subset_len_lt) defer defer apply fastforce by (auto intro:sorted_less_rev_set_unique) lemma wf_nodes_le_auxi2: "wf {(y, x). sorted (rev y) \ sorted (rev x) \ set y \ set x \ rev x \ rev y}" using wf_nodes_le_auxi by auto lemma wf_nodes_le: "wf {(y, x). nodes_le x y \ x \ y}" proof - have eq_set: "{(y, x). (sorted (rev y) \ sorted (rev x) \ set y \ set x) \ x \ y} = {(y, x). nodes_le x y \ x \ y}" by (unfold nodes_le_def) auto have "{(y, x). (sorted (rev y) \ sorted (rev x) \ set y \ set x) \ x \ y} = {(y, x). (sorted (rev y) \ sorted (rev x) \ set y \ set x) \ x \ y}" by (auto simp add:sorted_less_rev_set_unique) from this wf_nodes_le_auxi have "wf {(y, x). (sorted (rev y) \ sorted (rev x) \ set y \ set x) \ x \ y}" by (rule subst) with eq_set show ?thesis by (rule subst) qed lemma acc_nodes_le: "acc nodes_le" apply (unfold acc_def lesssub_def lesub_def) apply (rule wf_nodes_le) done lemma acc_nodes_le2: "acc (fst (snd nodes_semi))" apply (unfold nodes_semi_def) apply (auto simp add: lesssub_def lesub_def intro: acc_nodes_le) done -lemma "top nodes_le [] (fst nodes_semi)" - by(auto simp add: top_def nodes_semi_def nodes_le_def lesssub_def lesub_def sorted_less_sorted_list_of_set) +lemma nodes_le_refl [iff] : "nodes_le s s" + apply (unfold nodes_le_def lesssub_def lesub_def) + apply (auto) + done end end diff --git a/thys/Dominance_CHK/ROOT b/thys/Dominance_CHK/ROOT --- a/thys/Dominance_CHK/ROOT +++ b/thys/Dominance_CHK/ROOT @@ -1,20 +1,20 @@ chapter AFP session "Dominance_CHK" (AFP) = HOL + options [timeout=300] sessions "HOL-Data_Structures" "HOL-Library" + "List-Index" + "Jinja" theories Cfg Dom_Semi_List Dom_Kildall Dom_Kildall_Property Dom_Kildall_Correct - Semilat - Listn Sorted_List_Operations2 Sorted_Less2 document_files "root.bib" "root.tex" \ No newline at end of file diff --git a/thys/Jinja/BV/BVConform.thy b/thys/Jinja/BV/BVConform.thy --- a/thys/Jinja/BV/BVConform.thy +++ b/thys/Jinja/BV/BVConform.thy @@ -1,154 +1,152 @@ (* Title: HOL/MicroJava/BV/Correct.thy Author: Cornelia Pusch, Gerwin Klein Copyright 1999 Technische Universitaet Muenchen The invariant for the type safety proof. *) section \BV Type Safety Invariant\ theory BVConform imports BVSpec "../JVM/JVMExec" "../Common/Conform" begin definition confT :: "'c prog \ heap \ val \ ty err \ bool" ("_,_ \ _ :\\<^sub>\ _" [51,51,51,51] 50) where "P,h \ v :\\<^sub>\ E \ case E of Err \ True | OK T \ P,h \ v :\ T" notation (ASCII) confT ("_,_ |- _ :<=T _" [51,51,51,51] 50) abbreviation confTs :: "'c prog \ heap \ val list \ ty\<^sub>l \ bool" ("_,_ \ _ [:\\<^sub>\] _" [51,51,51,51] 50) where "P,h \ vs [:\\<^sub>\] Ts \ list_all2 (confT P h) vs Ts" notation (ASCII) confTs ("_,_ |- _ [:<=T] _" [51,51,51,51] 50) definition conf_f :: "jvm_prog \ heap \ ty\<^sub>i \ bytecode \ frame \ bool" where "conf_f P h \ \(ST,LT) is (stk,loc,C,M,pc). P,h \ stk [:\] ST \ P,h \ loc [:\\<^sub>\] LT \ pc < size is" lemma conf_f_def2: "conf_f P h (ST,LT) is (stk,loc,C,M,pc) \ P,h \ stk [:\] ST \ P,h \ loc [:\\<^sub>\] LT \ pc < size is" by (simp add: conf_f_def) primrec conf_fs :: "[jvm_prog,heap,ty\<^sub>P,mname,nat,ty,frame list] \ bool" where "conf_fs P h \ M\<^sub>0 n\<^sub>0 T\<^sub>0 [] = True" | "conf_fs P h \ M\<^sub>0 n\<^sub>0 T\<^sub>0 (f#frs) = (let (stk,loc,C,M,pc) = f in (\ST LT Ts T mxs mxl\<^sub>0 is xt. \ C M ! pc = Some (ST,LT) \ (P \ C sees M:Ts \ T = (mxs,mxl\<^sub>0,is,xt) in C) \ (\D Ts' T' m D'. is!pc = (Invoke M\<^sub>0 n\<^sub>0) \ ST!n\<^sub>0 = Class D \ P \ D sees M\<^sub>0:Ts' \ T' = m in D' \ P \ T\<^sub>0 \ T') \ conf_f P h (ST, LT) is f \ conf_fs P h \ M (size Ts) T frs))" definition correct_state :: "[jvm_prog,ty\<^sub>P,jvm_state] \ bool" ("_,_ \ _ \" [61,0,0] 61) where "correct_state P \ \ \(xp,h,frs). case xp of None \ (case frs of [] \ True | (f#fs) \ P\ h\ \ (let (stk,loc,C,M,pc) = f in \Ts T mxs mxl\<^sub>0 is xt \. (P \ C sees M:Ts\T = (mxs,mxl\<^sub>0,is,xt) in C) \ \ C M ! pc = Some \ \ conf_f P h \ is f \ conf_fs P h \ M (size Ts) T fs)) | Some x \ frs = []" notation correct_state ("_,_ |- _ [ok]" [61,0,0] 61) subsection \Values and \\\\ lemma confT_Err [iff]: "P,h \ x :\\<^sub>\ Err" by (simp add: confT_def) lemma confT_OK [iff]: "P,h \ x :\\<^sub>\ OK T = (P,h \ x :\ T)" by (simp add: confT_def) lemma confT_cases: "P,h \ x :\\<^sub>\ X = (X = Err \ (\T. X = OK T \ P,h \ x :\ T))" by (cases X) auto lemma confT_hext [intro?, trans]: "\ P,h \ x :\\<^sub>\ T; h \ h' \ \ P,h' \ x :\\<^sub>\ T" by (cases T) (blast intro: conf_hext)+ lemma confT_widen [intro?, trans]: "\ P,h \ x :\\<^sub>\ T; P \ T \\<^sub>\ T' \ \ P,h \ x :\\<^sub>\ T'" by (cases T', auto intro: conf_widen) subsection \Stack and Registers\ lemmas confTs_Cons1 [iff] = list_all2_Cons1 [of "confT P h"] for P h lemma confTs_confT_sup: -assumes confTs: "P,h \ loc [:\\<^sub>\] LT" and n: "n < size LT" and - LTn: "LT!n = OK T" and subtype: "P \ T \ T'" -shows "P,h \ (loc!n) :\ T'" + "\ P,h \ loc [:\\<^sub>\] LT; n < size LT; LT!n = OK T; P \ T \ T' \ + \ P,h \ (loc!n) :\ T'" (*<*) -proof - - have len: "n < length loc" using list_all2_lengthD[OF confTs] n - by simp - show ?thesis - using list_all2_nthD[OF confTs len] conf_widen[OF _ subtype] LTn - by simp -qed + apply (frule list_all2_lengthD) + apply (drule list_all2_nthD, simp) + apply simp + apply (erule conf_widen, assumption+) + done (*>*) lemma confTs_hext [intro?]: "P,h \ loc [:\\<^sub>\] LT \ h \ h' \ P,h' \ loc [:\\<^sub>\] LT" by (fast elim: list_all2_mono confT_hext) lemma confTs_widen [intro?, trans]: "P,h \ loc [:\\<^sub>\] LT \ P \ LT [\\<^sub>\] LT' \ P,h \ loc [:\\<^sub>\] LT'" by (rule list_all2_trans, rule confT_widen) lemma confTs_map [iff]: "\vs. (P,h \ vs [:\\<^sub>\] map OK Ts) = (P,h \ vs [:\] Ts)" by (induct Ts) (auto simp add: list_all2_Cons2) lemma reg_widen_Err [iff]: "\LT. (P \ replicate n Err [\\<^sub>\] LT) = (LT = replicate n Err)" by (induct n) (auto simp add: list_all2_Cons1) lemma confTs_Err [iff]: "P,h \ replicate n v [:\\<^sub>\] replicate n Err" by (induct n) auto subsection \correct-frames\ lemmas [simp del] = fun_upd_apply lemma conf_fs_hext: "\M n T\<^sub>r. \ conf_fs P h \ M n T\<^sub>r frs; h \ h' \ \ conf_fs P h' \ M n T\<^sub>r frs" (*<*) -proof(induct frs) - case (Cons fr frs) - obtain stk ls C M pc where fr: "fr = (stk, ls, C, M, pc)" by(cases fr) simp - moreover obtain ST LT where \: "\ C M ! pc = \(ST, LT)\" and - ST: "P,h \ stk [:\] ST" and LT: "P,h \ ls [:\\<^sub>\] LT" - using Cons.prems(1) fr by(auto simp: conf_f_def) - ultimately show ?case using Cons confs_hext[OF ST Cons(3)] confTs_hext[OF LT Cons(3)] - by (fastforce simp: conf_f_def) -qed simp +apply (induct frs) + apply simp +apply clarify +apply (simp (no_asm_use)) +apply clarify +apply (unfold conf_f_def) +apply (simp (no_asm_use)) +apply clarify +apply (fast elim!: confs_hext confTs_hext) +done (*>*) end diff --git a/thys/Jinja/BV/BVExec.thy b/thys/Jinja/BV/BVExec.thy --- a/thys/Jinja/BV/BVExec.thy +++ b/thys/Jinja/BV/BVExec.thy @@ -1,247 +1,460 @@ (* Title: HOL/MicroJava/BV/JVM.thy Author: Tobias Nipkow, Gerwin Klein Copyright 2000 TUM *) section \Kildall for the JVM \label{sec:JVM}\ theory BVExec -imports "../DFA/Abstract_BV" TF_JVM +imports "../DFA/Abstract_BV" TF_JVM "../DFA/Typing_Framework_2" begin definition kiljvm :: "jvm_prog \ nat \ nat \ ty \ instr list \ ex_table \ ty\<^sub>i' err list \ ty\<^sub>i' err list" where "kiljvm P mxs mxl T\<^sub>r is xt \ kildall (JVM_SemiType.le P mxs mxl) (JVM_SemiType.sup P mxs mxl) (exec P mxs T\<^sub>r xt is)" definition wt_kildall :: "jvm_prog \ cname \ ty list \ ty \ nat \ nat \ instr list \ ex_table \ bool" where "wt_kildall P C' Ts T\<^sub>r mxs mxl\<^sub>0 is xt \ 0 < size is \ (let first = Some ([],[OK (Class C')]@(map OK Ts)@(replicate mxl\<^sub>0 Err)); start = OK first#(replicate (size is - 1) (OK None)); result = kiljvm P mxs (1+size Ts+mxl\<^sub>0) T\<^sub>r is xt start in \n < size is. result!n \ Err)" definition wf_jvm_prog\<^sub>k :: "jvm_prog \ bool" where "wf_jvm_prog\<^sub>k P \ wf_prog (\P C' (M,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,is,xt)). wt_kildall P C' Ts T\<^sub>r mxs mxl\<^sub>0 is xt) P" +context start_context +begin -theorem (in start_context) is_bcv_kiljvm: - "is_bcv r Err step (size is) A (kiljvm P mxs mxl T\<^sub>r is xt)" -(*<*) +lemma Cons_less_Conss3 [simp]: + "x#xs [\\<^bsub>r\<^esub>] y#ys = (x \\<^bsub>r\<^esub> y \ xs [\\<^bsub>r\<^esub>] ys \ x = y \ xs [\\<^bsub>r\<^esub>] ys)" + apply (unfold lesssub_def ) + apply auto + apply (insert sup_state_opt_err) + apply (unfold lesssub_def lesub_def sup_state_opt_def sup_state_def sup_ty_opt_def) + apply (simp only: JVM_le_unfold ) + apply fastforce + done + +lemma acc_le_listI3 [intro!]: + " acc r \ acc (Listn.le r)" +apply (unfold acc_def) +apply (subgoal_tac + "wf(UN n. {(ys,xs). size xs = n \ size ys = n \ xs <_(Listn.le r) ys})") + apply (erule wf_subset) + apply (blast intro: lesssub_lengthD) +apply (rule wf_UN) + prefer 2 + apply (rename_tac m n) + apply (case_tac "m=n") + apply simp + apply (fast intro!: equals0I dest: not_sym) +apply (rename_tac n) +apply (induct_tac n) + apply (simp add: lesssub_def cong: conj_cong) +apply (rename_tac k) +apply (simp add: wf_eq_minimal del: r_def f_def step_def A_def) +apply (simp (no_asm) add: length_Suc_conv cong: conj_cong del: r_def f_def step_def A_def) +apply clarify +apply (rename_tac M m) +apply (case_tac "\x xs. size xs = k \ x#xs \ M") + prefer 2 + apply (erule thin_rl) + apply (erule thin_rl) + apply blast +apply (erule_tac x = "{a. \xs. size xs = k \ a#xs:M}" in allE) +apply (erule impE) + apply blast +apply (thin_tac "\x xs. P x xs" for P) +apply clarify +apply (rename_tac maxA xs) +apply (erule_tac x = "{ys. size ys = size xs \ maxA#ys \ M}" in allE) +apply (erule impE) + apply blast +apply clarify +apply (thin_tac "m \ M") +apply (thin_tac "maxA#xs \ M") +apply (rule bexI) + prefer 2 + apply assumption +apply clarify +apply (simp del: r_def f_def step_def A_def) +apply blast + done + + +lemma wf_jvm: " wf {(ss', ss). ss [\\<^bsub>r\<^esub>] ss'}" + apply (insert acc_le_listI3 acc_JVM [OF wf]) + by (simp add: acc_def r_def) + +lemma iter_properties_bv[rule_format]: +shows "\ \p\w0. p < n; ss0 \ list n A; \p w0 \ stable r step ss0 p \ \ + iter f step ss0 w0 = (ss',w') \ + ss' \ list n A \ stables r step ss' \ ss0 [\\<^sub>r] ss' \ + (\ts\list n A. ss0 [\\<^sub>r] ts \ stables r step ts \ ss' [\\<^sub>r] ts)" +(*<*) (is "PROP ?P") + proof - - let ?n = "length is" - have "Semilat A r f" using semilat_JVM[OF wf] - by (simp add: Semilat.intro sl_def2) - moreover have "acc r" using wf by simp blast - moreover have "top r Err" by (simp add: JVM_le_unfold) - moreover have "pres_type step ?n A" by (rule exec_pres_type) - moreover have "bounded step ?n" by (rule bounded_step) - moreover have "mono r step ?n A" using step_mono[OF wf] by simp - ultimately have "is_bcv r Err step ?n A (kildall r f step)" - by(rule is_bcv_kildall) - moreover have kileq: "kiljvm P mxs mxl T\<^sub>r is xt = kildall r f step" - using f_def kiljvm_def r_def step_def_exec by blast - ultimately show ?thesis by simp + show "PROP ?P" + apply (insert semi bounded_step exec_pres_type step_mono[OF wf]) + apply (unfold iter_def stables_def) + + apply (rule_tac P = "\(ss,w). + ss \ list n A \ (\p w \ stable r step ss p) \ ss0 [\\<^sub>r] ss \ + (\ts\list n A. ss0 [\\<^sub>r] ts \ stables r step ts \ ss [\\<^sub>r] ts) \ + (\p\w. p < n)" and + r = "{(ss',ss) . ss [\\<^sub>r] ss'} <*lex*> finite_psubset" + in while_rule) + + \ \Invariant holds initially:\ + apply (simp add:stables_def semilat_Def del: n_def A_def r_def f_def step_def) + apply (blast intro:le_list_refl') + + \ \Invariant is preserved:\ + apply(simp add: stables_def split_paired_all del: A_def r_def f_def step_def n_def) + apply(rename_tac ss w) + apply(subgoal_tac "(SOME p. p \ w) \ w") + prefer 2 apply (fast intro: someI) + apply(subgoal_tac "\(q,t) \ set (step (SOME p. p \ w) (ss ! (SOME p. p \ w))). q < length ss \ t \ A") + prefer 2 + apply clarify + apply (rule conjI) + apply(clarsimp, blast dest!: boundedD) + apply (subgoal_tac "(SOME p. p \ w) < n") + apply (subgoal_tac "(ss ! (SOME p. p \ w)) \ A") + apply (fastforce simp only:n_def dest:pres_typeD ) + apply simp + apply simp + apply (subst decomp_propa) + apply blast + apply (simp del:A_def r_def f_def step_def n_def) + apply (rule conjI) + apply (rule Semilat.merges_preserves_type[OF Semilat.intro, OF semi]) + apply blast + apply clarify + apply (rule conjI) + apply(clarsimp, blast dest!: boundedD) + apply (erule pres_typeD) + prefer 3 + apply assumption + apply (erule listE_nth_in) + apply blast + apply (simp only:n_def) + apply (rule conjI) + apply clarify + apply (subgoal_tac "ss \ list (length is) A" "\p\w. p < (length is) " "\p<(length is). p \ w \ stable r step ss p " + "p < length is") + apply (blast intro!: Semilat.stable_pres_lemma[OF Semilat.intro, OF semi]) + apply (simp only:n_def) + apply (simp only:n_def) + apply (simp only:n_def) + apply (simp only:n_def) + apply (rule conjI) + apply (subgoal_tac "ss \ list (length is) A" + "\(q,t)\set (step (SOME p. p \ w) (ss ! (SOME p. p \ w))). q < length is \ t \ A" + "ss [\\<^bsub>r\<^esub>] merges f (step (SOME p. p \ w) (ss ! (SOME p. p \ w))) ss" "ss0\ list (size is) A" + "merges f (step (SOME p. p \ w) (ss ! (SOME p. p \ w))) ss \ list (size is) A" + "ss \list (size is) A" "order r A" "ss0 [\\<^bsub>r\<^esub>] ss ") + apply (blast dest: le_list_trans) + apply simp + apply (simp only:semilat_Def) + apply (simp only:n_def) + apply (fastforce simp only: n_def dest:Semilat.merges_preserves_type[OF Semilat.intro, OF semi] ) + apply (simp only:n_def) + apply (blast intro:Semilat.merges_incr[OF Semilat.intro, OF semi]) + apply (subgoal_tac "length ss = n") + apply (simp only:n_def) + apply (subgoal_tac "ss \list n A") + defer + apply simp + apply (simp only:n_def) + prefer 5 + apply (simp only:listE_length n_def) + apply(rule conjI) + apply (clarsimp simp del: A_def r_def f_def step_def) + apply (blast intro!: Semilat.merges_bounded_lemma[OF Semilat.intro, OF semi]) + apply (subgoal_tac "bounded step n") + apply (blast dest!: boundedD) + apply (simp only:n_def) + + \ \Postcondition holds upon termination:\ + apply(clarsimp simp add: stables_def split_paired_all) + + \ \Well-foundedness of the termination relation:\ + apply (rule wf_lex_prod) + apply (simp only:wf_jvm) + apply (rule wf_finite_psubset) + + \ \Loop decreases along termination relation:\ + apply(simp add: stables_def split_paired_all del: A_def r_def f_def step_def) + apply(rename_tac ss w) + apply(subgoal_tac "(SOME p. p \ w) \ w") + prefer 2 apply (fast intro: someI) + apply(subgoal_tac "\(q,t) \ set (step (SOME p. p \ w) (ss ! (SOME p. p \ w))). q < length ss \ t \ A") + prefer 2 + apply clarify + apply (rule conjI) + apply(clarsimp, blast dest!: boundedD) + apply (erule pres_typeD) + prefer 3 + apply assumption + apply (erule listE_nth_in) + apply blast + apply blast + apply (subst decomp_propa) + apply blast + apply clarify + apply (simp del: listE_length A_def r_def f_def step_def + add: lex_prod_def finite_psubset_def + bounded_nat_set_is_finite) + apply (rule termination_lemma) + apply (insert Semilat.intro) + apply assumption+ + defer + apply assumption + defer + apply clarsimp + done qed + (*>*) + +lemma kildall_properties_bv: +shows "\ ss0 \ list n A \ \ + kildall r f step ss0 \ list n A \ + stables r step (kildall r f step ss0) \ + ss0 [\\<^sub>r] kildall r f step ss0 \ + (\ts\list n A. ss0 [\\<^sub>r] ts \ stables r step ts \ + kildall r f step ss0 [\\<^sub>r] ts)" +(*<*) (is "PROP ?P") +proof - + show "PROP ?P" + apply (unfold kildall_def) + apply(case_tac "iter f step ss0 (unstables r step ss0)") + apply (simp del:r_def f_def n_def step_def A_def) + apply (rule iter_properties_bv) + apply (simp_all add: unstables_def stable_def) + done +qed + +end + +lemma (in start_context) is_bcv_kiljvm: +shows "is_bcv r Err step (size is) A (kiljvm P mxs mxl T\<^sub>r is xt)" + apply (insert wf) + apply (unfold kiljvm_def) + apply (fold r_def f_def step_def_exec n_def) + apply(unfold is_bcv_def wt_step_def) + apply(insert semi kildall_properties_bv) + apply(simp only:stables_def) + apply clarify + apply(subgoal_tac "kildall r f step \s\<^sub>0 \ list n A") + prefer 2 + apply (fastforce intro: kildall_properties_bv) + apply (rule iffI) + apply (rule_tac x = "kildall r f step \s\<^sub>0" in bexI) + apply (rule conjI) + apply (fastforce intro: kildall_properties_bv) + apply (force intro: kildall_properties_bv) + apply simp + apply clarify + apply(subgoal_tac "kildall r f step \s\<^sub>0!pa <=_r \s!pa") + defer + apply (blast intro!: le_listD less_lengthI) + apply (subgoal_tac "\s!pa \ Err") + defer + apply simp + apply (rule ccontr) + apply (simp only:top_def r_def JVM_le_unfold) + apply fastforce + done + (* FIXME: move? *) lemma subset_replicate [intro?]: "set (replicate n x) \ {x}" by (induct n) auto lemma in_set_replicate: assumes "x \ set (replicate n y)" shows "x = y" (*<*) proof - note assms also have "set (replicate n y) \ {y}" .. finally show ?thesis by simp qed (*>*) lemma (in start_context) start_in_A [intro?]: "0 < size is \ start \ list (size is) A" using Ts C (*<*) - by (force simp add: JVM_states_unfold intro!: listI list_appendI - dest!: in_set_replicate) + apply (simp add: JVM_states_unfold) + apply (force intro!: listI list_appendI dest!: in_set_replicate) + done (*>*) theorem (in start_context) wt_kil_correct: assumes wtk: "wt_kildall P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt" shows "\\s. wt_method P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt \s" (*<*) proof - from wtk obtain res where result: "res = kiljvm P mxs mxl T\<^sub>r is xt start" and success: "\n < size is. res!n \ Err" and instrs: "0 < size is" by (unfold wt_kildall_def) simp have bcv: "is_bcv r Err step (size is) A (kiljvm P mxs mxl T\<^sub>r is xt)" by (rule is_bcv_kiljvm) from instrs have "start \ list (size is) A" .. with bcv success result have "\ts\list (size is) A. start [\\<^sub>r] ts \ wt_step r Err step ts" by (unfold is_bcv_def) blast then obtain \s' where in_A: "\s' \ list (size is) A" and s: "start [\\<^sub>r] \s'" and w: "wt_step r Err step \s'" by blast hence wt_err_step: "wt_err_step (sup_state_opt P) step \s'" by (simp add: wt_err_step_def JVM_le_Err_conv) from in_A have l: "size \s' = size is" by simp moreover { from in_A have "check_types P mxs mxl \s'" by (simp add: check_types_def) also from w have "\x \ set \s'. x \ Err" by (auto simp add: wt_step_def all_set_conv_all_nth) hence [symmetric]: "map OK (map ok_val \s') = \s'" by (auto intro!: map_idI simp add: wt_step_def) finally have "check_types P mxs mxl (map OK (map ok_val \s'))" . } moreover { from s have "start!0 \\<^sub>r \s'!0" by (rule le_listD) simp moreover from instrs w l have "\s'!0 \ Err" by (unfold wt_step_def) simp then obtain \s0 where "\s'!0 = OK \s0" by auto ultimately have "wt_start P C Ts mxl\<^sub>0 (map ok_val \s')" using l instrs by (unfold wt_start_def) (simp add: lesub_def JVM_le_Err_conv Err.le_def) } moreover from in_A have "set \s' \ A" by simp with wt_err_step bounded_step have "wt_app_eff (sup_state_opt P) app eff (map ok_val \s')" by (auto intro: wt_err_imp_wt_app_eff simp add: l) ultimately have "wt_method P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt (map ok_val \s')" using instrs by (simp add: wt_method_def2 check_types_def del: map_map) thus ?thesis by blast qed (*>*) theorem (in start_context) wt_kil_complete: assumes wtm: "wt_method P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt \s" shows "wt_kildall P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt" (*<*) proof - from wtm obtain instrs: "0 < size is" and length: "length \s = length is" and ck_type: "check_types P mxs mxl (map OK \s)" and wt_start: "wt_start P C Ts mxl\<^sub>0 \s" and app_eff: "wt_app_eff (sup_state_opt P) app eff \s" by (simp add: wt_method_def2 check_types_def) from ck_type have in_A: "set (map OK \s) \ A" by (simp add: check_types_def) with app_eff in_A bounded_step have "wt_err_step (sup_state_opt P) (err_step (size \s) app eff) (map OK \s)" by - (erule wt_app_eff_imp_wt_err, auto simp add: exec_def length states_def) hence wt_err: "wt_err_step (sup_state_opt P) step (map OK \s)" by (simp add: length) have is_bcv: "is_bcv r Err step (size is) A (kiljvm P mxs mxl T\<^sub>r is xt)" by (rule is_bcv_kiljvm) moreover from instrs have "start \ list (size is) A" .. moreover let ?\s = "map OK \s" have less_\s: "start [\\<^sub>r] ?\s" proof (rule le_listI) from length instrs show "length start = length (map OK \s)" by simp next fix n from wt_start have "P \ ok_val (start!0) \' \s!0" by (simp add: wt_start_def) moreover from instrs length have "0 < length \s" by simp ultimately have "start!0 \\<^sub>r ?\s!0" by (simp add: JVM_le_Err_conv lesub_def) moreover { fix n' have "OK None \\<^sub>r ?\s!n" by (auto simp add: JVM_le_Err_conv Err.le_def lesub_def split: err.splits) hence "\n = Suc n'; n < size start\ \ start!n \\<^sub>r ?\s!n" by simp } ultimately show "n < size start \ start!n \\<^sub>r ?\s!n" by (cases n, blast+) qed moreover from ck_type length have "?\s \ list (size is) A" by (auto intro!: listI simp add: check_types_def) moreover from wt_err have "wt_step r Err step ?\s" by (simp add: wt_err_step_def JVM_le_Err_conv) ultimately have "\p. p < size is \ kiljvm P mxs mxl T\<^sub>r is xt start ! p \ Err" by (unfold is_bcv_def) blast with instrs show "wt_kildall P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt" by (unfold wt_kildall_def) simp qed (*>*) theorem jvm_kildall_correct: "wf_jvm_prog\<^sub>k P = wf_jvm_prog P" (*<*) -proof - +proof let ?\ = "\C M. let (C,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,is,xt)) = method P C M in SOME \s. wt_method P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt \s" - let ?A = "\P C' (M, Ts, T\<^sub>r, mxs, mxl\<^sub>0, is, xt). wt_kildall P C' Ts T\<^sub>r mxs mxl\<^sub>0 is xt" - let ?B\<^sub>\ = "\\. (\P C (M,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,is,xt)). - wt_method P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\ C M))" - show ?thesis proof(rule iffI) - \ \soundness\ - assume wt: "wf_jvm_prog\<^sub>k P" - then have "wf_prog ?A P" by(simp add: wf_jvm_prog\<^sub>k_def) - moreover { - fix wf_md C M Ts Ca T m bd - assume "wf_prog wf_md P" and sees: "P \ Ca sees M : Ts\T = m in Ca" and - "set Ts \ types P" and "bd = (M, Ts, T, m)" and - "?A P Ca bd" - then have "(?B\<^sub>\ ?\) P Ca bd" using sees_method_is_class[OF sees] - by (auto dest!: start_context.wt_kil_correct [OF start_context.intro] - intro: someI) - } - ultimately have "wf_prog (?B\<^sub>\ ?\) P" by(rule wf_prog_lift) - then have "wf_jvm_prog\<^bsub>?\\<^esub> P" by (simp add: wf_jvm_prog_phi_def) - thus "wf_jvm_prog P" by (unfold wf_jvm_prog_def) fast - next - \ \completeness\ - assume wt: "wf_jvm_prog P" - then obtain \ where "wf_prog (?B\<^sub>\ \) P" by(clarsimp simp: wf_jvm_prog_def wf_jvm_prog_phi_def) - moreover { - fix wf_md C M b Ts Ca T m bd - assume "wf_prog wf_md P" and sees: "P \ Ca sees M: Ts\T = m in Ca" and - "set Ts \ types P" and "bd = (M, Ts, T, m)" and - "(?B\<^sub>\ \) P Ca bd" - then have "?A P Ca bd" using sees_method_is_class[OF sees] - by (auto intro!: start_context.wt_kil_complete start_context.intro) - } - ultimately have "wf_prog ?A P" by(rule wf_prog_lift) - thus "wf_jvm_prog\<^sub>k P" by (simp add: wf_jvm_prog\<^sub>k_def) - qed + \ \soundness\ + assume wt: "wf_jvm_prog\<^sub>k P" + hence "wf_jvm_prog\<^bsub>?\\<^esub> P" + apply (unfold wf_jvm_prog_phi_def wf_jvm_prog\<^sub>k_def) + apply (erule wf_prog_lift) + apply (auto dest!: start_context.wt_kil_correct [OF start_context.intro] + intro: someI) + apply (erule sees_method_is_class) + done + thus "wf_jvm_prog P" by (unfold wf_jvm_prog_def) fast +next + \ \completeness\ + assume wt: "wf_jvm_prog P" + thus "wf_jvm_prog\<^sub>k P" + apply (unfold wf_jvm_prog_def wf_jvm_prog_phi_def wf_jvm_prog\<^sub>k_def) + apply (clarify) + apply (erule wf_prog_lift) + apply (auto intro!: start_context.wt_kil_complete start_context.intro) + apply (erule sees_method_is_class) + done qed (*>*) end diff --git a/thys/Jinja/BV/BVNoTypeError.thy b/thys/Jinja/BV/BVNoTypeError.thy --- a/thys/Jinja/BV/BVNoTypeError.thy +++ b/thys/Jinja/BV/BVNoTypeError.thy @@ -1,309 +1,312 @@ (* Title: HOL/MicroJava/BV/BVNoTypeErrors.thy Author: Gerwin Klein Copyright GPL *) section \Welltyped Programs produce no Type Errors\ theory BVNoTypeError imports "../JVM/JVMDefensive" BVSpecTypeSafe begin lemma has_methodI: "P \ C sees M:Ts\T = m in D \ P \ C has M" by (unfold has_method_def) blast text \ Some simple lemmas about the type testing functions of the defensive JVM: \ lemma typeof_NoneD [simp,dest]: "typeof v = Some x \ \is_Addr v" by (cases v) auto lemma is_Ref_def2: "is_Ref v = (v = Null \ (\a. v = Addr a))" by (cases v) (auto simp add: is_Ref_def) lemma [iff]: "is_Ref Null" by (simp add: is_Ref_def2) lemma is_RefI [intro, simp]: "P,h \ v :\ T \ is_refT T \ is_Ref v" (*<*) -proof (cases T) -qed (auto simp add: is_refT_def is_Ref_def dest: conf_ClassD) + apply (cases T) + apply (auto simp add: is_refT_def is_Ref_def dest: conf_ClassD) + done (*>*) lemma is_IntgI [intro, simp]: "P,h \ v :\ Integer \ is_Intg v" -(*<*)by (unfold conf_def) auto(*>*) +(*<*) + apply (unfold conf_def) + apply auto + done +(*>*) lemma is_BoolI [intro, simp]: "P,h \ v :\ Boolean \ is_Bool v" -(*<*)by (unfold conf_def) auto(*>*) +(*<*) + apply (unfold conf_def) + apply auto + done +(*>*) declare defs1 [simp del] lemma wt_jvm_prog_states: -assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" - and mC: "P \ C sees M: Ts\T = (mxs, mxl, ins, et) in C" - and \: "\ C M ! pc = \" and pc: "pc < size ins" -shows "OK \ \ states P mxs (1+size Ts+mxl)" + "\ wf_jvm_prog\<^bsub>\\<^esub> P; P \ C sees M: Ts\T = (mxs, mxl, ins, et) in C; + \ C M ! pc = \; pc < size ins \ + \ OK \ \ states P mxs (1+size Ts+mxl)" (*<*) -proof - - let ?wf_md = "(\P C (M, Ts, T\<^sub>r, mxs, mxl\<^sub>0, is, xt). - wt_method P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\ C M))" - have wfmd: "wf_prog ?wf_md P" using wf - by (unfold wf_jvm_prog_phi_def) assumption - show ?thesis using sees_wf_mdecl[OF wfmd mC] \ pc - by (simp add: wf_mdecl_def wt_method_def check_types_def) - (blast intro: nth_in) -qed + apply (unfold wf_jvm_prog_phi_def) + apply (drule (1) sees_wf_mdecl) + apply (simp add: wf_mdecl_def wt_method_def check_types_def) + apply (blast intro: nth_in) + done (*>*) text \ The main theorem: welltyped programs do not produce type errors if they are started in a conformant state. \ theorem no_type_error: fixes \ :: jvm_state assumes welltyped: "wf_jvm_prog\<^bsub>\\<^esub> P" and conforms: "P,\ \ \ \" shows "exec_d P \ \ TypeError" (*<*) proof - from welltyped obtain mb where wf: "wf_prog mb P" by (fast dest: wt_jvm_progD) obtain xcp h frs where s [simp]: "\ = (xcp, h, frs)" by (cases \) from conforms have "xcp \ None \ frs = [] \ check P \" by (unfold correct_state_def check_def) auto moreover { assume "\(xcp \ None \ frs = [])" then obtain stk reg C M pc frs' where xcp [simp]: "xcp = None" and frs [simp]: "frs = (stk,reg,C,M,pc)#frs'" by (clarsimp simp add: neq_Nil_conv) from conforms obtain ST LT Ts T mxs mxl ins xt where hconf: "P \ h \" and meth: "P \ C sees M:Ts\T = (mxs, mxl, ins, xt) in C" and \: "\ C M ! pc = Some (ST,LT)" and frame: "conf_f P h (ST,LT) ins (stk,reg,C,M,pc)" and frames: "conf_fs P h \ M (size Ts) T frs'" by (fastforce simp add: correct_state_def dest: sees_method_fun) from frame obtain stk: "P,h \ stk [:\] ST" and reg: "P,h \ reg [:\\<^sub>\] LT" and pc: "pc < size ins" by (simp add: conf_f_def) from welltyped meth \ pc have "OK (Some (ST, LT)) \ states P mxs (1+size Ts+mxl)" by (rule wt_jvm_prog_states) hence "size ST \ mxs" by (auto simp add: JVM_states_unfold) with stk have mxs: "size stk \ mxs" by (auto dest: list_all2_lengthD) from welltyped meth pc have "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" by (rule wt_jvm_prog_impl_wt_instr) hence app\<^sub>0: "app (ins!pc) P mxs T pc (size ins) xt (\ C M!pc) " by (simp add: wt_instr_def) with \ have eff: "\(pc',s')\set (eff (ins ! pc) P pc xt (\ C M ! pc)). pc' < size ins" by (unfold app_def) simp from app\<^sub>0 \ have app: "xcpt_app (ins!pc) P pc mxs xt (ST,LT) \ app\<^sub>i (ins!pc, P, pc, mxs, T, (ST,LT))" by (clarsimp simp add: app_def) with eff stk reg have "check_instr (ins!pc) P h stk reg C M pc frs'" proof (cases "ins!pc") case (Getfield F C) with app stk reg \ obtain v vT stk' where field: "P \ C sees F:vT in C" and stk: "stk = v # stk'" and conf: "P,h \ v :\ Class C" by auto from conf have is_Ref: "is_Ref v" by auto moreover { assume "v \ Null" with conf field is_Ref wf have "\D vs. h (the_Addr v) = Some (D,vs) \ P \ D \\<^sup>* C" by (auto dest!: non_npD) } - ultimately show ?thesis using Getfield field stk - has_field_mono[OF has_visible_field[OF field]] hconfD[OF hconf] - by (unfold oconf_def has_field_def) (fastforce dest: has_fields_fun) + ultimately show ?thesis using Getfield field stk hconf + apply clarsimp + apply (rule conjI, fastforce) + apply clarsimp + apply (drule has_visible_field) + apply (drule (1) has_field_mono) + apply (drule (1) hconfD) + apply (unfold oconf_def has_field_def) + apply clarsimp + apply (fastforce dest: has_fields_fun) + done next case (Putfield F C) with app stk reg \ obtain v ref vT stk' where field: "P \ C sees F:vT in C" and stk: "stk = v # ref # stk'" and confv: "P,h \ v :\ vT" and confr: "P,h \ ref :\ Class C" by fastforce from confr have is_Ref: "is_Ref ref" by simp moreover { assume "ref \ Null" with confr field is_Ref wf have "\D vs. h (the_Addr ref) = Some (D,vs) \ P \ D \\<^sup>* C" by (auto dest: non_npD) } ultimately show ?thesis using Putfield field stk confv by fastforce next case (Invoke M' n) with app have n: "n < size ST" by simp from stk have [simp]: "size stk = size ST" by (rule list_all2_lengthD) { assume "stk!n = Null" with n Invoke have ?thesis by simp } moreover { assume "ST!n = NT" with n stk have "stk!n = Null" by (auto simp: list_all2_conv_all_nth) with n Invoke have ?thesis by simp } moreover { assume Null: "stk!n \ Null" and NT: "ST!n \ NT" from NT app Invoke obtain D D' Ts T m where D: "ST!n = Class D" and M': "P \ D sees M': Ts\T = m in D'" and Ts: "P \ rev (take n ST) [\] Ts" by auto from D stk n have "P,h \ stk!n :\ Class D" by (auto simp: list_all2_conv_all_nth) with Null obtain a C' fs where [simp]: "stk!n = Addr a" "h a = Some (C',fs)" and "P \ C' \\<^sup>* D" by (fastforce dest!: conf_ClassD) with M' wf obtain m' Ts' T' D'' where C': "P \ C' sees M': Ts'\T' = m' in D''" and Ts': "P \ Ts [\] Ts'" by (auto dest!: sees_method_mono) from stk have "P,h \ take n stk [:\] take n ST" .. hence "P,h \ rev (take n stk) [:\] rev (take n ST)" .. also note Ts also note Ts' finally have "P,h \ rev (take n stk) [:\] Ts'" . with Invoke Null n C' have ?thesis by (auto simp add: is_Ref_def2 has_methodI) } ultimately show ?thesis by blast next case Return with stk app \ meth frames show ?thesis by (auto simp add: has_methodI) qed (auto simp add: list_all2_lengthD) hence "check P \" using meth pc mxs by (simp add: check_def has_methodI) } ultimately have "check P \" by blast thus "exec_d P \ \ TypeError" .. qed (*>*) text \ The theorem above tells us that, in welltyped programs, the defensive machine reaches the same result as the aggressive one (after arbitrarily many steps). \ theorem welltyped_aggressive_imp_defensive: "wf_jvm_prog\<^bsub>\\<^esub> P \ P,\ \ \ \ \ P \ \ -jvm\ \' \ P \ (Normal \) -jvmd\ (Normal \')" (*<*) -proof - - assume wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and cf: "P,\ \ \ \" and exec: "P \ \ -jvm\ \'" - then have "(\, \') \ {(\, \'). exec (P, \) = \\'\}\<^sup>*" by(simp only: exec_all_def) - then show ?thesis proof(induct rule: rtrancl_induct) - case base - then show ?case by (simp add: exec_all_d_def1) - next - case (step y z) - then have \y: "P \ \ -jvm\ y" by (simp only: exec_all_def [symmetric]) - have exec_d: "exec_d P y = Normal \z\" using step - no_type_error_commutes[OF no_type_error[OF wf BV_correct[OF wf \y cf]]] - by (simp add: exec_all_d_def1) - show ?case using step.hyps(3) exec_1_d_NormalI[OF exec_d] - by (simp add: exec_all_d_def1) - qed -qed + apply (simp only: exec_all_def) + apply (erule rtrancl_induct) + apply (simp add: exec_all_d_def1) + apply simp + apply (simp only: exec_all_def [symmetric]) + apply (frule BV_correct, assumption+) + apply (drule no_type_error, assumption, drule no_type_error_commutes, simp) + apply (simp add: exec_all_d_def1) + apply (rule rtrancl_trans, assumption) + apply (drule exec_1_d_NormalI) + apply auto + done (*>*) text \ As corollary we get that the aggressive and the defensive machine are equivalent for welltyped programs (if started in a conformant state or in the canonical start state) \ corollary welltyped_commutes: fixes \ :: jvm_state assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and conforms: "P,\ \ \ \" shows "P \ (Normal \) -jvmd\ (Normal \') = P \ \ -jvm\ \'" -proof(rule iffI) - assume "P \ Normal \ -jvmd\ Normal \'" then show "P \ \ -jvm\ \'" - by (rule defensive_imp_aggressive) -next - assume "P \ \ -jvm\ \'" then show "P \ Normal \ -jvmd\ Normal \'" - by (rule welltyped_aggressive_imp_defensive [OF wf conforms]) -qed + apply rule + apply (erule defensive_imp_aggressive) + apply (erule welltyped_aggressive_imp_defensive [OF wf conforms]) + done corollary welltyped_initial_commutes: assumes wf: "wf_jvm_prog P" assumes meth: "P \ C sees M:[]\T = b in C" defines start: "\ \ start_state P C M" shows "P \ (Normal \) -jvmd\ (Normal \') = P \ \ -jvm\ \'" proof - from wf obtain \ where wf': "wf_jvm_prog\<^bsub>\\<^esub> P" by (auto simp: wf_jvm_prog_def) from this meth have "P,\ \ \ \" unfolding start by (rule BV_correct_initial) with wf' show ?thesis by (rule welltyped_commutes) qed lemma not_TypeError_eq [iff]: "x \ TypeError = (\t. x = Normal t)" by (cases x) auto locale cnf = fixes P and \ and \ assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" assumes cnf: "correct_state P \ \" theorem (in cnf) no_type_errors: "P \ (Normal \) -jvmd\ \' \ \' \ TypeError" -proof - - assume "P \ (Normal \) -jvmd\ \'" - then have "(Normal \, \') \ (exec_1_d P)\<^sup>*" by (unfold exec_all_d_def1) simp - then show ?thesis proof(induct rule: rtrancl_induct) - case (step y z) - then obtain y\<^sub>n where [simp]: "y = Normal y\<^sub>n" by clarsimp - have n\y: "P \ Normal \ -jvmd\ Normal y\<^sub>n" using step.hyps(1) - by (fold exec_all_d_def1) simp - have \y: "P \ \ -jvm\ y\<^sub>n" using defensive_imp_aggressive[OF n\y] by simp - show ?case using step no_type_error[OF wf BV_correct[OF wf \y cnf]] - by (auto simp add: exec_1_d_eq) - qed simp -qed + apply (unfold exec_all_d_def1) + apply (erule rtrancl_induct) + apply simp + apply (fold exec_all_d_def1) + apply (insert cnf wf) + apply clarsimp + apply (drule defensive_imp_aggressive) + apply (frule (2) BV_correct) + apply (drule (1) no_type_error) back + apply (auto simp add: exec_1_d_eq) + done locale start = fixes P and C and M and \ and T and b assumes wf: "wf_jvm_prog P" assumes sees: "P \ C sees M:[]\T = b in C" defines "\ \ Normal (start_state P C M)" corollary (in start) bv_no_type_error: shows "P \ \ -jvmd\ \' \ \' \ TypeError" proof - from wf obtain \ where "wf_jvm_prog\<^bsub>\\<^esub> P" by (auto simp: wf_jvm_prog_def) moreover with sees have "correct_state P \ (start_state P C M)" by - (rule BV_correct_initial) ultimately have "cnf P \ (start_state P C M)" by (rule cnf.intro) moreover assume "P \ \ -jvmd\ \'" ultimately show ?thesis by (unfold \_def) (rule cnf.no_type_errors) qed end diff --git a/thys/Jinja/BV/BVSpec.thy b/thys/Jinja/BV/BVSpec.thy --- a/thys/Jinja/BV/BVSpec.thy +++ b/thys/Jinja/BV/BVSpec.thy @@ -1,102 +1,102 @@ (* Title: HOL/MicroJava/BV/BVSpec.thy Author: Cornelia Pusch, Gerwin Klein Copyright 1999 Technische Universitaet Muenchen *) section \The Bytecode Verifier \label{sec:BVSpec}\ theory BVSpec imports Effect begin text \ This theory contains a specification of the BV. The specification describes correct typings of method bodies; it corresponds to type \emph{checking}. \ definition \ \The method type only contains declared classes:\ check_types :: "'m prog \ nat \ nat \ ty\<^sub>i' err list \ bool" where "check_types P mxs mxl \s \ set \s \ states P mxs mxl" \ \An instruction is welltyped if it is applicable and its effect\ \ \is compatible with the type at all successor instructions:\ definition wt_instr :: "['m prog,ty,nat,pc,ex_table,instr,pc,ty\<^sub>m] \ bool" ("_,_,_,_,_ \ _,_ :: _" [60,0,0,0,0,0,0,61] 60) where "P,T,mxs,mpc,xt \ i,pc :: \s \ app i P mxs T pc mpc xt (\s!pc) \ (\(pc',\') \ set (eff i P pc xt (\s!pc)). P \ \' \' \s!pc')" \ \The type at @{text "pc=0"} conforms to the method calling convention:\ definition wt_start :: "['m prog,cname,ty list,nat,ty\<^sub>m] \ bool" where "wt_start P C Ts mxl\<^sub>0 \s \ P \ Some ([],OK (Class C)#map OK Ts@replicate mxl\<^sub>0 Err) \' \s!0" \ \A method is welltyped if the body is not empty,\ \ \if the method type covers all instructions and mentions\ \ \declared classes only, if the method calling convention is respected, and\ \ \if all instructions are welltyped.\ definition wt_method :: "['m prog,cname,ty list,ty,nat,nat,instr list, ex_table,ty\<^sub>m] \ bool" where "wt_method P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt \s \ 0 < size is \ size \s = size is \ check_types P mxs (1+size Ts+mxl\<^sub>0) (map OK \s) \ wt_start P C Ts mxl\<^sub>0 \s \ (\pc < size is. P,T\<^sub>r,mxs,size is,xt \ is!pc,pc :: \s)" \ \A program is welltyped if it is wellformed and all methods are welltyped\ definition wf_jvm_prog_phi :: "ty\<^sub>P \ jvm_prog \ bool" ("wf'_jvm'_prog\<^bsub>_\<^esub>") where "wf_jvm_prog\<^bsub>\\<^esub> \ wf_prog (\P C (M,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,is,xt)). wt_method P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\ C M))" definition wf_jvm_prog :: "jvm_prog \ bool" where "wf_jvm_prog P \ \\. wf_jvm_prog\<^bsub>\\<^esub> P" lemma wt_jvm_progD: "wf_jvm_prog\<^bsub>\\<^esub> P \ \wt. wf_prog wt P" (*<*) by (unfold wf_jvm_prog_phi_def, blast) (*>*) lemma wt_jvm_prog_impl_wt_instr: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and sees: "P \ C sees M:Ts \ T = (mxs,mxl\<^sub>0,ins,xt) in C" and pc: "pc < size ins" shows "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" (*<*) proof - have wfm: "wf_prog (\P C (M, Ts, T\<^sub>r, mxs, mxl\<^sub>0, is, xt). wt_method P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\ C M)) P" using wf by (unfold wf_jvm_prog_phi_def) show ?thesis using sees_wf_mdecl[OF wfm sees] pc by (simp add: wf_mdecl_def wt_method_def) qed (*>*) lemma wt_jvm_prog_impl_wt_start: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and sees: "P \ C sees M:Ts \ T = (mxs,mxl\<^sub>0,ins,xt) in C" shows "0 < size ins \ wt_start P C Ts mxl\<^sub>0 (\ C M)" (*<*) proof - have wfm: "wf_prog (\P C (M, Ts, T\<^sub>r, mxs, mxl\<^sub>0, is, xt). wt_method P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\ C M)) P" using wf by (unfold wf_jvm_prog_phi_def) show ?thesis using sees_wf_mdecl[OF wfm sees] by (simp add: wf_mdecl_def wt_method_def) qed -(*>*) + end diff --git a/thys/Jinja/BV/BVSpecTypeSafe.thy b/thys/Jinja/BV/BVSpecTypeSafe.thy --- a/thys/Jinja/BV/BVSpecTypeSafe.thy +++ b/thys/Jinja/BV/BVSpecTypeSafe.thy @@ -1,1102 +1,1098 @@ (* Title: HOL/MicroJava/BV/BVSpecTypeSafe.thy Author: Cornelia Pusch, Gerwin Klein Copyright 1999 Technische Universitaet Muenchen *) section \BV Type Safety Proof \label{sec:BVSpecTypeSafe}\ theory BVSpecTypeSafe imports BVConform begin text \ This theory contains proof that the specification of the bytecode verifier only admits type safe programs. \ subsection \Preliminaries\ text \ Simp and intro setup for the type safety proof: \ lemmas defs1 = correct_state_def conf_f_def wt_instr_def eff_def norm_eff_def app_def xcpt_app_def lemmas widen_rules [intro] = conf_widen confT_widen confs_widens confTs_widen subsection \Exception Handling\ text \ For the \Invoke\ instruction the BV has checked all handlers that guard the current \pc\. \ lemma Invoke_handlers: "match_ex_table P C pc xt = Some (pc',d') \ \(f,t,D,h,d) \ set (relevant_entries P (Invoke n M) pc xt). P \ C \\<^sup>* D \ pc \ {f.. pc' = h \ d' = d" by (induct xt) (auto simp add: relevant_entries_def matches_ex_entry_def is_relevant_entry_def split: if_split_asm) text \ We can prove separately that the recursive search for exception handlers (\find_handler\) in the frame stack results in a conforming state (if there was no matching exception handler in the current frame). We require that the exception is a valid heap address, and that the state before the exception occurred conforms. \ term find_handler lemma uncaught_xcpt_correct: assumes wt: "wf_jvm_prog\<^bsub>\\<^esub> P" assumes h: "h xcp = Some obj" shows "\f. P,\ \ (None, h, f#frs)\ \ P,\ \ (find_handler P xcp h frs) \" (is "\f. ?correct (None, h, f#frs) \ ?correct (?find frs)") (*<*) proof (induct frs) \ \the base case is trivial as it should be\ show "?correct (?find [])" by (simp add: correct_state_def) next \ \we will need both forms @{text wf_jvm_prog} and @{text wf_prog} later\ from wt obtain mb where wf: "wf_prog mb P" by (simp add: wf_jvm_prog_phi_def) \ \the assumption for the cons case:\ fix f f' frs' assume cr: "?correct (None, h, f#f'#frs')" \ \the induction hypothesis:\ assume IH: "\f. ?correct (None, h, f#frs') \ ?correct (?find frs')" from cr have cr': "?correct (None, h, f'#frs')" by (fastforce simp add: correct_state_def) obtain stk loc C M pc where [simp]: "f' = (stk,loc,C,M,pc)" by (cases f') from cr obtain Ts T mxs mxl\<^sub>0 ins xt where meth: "P \ C sees M:Ts \ T = (mxs,mxl\<^sub>0,ins,xt) in C" by (simp add: correct_state_def, blast) hence [simp]: "ex_table_of P C M = xt" by simp show "?correct (?find (f'#frs'))" proof (cases "match_ex_table P (cname_of h xcp) pc xt") case None with cr' IH [of f'] show ?thesis by fastforce next fix pc_d assume "match_ex_table P (cname_of h xcp) pc xt = Some pc_d" then obtain pc' d' where match: "match_ex_table P (cname_of h xcp) pc xt = Some (pc',d')" (is "?match (cname_of h xcp) = _") by (cases pc_d) auto from wt meth cr' [simplified] have wti: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" by (fastforce simp add: correct_state_def conf_f_def dest: sees_method_fun elim!: wt_jvm_prog_impl_wt_instr) from cr meth obtain n M' ST LT where ins: "ins!pc = Invoke n M'" (is "_ = ?i") and \: "\ C M ! pc = Some (ST, LT)" by (fastforce dest: sees_method_fun simp add: correct_state_def) from ins match obtain f t D where rel: "(f,t,D,pc',d') \ set (relevant_entries P (ins!pc) pc xt)" and D: "P \ cname_of h xcp \\<^sup>* D" by (fastforce dest: Invoke_handlers) from rel have "(pc', Some (Class D # drop (size ST - d') ST, LT)) \ set (xcpt_eff (ins!pc) P pc (ST,LT) xt)" (is "(_, Some (?ST',_)) \ _") by (force simp add: xcpt_eff_def image_def) with wti \ obtain pc: "pc' < size ins" and "P \ Some (?ST', LT) \' \ C M ! pc'" by (clarsimp simp add: defs1) blast then obtain ST' LT' where \': "\ C M ! pc' = Some (ST',LT')" and less: "P \ (?ST', LT) \\<^sub>i (ST',LT')" by (auto simp add: sup_state_opt_any_Some) from cr' \ meth have "conf_f P h (ST, LT) ins f'" by (unfold correct_state_def) (fastforce dest: sees_method_fun) hence loc: "P,h \ loc [:\\<^sub>\] LT" and stk: "P,h \ stk [:\] ST" by (unfold conf_f_def) auto hence [simp]: "size stk = size ST" by (simp add: list_all2_lengthD) let ?f = "(Addr xcp # drop (length stk - d') stk, loc, C, M, pc')" have "conf_f P h (ST',LT') ins ?f" proof - from wf less loc have "P,h \ loc [:\\<^sub>\] LT'" by simp blast moreover from D h have "P,h \ Addr xcp :\ Class D" by (simp add: conf_def obj_ty_def case_prod_unfold) with less stk have "P,h \ Addr xcp # drop (length stk - d') stk [:\] ST'" by (auto intro!: list_all2_dropI) ultimately show ?thesis using pc by (simp add: conf_f_def) qed with cr' match \' meth pc show ?thesis by (unfold correct_state_def) (fastforce dest: sees_method_fun) qed qed (*>*) text \ The requirement of lemma \uncaught_xcpt_correct\ (that the exception is a valid reference on the heap) is always met for welltyped instructions and conformant states: \ lemma exec_instr_xcpt_h: "\ fst (exec_instr (ins!pc) P h stk vars Cl M pc frs) = Some xcp; P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\ \ \ \obj. h xcp = Some obj" (is "\ ?xcpt; ?wt; ?correct \ \ ?thesis") (*<*) proof - note [simp] = split_beta note [split] = if_split_asm option.split_asm assume wt: ?wt ?correct hence pre: "preallocated h" by (simp add: correct_state_def hconf_def) assume xcpt: ?xcpt with pre show ?thesis proof (cases "ins!pc") case Throw with xcpt wt pre show ?thesis by (clarsimp iff: list_all2_Cons2 simp add: defs1) (auto dest: non_npD simp: is_refT_def elim: preallocatedE) qed (auto elim: preallocatedE) qed (*>*) lemma conf_sys_xcpt: "\preallocated h; C \ sys_xcpts\ \ P,h \ Addr (addr_of_sys_xcpt C) :\ Class C" by (auto elim: preallocatedE) lemma match_ex_table_SomeD: "match_ex_table P C pc xt = Some (pc',d') \ \(f,t,D,h,d) \ set xt. matches_ex_entry P C pc (f,t,D,h,d) \ h = pc' \ d=d'" by (induct xt) (auto split: if_split_asm) text \ Finally we can state that, whenever an exception occurs, the next state always conforms: \ lemma xcpt_correct: fixes \' :: jvm_state assumes wtp: "wf_jvm_prog\<^bsub>\\<^esub> P" assumes meth: "P \ C sees M:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" assumes wt: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" assumes xp: "fst (exec_instr (ins!pc) P h stk loc C M pc frs) = Some xcp" assumes s': "Some \' = exec (P, None, h, (stk,loc,C,M,pc)#frs)" assumes correct: "P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\" shows "P,\ \ \'\" (*<*) proof - from wtp obtain wfmb where wf: "wf_prog wfmb P" by (simp add: wf_jvm_prog_phi_def) note conf_sys_xcpt [elim!] note xp' = meth s' xp note wtp moreover from xp wt correct obtain obj where h: "h xcp = Some obj" by (blast dest: exec_instr_xcpt_h) moreover note correct ultimately have "P,\ \ find_handler P xcp h frs \" by (rule uncaught_xcpt_correct) with xp' have "match_ex_table P (cname_of h xcp) pc xt = None \ ?thesis" (is "?m (cname_of h xcp) = _ \ _" is "?match = _ \ _") by (simp add: split_beta) moreover { fix pc_d assume "?match = Some pc_d" then obtain pc' d' where some_handler: "?match = Some (pc',d')" by (cases pc_d) auto from correct meth obtain ST LT where h_ok: "P \ h \" and \_pc: "\ C M ! pc = Some (ST, LT)" and frame: "conf_f P h (ST,LT) ins (stk,loc,C,M,pc)" and frames: "conf_fs P h \ M (size Ts) T frs" by (unfold correct_state_def) (fastforce dest: sees_method_fun) from h_ok have preh: "preallocated h" by (simp add: hconf_def) from frame obtain stk: "P,h \ stk [:\] ST" and loc: "P,h \ loc [:\\<^sub>\] LT" and pc: "pc < size ins" by (unfold conf_f_def) auto from stk have [simp]: "size stk = size ST" .. from wt \_pc have eff: "\(pc', s')\set (xcpt_eff (ins!pc) P pc (ST,LT) xt). pc' < size ins \ P \ s' \' \ C M!pc'" by (auto simp add: defs1) let ?stk' = "Addr xcp # drop (length stk - d') stk" let ?f = "(?stk', loc, C, M, pc')" from some_handler xp' have \': "\' = (None, h, ?f#frs)" by (simp add: split_beta) from some_handler obtain f t D where xt: "(f,t,D,pc',d') \ set xt" and "matches_ex_entry P (cname_of h xcp) pc (f,t,D,pc',d')" by (auto dest: match_ex_table_SomeD) hence match: "P \ cname_of h xcp \\<^sup>* D" "pc \ {f.. set (relevant_entries P (ins!pc) pc xt)" and conf: "P,h \ Addr xcp :\ Class D" proof (cases "ins!pc") case Return with xp have False by (auto simp: split_beta split: if_split_asm) thus ?thesis .. next case New with xp have [simp]: "xcp = addr_of_sys_xcpt OutOfMemory" by simp with New match preh have "is_relevant_entry P (ins!pc) pc (f,t,D,pc',d')" by (simp add: is_relevant_entry_def) with match preh xt show ?thesis by (fastforce simp add: relevant_entries_def intro: that) next case Getfield with xp have [simp]: "xcp = addr_of_sys_xcpt NullPointer" by (simp add: split_beta split: if_split_asm) with Getfield match preh have "is_relevant_entry P (ins!pc) pc (f,t,D,pc',d')" by (simp add: is_relevant_entry_def) with match preh xt show ?thesis by (fastforce simp add: relevant_entries_def intro: that) next case Putfield with xp have [simp]: "xcp = addr_of_sys_xcpt NullPointer" by (simp add: split_beta split: if_split_asm) with Putfield match preh have "is_relevant_entry P (ins!pc) pc (f,t,D,pc',d')" by (simp add: is_relevant_entry_def) with match preh xt show ?thesis by (fastforce simp add: relevant_entries_def intro: that) next case Checkcast with xp have [simp]: "xcp = addr_of_sys_xcpt ClassCast" by (simp add: split_beta split: if_split_asm) with Checkcast match preh have "is_relevant_entry P (ins!pc) pc (f,t,D,pc',d')" by (simp add: is_relevant_entry_def) with match preh xt show ?thesis by (fastforce simp add: relevant_entries_def intro: that) next case Invoke with xp match have "is_relevant_entry P (ins!pc) pc (f,t,D,pc',d')" by (simp add: is_relevant_entry_def) moreover from xp wt correct obtain obj where xcp: "h xcp = Some obj" by (blast dest: exec_instr_xcpt_h) ultimately show ?thesis using xt match by (auto simp add: relevant_entries_def conf_def case_prod_unfold intro: that) next case Throw with xp match preh have "is_relevant_entry P (ins!pc) pc (f,t,D,pc',d')" by (simp add: is_relevant_entry_def) moreover from xp wt correct obtain obj where xcp: "h xcp = Some obj" by (blast dest: exec_instr_xcpt_h) ultimately show ?thesis using xt match by (auto simp add: relevant_entries_def conf_def case_prod_unfold intro: that) qed auto with eff obtain ST' LT' where \_pc': "\ C M ! pc' = Some (ST', LT')" and pc': "pc' < size ins" and less: "P \ (Class D # drop (size ST - d') ST, LT) \\<^sub>i (ST', LT')" by (fastforce simp add: xcpt_eff_def sup_state_opt_any_Some) with conf loc stk have "conf_f P h (ST',LT') ins ?f" by (auto simp add: defs1 intro: list_all2_dropI) with meth h_ok frames \_pc' \' have ?thesis by (unfold correct_state_def) (fastforce dest: sees_method_fun) } ultimately show ?thesis by (cases "?match") blast+ qed (*>*) subsection \Single Instructions\ text \ In this section we prove for each single (welltyped) instruction that the state after execution of the instruction still conforms. Since we have already handled exceptions above, we can now assume that no exception occurs in this step. \ declare defs1 [simp] lemma Invoke_correct: fixes \' :: jvm_state assumes wtprog: "wf_jvm_prog\<^bsub>\\<^esub> P" assumes meth_C: "P \ C sees M:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" assumes ins: "ins ! pc = Invoke M' n" assumes wti: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" assumes \': "Some \' = exec (P, None, h, (stk,loc,C,M,pc)#frs)" assumes approx: "P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\" assumes no_xcp: "fst (exec_instr (ins!pc) P h stk loc C M pc frs) = None" shows "P,\ \ \'\" (*<*) proof - note split_paired_Ex [simp del] from wtprog obtain wfmb where wfprog: "wf_prog wfmb P" by (simp add: wf_jvm_prog_phi_def) from ins meth_C approx obtain ST LT where heap_ok: "P\ h\" and \_pc: "\ C M!pc = Some (ST,LT)" and frame: "conf_f P h (ST,LT) ins (stk,loc,C,M,pc)" and frames: "conf_fs P h \ M (size Ts) T frs" by (fastforce dest: sees_method_fun) from ins wti \_pc have n: "n < size ST" by simp { assume "stk!n = Null" with ins no_xcp have False by (simp add: split_beta) hence ?thesis .. } moreover { assume "ST!n = NT" moreover from frame have "P,h \ stk [:\] ST" by simp with n have "P,h \ stk!n :\ ST!n" by (simp add: list_all2_conv_all_nth) ultimately have "stk!n = Null" by simp with ins no_xcp have False by (simp add: split_beta) hence ?thesis .. } moreover { assume NT: "ST!n \ NT" and Null: "stk!n \ Null" from NT ins wti \_pc obtain D D' Ts T m ST' LT' where D: "ST!n = Class D" and pc': "pc+1 < size ins" and m_D: "P \ D sees M': Ts\T = m in D'" and Ts: "P \ rev (take n ST) [\] Ts" and \': "\ C M ! (pc+1) = Some (ST', LT')" and LT': "P \ LT [\\<^sub>\] LT'" and ST': "P \ (T # drop (n+1) ST) [\] ST'" by (clarsimp simp add: sup_state_opt_any_Some) from frame obtain stk: "P,h \ stk [:\] ST" and loc: "P,h \ loc [:\\<^sub>\] LT" by simp from n stk D have "P,h \ stk!n :\ Class D" by (auto simp add: list_all2_conv_all_nth) with Null obtain a C' fs where Addr: "stk!n = Addr a" and obj: "h a = Some (C',fs)" and C'subD: "P \ C' \\<^sup>* D" by (fastforce dest!: conf_ClassD) with wfprog m_D obtain Ts' T' m' D'' mxs' mxl' ins' xt' where m_C': "P \ C' sees M': Ts'\T' = (mxs',mxl',ins',xt') in D''" and T': "P \ T' \ T" and Ts': "P \ Ts [\] Ts'" by (auto dest: sees_method_mono) let ?loc' = "Addr a # rev (take n stk) @ replicate mxl' undefined" let ?f' = "([], ?loc', D'', M', 0)" let ?f = "(stk, loc, C, M, pc)" from Addr obj m_C' ins \' meth_C have s': "\' = (None, h, ?f' # ?f # frs)" by simp from Ts n have [simp]: "size Ts = n" by (auto dest: list_all2_lengthD simp: min_def) with Ts' have [simp]: "size Ts' = n" by (auto dest: list_all2_lengthD) from m_C' wfprog obtain mD'': "P \ D'' sees M':Ts'\T'=(mxs',mxl',ins',xt') in D''" by (fast dest: sees_method_idemp) moreover with wtprog obtain start: "wt_start P D'' Ts' mxl' (\ D'' M')" and ins': "ins' \ []" by (auto dest: wt_jvm_prog_impl_wt_start) then obtain LT\<^sub>0 where LT\<^sub>0: "\ D'' M' ! 0 = Some ([], LT\<^sub>0)" by (clarsimp simp add: wt_start_def defs1 sup_state_opt_any_Some) moreover have "conf_f P h ([], LT\<^sub>0) ins' ?f'" proof - let ?LT = "OK (Class D'') # (map OK Ts') @ (replicate mxl' Err)" from stk have "P,h \ take n stk [:\] take n ST" .. hence "P,h \ rev (take n stk) [:\] rev (take n ST)" by simp also note Ts also note Ts' finally have "P,h \ rev (take n stk) [:\\<^sub>\] map OK Ts'" by simp also have "P,h \ replicate mxl' undefined [:\\<^sub>\] replicate mxl' Err" by simp also from m_C' have "P \ C' \\<^sup>* D''" by (rule sees_method_decl_above) with obj have "P,h \ Addr a :\ Class D''" by (simp add: conf_def) ultimately have "P,h \ ?loc' [:\\<^sub>\] ?LT" by simp also from start LT\<^sub>0 have "P \ \ [\\<^sub>\] LT\<^sub>0" by (simp add: wt_start_def) finally have "P,h \ ?loc' [:\\<^sub>\] LT\<^sub>0" . thus ?thesis using ins' by simp qed ultimately have ?thesis using s' \_pc approx meth_C m_D T' ins D by (fastforce dest: sees_method_fun [of _ C]) } ultimately show ?thesis by blast qed (*>*) declare list_all2_Cons2 [iff] lemma Return_correct: fixes \' :: jvm_state assumes wt_prog: "wf_jvm_prog\<^bsub>\\<^esub> P" assumes meth: "P \ C sees M:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" assumes ins: "ins ! pc = Return" assumes wt: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" assumes s': "Some \' = exec (P, None, h, (stk,loc,C,M,pc)#frs)" assumes correct: "P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\" shows "P,\ \ \'\" (*<*) proof - from wt_prog obtain wfmb where wf: "wf_prog wfmb P" by (simp add: wf_jvm_prog_phi_def) from meth ins s' have "frs = [] \ ?thesis" by (simp add: correct_state_def) moreover { fix f frs' assume frs': "frs = f#frs'" moreover obtain stk' loc' C' M' pc' where f: "f = (stk',loc',C',M',pc')" by (cases f) moreover note meth ins s' ultimately have \': "\' = (None,h,(hd stk#(drop (1+size Ts) stk'),loc',C',M',pc'+1)#frs')" (is "\' = (None,h,?f'#frs')") by simp from correct meth obtain ST LT where h_ok: "P \ h \" and \_pc: "\ C M ! pc = Some (ST, LT)" and frame: "conf_f P h (ST, LT) ins (stk,loc,C,M,pc)" and frames: "conf_fs P h \ M (size Ts) T frs" by (auto dest: sees_method_fun simp add: correct_state_def) from \_pc ins wt obtain U ST\<^sub>0 where "ST = U # ST\<^sub>0" "P \ U \ T" by (simp add: wt_instr_def app_def) blast with wf frame have hd_stk: "P,h \ hd stk :\ T" by (auto simp add: conf_f_def) from f frs' frames obtain ST' LT' Ts'' T'' mxs' mxl\<^sub>0' ins' xt' D Ts' T' m D' where \': "\ C' M' ! pc' = Some (ST', LT')" and meth_C': "P \ C' sees M':Ts''\T''=(mxs',mxl\<^sub>0',ins',xt') in C'" and ins': "ins' ! pc' = Invoke M (size Ts)" and D: "ST' ! (size Ts) = Class D" and meth_D: "P \ D sees M: Ts'\T' = m in D'" and T': "P \ T \ T'" and frame': "conf_f P h (ST',LT') ins' f" and conf_fs: "conf_fs P h \ M' (size Ts'') T'' frs'" by clarsimp from f frame' obtain stk': "P,h \ stk' [:\] ST'" and loc': "P,h \ loc' [:\\<^sub>\] LT'" and pc': "pc' < size ins'" by (simp add: conf_f_def) from wt_prog meth_C' pc' have "P,T'',mxs',size ins',xt' \ ins'!pc',pc' :: \ C' M'" by (rule wt_jvm_prog_impl_wt_instr) with ins' \' D meth_D obtain aTs ST'' LT'' where \_suc: "\ C' M' ! Suc pc' = Some (ST'', LT'')" and less: "P \ (T' # drop (size Ts+1) ST', LT') \\<^sub>i (ST'', LT'')" and suc_pc': "Suc pc' < size ins'" by (clarsimp simp add: sup_state_opt_any_Some) from hd_stk T' have hd_stk': "P,h \ hd stk :\ T'" .. have frame'': "conf_f P h (ST'',LT'') ins' ?f'" proof - from stk' have "P,h \ drop (1+size Ts) stk' [:\] drop (1+size Ts) ST'" .. moreover with hd_stk' less have "P,h \ hd stk # drop (1+size Ts) stk' [:\] ST''" by auto moreover from wf loc' less have "P,h \ loc' [:\\<^sub>\] LT''" by auto moreover note suc_pc' ultimately show ?thesis by (simp add: conf_f_def) qed with \' frs' f meth h_ok hd_stk \_suc frames meth_C' \' have ?thesis by (fastforce dest: sees_method_fun [of _ C']) } ultimately show ?thesis by (cases frs) blast+ qed (*>*) declare sup_state_opt_any_Some [iff] declare not_Err_eq [iff] lemma Load_correct: "\ wf_prog wt P; P \ C sees M:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C; ins!pc = Load idx; P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; Some \' = exec (P, None, h, (stk,loc,C,M,pc)#frs); P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\ \ \ P,\ \ \'\" by (fastforce dest: sees_method_fun [of _ C] elim!: confTs_confT_sup) declare [[simproc del: list_to_set_comprehension]] lemma Store_correct: "\ wf_prog wt P; P \ C sees M:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C; ins!pc = Store idx; P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; Some \' = exec (P, None, h, (stk,loc,C,M,pc)#frs); P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\ \ \ P,\ \ \'\" -(*<*)by clarsimp (blast dest: sees_method_fun intro!: list_all2_update_cong)(*>*) +(*<*) + apply clarsimp + apply (drule (1) sees_method_fun) + apply clarsimp + apply (blast intro!: list_all2_update_cong) + done +(*>*) lemma Push_correct: "\ wf_prog wt P; P \ C sees M:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C; ins!pc = Push v; P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; Some \' = exec (P, None, h, (stk,loc,C,M,pc)#frs); P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\ \ \ P,\ \ \'\" -(*<*)by clarsimp (blast dest: sees_method_fun typeof_lit_conf)(*>*) +(*<*) + apply clarsimp + apply (drule (1) sees_method_fun) + apply clarsimp + apply (blast dest: typeof_lit_conf) + done +(*>*) lemma Cast_conf2: "\ wf_prog ok P; P,h \ v :\ T; is_refT T; cast_ok P C h v; P \ Class C \ T'; is_class P C\ \ P,h \ v :\ T'" (*<*) -proof - - assume "wf_prog ok P" and "P,h \ v :\ T" and "is_refT T" and - "cast_ok P C h v" and wid: "P \ Class C \ T'" and "is_class P C" - then show "P,h \ v :\ T'" using Class_widen[OF wid] - by(cases v) - (auto simp: cast_ok_def is_refT_def conf_def obj_ty_def - intro: rtrancl_trans) -qed + apply (unfold cast_ok_def is_refT_def) + apply (frule Class_widen) + apply (elim exE disjE) + apply simp + apply simp + apply simp + apply (clarsimp simp add: conf_def obj_ty_def) + apply (cases v) + apply (auto intro: rtrancl_trans) + done (*>*) lemma Checkcast_correct: "\ wf_jvm_prog\<^bsub>\\<^esub> P; P \ C sees M:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C; ins!pc = Checkcast D; P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; Some \' = exec (P, None, h, (stk,loc,C,M,pc)#frs) ; P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\; fst (exec_instr (ins!pc) P h stk loc C M pc frs) = None \ \ P,\ \ \'\" (*<*) -by (clarsimp simp: wf_jvm_prog_phi_def split: if_split_asm) - (blast intro: Cast_conf2 dest: sees_method_fun) + apply (clarsimp simp add: wf_jvm_prog_phi_def split: if_split_asm) + apply (drule (1) sees_method_fun) + apply (blast intro: Cast_conf2) + done (*>*) declare split_paired_All [simp del] lemmas widens_Cons [iff] = list_all2_Cons1 [of "widen P"] for P lemma Getfield_correct: fixes \' :: jvm_state assumes wf: "wf_prog wt P" assumes mC: "P \ C sees M:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" assumes i: "ins!pc = Getfield F D" assumes wt: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" assumes s': "Some \' = exec (P, None, h, (stk,loc,C,M,pc)#frs)" assumes cf: "P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\" assumes xc: "fst (exec_instr (ins!pc) P h stk loc C M pc frs) = None" shows "P,\ \ \'\" (*<*) proof - from mC cf obtain ST LT where "h\": "P \ h \" and \: "\ C M ! pc = Some (ST,LT)" and stk: "P,h \ stk [:\] ST" and loc: "P,h \ loc [:\\<^sub>\] LT" and pc: "pc < size ins" and fs: "conf_fs P h \ M (size Ts) T frs" by (fastforce dest: sees_method_fun) from i \ wt obtain oT ST'' vT ST' LT' vT' where oT: "P \ oT \ Class D" and ST: "ST = oT # ST''" and F: "P \ D sees F:vT in D" and pc': "pc+1 < size ins" and \': "\ C M ! (pc+1) = Some (vT'#ST', LT')" and ST': "P \ ST'' [\] ST'" and LT': "P \ LT [\\<^sub>\] LT'" and vT': "P \ vT \ vT'" by fastforce from stk ST obtain ref stk' where stk': "stk = ref#stk'" and ref: "P,h \ ref :\ oT" and ST'': "P,h \ stk' [:\] ST''" by auto from stk' i mC s' xc have "ref \ Null" by (simp add: split_beta split:if_split_asm) moreover from ref oT have "P,h \ ref :\ Class D" .. ultimately obtain a D' fs where a: "ref = Addr a" and h: "h a = Some (D', fs)" and D': "P \ D' \\<^sup>* D" by (blast dest: non_npD) from D' F have has_field: "P \ D' has F:vT in D" by (blast intro: has_field_mono has_visible_field) moreover from "h\" h have "P,h \ (D', fs) \" by (rule hconfD) ultimately obtain v where v: "fs (F, D) = Some v" "P,h \ v :\ vT" by (clarsimp simp add: oconf_def has_field_def) (blast dest: has_fields_fun) from a h i mC s' stk' v have "\' = (None, h, (v#stk',loc,C,M,pc+1)#frs)" by simp moreover from ST'' ST' have "P,h \ stk' [:\] ST'" .. moreover from v vT' have "P,h \ v :\ vT'" by blast moreover from loc LT' have "P,h \ loc [:\\<^sub>\] LT'" .. moreover note "h\" mC \' pc' v fs ultimately show "P,\ \ \' \" by fastforce qed (*>*) lemma Putfield_correct: fixes \' :: jvm_state assumes wf: "wf_prog wt P" assumes mC: "P \ C sees M:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" assumes i: "ins!pc = Putfield F D" assumes wt: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" assumes s': "Some \' = exec (P, None, h, (stk,loc,C,M,pc)#frs)" assumes cf: "P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\" assumes xc: "fst (exec_instr (ins!pc) P h stk loc C M pc frs) = None" shows "P,\ \ \'\" (*<*) proof - from mC cf obtain ST LT where "h\": "P \ h \" and \: "\ C M ! pc = Some (ST,LT)" and stk: "P,h \ stk [:\] ST" and loc: "P,h \ loc [:\\<^sub>\] LT" and pc: "pc < size ins" and fs: "conf_fs P h \ M (size Ts) T frs" by (fastforce dest: sees_method_fun) from i \ wt obtain vT vT' oT ST'' ST' LT' where ST: "ST = vT # oT # ST''" and field: "P \ D sees F:vT' in D" and oT: "P \ oT \ Class D" and vT: "P \ vT \ vT'" and pc': "pc+1 < size ins" and \': "\ C M!(pc+1) = Some (ST',LT')" and ST': "P \ ST'' [\] ST'" and LT': "P \ LT [\\<^sub>\] LT'" by clarsimp from stk ST obtain v ref stk' where stk': "stk = v#ref#stk'" and v: "P,h \ v :\ vT" and ref: "P,h \ ref :\ oT" and ST'': "P,h \ stk' [:\] ST''" by auto from stk' i mC s' xc have "ref \ Null" by (auto simp add: split_beta) moreover from ref oT have "P,h \ ref :\ Class D" .. ultimately obtain a D' fs where a: "ref = Addr a" and h: "h a = Some (D', fs)" and D': "P \ D' \\<^sup>* D" by (blast dest: non_npD) from v vT have vT': "P,h \ v :\ vT'" .. from field D' have has_field: "P \ D' has F:vT' in D" by (blast intro: has_field_mono has_visible_field) let ?h' = "h(a\(D', fs((F, D)\v)))" and ?f' = "(stk',loc,C,M,pc+1)" from h have hext: "h \ ?h'" by (rule hext_upd_obj) from a h i mC s' stk' have "\' = (None, ?h', ?f'#frs)" by simp moreover from "h\" h have "P,h \ (D',fs)\" by (rule hconfD) with has_field vT' have "P,h \ (D',fs((F, D)\v))\" .. with "h\" h have "P \ ?h'\" by (rule hconf_upd_obj) moreover from ST'' ST' have "P,h \ stk' [:\] ST'" .. from this hext have "P,?h' \ stk' [:\] ST'" by (rule confs_hext) moreover from loc LT' have "P,h \ loc [:\\<^sub>\] LT'" .. from this hext have "P,?h' \ loc [:\\<^sub>\] LT'" by (rule confTs_hext) moreover from fs hext have "conf_fs P ?h' \ M (size Ts) T frs" by (rule conf_fs_hext) moreover note mC \' pc' ultimately show "P,\ \ \' \" by fastforce qed (*>*) (* FIXME: move *) lemma has_fields_b_fields: "P \ C has_fields FDTs \ fields P C = FDTs" (*<*) -by (unfold fields_def) - (blast intro: the_equality has_fields_fun) + apply (unfold fields_def) + apply (blast intro: the_equality has_fields_fun) + done (*>*) (* FIXME: move *) lemma oconf_blank [intro, simp]: "\is_class P C; wf_prog wt P\ \ P,h \ blank P C \" (*<*) by (fastforce simp add: blank_def has_fields_b_fields oconf_init_fields dest: wf_Fields_Ex) (*>*) lemma obj_ty_blank [iff]: "obj_ty (blank P C) = Class C" by (simp add: blank_def) lemma New_correct: fixes \' :: jvm_state assumes wf: "wf_prog wt P" assumes meth: "P \ C sees M:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" assumes ins: "ins!pc = New X" assumes wt: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" assumes exec: "Some \' = exec (P, None, h, (stk,loc,C,M,pc)#frs)" assumes conf: "P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\" assumes no_x: "fst (exec_instr (ins!pc) P h stk loc C M pc frs) = None" shows "P,\ \ \'\" (*<*) proof - from ins conf meth obtain ST LT where heap_ok: "P\ h\" and \_pc: "\ C M!pc = Some (ST,LT)" and frame: "conf_f P h (ST,LT) ins (stk,loc,C,M,pc)" and frames: "conf_fs P h \ M (size Ts) T frs" by (auto dest: sees_method_fun) from \_pc ins wt obtain ST' LT' where is_class_X: "is_class P X" and mxs: "size ST < mxs" and suc_pc: "pc+1 < size ins" and \_suc: "\ C M!(pc+1) = Some (ST', LT')" and less: "P \ (Class X # ST, LT) \\<^sub>i (ST', LT')" by auto from ins no_x obtain oref where new_Addr: "new_Addr h = Some oref" by auto hence h: "h oref = None" by (rule new_Addr_SomeD) with exec ins meth new_Addr have \': "\' = (None, h(oref \ blank P X), (Addr oref#stk,loc,C,M,pc+1)#frs)" (is "\' = (None, ?h', ?f # frs)") by simp moreover from wf h heap_ok is_class_X have h': "P \ ?h' \" by (auto intro: hconf_new) moreover from h frame less suc_pc wf have "conf_f P ?h' (ST', LT') ins ?f" - by (clarsimp simp add: fun_upd_apply conf_def blank_def split_beta) - (auto intro: confs_hext confTs_hext) + apply (clarsimp simp add: fun_upd_apply conf_def blank_def split_beta) + apply (auto intro: confs_hext confTs_hext) + done moreover from h have "h \ ?h'" by simp with frames have "conf_fs P ?h' \ M (size Ts) T frs" by (rule conf_fs_hext) ultimately show ?thesis using meth \_suc by fastforce qed (*>*) lemma Goto_correct: "\ wf_prog wt P; P \ C sees M:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C; ins ! pc = Goto branch; P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; Some \' = exec (P, None, h, (stk,loc,C,M,pc)#frs) ; P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\ \ \ P,\ \ \'\" -(*<*)by clarsimp (fastforce dest: sees_method_fun)(*>*) +(*<*) +apply clarsimp +apply (drule (1) sees_method_fun) +apply fastforce +done +(*>*) lemma IfFalse_correct: "\ wf_prog wt P; P \ C sees M:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C; ins ! pc = IfFalse branch; P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; Some \' = exec (P, None, h, (stk,loc,C,M,pc)#frs) ; P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\ \ \ P,\ \ \'\" -(*<*)by clarsimp (fastforce dest: sees_method_fun)(*>*) +(*<*) +apply clarsimp +apply (drule (1) sees_method_fun) +apply fastforce +done +(*>*) lemma CmpEq_correct: "\ wf_prog wt P; P \ C sees M:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C; ins ! pc = CmpEq; P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; Some \' = exec (P, None, h, (stk,loc,C,M,pc)#frs) ; P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\ \ \ P,\ \ \'\" -(*<*)by clarsimp (fastforce dest: sees_method_fun)(*>*) +(*<*) +apply clarsimp +apply (drule (1) sees_method_fun) +apply fastforce +done +(*>*) lemma Pop_correct: "\ wf_prog wt P; P \ C sees M:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C; ins ! pc = Pop; P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; Some \' = exec (P, None, h, (stk,loc,C,M,pc)#frs) ; P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\ \ \ P,\ \ \'\" -(*<*)by clarsimp (fastforce dest: sees_method_fun)(*>*) +(*<*) +apply clarsimp +apply (drule (1) sees_method_fun) +apply fastforce +done +(*>*) lemma IAdd_correct: "\ wf_prog wt P; P \ C sees M:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C; ins ! pc = IAdd; P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M; Some \' = exec (P, None, h, (stk,loc,C,M,pc)#frs) ; P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\ \ \ P,\ \ \'\" -(*<*)by clarsimp (fastforce dest: sees_method_fun)(*>*) +(*<*) +apply (clarsimp simp add: conf_def) +apply (drule (1) sees_method_fun) +apply fastforce +done +(*>*) lemma Throw_correct: "\ wf_prog wt P; P \ C sees M:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C; ins ! pc = Throw; Some \' = exec (P, None, h, (stk,loc,C,M,pc)#frs) ; P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\; fst (exec_instr (ins!pc) P h stk loc C M pc frs) = None \ \ P,\ \ \'\" by simp text \ The next theorem collects the results of the sections above, i.e.~exception handling and the execution step for each instruction. It states type safety for single step execution: in welltyped programs, a conforming state is transformed into another conforming state when one instruction is executed. \ theorem instr_correct: -fixes \' :: jvm_state -assumes wtp: "wf_jvm_prog\<^bsub>\\<^esub> P" - and meth: "P \ C sees M:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C" - and exec: "Some \' = exec (P, None, h, (stk,loc,C,M,pc)#frs)" - and conf: "P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\" -shows "P,\ \ \'\" +"\ wf_jvm_prog\<^bsub>\\<^esub> P; + P \ C sees M:Ts\T=(mxs,mxl\<^sub>0,ins,xt) in C; + Some \' = exec (P, None, h, (stk,loc,C,M,pc)#frs); + P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\ \ +\ P,\ \ \'\" (*<*) -proof - - from assms have pc: "pc < length ins" by(auto dest: sees_method_fun) - with wt_jvm_prog_impl_wt_instr[OF wtp meth] have wt: "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M" - by simp - - from conf obtain ST LT where \: "\ C M ! pc = Some(ST,LT)" by clarsimp - - show ?thesis - proof(cases "fst (exec_instr (ins!pc) P h stk loc C M pc frs)") - case Some show ?thesis by(rule xcpt_correct[OF wtp meth wt Some exec conf]) - next - case None - from wt_jvm_progD[OF wtp] obtain wf_md where wf: "wf_prog wf_md P" by clarify - - from exec conf None obtain - exec': "Some \' = exec (P, None, h, (stk,loc,C,M,pc)#frs)" - and conf': "P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\" - and None': "fst (exec_instr (ins!pc) P h stk loc C M pc frs) = None" by simp - - show ?thesis - proof(cases "ins!pc") - case Load from Load_correct[OF wf meth this wt exec conf] show ?thesis by simp - next - case Store from Store_correct[OF wf meth this wt exec conf] show ?thesis by simp - next - case Push from Push_correct[OF wf meth this wt exec conf] show ?thesis by simp - next - case (New C) from New_correct[OF wf meth this wt exec conf None'] show ?thesis by simp - next - case Getfield from Getfield_correct[OF wf meth this wt exec conf None] - show ?thesis by simp - next - case Putfield from Putfield_correct[OF wf meth this wt exec conf None] - show ?thesis by simp - next - case Checkcast from Checkcast_correct[OF wtp meth this wt exec conf None] - show ?thesis by simp - next - case Invoke with Invoke_correct[OF wtp meth this wt exec conf None] show ?thesis by simp - next - case Return with Return_correct[OF wtp meth this wt exec conf] show ?thesis by simp - next - case Pop with Pop_correct[OF wf meth this wt exec conf] show ?thesis by simp - next - case IAdd with IAdd_correct[OF wf meth this wt exec conf] show ?thesis by simp - next - case Goto with Goto_correct[OF wf meth this wt exec conf] show ?thesis by simp - next - case CmpEq with CmpEq_correct[OF wf meth this wt exec conf] show ?thesis by simp - next - case IfFalse with IfFalse_correct[OF wf meth this wt exec conf] show ?thesis by simp - next - case Throw with Throw_correct[OF wf meth this exec conf None] show ?thesis by simp - qed - qed -qed +apply (subgoal_tac "P,T,mxs,size ins,xt \ ins!pc,pc :: \ C M") + prefer 2 + apply (erule wt_jvm_prog_impl_wt_instr, assumption) + apply clarsimp + apply (drule (1) sees_method_fun) + apply simp +apply (cases "fst (exec_instr (ins!pc) P h stk loc C M pc frs)") + prefer 2 + apply (erule xcpt_correct, assumption+) +apply (frule wt_jvm_progD, erule exE) +apply (cases "ins!pc") +apply (rule Load_correct, assumption+) +apply (rule Store_correct, assumption+) +apply (rule Push_correct, assumption+) +apply (rule New_correct, assumption+) +apply (rule Getfield_correct, assumption+) +apply (rule Putfield_correct, assumption+) +apply (rule Checkcast_correct, assumption+) +apply (rule Invoke_correct, assumption+) +apply (rule Return_correct, assumption+) +apply (rule Pop_correct, assumption+) +apply (rule IAdd_correct, assumption+) +apply (rule Goto_correct, assumption+) +apply (rule CmpEq_correct, assumption+) +apply (rule IfFalse_correct, assumption+) +apply (rule Throw_correct, assumption+) +done (*>*) subsection \Main\ lemma correct_state_impl_Some_method: "P,\ \ (None, h, (stk,loc,C,M,pc)#frs)\ \ \m Ts T. P \ C sees M:Ts\T = m in C" by fastforce lemma BV_correct_1 [rule_format]: "\\. \ wf_jvm_prog\<^bsub>\\<^esub> P; P,\ \ \\\ \ P \ \ -jvm\\<^sub>1 \' \ P,\ \ \'\" (*<*) -proof - - fix \ assume wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and cf: "P,\ \ \\" - obtain xp h frs where \[simp]: "\ = (xp, h, frs)" by(cases \) simp - have "exec (P, xp, h, frs) = \\'\ \ P,\ |- \' [ok]" - proof(cases xp) - case None - with wf cf show ?thesis - proof(cases frs) - case (Cons fr frs') - obtain stk loc C M pc where [simp]: "fr = (stk,loc,C,M,pc)" by(cases fr) simp - then have cf': "P,\ |- (None, h, (stk,loc,C,M,pc) # frs') [ok]" - using Cons None cf by simp - then obtain mxs mxl\<^sub>0 ins xt Ts T - where mC: "P \ C sees M : Ts\T = (mxs, mxl\<^sub>0, ins, xt) in C" - using correct_state_impl_Some_method[OF cf'] by clarify - show ?thesis using Cons None instr_correct[OF wf mC _ cf'] by simp - qed simp - next - case (Some a) - then show ?thesis using wf cf by (case_tac frs) simp_all - qed - then show "P \ \ -jvm\\<^sub>1 \' \ P,\ \ \'\" by(simp add: exec_1_iff) -qed +apply (simp only: split_tupled_all exec_1_iff) +apply (rename_tac xp h frs) +apply (case_tac xp) + apply (case_tac frs) + apply simp + apply (simp only: split_tupled_all) + apply hypsubst + apply (frule correct_state_impl_Some_method) + apply clarify + apply (rule instr_correct) + apply assumption+ + apply (rule sym) + apply assumption+ +apply (case_tac frs) +apply simp_all +done (*>*) theorem progress: "\ xp=None; frs\[] \ \ \\'. P \ (xp,h,frs) -jvm\\<^sub>1 \'" by (clarsimp simp add: exec_1_iff neq_Nil_conv split_beta simp del: split_paired_Ex) lemma progress_conform: "\wf_jvm_prog\<^bsub>\\<^esub> P; P,\ \ (xp,h,frs)\; xp=None; frs\[]\ \ \\'. P \ (xp,h,frs) -jvm\\<^sub>1 \' \ P,\ \ \'\" -(*<*)by (drule (1) progress) (fast intro: BV_correct_1)(*>*) +(*<*) +apply (drule progress) +apply assumption +apply (fast intro: BV_correct_1) +done +(*>*) theorem BV_correct [rule_format]: "\ wf_jvm_prog\<^bsub>\\<^esub> P; P \ \ -jvm\ \' \ \ P,\ \ \\ \ P,\ \ \'\" (*<*) -proof - - assume wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and "P \ \ -jvm\ \'" - then have "(\, \') \ (exec_1 P)\<^sup>*" by (simp only: exec_all_def1) - then show ?thesis proof(induct rule: rtrancl_induct) - case (step y z) - then show ?case by clarify (erule (1) BV_correct_1[OF wf]) - qed simp -qed +apply (simp only: exec_all_def1) +apply (erule rtrancl_induct) + apply simp +apply clarify +apply (erule (2) BV_correct_1) +done (*>*) lemma hconf_start: assumes wf: "wf_prog wf_mb P" shows "P \ (start_heap P) \" (*<*) -proof - - { fix a obj assume assm: "start_heap P a = \obj\" - have "P,start_heap P \ obj \" using wf assm[THEN sym] - by (unfold start_heap_def) - (auto simp: fun_upd_apply is_class_xcpt split: if_split_asm) - } - then show ?thesis using preallocated_start[of P] - by (unfold hconf_def) simp -qed + apply (unfold hconf_def) + apply (simp add: preallocated_start) + apply (clarify) + apply (drule sym) + apply (unfold start_heap_def) + apply (insert wf) + apply (auto simp add: fun_upd_apply is_class_xcpt split: if_split_asm) + done (*>*) lemma BV_correct_initial: -assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" - and mC: "P \ C sees M:[]\T = m in C" -shows "P,\ \ start_state P C M \" + shows "\ wf_jvm_prog\<^bsub>\\<^esub> P; P \ C sees M:[]\T = m in C \ + \ P,\ \ start_state P C M \" (*<*) -proof - - obtain mxs mxl\<^sub>0 ins xt where - mC': "P \ C sees M:[]\T = (mxs,mxl\<^sub>0,ins,xt) in C" - using mC by (cases m) simp - then have method: "(C,[],T,mxs,mxl\<^sub>0,ins,xt) = method P C M" by simp - let ?h = "start_heap P" and ?pc = 0 - let ?f = "([], Null # replicate mxl\<^sub>0 undefined, C, M, ?pc)" and ?fs = "[]" - let ?frs = "?f#?fs" - have "is_class P Object" using wf by(simp add: wf_jvm_prog_phi_def) - then obtain Mm where Mm: "P \ Object sees_methods Mm" - by(fastforce simp: is_class_def dest: sees_methods_Object) - have "P\ ?h\" using wf by (simp add: wf_jvm_prog_phi_def hconf_start) - moreover have "\b Ts T mxs mxl\<^sub>0 is xt \. - (P \ C sees M:Ts\T = (mxs,mxl\<^sub>0,is,xt) in C) - \ \ C M ! ?pc = Some \ - \ conf_f P ?h \ is ?f \ conf_fs P ?h \ M (size Ts) T ?fs" - using wt_jvm_prog_impl_wt_start[OF wf mC'] mC' - by (unfold conf_f_def wt_start_def) fastforce - ultimately show ?thesis using method - by (fastforce simp del: defs1 simp: start_state_def correct_state_def) -qed + apply (cases m) + apply (unfold start_state_def) + apply (unfold correct_state_def) + apply (simp del: defs1) + apply (rule conjI) + apply (simp add: wf_jvm_prog_phi_def hconf_start) + apply (drule wt_jvm_prog_impl_wt_start, assumption+) + apply (unfold conf_f_def wt_start_def) + apply fastforce + done declare [[simproc add: list_to_set_comprehension]] (*>*) theorem typesafe: assumes welltyped: "wf_jvm_prog\<^bsub>\\<^esub> P" assumes main_method: "P \ C sees M:[]\T = m in C" shows "P \ start_state P C M -jvm\ \ \ P,\ \ \ \" (*<*) proof - from welltyped main_method have "P,\ \ start_state P C M \" by (rule BV_correct_initial) moreover assume "P \ start_state P C M -jvm\ \" ultimately show "P,\ \ \ \" using welltyped by - (rule BV_correct) qed (*>*) end diff --git a/thys/Jinja/BV/JVM_SemiType.thy b/thys/Jinja/BV/JVM_SemiType.thy --- a/thys/Jinja/BV/JVM_SemiType.thy +++ b/thys/Jinja/BV/JVM_SemiType.thy @@ -1,256 +1,375 @@ (* Title: HOL/MicroJava/BV/JVM.thy Author: Gerwin Klein Copyright 2000 TUM *) section \The JVM Type System as Semilattice\ theory JVM_SemiType imports SemiType begin type_synonym ty\<^sub>l = "ty err list" type_synonym ty\<^sub>s = "ty list" type_synonym ty\<^sub>i = "ty\<^sub>s \ ty\<^sub>l" type_synonym ty\<^sub>i' = "ty\<^sub>i option" type_synonym ty\<^sub>m = "ty\<^sub>i' list" type_synonym ty\<^sub>P = "mname \ cname \ ty\<^sub>m" definition stk_esl :: "'c prog \ nat \ ty\<^sub>s esl" where "stk_esl P mxs \ upto_esl mxs (SemiType.esl P)" definition loc_sl :: "'c prog \ nat \ ty\<^sub>l sl" where "loc_sl P mxl \ Listn.sl mxl (Err.sl (SemiType.esl P))" definition sl :: "'c prog \ nat \ nat \ ty\<^sub>i' err sl" where "sl P mxs mxl \ Err.sl(Opt.esl(Product.esl (stk_esl P mxs) (Err.esl(loc_sl P mxl))))" definition states :: "'c prog \ nat \ nat \ ty\<^sub>i' err set" where "states P mxs mxl \ fst(sl P mxs mxl)" definition le :: "'c prog \ nat \ nat \ ty\<^sub>i' err ord" where "le P mxs mxl \ fst(snd(sl P mxs mxl))" definition sup :: "'c prog \ nat \ nat \ ty\<^sub>i' err binop" where "sup P mxs mxl \ snd(snd(sl P mxs mxl))" definition sup_ty_opt :: "['c prog,ty err,ty err] \ bool" ("_ \ _ \\<^sub>\ _" [71,71,71] 70) where "sup_ty_opt P \ Err.le (subtype P)" definition sup_state :: "['c prog,ty\<^sub>i,ty\<^sub>i] \ bool" ("_ \ _ \\<^sub>i _" [71,71,71] 70) where "sup_state P \ Product.le (Listn.le (subtype P)) (Listn.le (sup_ty_opt P))" definition sup_state_opt :: "['c prog,ty\<^sub>i',ty\<^sub>i'] \ bool" ("_ \ _ \'' _" [71,71,71] 70) where "sup_state_opt P \ Opt.le (sup_state P)" abbreviation sup_loc :: "['c prog,ty\<^sub>l,ty\<^sub>l] \ bool" ("_ \ _ [\\<^sub>\] _" [71,71,71] 70) where "P \ LT [\\<^sub>\] LT' \ list_all2 (sup_ty_opt P) LT LT'" notation (ASCII) sup_ty_opt ("_ |- _ <=T _" [71,71,71] 70) and sup_state ("_ |- _ <=i _" [71,71,71] 70) and sup_state_opt ("_ |- _ <=' _" [71,71,71] 70) and sup_loc ("_ |- _ [<=T] _" [71,71,71] 70) subsection "Unfolding" lemma JVM_states_unfold: "states P mxs mxl \ err(opt((Union {list n (types P) |n. n <= mxs}) \ list mxl (err(types P))))" (*<*) apply (unfold states_def sl_def Opt.esl_def Err.sl_def stk_esl_def loc_sl_def Product.esl_def Listn.sl_def upto_esl_def SemiType.esl_def Err.esl_def) apply simp done (*>*) lemma JVM_le_unfold: "le P m n \ Err.le(Opt.le(Product.le(Listn.le(subtype P))(Listn.le(Err.le(subtype P)))))" (*<*) apply (unfold le_def sl_def Opt.esl_def Err.sl_def stk_esl_def loc_sl_def Product.esl_def Listn.sl_def upto_esl_def SemiType.esl_def Err.esl_def) apply simp done (*>*) lemma sl_def2: "JVM_SemiType.sl P mxs mxl \ (states P mxs mxl, JVM_SemiType.le P mxs mxl, JVM_SemiType.sup P mxs mxl)" (*<*) by (unfold JVM_SemiType.sup_def states_def JVM_SemiType.le_def) simp (*>*) lemma JVM_le_conv: "le P m n (OK t1) (OK t2) = P \ t1 \' t2" (*<*) by (simp add: JVM_le_unfold Err.le_def lesub_def sup_state_opt_def sup_state_def sup_ty_opt_def) (*>*) lemma JVM_le_Err_conv: "le P m n = Err.le (sup_state_opt P)" (*<*) by (unfold sup_state_opt_def sup_state_def sup_ty_opt_def JVM_le_unfold) simp (*>*) lemma err_le_unfold [iff]: "Err.le r (OK a) (OK b) = r a b" (*<*) by (simp add: Err.le_def lesub_def) (*>*) subsection \Semilattice\ +lemma order_sup_state_opt' [intro, simp]: + "wf_prog wf_mb P \ + order (sup_state_opt P) (opt ((\ {list n (types P) |n. n \ mxs} ) \ list (Suc (length Ts + mxl\<^sub>0)) (err (types P))))" +(*<*) + apply (unfold sup_state_opt_def sup_state_def sup_ty_opt_def ) + apply (blast intro:order_le_prodI) \\ use Listn.thy.order_listI2 \ + done +(*<*) -lemma order_sup_state_opt [intro, simp]: - "wf_prog wf_mb P \ order (sup_state_opt P)" -(*<*) by (unfold sup_state_opt_def sup_state_def sup_ty_opt_def) blast (*>*) lemma semilat_JVM [intro?]: "wf_prog wf_mb P \ semilat (JVM_SemiType.sl P mxs mxl)" (*<*) apply (unfold JVM_SemiType.sl_def stk_esl_def loc_sl_def) apply (blast intro: err_semilat_Product_esl err_semilat_upto_esl Listn_sl err_semilat_JType_esl) done (*>*) -lemma acc_JVM [intro]: - "wf_prog wf_mb P \ acc (JVM_SemiType.le P mxs mxl)" -(*<*) by (unfold JVM_le_unfold) blast (*>*) - - subsection \Widening with \\\\ -lemma subtype_refl[iff]: "subtype P t t" (*<*) by (simp add: fun_of_def) (*>*) +lemma subtype_refl: "subtype P t t" (*<*) by (simp add: fun_of_def) (*>*) lemma sup_ty_opt_refl [iff]: "P \ T \\<^sub>\ T" (*<*) apply (unfold sup_ty_opt_def) apply (fold lesub_def) apply (rule le_err_refl) apply (simp add: lesub_def) done (*>*) lemma Err_any_conv [iff]: "P \ Err \\<^sub>\ T = (T = Err)" (*<*) by (unfold sup_ty_opt_def) (rule Err_le_conv [simplified lesub_def]) (*>*) lemma any_Err [iff]: "P \ T \\<^sub>\ Err" (*<*) by (unfold sup_ty_opt_def) (rule le_Err [simplified lesub_def]) (*>*) lemma OK_OK_conv [iff]: "P \ OK T \\<^sub>\ OK T' = P \ T \ T'" (*<*) by (simp add: sup_ty_opt_def fun_of_def) (*>*) lemma any_OK_conv [iff]: "P \ X \\<^sub>\ OK T' = (\T. X = OK T \ P \ T \ T')" (*<*) apply (unfold sup_ty_opt_def) apply (rule le_OK_conv [simplified lesub_def]) done (*>*) lemma OK_any_conv: "P \ OK T \\<^sub>\ X = (X = Err \ (\T'. X = OK T' \ P \ T \ T'))" (*<*) apply (unfold sup_ty_opt_def) apply (rule OK_le_conv [simplified lesub_def]) done (*>*) lemma sup_ty_opt_trans [intro?, trans]: "\P \ a \\<^sub>\ b; P \ b \\<^sub>\ c\ \ P \ a \\<^sub>\ c" (*<*) by (auto intro: widen_trans simp add: sup_ty_opt_def Err.le_def lesub_def fun_of_def split: err.splits) (*>*) subsection "Stack and Registers" lemma stk_convert: "P \ ST [\] ST' = Listn.le (subtype P) ST ST'" (*<*) by (simp add: Listn.le_def lesub_def) (*>*) lemma sup_loc_refl [iff]: "P \ LT [\\<^sub>\] LT" (*<*) by (rule list_all2_refl) simp (*>*) lemmas sup_loc_Cons1 [iff] = list_all2_Cons1 [of "sup_ty_opt P"] for P lemma sup_loc_def: "P \ LT [\\<^sub>\] LT' \ Listn.le (sup_ty_opt P) LT LT'" (*<*) by (simp add: Listn.le_def lesub_def) (*>*) lemma sup_loc_widens_conv [iff]: "P \ map OK Ts [\\<^sub>\] map OK Ts' = P \ Ts [\] Ts'" (*<*) by (simp add: list_all2_map1 list_all2_map2) (*>*) lemma sup_loc_trans [intro?, trans]: "\P \ a [\\<^sub>\] b; P \ b [\\<^sub>\] c\ \ P \ a [\\<^sub>\] c" (*<*) by (rule list_all2_trans, rule sup_ty_opt_trans) (*>*) subsection "State Type" lemma sup_state_conv [iff]: "P \ (ST,LT) \\<^sub>i (ST',LT') = (P \ ST [\] ST' \ P \ LT [\\<^sub>\] LT')" (*<*) by (auto simp add: sup_state_def stk_convert lesub_def Product.le_def sup_loc_def) (*>*) lemma sup_state_conv2: "P \ s1 \\<^sub>i s2 = (P \ fst s1 [\] fst s2 \ P \ snd s1 [\\<^sub>\] snd s2)" (*<*) by (cases s1, cases s2) simp (*>*) lemma sup_state_refl [iff]: "P \ s \\<^sub>i s" (*<*) by (auto simp add: sup_state_conv2) (*>*) lemma sup_state_trans [intro?, trans]: "\P \ a \\<^sub>i b; P \ b \\<^sub>i c\ \ P \ a \\<^sub>i c" (*<*) by (auto intro: sup_loc_trans widens_trans simp add: sup_state_conv2) (*>*) lemma sup_state_opt_None_any [iff]: "P \ None \' s" (*<*) by (simp add: sup_state_opt_def Opt.le_def) (*>*) lemma sup_state_opt_any_None [iff]: "P \ s \' None = (s = None)" (*<*) by (simp add: sup_state_opt_def Opt.le_def) (*>*) lemma sup_state_opt_Some_Some [iff]: "P \ Some a \' Some b = P \ a \\<^sub>i b" (*<*) by (simp add: sup_state_opt_def Opt.le_def lesub_def) (*>*) lemma sup_state_opt_any_Some: "P \ (Some s) \' X = (\s'. X = Some s' \ P \ s \\<^sub>i s')" (*<*) by (simp add: sup_state_opt_def Opt.le_def lesub_def) (*>*) lemma sup_state_opt_refl [iff]: "P \ s \' s" (*<*) by (simp add: sup_state_opt_def Opt.le_def lesub_def) (*>*) lemma sup_state_opt_trans [intro?, trans]: "\P \ a \' b; P \ b \' c\ \ P \ a \' c" (*<*) apply (unfold sup_state_opt_def Opt.le_def lesub_def) apply (simp del: split_paired_All) apply (rule sup_state_trans, assumption+) done (*>*) +lemma sup_state_opt_err : "(Err.le (sup_state_opt P)) s s" + apply (unfold JVM_le_unfold Product.le_def Opt.le_def Err.le_def lesssub_def lesub_def Listn.le_def) + apply (auto split: err.splits) + done + +lemma Cons_less_Conss1 [simp]: + "x#xs [\\<^bsub>subtype P\<^esub>] y#ys = (x \\<^bsub>subtype P\<^esub> y \ xs [\\<^bsub>subtype P\<^esub>] ys \ x = y \ xs [\\<^bsub>subtype P\<^esub>] ys)" + apply (unfold lesssub_def ) + apply auto + apply (simp add:lesssub_def lesub_def) \\widen_refl, subtype_refl \ + done + +lemma Cons_less_Conss2 [simp]: + "x#xs [\\<^bsub>Err.le (subtype P)\<^esub>] y#ys = (x \\<^bsub>Err.le (subtype P)\<^esub> y \ xs [\\<^bsub>Err.le (subtype P)\<^esub>] ys \ x = y \ xs [\\<^bsub>Err.le (subtype P)\<^esub>] ys)" + apply (unfold lesssub_def ) + apply auto + apply (simp add:lesssub_def lesub_def Err.le_def split: err.splits) + done + +lemma acc_le_listI1 [intro!]: + " acc (subtype P) \ acc (Listn.le (subtype P))" + (*<*) +apply (unfold acc_def) +apply (subgoal_tac + "wf(UN n. {(ys,xs). size xs = n \ size ys = n \ xs <_(Listn.le (subtype P)) ys})") + apply (erule wf_subset) + + apply (blast intro: lesssub_lengthD) +apply (rule wf_UN) + prefer 2 + apply (rename_tac m n) + apply (case_tac "m=n") + apply simp + apply (fast intro!: equals0I dest: not_sym) +apply (rename_tac n) +apply (induct_tac n) + apply (simp add: lesssub_def cong: conj_cong) +apply (rename_tac k) +apply (simp add: wf_eq_minimal) +apply (simp (no_asm) add: length_Suc_conv cong: conj_cong) +apply clarify +apply (rename_tac M m) +apply (case_tac "\x xs. size xs = k \ x#xs \ M") + prefer 2 + apply (erule thin_rl) + apply (erule thin_rl) + apply blast +apply (erule_tac x = "{a. \xs. size xs = k \ a#xs:M}" in allE) +apply (erule impE) + apply blast +apply (thin_tac "\x xs. P x xs" for P) +apply clarify +apply (rename_tac maxA xs) +apply (erule_tac x = "{ys. size ys = size xs \ maxA#ys \ M}" in allE) +apply (erule impE) + apply blast +apply clarify +apply (thin_tac "m \ M") +apply (thin_tac "maxA#xs \ M") +apply (rule bexI) + prefer 2 + apply assumption +apply clarify +apply simp +apply blast + done + +lemma acc_le_listI2 [intro!]: + " acc (Err.le (subtype P)) \ acc (Listn.le (Err.le (subtype P)))" + (*<*) +apply (unfold acc_def) +apply (subgoal_tac + "wf(UN n. {(ys,xs). size xs = n \ size ys = n \ xs <_(Listn.le (Err.le (subtype P))) ys})") + apply (erule wf_subset) + + apply (blast intro: lesssub_lengthD) +apply (rule wf_UN) + prefer 2 + apply (rename_tac m n) + apply (case_tac "m=n") + apply simp + apply (fast intro!: equals0I dest: not_sym) +apply (rename_tac n) +apply (induct_tac n) + apply (simp add: lesssub_def cong: conj_cong) +apply (rename_tac k) +apply (simp add: wf_eq_minimal) +apply (simp (no_asm) add: length_Suc_conv cong: conj_cong) +apply clarify +apply (rename_tac M m) +apply (case_tac "\x xs. size xs = k \ x#xs \ M") + prefer 2 + apply (erule thin_rl) + apply (erule thin_rl) + apply blast +apply (erule_tac x = "{a. \xs. size xs = k \ a#xs:M}" in allE) +apply (erule impE) + apply blast +apply (thin_tac "\x xs. P x xs" for P) +apply clarify +apply (rename_tac maxA xs) +apply (erule_tac x = "{ys. size ys = size xs \ maxA#ys \ M}" in allE) +apply (erule impE) + apply blast +apply clarify +apply (thin_tac "m \ M") +apply (thin_tac "maxA#xs \ M") +apply (rule bexI) + prefer 2 + apply assumption +apply clarify +apply simp +apply blast + done + +lemma acc_JVM [intro]: + "wf_prog wf_mb P \ acc (JVM_SemiType.le P mxs mxl)" +(*<*) by (unfold JVM_le_unfold) blast (*>*) \\ use acc_listI1, acc_listI2 \ + end diff --git a/thys/Jinja/BV/LBVJVM.thy b/thys/Jinja/BV/LBVJVM.thy --- a/thys/Jinja/BV/LBVJVM.thy +++ b/thys/Jinja/BV/LBVJVM.thy @@ -1,242 +1,223 @@ (* Title: HOL/MicroJava/BV/JVM.thy Author: Tobias Nipkow, Gerwin Klein Copyright 2000 TUM *) section \LBV for the JVM \label{sec:JVM}\ theory LBVJVM imports "../DFA/Abstract_BV" TF_JVM begin type_synonym prog_cert = "cname \ mname \ ty\<^sub>i' err list" definition check_cert :: "jvm_prog \ nat \ nat \ nat \ ty\<^sub>i' err list \ bool" where "check_cert P mxs mxl n cert \ check_types P mxs mxl cert \ size cert = n+1 \ (\i Err) \ cert!n = OK None" definition lbvjvm :: "jvm_prog \ nat \ nat \ ty \ ex_table \ ty\<^sub>i' err list \ instr list \ ty\<^sub>i' err \ ty\<^sub>i' err" where "lbvjvm P mxs maxr T\<^sub>r et cert bs \ wtl_inst_list bs cert (JVM_SemiType.sup P mxs maxr) (JVM_SemiType.le P mxs maxr) Err (OK None) (exec P mxs T\<^sub>r et bs) 0" definition wt_lbv :: "jvm_prog \ cname \ ty list \ ty \ nat \ nat \ ex_table \ ty\<^sub>i' err list \ instr list \ bool" where "wt_lbv P C Ts T\<^sub>r mxs mxl\<^sub>0 et cert ins \ check_cert P mxs (1+size Ts+mxl\<^sub>0) (size ins) cert \ 0 < size ins \ (let start = Some ([],(OK (Class C))#((map OK Ts))@(replicate mxl\<^sub>0 Err)); result = lbvjvm P mxs (1+size Ts+mxl\<^sub>0) T\<^sub>r et cert ins (OK start) in result \ Err)" definition wt_jvm_prog_lbv :: "jvm_prog \ prog_cert \ bool" where "wt_jvm_prog_lbv P cert \ wf_prog (\P C (mn,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,b,et)). wt_lbv P C Ts T\<^sub>r mxs mxl\<^sub>0 et (cert C mn) b) P" definition mk_cert :: "jvm_prog \ nat \ ty \ ex_table \ instr list \ ty\<^sub>m \ ty\<^sub>i' err list" where "mk_cert P mxs T\<^sub>r et bs phi \ make_cert (exec P mxs T\<^sub>r et bs) (map OK phi) (OK None)" definition prg_cert :: "jvm_prog \ ty\<^sub>P \ prog_cert" where "prg_cert P phi C mn \ let (C,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,ins,et)) = method P C mn in mk_cert P mxs T\<^sub>r et ins (phi C mn)" lemma check_certD [intro?]: "check_cert P mxs mxl n cert \ cert_ok cert n Err (OK None) (states P mxs mxl)" by (unfold cert_ok_def check_cert_def check_types_def) auto lemma (in start_context) wt_lbv_wt_step: assumes lbv: "wt_lbv P C Ts T\<^sub>r mxs mxl\<^sub>0 xt cert is" shows "\\s \ list (size is) A. wt_step r Err step \s \ OK first \\<^sub>r \s!0" (*<*) proof - from wf have "semilat (JVM_SemiType.sl P mxs mxl)" .. hence "semilat (A, r, f)" by (simp add: sl_def2) moreover have "top r Err" by (simp add: JVM_le_Err_conv) moreover have "Err \ A" by (simp add: JVM_states_unfold) moreover have "bottom r (OK None)" by (simp add: JVM_le_Err_conv bottom_def lesub_def Err.le_def split: err.split) moreover have "OK None \ A" by (simp add: JVM_states_unfold) moreover note bounded_step moreover from lbv have "cert_ok cert (size is) Err (OK None) A" by (unfold wt_lbv_def) (auto dest: check_certD) moreover note exec_pres_type moreover from lbv have "wtl_inst_list is cert f r Err (OK None) step 0 (OK first) \ Err" by (simp add: wt_lbv_def lbvjvm_def step_def_exec [symmetric]) moreover note first_in_A moreover from lbv have "0 < size is" by (simp add: wt_lbv_def) ultimately show ?thesis by (rule lbvs.wtl_sound_strong [OF lbvs.intro, OF lbv.intro lbvs_axioms.intro, OF Semilat.intro lbv_axioms.intro]) qed (*>*) lemma (in start_context) wt_lbv_wt_method: assumes lbv: "wt_lbv P C Ts T\<^sub>r mxs mxl\<^sub>0 xt cert is" shows "\\s. wt_method P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt \s" (*<*) proof - from lbv have l: "is \ []" by (simp add: wt_lbv_def) moreover from wf lbv C Ts obtain \s where list: "\s \ list (size is) A" and step: "wt_step r Err step \s" and start: "OK first \\<^sub>r \s!0" by (blast dest: wt_lbv_wt_step) from list have [simp]: "size \s = size is" by simp have "size (map ok_val \s) = size is" by simp moreover from l have 0: "0 < size \s" by simp with step obtain \s0 where "\s!0 = OK \s0" by (unfold wt_step_def) blast with start 0 have "wt_start P C Ts mxl\<^sub>0 (map ok_val \s)" by (simp add: wt_start_def JVM_le_Err_conv lesub_def Err.le_def) moreover { from list have "check_types P mxs mxl \s" by (simp add: check_types_def) also from step have "\x \ set \s. x \ Err" by (auto simp add: all_set_conv_all_nth wt_step_def) hence [symmetric]: "map OK (map ok_val \s) = \s" by (auto intro!: map_idI) finally have "check_types P mxs mxl (map OK (map ok_val \s))" . } moreover { note bounded_step moreover from list have "set \s \ A" by simp moreover from step have "wt_err_step (sup_state_opt P) step \s" by (simp add: wt_err_step_def JVM_le_Err_conv) ultimately have "wt_app_eff (sup_state_opt P) app eff (map ok_val \s)" by (auto intro: wt_err_imp_wt_app_eff simp add: exec_def states_def) } ultimately have "wt_method P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt (map ok_val \s)" by (simp add: wt_method_def2 check_types_def del: map_map) thus ?thesis .. qed (*>*) lemma (in start_context) wt_method_wt_lbv: assumes wt: "wt_method P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt \s" defines [simp]: "cert \ mk_cert P mxs T\<^sub>r xt is \s" shows "wt_lbv P C Ts T\<^sub>r mxs mxl\<^sub>0 xt cert is" (*<*) proof - let ?\s = "map OK \s" let ?cert = "make_cert step ?\s (OK None)" from wt obtain 0: "0 < size is" and size: "size is = size ?\s" and ck_types: "check_types P mxs mxl ?\s" and wt_start: "wt_start P C Ts mxl\<^sub>0 \s" and app_eff: "wt_app_eff (sup_state_opt P) app eff \s" by (force simp add: wt_method_def2 check_types_def) from wf have "semilat (JVM_SemiType.sl P mxs mxl)" .. hence "semilat (A, r, f)" by (simp add: sl_def2) moreover have "top r Err" by (simp add: JVM_le_Err_conv) moreover have "Err \ A" by (simp add: JVM_states_unfold) moreover have "bottom r (OK None)" by (simp add: JVM_le_Err_conv bottom_def lesub_def Err.le_def split: err.split) moreover have "OK None \ A" by (simp add: JVM_states_unfold) moreover from wf have "mono r step (size is) A" by (rule step_mono) hence "mono r step (size ?\s) A" by (simp add: size) moreover from exec_pres_type have "pres_type step (size ?\s) A" by (simp add: size) moreover from ck_types have \s_in_A: "set ?\s \ A" by (simp add: check_types_def) hence "\pc. pc < size ?\s \ ?\s!pc \ A \ ?\s!pc \ Err" by auto moreover from bounded_step have "bounded step (size ?\s)" by (simp add: size) moreover have "OK None \ Err" by simp moreover from bounded_step size \s_in_A app_eff have "wt_err_step (sup_state_opt P) step ?\s" by (auto intro: wt_app_eff_imp_wt_err simp add: exec_def states_def) hence "wt_step r Err step ?\s" by (simp add: wt_err_step_def JVM_le_Err_conv) moreover from 0 size have "0 < size \s" by auto hence "?\s!0 = OK (\s!0)" by simp with wt_start have "OK first \\<^sub>r ?\s!0" by (clarsimp simp add: wt_start_def lesub_def Err.le_def JVM_le_Err_conv) moreover note first_in_A moreover have "OK first \ Err" by simp moreover note size ultimately have "wtl_inst_list is ?cert f r Err (OK None) step 0 (OK first) \ Err" by (rule lbvc.wtl_complete [OF lbvc.intro, OF lbv.intro lbvc_axioms.intro, OF Semilat.intro lbv_axioms.intro]) moreover from 0 size have "\s \ []" by auto moreover from ck_types have "check_types P mxs mxl ?cert" - by (fastforce simp: make_cert_def check_types_def JVM_states_unfold - dest!: nth_mem) + apply (auto simp add: make_cert_def check_types_def JVM_states_unfold) + apply (subst Ok_in_err [symmetric]) + apply (drule nth_mem) + apply auto + done moreover note 0 size ultimately show ?thesis by (simp add: wt_lbv_def lbvjvm_def mk_cert_def step_def_exec [symmetric] check_cert_def make_cert_def nth_append) qed (*>*) theorem jvm_lbv_correct: "wt_jvm_prog_lbv P Cert \ wf_jvm_prog P" (*<*) proof - let ?\ = "\C mn. let (C,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,is,xt)) = method P C mn in SOME \s. wt_method P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt \s" - - let ?A = "\P C (mn,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,ins,et)). wt_lbv P C Ts T\<^sub>r mxs mxl\<^sub>0 et (Cert C mn) ins" - let ?B = "\P C (M,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,is,xt)). - wt_method P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt (?\ C M)" - + assume wt: "wt_jvm_prog_lbv P Cert" - then have "wf_prog ?A P" by(simp add: wt_jvm_prog_lbv_def) - moreover { - fix wf_md C M b Ts Ca T m bd - assume "wf_prog wf_md P" and sees: "P \ Ca sees M: Ts\T = m in Ca" and - "set Ts \ types P" and "bd = (M, Ts, T, m)" and - "?A P Ca bd" - then have "?B P Ca bd" using sees_method_is_class[OF sees] - by (auto dest!: start_context.wt_lbv_wt_method [OF start_context.intro] - intro: someI) - } - ultimately have "wf_prog ?B P" by(rule wf_prog_lift) - hence "wf_jvm_prog\<^bsub>?\\<^esub> P" by (simp add: wf_jvm_prog_phi_def) + hence "wf_jvm_prog\<^bsub>?\\<^esub> P" + apply (unfold wf_jvm_prog_phi_def wt_jvm_prog_lbv_def) + apply (erule wf_prog_lift) + apply (auto dest!: start_context.wt_lbv_wt_method [OF start_context.intro] + intro: someI) + apply (erule sees_method_is_class) + done thus ?thesis by (unfold wf_jvm_prog_def) blast qed (*>*) theorem jvm_lbv_complete: assumes wt: "wf_jvm_prog\<^bsub>\\<^esub> P" shows "wt_jvm_prog_lbv P (prg_cert P \)" (*<*) -proof - - let ?cert = "prg_cert P \" - let ?A = "\P C (M,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,is,xt)). - wt_method P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt (\ C M)" - let ?B = "\P C (mn,Ts,T\<^sub>r,(mxs,mxl\<^sub>0,ins,et)). - wt_lbv P C Ts T\<^sub>r mxs mxl\<^sub>0 et (?cert C mn) ins" - - from wt have "wf_prog ?A P" by(clarsimp simp: wf_jvm_prog_def wf_jvm_prog_phi_def) - moreover { - fix wf_md C M Ts Ca T m bd - assume "wf_prog wf_md P" and sees: "P \ Ca sees M: Ts\T = m in Ca" and - "set Ts \ types P" and "bd = (M, Ts, T, m)" and - "?A P Ca bd" - then have "?B P Ca bd" using sees_method_is_class[OF sees] - by (auto simp add: prg_cert_def - intro!: start_context.wt_method_wt_lbv start_context.intro) - } - ultimately have "wf_prog ?B P" by(rule wf_prog_lift) - thus "wt_jvm_prog_lbv P (prg_cert P \)" by (simp add: wt_jvm_prog_lbv_def) -qed + using wt + apply (unfold wf_jvm_prog_phi_def wt_jvm_prog_lbv_def) + apply (erule wf_prog_lift) + apply (auto simp add: prg_cert_def + intro!: start_context.wt_method_wt_lbv start_context.intro) + apply (erule sees_method_is_class) + done (*>*) -end +end diff --git a/thys/Jinja/BV/SemiType.thy b/thys/Jinja/BV/SemiType.thy --- a/thys/Jinja/BV/SemiType.thy +++ b/thys/Jinja/BV/SemiType.thy @@ -1,280 +1,280 @@ (* Title: Jinja/BV/SemiType.thy Author: Tobias Nipkow, Gerwin Klein Copyright 2000 TUM *) section \The Jinja Type System as a Semilattice\ theory SemiType imports "../Common/WellForm" "../DFA/Semilattices" begin definition super :: "'a prog \ cname \ cname" where "super P C \ fst (the (class P C))" lemma superI: "(C,D) \ subcls1 P \ super P C = D" by (unfold super_def) (auto dest: subcls1D) primrec the_Class :: "ty \ cname" where "the_Class (Class C) = C" definition sup :: "'c prog \ ty \ ty \ ty err" where "sup P T\<^sub>1 T\<^sub>2 \ if is_refT T\<^sub>1 \ is_refT T\<^sub>2 then OK (if T\<^sub>1 = NT then T\<^sub>2 else if T\<^sub>2 = NT then T\<^sub>1 else (Class (exec_lub (subcls1 P) (super P) (the_Class T\<^sub>1) (the_Class T\<^sub>2)))) else (if T\<^sub>1 = T\<^sub>2 then OK T\<^sub>1 else Err)" lemma sup_def': "sup P = (\T\<^sub>1 T\<^sub>2. if is_refT T\<^sub>1 \ is_refT T\<^sub>2 then OK (if T\<^sub>1 = NT then T\<^sub>2 else if T\<^sub>2 = NT then T\<^sub>1 else (Class (exec_lub (subcls1 P) (super P) (the_Class T\<^sub>1) (the_Class T\<^sub>2)))) else (if T\<^sub>1 = T\<^sub>2 then OK T\<^sub>1 else Err))" by (simp add: sup_def fun_eq_iff) abbreviation subtype :: "'c prog \ ty \ ty \ bool" where "subtype P \ widen P" definition esl :: "'c prog \ ty esl" where "esl P \ (types P, subtype P, sup P)" (* FIXME: move to wellform *) lemma is_class_is_subcls: "wf_prog m P \ is_class P C = P \ C \\<^sup>* Object" (*<*)by (fastforce simp:is_class_def elim: subcls_C_Object converse_rtranclE subcls1I dest: subcls1D) (*>*) (* FIXME: move to wellform *) lemma subcls_antisym: "\wf_prog m P; P \ C \\<^sup>* D; P \ D \\<^sup>* C\ \ C = D" (*<*) by (auto dest: acyclic_subcls1 acyclic_impl_antisym_rtrancl antisymD) (*>*) (* FIXME: move to wellform *) lemma widen_antisym: "\ wf_prog m P; P \ T \ U; P \ U \ T \ \ T = U" (*<*) apply (cases T) apply (cases U) apply auto apply (cases U) apply (auto elim!: subcls_antisym) done (*>*) lemma order_widen [intro,simp]: - "wf_prog m P \ order (subtype P)" + "wf_prog m P \ order (subtype P) (types P)" (*<*) apply (unfold Semilat.order_def lesub_def) apply (auto intro: widen_trans widen_antisym) done (*>*) (* FIXME: move to TypeRel *) lemma NT_widen: "P \ NT \ T = (T = NT \ (\C. T = Class C))" (*<*) by (cases T) auto (*>*) (* FIXME: move to TypeRel *) lemma Class_widen2: "P \ Class C \ T = (\D. T = Class D \ P \ C \\<^sup>* D)" (*<*) by (cases T) auto (*>*) lemma wf_converse_subcls1_impl_acc_subtype: "wf ((subcls1 P)^-1) \ acc (subtype P)" (*<*) apply (unfold Semilat.acc_def lesssub_def) apply (drule_tac p = "(subcls1 P)^-1 - Id" in wf_subset) apply blast apply (drule wf_trancl) apply (simp add: wf_eq_minimal) apply clarify apply (unfold lesub_def) apply (rename_tac M T) apply (case_tac "\C. Class C \ M") prefer 2 apply (case_tac T) apply fastforce apply fastforce apply fastforce apply simp apply (rule_tac x = NT in bexI) apply (rule allI) apply (rule impI, erule conjE) apply (clarsimp simp add: NT_widen) apply simp apply clarsimp apply (erule_tac x = "{C. Class C : M}" in allE) apply auto apply (rename_tac D) apply (rule_tac x = "Class D" in bexI) prefer 2 apply assumption apply clarify apply (clarsimp simp: Class_widen2) apply (insert rtrancl_r_diff_Id [symmetric, of "subcls1 P"]) apply simp apply (erule rtranclE) apply blast apply (drule rtrancl_converseI) apply (subgoal_tac "((subcls1 P)-Id)^-1 = ((subcls1 P)^-1 - Id)") prefer 2 apply blast apply simp apply (blast intro: rtrancl_into_trancl2) done (*>*) lemma wf_subtype_acc [intro, simp]: "wf_prog wf_mb P \ acc (subtype P)" (*<*) by (rule wf_converse_subcls1_impl_acc_subtype, rule wf_subcls1) (*>*) lemma exec_lub_refl [simp]: "exec_lub r f T T = T" (*<*) by (simp add: exec_lub_def while_unfold) (*>*) lemma closed_err_types: "wf_prog wf_mb P \ closed (err (types P)) (lift2 (sup P))" (*<*) apply (unfold closed_def plussub_def lift2_def sup_def') apply (frule acyclic_subcls1) apply (frule single_valued_subcls1) apply (auto simp: is_type_def is_refT_def is_class_is_subcls split: err.split ty.splits) apply (blast dest!: is_lub_exec_lub is_lubD is_ubD intro!: is_ubI superI) done (*>*) lemma sup_subtype_greater: "\ wf_prog wf_mb P; is_type P t1; is_type P t2; sup P t1 t2 = OK s \ \ subtype P t1 s \ subtype P t2 s" (*<*) proof - assume wf_prog: "wf_prog wf_mb P" { fix c1 c2 assume is_class: "is_class P c1" "is_class P c2" with wf_prog obtain "P \ c1 \\<^sup>* Object" "P \ c2 \\<^sup>* Object" by (blast intro: subcls_C_Object) with single_valued_subcls1[OF wf_prog] obtain u where "is_lub ((subcls1 P)^* ) c1 c2 u" by (blast dest: single_valued_has_lubs) moreover note acyclic_subcls1[OF wf_prog] moreover have "\x y. (x, y) \ subcls1 P \ super P x = y" by (blast intro: superI) ultimately have "P \ c1 \\<^sup>* exec_lub (subcls1 P) (super P) c1 c2 \ P \ c2 \\<^sup>* exec_lub (subcls1 P) (super P) c1 c2" by (simp add: exec_lub_conv) (blast dest: is_lubD is_ubD) } note this [simp] assume "is_type P t1" "is_type P t2" "sup P t1 t2 = OK s" thus ?thesis apply (unfold sup_def) apply (cases s) apply (auto simp add: is_refT_def split: if_split_asm) done qed (*>*) lemma sup_subtype_smallest: "\ wf_prog wf_mb P; is_type P a; is_type P b; is_type P c; subtype P a c; subtype P b c; sup P a b = OK d \ \ subtype P d c" (*<*) proof - assume wf_prog: "wf_prog wf_mb P" { fix c1 c2 D assume is_class: "is_class P c1" "is_class P c2" assume le: "P \ c1 \\<^sup>* D" "P \ c2 \\<^sup>* D" from wf_prog is_class obtain "P \ c1 \\<^sup>* Object" "P \ c2 \\<^sup>* Object" by (blast intro: subcls_C_Object) with single_valued_subcls1[OF wf_prog] obtain u where lub: "is_lub ((subcls1 P)^* ) c1 c2 u" by (blast dest: single_valued_has_lubs) with acyclic_subcls1[OF wf_prog] have "exec_lub (subcls1 P) (super P) c1 c2 = u" by (blast intro: superI exec_lub_conv) moreover from lub le have "P \ u \\<^sup>* D" by (simp add: is_lub_def is_ub_def) ultimately have "P \ exec_lub (subcls1 P) (super P) c1 c2 \\<^sup>* D" by blast } note this [intro] have [dest!]: "\C T. P \ Class C \ T \ \D. T=Class D \ P \ C \\<^sup>* D" by (frule Class_widen, auto) assume "is_type P a" "is_type P b" "is_type P c" "subtype P a c" "subtype P b c" "sup P a b = OK d" thus ?thesis by (auto simp add: sup_def is_refT_def split: if_split_asm) qed (*>*) lemma sup_exists: "\ subtype P a c; subtype P b c \ \ \T. sup P a b = OK T" (*<*) apply (unfold sup_def) apply (cases b) apply auto apply (cases a) apply auto apply (cases a) apply auto done (*>*) lemma err_semilat_JType_esl: "wf_prog wf_mb P \ err_semilat (esl P)" (*<*) proof - assume wf_prog: "wf_prog wf_mb P" - hence "order (subtype P)".. + hence "order (subtype P) (types P)".. moreover from wf_prog have "closed (err (types P)) (lift2 (sup P))" by (rule closed_err_types) moreover from wf_prog have "(\x\err (types P). \y\err (types P). x \\<^bsub>Err.le (subtype P)\<^esub> x \\<^bsub>lift2 (sup P)\<^esub> y) \ (\x\err (types P). \y\err (types P). y \\<^bsub>Err.le (subtype P)\<^esub> x \\<^bsub>lift2 (sup P)\<^esub> y)" by (auto simp add: lesub_def plussub_def Err.le_def lift2_def sup_subtype_greater split: err.split) moreover from wf_prog have "\x\err (types P). \y\err (types P). \z\err (types P). x \\<^bsub>Err.le (subtype P)\<^esub> z \ y \\<^bsub>Err.le (subtype P)\<^esub> z \ x \\<^bsub>lift2 (sup P)\<^esub> y \\<^bsub>Err.le (subtype P)\<^esub> z" by (unfold lift2_def plussub_def lesub_def Err.le_def) (auto intro: sup_subtype_smallest dest:sup_exists split: err.split) ultimately show ?thesis by (simp add: esl_def semilat_def sl_def Err.sl_def) qed (*>*) end diff --git a/thys/Jinja/BV/TF_JVM.thy b/thys/Jinja/BV/TF_JVM.thy --- a/thys/Jinja/BV/TF_JVM.thy +++ b/thys/Jinja/BV/TF_JVM.thy @@ -1,292 +1,278 @@ (* Title: HOL/MicroJava/BV/JVM.thy Author: Tobias Nipkow, Gerwin Klein Copyright 2000 TUM *) section \The Typing Framework for the JVM \label{sec:JVM}\ theory TF_JVM imports "../DFA/Typing_Framework_err" EffectMono BVSpec begin definition exec :: "jvm_prog \ nat \ ty \ ex_table \ instr list \ ty\<^sub>i' err step_type" where "exec G maxs rT et bs \ err_step (size bs) (\pc. app (bs!pc) G maxs rT pc (size bs) et) (\pc. eff (bs!pc) G pc et)" locale JVM_sl = - fixes P :: jvm_prog and mxs and mxl\<^sub>0 + fixes P :: jvm_prog and mxs and mxl\<^sub>0 and n fixes Ts :: "ty list" and "is" and xt and T\<^sub>r fixes mxl and A and r and f and app and eff and step defines [simp]: "mxl \ 1+size Ts+mxl\<^sub>0" defines [simp]: "A \ states P mxs mxl" defines [simp]: "r \ JVM_SemiType.le P mxs mxl" defines [simp]: "f \ JVM_SemiType.sup P mxs mxl" defines [simp]: "app \ \pc. Effect.app (is!pc) P mxs T\<^sub>r pc (size is) xt" defines [simp]: "eff \ \pc. Effect.eff (is!pc) P pc xt" defines [simp]: "step \ err_step (size is) app eff" + defines [simp]: "n \ size is" locale start_context = JVM_sl + fixes p and C assumes wf: "wf_prog p P" assumes C: "is_class P C" assumes Ts: "set Ts \ types P" fixes first :: ty\<^sub>i' and start defines [simp]: "first \ Some ([],OK (Class C) # map OK Ts @ replicate mxl\<^sub>0 Err)" defines [simp]: "start \ OK first # replicate (size is - 1) (OK None)" - - subsection \Connecting JVM and Framework\ - +lemma (in start_context) semi: "semilat (A, r, f)" + apply (insert semilat_JVM[OF wf]) + apply (unfold A_def r_def f_def JVM_SemiType.le_def JVM_SemiType.sup_def states_def) + apply auto + done lemma (in JVM_sl) step_def_exec: "step \ exec P mxs T\<^sub>r xt is" by (simp add: exec_def) lemma special_ex_swap_lemma [iff]: "(? X. (? n. X = A n & P n) & Q X) = (? n. Q(A n) & P n)" by blast lemma ex_in_list [iff]: "(\n. ST \ list n A \ n \ mxs) = (set ST \ A \ size ST \ mxs)" by (unfold list_def) auto lemma singleton_list: "(\n. [Class C] \ list n (types P) \ n \ mxs) = (is_class P C \ 0 < mxs)" by auto lemma set_drop_subset: "set xs \ A \ set (drop n xs) \ A" by (auto dest: in_set_dropD) lemma Suc_minus_minus_le: "n < mxs \ Suc (n - (n - b)) \ mxs" by arith lemma in_listE: "\ xs \ list n A; \size xs = n; set xs \ A\ \ P \ \ P" by (unfold list_def) blast declare is_relevant_entry_def [simp] declare set_drop_subset [simp] theorem (in start_context) exec_pres_type: "pres_type step (size is) A" (*<*) -proof - - let ?n = "size is" and ?app = app and ?step = eff - let ?mxl = "1 + length Ts + mxl\<^sub>0" - let ?A = "opt((Union {list n (types P) |n. n <= mxs}) \ - list ?mxl (err(types P)))" - have "pres_type (err_step ?n ?app ?step) ?n (err ?A)" - proof(rule pres_type_lift) - have "\s pc pc' s'. s\?A \ pc < ?n \ ?app pc s - \ (pc', s')\set (?step pc s) \ s' \ ?A" - proof - - fix s pc pc' s' - assume asms: "s\?A" "pc < ?n" "?app pc s" "(pc', s')\set (?step pc s)" - show "s' \ ?A" - proof(cases s) - case None - then show ?thesis using asms by (fastforce dest: effNone) - next - case (Some ab) - then show ?thesis using asms proof(cases "is!pc") - case Load - then show ?thesis using asms - by (fastforce simp: Effect.app_def xcpt_app_def Effect.eff_def - xcpt_eff_def norm_eff_def - dest: listE_nth_in) - next - case Push - then show ?thesis using asms Some - by (fastforce simp: Effect.app_def xcpt_app_def Effect.eff_def - xcpt_eff_def norm_eff_def typeof_lit_is_type) - next - case Getfield - then show ?thesis using asms wf - by (fastforce simp: Effect.app_def xcpt_app_def Effect.eff_def - xcpt_eff_def norm_eff_def - dest: sees_field_is_type) - next - case (Invoke M n) - obtain a b where [simp]: "s = \(a,b)\" using Some asms(1) by blast - show ?thesis - proof(cases "a!n = NT") - case True - then show ?thesis using Invoke asms wf - by (fastforce simp: Effect.app_def xcpt_app_def Effect.eff_def - xcpt_eff_def norm_eff_def) - next - case False - have "(pc', s') \ set (norm_eff (Invoke M n) P pc (a, b)) \ - (pc', s') \ set (xcpt_eff (Invoke M n) P pc (a, b) xt)" - using Invoke asms(4) by (simp add: Effect.eff_def) - then show ?thesis proof(rule disjE) - assume "(pc', s') \ set (xcpt_eff (Invoke M n) P pc (a, b) xt)" - then show ?thesis using Invoke asms(1-3) - by (fastforce simp: Effect.app_def xcpt_app_def xcpt_eff_def) - next - assume norm: "(pc', s') \ set (norm_eff (Invoke M n) P pc (a, b))" - also have "Suc (length a - Suc n) \ mxs" using Invoke asms(1,3) - by (simp add: Effect.app_def xcpt_app_def) arith - ultimately show ?thesis using False Invoke asms(1-3) wf - by (auto simp: Effect.app_def xcpt_app_def norm_eff_def wf_mdecl_def - dest!: sees_wf_mdecl) - qed - qed - qed (fastforce simp: Effect.app_def xcpt_app_def Effect.eff_def - xcpt_eff_def norm_eff_def)+ - qed - qed - then show "\s\?A. \p. p < ?n \ ?app p s \ (\(q, s')\set (?step p s). s' \ ?A)" - by clarsimp - qed - then show ?thesis by (simp add: JVM_states_unfold) -qed + apply (insert wf) + apply simp + apply (unfold JVM_states_unfold) + apply (rule pres_type_lift) + apply clarify + apply (rename_tac s pc pc' s') + apply (case_tac s) + apply simp + apply (drule effNone) + apply simp + apply (simp add: Effect.app_def xcpt_app_def Effect.eff_def + xcpt_eff_def norm_eff_def relevant_entries_def) + apply (case_tac "is!pc") + + \ \Load\ + apply clarsimp + apply (frule listE_nth_in, assumption) + apply fastforce + + \ \Store\ + apply fastforce + + \ \Push\ + apply (fastforce simp add: typeof_lit_is_type) + + \ \New\ + apply clarsimp apply fastforce + + \ \Getfield\ + apply clarsimp apply (fastforce dest: sees_field_is_type) + + \ \Putfield\ + apply clarsimp apply fastforce + + \ \Checkcast\ + apply clarsimp apply fastforce + + defer + + \ \Return\ + apply fastforce + + \ \Pop\ + apply fastforce + + \ \IAdd\ + apply fastforce + + \ \Goto\ + apply fastforce + + \ \CmpEq\ + apply fastforce + + \ \IfFalse\ + apply fastforce + + \ \Throw\ + apply clarsimp apply fastforce + + \ \Invoke\ + apply (clarsimp split!: if_splits) + apply fastforce + apply (erule disjE) + prefer 2 + apply fastforce + apply clarsimp + apply (rule conjI) + apply (drule (1) sees_wf_mdecl) + apply (clarsimp simp add: wf_mdecl_def) + apply arith + done (*>*) declare is_relevant_entry_def [simp del] declare set_drop_subset [simp del] lemma lesubstep_type_simple: "xs [\\<^bsub>Product.le (=) r\<^esub>] ys \ set xs {\\<^bsub>r\<^esub>} set ys" (*<*) -proof - - assume assm: "xs [\\<^bsub>Product.le (=) r\<^esub>] ys" - have "\a b i y. (a, b) = xs ! i \ i < length xs - \ \\'. (\i. (a, \') = ys ! i \ i < length xs) \ b \\<^bsub>r\<^esub> \'" - proof - - fix a b i assume ith: "(a, b) = xs ! i" and len: "i < length xs" - obtain \ where "ys ! i = (a, \) \ r b \" - using le_listD[OF assm len] ith - by (clarsimp simp: lesub_def Product.le_def) - then have "(a, \) = ys ! i \ b \\<^bsub>r\<^esub> \" - by (clarsimp simp: lesub_def) - with len show "\\'. (\i. (a, \') = ys ! i \ i < length xs) \ b \\<^bsub>r\<^esub> \'" - by fastforce - qed - then show "set xs {\\<^bsub>r\<^esub>} set ys" using assm - by (clarsimp simp: lesubstep_type_def set_conv_nth) -qed + apply (unfold lesubstep_type_def) + apply clarify + apply (simp add: set_conv_nth) + apply clarify + apply (drule le_listD, assumption) + apply (clarsimp simp add: lesub_def Product.le_def) + apply (rule exI) + apply (rule conjI) + apply (rule exI) + apply (rule conjI) + apply (rule sym) + apply assumption + apply assumption + apply assumption + done (*>*) declare is_relevant_entry_def [simp del] lemma conjI2: "\ A; A \ B \ \ A \ B" by blast lemma (in JVM_sl) eff_mono: -assumes wf: "wf_prog p P" and "pc < length is" and - lesub: "s \\<^bsub>sup_state_opt P\<^esub> t" and app: "app pc t" -shows "set (eff pc s) {\\<^bsub>sup_state_opt P\<^esub>} set (eff pc t)" + "\wf_prog p P; pc < length is; s \\<^bsub>sup_state_opt P\<^esub> t; app pc t\ + \ set (eff pc s) {\\<^bsub>sup_state_opt P\<^esub>} set (eff pc t)" (*<*) -proof(cases t) - case None then show ?thesis using lesub - by (simp add: Effect.eff_def lesub_def) -next - case tSome: (Some a) - show ?thesis proof(cases s) - case None then show ?thesis using lesub - by (simp add: Effect.eff_def lesub_def) - next - case (Some b) - let ?norm = "\x. norm_eff (is ! pc) P pc x" - let ?xcpt = "\x. xcpt_eff (is ! pc) P pc x xt" - let ?r = "Product.le (=) (sup_state_opt P)" - let ?\' = "\eff\<^sub>i (is ! pc, P, a)\" - { - fix x assume xb: "x \ set (succs (is ! pc) b pc)" - then have appi: "app\<^sub>i (is ! pc, P, pc, mxs, T\<^sub>r, a)" and - bia: "P \ b \\<^sub>i a" and appa: "app pc \a\" - using lesub app tSome Some by (auto simp add: lesub_def Effect.app_def) - have xa: "x \ set (succs (is ! pc) a pc)" - using xb succs_mono[OF wf appi bia] by auto - then have "(x, ?\') \ (\pc'. (pc', ?\')) ` set (succs (is ! pc) a pc)" - by (rule imageI) - moreover have "P \ \eff\<^sub>i (is ! pc, P, b)\ \' ?\'" - using xb xa eff\<^sub>i_mono[OF wf bia] appa by fastforce - ultimately have "\\'. (x, \') \ (\pc'. (pc', \eff\<^sub>i (is ! pc, P, a)\)) ` set (succs (is ! pc) a pc) \ - P \ \eff\<^sub>i (is ! pc, P, b)\ \' \'" by blast - } - then have norm: "set (?norm b) {\\<^bsub>sup_state_opt P\<^esub>} set (?norm a)" - using tSome Some by (clarsimp simp: norm_eff_def lesubstep_type_def lesub_def) - obtain a1 b1 a2 b2 where a: "a = (a1, b1)" and b: "b = (a2, b2)" - using tSome Some by fastforce - then have a12: "size a2 = size a1" using lesub tSome Some - by (clarsimp simp: lesub_def list_all2_lengthD) - have "length (?xcpt b) = length (?xcpt a)" - by (simp add: xcpt_eff_def split_beta) - moreover have "\n. n < length (?xcpt b) \ (?xcpt b) ! n \\<^bsub>?r\<^esub> (?xcpt a) ! n" - using lesub tSome Some a b a12 - by (simp add: xcpt_eff_def split_beta fun_of_def) (simp add: lesub_def) - ultimately have "?xcpt b [\\<^bsub>?r\<^esub>] ?xcpt a" - by(rule le_listI) - then have "set (?xcpt b) {\\<^bsub>sup_state_opt P\<^esub>} set (?xcpt a)" - by (rule lesubstep_type_simple) - moreover note norm - ultimately have - "set (?norm b) \ set (?xcpt b) {\\<^bsub>sup_state_opt P\<^esub>} set (?norm a) \ set (?xcpt a)" - by(intro lesubstep_union) - then show ?thesis using tSome Some by(simp add: Effect.eff_def) - qed -qed + apply simp + apply (unfold Effect.eff_def) + apply (cases t) + apply (simp add: lesub_def) + apply (rename_tac a) + apply (cases s) + apply simp + apply (rename_tac b) + apply simp + apply (rule lesubstep_union) + prefer 2 + apply (rule lesubstep_type_simple) + apply (simp add: xcpt_eff_def) + apply (rule le_listI) + apply (simp add: split_beta) + apply (simp add: split_beta) + apply (simp add: lesub_def fun_of_def) + apply (case_tac a) + apply (case_tac b) + apply simp + apply (subgoal_tac "size ab = size aa") + prefer 2 + apply (clarsimp simp add: list_all2_lengthD) + apply simp + apply (clarsimp simp add: norm_eff_def lesubstep_type_def lesub_def iff del: sup_state_conv) + apply (rule exI) + apply (rule conjI2) + apply (rule imageI) + apply (clarsimp simp add: Effect.app_def iff del: sup_state_conv) + apply (drule (2) succs_mono) + apply blast + apply simp + apply (erule eff\<^sub>i_mono) + apply simp + apply assumption + apply clarsimp + apply clarsimp + done (*>*) lemma (in JVM_sl) bounded_step: "bounded step (size is)" (*<*) - by (auto simp: bounded_def err_step_def Effect.app_def Effect.eff_def - error_def map_snd_def - split: err.splits option.splits) + apply simp + apply (unfold bounded_def err_step_def Effect.app_def Effect.eff_def) + apply (auto simp add: error_def map_snd_def split: err.splits option.splits) + done (*>*) theorem (in JVM_sl) step_mono: "wf_prog wf_mb P \ mono r step (size is) A" (*<*) -proof - - assume wf: "wf_prog wf_mb P" - let ?r = "sup_state_opt P" and ?n = "length is" and ?app = app and ?step = eff - let ?A = "opt (\ {list n (types P) |n. n \ mxs} \ - list (1 + length Ts + mxl\<^sub>0) (err (types P)))" - have "order ?r" using wf by simp - moreover have "app_mono ?r ?app ?n ?A" using app_mono[OF wf] - by (clarsimp simp: app_mono_def lesub_def) - moreover have "bounded (err_step ?n ?app ?step) ?n" using bounded_step - by simp - moreover have "\s p t. s \ ?A \ p < ?n \ s \\<^bsub>?r\<^esub> t \ - ?app p t \ set (?step p s) {\\<^bsub>?r\<^esub>} set (?step p t)" - using eff_mono[OF wf] by simp - ultimately have "mono (Err.le ?r) (err_step ?n ?app ?step) ?n (err ?A)" - by(rule mono_lift) - then show "mono r step (size is) A" using bounded_step - by (simp add: JVM_le_Err_conv JVM_states_unfold) -qed + apply (simp add: JVM_le_Err_conv) + apply (insert bounded_step) + apply (unfold JVM_states_unfold) + apply (rule mono_lift) + apply blast \\ order_sup_state_opt' \ + apply (unfold app_mono_def lesub_def) + apply clarsimp + apply (erule (2) app_mono) + apply simp + apply clarify + apply (drule eff_mono) + apply (auto simp add: lesub_def) + done (*>*) lemma (in start_context) first_in_A [iff]: "OK first \ A" using Ts C by (force intro!: list_appendI simp add: JVM_states_unfold) lemma (in JVM_sl) wt_method_def2: "wt_method P C' Ts T\<^sub>r mxs mxl\<^sub>0 is xt \s = (is \ [] \ size \s = size is \ OK ` set \s \ states P mxs mxl \ wt_start P C' Ts mxl\<^sub>0 \s \ wt_app_eff (sup_state_opt P) app eff \s)" (*<*) - by (unfold wt_method_def wt_app_eff_def wt_instr_def lesub_def - check_types_def) auto + apply (unfold wt_method_def wt_app_eff_def wt_instr_def lesub_def check_types_def) + apply auto + done (*>*) end diff --git a/thys/Jinja/Compiler/Hidden.thy b/thys/Jinja/Compiler/Hidden.thy --- a/thys/Jinja/Compiler/Hidden.thy +++ b/thys/Jinja/Compiler/Hidden.thy @@ -1,51 +1,51 @@ theory Hidden -imports "List-Index.List_Index" +imports "List-Index.List_Index" begin definition hidden :: "'a list \ nat \ bool" where "hidden xs i \ i < size xs \ xs!i \ set(drop (i+1) xs)" lemma hidden_last_index: "x \ set xs \ hidden (xs @ [x]) (last_index xs x)" by(auto simp add: hidden_def nth_append rev_nth[symmetric] dest: last_index_less[OF _ le_refl]) lemma hidden_inacc: "hidden xs i \ last_index xs x \ i" by(auto simp add: hidden_def last_index_drop last_index_less_size_conv) lemma [simp]: "hidden xs i \ hidden (xs@[x]) i" by(auto simp add:hidden_def nth_append) lemma fun_upds_apply: "(m(xs[\]ys)) x = (let xs' = take (size ys) xs in if x \ set xs' then Some(ys ! last_index xs' x) else m x)" proof(induct xs arbitrary: m ys) case Nil then show ?case by(simp add: Let_def) next case Cons show ?case proof(cases ys) case Nil then show ?thesis by(simp add:Let_def) next case Cons': Cons then show ?thesis using Cons by(simp add: Let_def last_index_Cons) qed qed lemma map_upds_apply_eq_Some: "((m(xs[\]ys)) x = Some y) = (let xs' = take (size ys) xs in if x \ set xs' then ys ! last_index xs' x = y else m x = Some y)" by(simp add:fun_upds_apply Let_def) lemma map_upds_upd_conv_last_index: "\x \ set xs; size xs \ size ys \ \ m(xs[\]ys)(x\y) = m(xs[\]ys[last_index xs x := y])" by(rule ext) (simp add:fun_upds_apply eq_sym_conv Let_def) end diff --git a/thys/Jinja/Compiler/TypeComp.thy b/thys/Jinja/Compiler/TypeComp.thy --- a/thys/Jinja/Compiler/TypeComp.thy +++ b/thys/Jinja/Compiler/TypeComp.thy @@ -1,1445 +1,1489 @@ (* Title: Jinja/Compiler/TypeComp.thy Author: Tobias Nipkow Copyright TUM 2003 *) section \Preservation of Well-Typedness\ theory TypeComp imports Compiler "../BV/BVSpec" begin (*<*) declare nth_append[simp] (*>*) locale TC0 = fixes P :: "J\<^sub>1_prog" and mxl :: nat begin definition "ty E e = (THE T. P,E \\<^sub>1 e :: T)" definition "ty\<^sub>l E A' = map (\i. if i \ A' \ i < size E then OK(E!i) else Err) [0..i' ST E A = (case A of None \ None | \A'\ \ Some(ST, ty\<^sub>l E A'))" definition "after E A ST e = ty\<^sub>i' (ty E e # ST) E (A \ \ e)" end lemma (in TC0) ty_def2 [simp]: "P,E \\<^sub>1 e :: T \ ty E e = T" (*<*)by(unfold ty_def) (blast intro: the_equality WT\<^sub>1_unique)(*>*) lemma (in TC0) [simp]: "ty\<^sub>i' ST E None = None" (*<*)by (simp add: ty\<^sub>i'_def)(*>*) lemma (in TC0) ty\<^sub>l_app_diff[simp]: "ty\<^sub>l (E@[T]) (A - {size E}) = ty\<^sub>l E A" (*<*)by(auto simp add:ty\<^sub>l_def hyperset_defs)(*>*) lemma (in TC0) ty\<^sub>i'_app_diff[simp]: "ty\<^sub>i' ST (E @ [T]) (A \ size E) = ty\<^sub>i' ST E A" (*<*)by(auto simp add:ty\<^sub>i'_def hyperset_defs)(*>*) lemma (in TC0) ty\<^sub>l_antimono: "A \ A' \ P \ ty\<^sub>l E A' [\\<^sub>\] ty\<^sub>l E A" (*<*)by(auto simp:ty\<^sub>l_def list_all2_conv_all_nth)(*>*) lemma (in TC0) ty\<^sub>i'_antimono: "A \ A' \ P \ ty\<^sub>i' ST E \A'\ \' ty\<^sub>i' ST E \A\" (*<*)by(auto simp:ty\<^sub>i'_def ty\<^sub>l_def list_all2_conv_all_nth)(*>*) lemma (in TC0) ty\<^sub>l_env_antimono: "P \ ty\<^sub>l (E@[T]) A [\\<^sub>\] ty\<^sub>l E A" (*<*)by(auto simp:ty\<^sub>l_def list_all2_conv_all_nth)(*>*) lemma (in TC0) ty\<^sub>i'_env_antimono: "P \ ty\<^sub>i' ST (E@[T]) A \' ty\<^sub>i' ST E A" (*<*)by(auto simp:ty\<^sub>i'_def ty\<^sub>l_def list_all2_conv_all_nth)(*>*) lemma (in TC0) ty\<^sub>i'_incr: "P \ ty\<^sub>i' ST (E @ [T]) \insert (size E) A\ \' ty\<^sub>i' ST E \A\" (*<*)by(auto simp:ty\<^sub>i'_def ty\<^sub>l_def list_all2_conv_all_nth)(*>*) lemma (in TC0) ty\<^sub>l_incr: "P \ ty\<^sub>l (E @ [T]) (insert (size E) A) [\\<^sub>\] ty\<^sub>l E A" (*<*)by(auto simp: hyperset_defs ty\<^sub>l_def list_all2_conv_all_nth)(*>*) lemma (in TC0) ty\<^sub>l_in_types: "set E \ types P \ ty\<^sub>l E A \ list mxl (err (types P))" (*<*)by(auto simp add:ty\<^sub>l_def intro!:listI dest!: nth_mem)(*>*) locale TC1 = TC0 begin primrec compT :: "ty list \ nat hyperset \ ty list \ expr\<^sub>1 \ ty\<^sub>i' list" and compTs :: "ty list \ nat hyperset \ ty list \ expr\<^sub>1 list \ ty\<^sub>i' list" where "compT E A ST (new C) = []" | "compT E A ST (Cast C e) = compT E A ST e @ [after E A ST e]" | "compT E A ST (Val v) = []" | "compT E A ST (e\<^sub>1 \bop\ e\<^sub>2) = (let ST\<^sub>1 = ty E e\<^sub>1#ST; A\<^sub>1 = A \ \ e\<^sub>1 in compT E A ST e\<^sub>1 @ [after E A ST e\<^sub>1] @ compT E A\<^sub>1 ST\<^sub>1 e\<^sub>2 @ [after E A\<^sub>1 ST\<^sub>1 e\<^sub>2])" | "compT E A ST (Var i) = []" | "compT E A ST (i := e) = compT E A ST e @ [after E A ST e, ty\<^sub>i' ST E (A \ \ e \ \{i}\)]" | "compT E A ST (e\F{D}) = compT E A ST e @ [after E A ST e]" | "compT E A ST (e\<^sub>1\F{D} := e\<^sub>2) = (let ST\<^sub>1 = ty E e\<^sub>1#ST; A\<^sub>1 = A \ \ e\<^sub>1; A\<^sub>2 = A\<^sub>1 \ \ e\<^sub>2 in compT E A ST e\<^sub>1 @ [after E A ST e\<^sub>1] @ compT E A\<^sub>1 ST\<^sub>1 e\<^sub>2 @ [after E A\<^sub>1 ST\<^sub>1 e\<^sub>2] @ [ty\<^sub>i' ST E A\<^sub>2])" | "compT E A ST {i:T; e} = compT (E@[T]) (A\i) ST e" | "compT E A ST (e\<^sub>1;;e\<^sub>2) = (let A\<^sub>1 = A \ \ e\<^sub>1 in compT E A ST e\<^sub>1 @ [after E A ST e\<^sub>1, ty\<^sub>i' ST E A\<^sub>1] @ compT E A\<^sub>1 ST e\<^sub>2)" | "compT E A ST (if (e) e\<^sub>1 else e\<^sub>2) = (let A\<^sub>0 = A \ \ e; \ = ty\<^sub>i' ST E A\<^sub>0 in compT E A ST e @ [after E A ST e, \] @ compT E A\<^sub>0 ST e\<^sub>1 @ [after E A\<^sub>0 ST e\<^sub>1, \] @ compT E A\<^sub>0 ST e\<^sub>2)" | "compT E A ST (while (e) c) = (let A\<^sub>0 = A \ \ e; A\<^sub>1 = A\<^sub>0 \ \ c; \ = ty\<^sub>i' ST E A\<^sub>0 in compT E A ST e @ [after E A ST e, \] @ compT E A\<^sub>0 ST c @ [after E A\<^sub>0 ST c, ty\<^sub>i' ST E A\<^sub>1, ty\<^sub>i' ST E A\<^sub>0])" | "compT E A ST (throw e) = compT E A ST e @ [after E A ST e]" | "compT E A ST (e\M(es)) = compT E A ST e @ [after E A ST e] @ compTs E (A \ \ e) (ty E e # ST) es" | "compT E A ST (try e\<^sub>1 catch(C i) e\<^sub>2) = compT E A ST e\<^sub>1 @ [after E A ST e\<^sub>1] @ [ty\<^sub>i' (Class C#ST) E A, ty\<^sub>i' ST (E@[Class C]) (A \ \{i}\)] @ compT (E@[Class C]) (A \ \{i}\) ST e\<^sub>2" | "compTs E A ST [] = []" | "compTs E A ST (e#es) = compT E A ST e @ [after E A ST e] @ compTs E (A \ (\ e)) (ty E e # ST) es" definition compT\<^sub>a :: "ty list \ nat hyperset \ ty list \ expr\<^sub>1 \ ty\<^sub>i' list" where "compT\<^sub>a E A ST e = compT E A ST e @ [after E A ST e]" end lemma compE\<^sub>2_not_Nil[simp]: "compE\<^sub>2 e \ []" (*<*)by(induct e) auto(*>*) lemma (in TC1) compT_sizes[simp]: shows "\E A ST. size(compT E A ST e) = size(compE\<^sub>2 e) - 1" and "\E A ST. size(compTs E A ST es) = size(compEs\<^sub>2 es)" (*<*) by(induct e and es rule: compE\<^sub>2.induct compEs\<^sub>2.induct) (auto split:bop.splits nat_diff_split) (*>*) lemma (in TC1) [simp]: "\ST E. \\\ \ set (compT E None ST e)" and [simp]: "\ST E. \\\ \ set (compTs E None ST es)" (*<*)by(induct e and es rule: compT.induct compTs.induct) (simp_all add:after_def)(*>*) lemma (in TC0) pair_eq_ty\<^sub>i'_conv: "(\(ST, LT)\ = ty\<^sub>i' ST\<^sub>0 E A) = (case A of None \ False | Some A \ (ST = ST\<^sub>0 \ LT = ty\<^sub>l E A))" (*<*)by(simp add:ty\<^sub>i'_def)(*>*) lemma (in TC0) pair_conv_ty\<^sub>i': "\(ST, ty\<^sub>l E A)\ = ty\<^sub>i' ST E \A\" (*<*)by(simp add:ty\<^sub>i'_def)(*>*) (*<*) declare (in TC0) ty\<^sub>i'_antimono [intro!] after_def[simp] pair_conv_ty\<^sub>i'[simp] pair_eq_ty\<^sub>i'_conv[simp] (*>*) lemma (in TC1) compT_LT_prefix: "\E A ST\<^sub>0. \ \(ST,LT)\ \ set(compT E A ST\<^sub>0 e); \ e (size E) \ \ P \ \(ST,LT)\ \' ty\<^sub>i' ST E A" and "\E A ST\<^sub>0. \ \(ST,LT)\ \ set(compTs E A ST\<^sub>0 es); \s es (size E) \ \ P \ \(ST,LT)\ \' ty\<^sub>i' ST E A" (*<*) proof(induct e and es rule: compT.induct compTs.induct) case FAss thus ?case by(fastforce simp:hyperset_defs elim!:sup_state_opt_trans) next case BinOp thus ?case by(fastforce simp:hyperset_defs elim!:sup_state_opt_trans split:bop.splits) next case Seq thus ?case by(fastforce simp:hyperset_defs elim!:sup_state_opt_trans) next case While thus ?case by(fastforce simp:hyperset_defs elim!:sup_state_opt_trans) next case Cond thus ?case by(fastforce simp:hyperset_defs elim!:sup_state_opt_trans) next case Block thus ?case by(force simp add:hyperset_defs ty\<^sub>i'_def simp del:pair_conv_ty\<^sub>i' elim!:sup_state_opt_trans) next case Call thus ?case by(fastforce simp:hyperset_defs elim!:sup_state_opt_trans) next case Cons_exp thus ?case by(fastforce simp:hyperset_defs elim!:sup_state_opt_trans) next case TryCatch thus ?case by(fastforce simp:hyperset_defs intro!:(* ty\<^sub>i'_env_antimono *) ty\<^sub>i'_incr elim!:sup_state_opt_trans) qed (auto simp:hyperset_defs) declare (in TC0) ty\<^sub>i'_antimono [rule del] after_def[simp del] pair_conv_ty\<^sub>i'[simp del] pair_eq_ty\<^sub>i'_conv[simp del] (*>*) lemma [iff]: "OK None \ states P mxs mxl" (*<*)by(simp add: JVM_states_unfold)(*>*) lemma (in TC0) after_in_states: assumes wf: "wf_prog p P" and wt: "P,E \\<^sub>1 e :: T" and Etypes: "set E \ types P" and STtypes: "set ST \ types P" and stack: "size ST + max_stack e \ mxs" shows "OK (after E A ST e) \ states P mxs mxl" (*<*) proof - have "size ST + 1 \ mxs" using max_stack1[of e] wt stack by fastforce then show ?thesis using assms by(simp add: after_def ty\<^sub>i'_def JVM_states_unfold ty\<^sub>l_in_types) (blast intro!:listI WT\<^sub>1_is_type) qed (*>*) lemma (in TC0) OK_ty\<^sub>i'_in_statesI[simp]: "\ set E \ types P; set ST \ types P; size ST \ mxs \ \ OK (ty\<^sub>i' ST E A) \ states P mxs mxl" (*<*) by(simp add:ty\<^sub>i'_def JVM_states_unfold ty\<^sub>l_in_types) (blast intro!:listI) (*>*) lemma is_class_type_aux: "is_class P C \ is_type P (Class C)" (*<*)by(simp)(*>*) (*<*) declare is_type_simps[simp del] subsetI[rule del] (*>*) theorem (in TC1) compT_states: assumes wf: "wf_prog p P" shows "\E T A ST. \ P,E \\<^sub>1 e :: T; set E \ types P; set ST \ types P; size ST + max_stack e \ mxs; size E + max_vars e \ mxl \ \ OK ` set(compT E A ST e) \ states P mxs mxl" (*<*)(is "\E T A ST. PROP ?P e E T A ST")(*>*) and "\E Ts A ST. \ P,E \\<^sub>1 es[::]Ts; set E \ types P; set ST \ types P; size ST + max_stacks es \ mxs; size E + max_varss es \ mxl \ \ OK ` set(compTs E A ST es) \ states P mxs mxl" (*<*)(is "\E Ts A ST. PROP ?Ps es E Ts A ST") proof(induct e and es rule: compT.induct compTs.induct) case new thus ?case by(simp) next case (Cast C e) thus ?case by (auto simp:after_in_states[OF wf]) next case Val thus ?case by(simp) next case Var thus ?case by(simp) next case LAss thus ?case by(auto simp:after_in_states[OF wf]) next case FAcc thus ?case by(auto simp:after_in_states[OF wf]) next case FAss thus ?case by(auto simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf]) next case Seq thus ?case by(auto simp:image_Un after_in_states[OF wf]) next case BinOp thus ?case by(auto simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf]) next case Cond thus ?case by(force simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf]) next case While thus ?case by(auto simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf]) next case Block thus ?case by(auto) next case (TryCatch e\<^sub>1 C i e\<^sub>2) moreover have "size ST + 1 \ mxs" using TryCatch.prems max_stack1[of e\<^sub>1] by auto ultimately show ?case by(auto simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf] is_class_type_aux) next case Nil_exp thus ?case by simp next case Cons_exp thus ?case by(auto simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf]) next case throw thus ?case by(auto simp: WT\<^sub>1_is_type[OF wf] after_in_states[OF wf]) next case Call thus ?case by(auto simp:image_Un WT\<^sub>1_is_type[OF wf] after_in_states[OF wf]) qed declare is_type_simps[simp] subsetI[intro!] (*>*) definition shift :: "nat \ ex_table \ ex_table" where "shift n xt \ map (\(from,to,C,handler,depth). (from+n,to+n,C,handler+n,depth)) xt" lemma [simp]: "shift 0 xt = xt" (*<*)by(induct xt)(auto simp:shift_def)(*>*) lemma [simp]: "shift n [] = []" (*<*)by(simp add:shift_def)(*>*) lemma [simp]: "shift n (xt\<^sub>1 @ xt\<^sub>2) = shift n xt\<^sub>1 @ shift n xt\<^sub>2" (*<*)by(simp add:shift_def)(*>*) lemma [simp]: "shift m (shift n xt) = shift (m+n) xt" (*<*)by(induct xt)(auto simp:shift_def)(*>*) lemma [simp]: "pcs (shift n xt) = {pc+n|pc. pc \ pcs xt}" (*<*) proof - { fix x f t C h d assume "(f,t,C,h,d) \ set xt" and "f + n \ x" and "x < t + n" then have "\pc. x = pc + n \ (\x\set xt. pc \ (case x of (f, t, C, h, d) \ {f..*) lemma shift_compxE\<^sub>2: shows "\pc pc' d. shift pc (compxE\<^sub>2 e pc' d) = compxE\<^sub>2 e (pc' + pc) d" and "\pc pc' d. shift pc (compxEs\<^sub>2 es pc' d) = compxEs\<^sub>2 es (pc' + pc) d" (*<*) by(induct e and es rule: compxE\<^sub>2.induct compxEs\<^sub>2.induct) (auto simp:shift_def ac_simps) (*>*) lemma compxE\<^sub>2_size_convs[simp]: shows "n \ 0 \ compxE\<^sub>2 e n d = shift n (compxE\<^sub>2 e 0 d)" and "n \ 0 \ compxEs\<^sub>2 es n d = shift n (compxEs\<^sub>2 es 0 d)" (*<*)by(simp_all add:shift_compxE\<^sub>2)(*>*) locale TC2 = TC1 + fixes T\<^sub>r :: ty and mxs :: pc begin definition wt_instrs :: "instr list \ ex_table \ ty\<^sub>i' list \ bool" ("(\ _, _ /[::]/ _)" [0,0,51] 50) where "\ is,xt [::] \s \ size is < size \s \ pcs xt \ {0.. (\pc< size is. P,T\<^sub>r,mxs,size \s,xt \ is!pc,pc :: \s)" end notation TC2.wt_instrs ("(_,_,_ \/ _, _ /[::]/ _)" [50,50,50,50,50,51] 50) (*<*) lemmas (in TC2) wt_defs = wt_instrs_def wt_instr_def app_def eff_def norm_eff_def (*>*) lemma (in TC2) [simp]: "\s \ [] \ \ [],[] [::] \s" (*<*) by (simp add: wt_defs) (*>*) lemma [simp]: "eff i P pc et None = []" (*<*)by (simp add: Effect.eff_def)(*>*) (*<*) declare split_comp_eq[simp del] (*>*) lemma wt_instr_appR: "\ P,T,m,mpc,xt \ is!pc,pc :: \s; pc < size is; size is < size \s; mpc \ size \s; mpc \ mpc' \ \ P,T,m,mpc',xt \ is!pc,pc :: \s@\s'" (*<*)by (fastforce simp:wt_instr_def app_def)(*>*) lemma relevant_entries_shift [simp]: "relevant_entries P i (pc+n) (shift n xt) = shift n (relevant_entries P i pc xt)" (*<*) proof(induct xt) case Nil then show ?case by (simp add: relevant_entries_def shift_def) next case (Cons a xt) then show ?case by (auto simp add: relevant_entries_def shift_def is_relevant_entry_def) qed (*>*) lemma [simp]: "xcpt_eff i P (pc+n) \ (shift n xt) = map (\(pc,\). (pc + n, \)) (xcpt_eff i P pc \ xt)" (*<*) proof - obtain ST LT where "\ = (ST, LT)" by(cases \) simp then show ?thesis by(simp add: xcpt_eff_def) (auto simp add: shift_def) qed (*>*) lemma [simp]: "app\<^sub>i (i, P, pc, m, T, \) \ eff i P (pc+n) (shift n xt) (Some \) = map (\(pc,\). (pc+n,\)) (eff i P pc xt (Some \))" (*<*)by(cases "i") (auto simp add:eff_def norm_eff_def)(*>*) lemma [simp]: "xcpt_app i P (pc+n) mxs (shift n xt) \ = xcpt_app i P pc mxs xt \" (*<*)by (simp add: xcpt_app_def) (auto simp add: shift_def)(*>*) lemma wt_instr_appL: assumes "P,T,m,mpc,xt \ i,pc :: \s" and "pc < size \s" and "mpc \ size \s" shows "P,T,m,mpc + size \s',shift (size \s') xt \ i,pc+size \s' :: \s'@\s" (*<*) proof - let ?t = "(\s'@\s)!(pc+size \s')" show ?thesis proof(cases ?t) case (Some \) obtain ST LT where [simp]: "\ = (ST, LT)" by(cases \) simp have "app\<^sub>i (i, P, pc + length \s', m, T, \)" using Some assms by(cases "i") (auto simp:wt_instr_def app_def) moreover { fix pc' \' assume "(pc',\') \ set (eff i P pc xt ?t)" then have "P \ \' \' \s!pc'" and "pc' < mpc" using Some assms by(auto simp:wt_instr_def app_def) } ultimately show ?thesis using Some assms by(fastforce simp:wt_instr_def app_def) qed (auto simp:wt_instr_def app_def) qed (*>*) lemma wt_instr_Cons: assumes wti: "P,T,m,mpc - 1,[] \ i,pc - 1 :: \s" and pcl: "0 < pc" and mpcl: "0 < mpc" and pcu: "pc < size \s + 1" and mpcu: "mpc \ size \s + 1" shows "P,T,m,mpc,[] \ i,pc :: \#\s" (*<*) proof - have "pc - 1 < length \s" using pcl pcu by arith moreover have "mpc - 1 \ length \s" using mpcl mpcu by arith ultimately have "P,T,m,mpc - 1 + length [\],shift (length [\]) [] \ i,pc - 1 + length [\] :: [\] @ \s" by(rule wt_instr_appL[where \s' = "[\]", OF wti]) then show ?thesis using pcl mpcl by (simp split:nat_diff_split_asm) qed (*>*) lemma wt_instr_append: assumes wti: "P,T,m,mpc - size \s',[] \ i,pc - size \s' :: \s" and pcl: "size \s' \ pc" and mpcl: "size \s' \ mpc" and pcu: "pc < size \s + size \s'" and mpcu: "mpc \ size \s + size \s'" shows "P,T,m,mpc,[] \ i,pc :: \s'@\s" (*<*) proof - have "pc - length \s' < length \s" using pcl pcu by arith moreover have "mpc - length \s' \ length \s" using mpcl mpcu by arith thm wt_instr_appL[where \s' = "\s'", OF wti] ultimately have "P,T,m,mpc - length \s' + length \s',shift (length \s') [] \ i,pc - length \s' + length \s' :: \s' @ \s" by(rule wt_instr_appL[where \s' = "\s'", OF wti]) then show ?thesis using pcl mpcl by (simp split:nat_diff_split_asm) qed (*>*) lemma xcpt_app_pcs: "pc \ pcs xt \ xcpt_app i P pc mxs xt \" (*<*) by (auto simp add: xcpt_app_def relevant_entries_def is_relevant_entry_def pcs_def) (*>*) lemma xcpt_eff_pcs: "pc \ pcs xt \ xcpt_eff i P pc \ xt = []" (*<*) by (cases \) (auto simp add: is_relevant_entry_def xcpt_eff_def relevant_entries_def pcs_def intro!: filter_False) (*>*) lemma pcs_shift: "pc < n \ pc \ pcs (shift n xt)" (*<*)by (auto simp add: shift_def pcs_def)(*>*) lemma wt_instr_appRx: "\ P,T,m,mpc,xt \ is!pc,pc :: \s; pc < size is; size is < size \s; mpc \ size \s \ \ P,T,m,mpc,xt @ shift (size is) xt' \ is!pc,pc :: \s" (*<*)by (auto simp:wt_instr_def eff_def app_def xcpt_app_pcs xcpt_eff_pcs)(*>*) lemma wt_instr_appLx: "\ P,T,m,mpc,xt \ i,pc :: \s; pc \ pcs xt' \ \ P,T,m,mpc,xt'@xt \ i,pc :: \s" (*<*)by (auto simp:wt_instr_def app_def eff_def xcpt_app_pcs xcpt_eff_pcs)(*>*) lemma (in TC2) wt_instrs_extR: "\ is,xt [::] \s \ \ is,xt [::] \s @ \s'" (*<*)by(auto simp add:wt_instrs_def wt_instr_appR)(*>*) lemma (in TC2) wt_instrs_ext: assumes wt\<^sub>1: "\ is\<^sub>1,xt\<^sub>1 [::] \s\<^sub>1@\s\<^sub>2" and wt\<^sub>2: "\ is\<^sub>2,xt\<^sub>2 [::] \s\<^sub>2" and \s_size: "size \s\<^sub>1 = size is\<^sub>1" shows "\ is\<^sub>1@is\<^sub>2, xt\<^sub>1 @ shift (size is\<^sub>1) xt\<^sub>2 [::] \s\<^sub>1@\s\<^sub>2" (*<*) proof - let ?is = "is\<^sub>1@is\<^sub>2" and ?xt = "xt\<^sub>1 @ shift (size is\<^sub>1) xt\<^sub>2" and ?\s = "\s\<^sub>1@\s\<^sub>2" have "size ?is < size ?\s" using wt\<^sub>2 \s_size by(fastforce simp:wt_instrs_def) moreover have "pcs ?xt \ {0..1 wt\<^sub>2 by(fastforce simp:wt_instrs_def) moreover { fix pc assume pc: "pcr,mxs,size ?\s,?xt \ ?is!pc,pc :: ?\s" proof(cases "pc < length is\<^sub>1") case True then show ?thesis using wt\<^sub>1 pc by(fastforce simp: wt_instrs_def wt_instr_appRx) next case False then have "pc - length is\<^sub>1 < length is\<^sub>2" using pc by fastforce then have "P,T\<^sub>r,mxs,length \s\<^sub>2,xt\<^sub>2 \ is\<^sub>2 ! (pc - length is\<^sub>1),pc - length is\<^sub>1 :: \s\<^sub>2" using wt\<^sub>2 by(clarsimp simp: wt_instrs_def) moreover have "pc - length is\<^sub>1 < length \s\<^sub>2" using pc wt\<^sub>2 by(clarsimp simp: wt_instrs_def) arith moreover have "length \s\<^sub>2 \ length \s\<^sub>2" by simp moreover have "pc - length is\<^sub>1 + length \s\<^sub>1 \ pcs xt\<^sub>1" using wt\<^sub>1 \s_size by(fastforce simp: wt_instrs_def) ultimately have "P,T\<^sub>r,mxs,length \s\<^sub>2 + length \s\<^sub>1,xt\<^sub>1 @ shift (length \s\<^sub>1) xt\<^sub>2 \ is\<^sub>2 ! (pc - length is\<^sub>1),pc - length is\<^sub>1 + length \s\<^sub>1 :: \s\<^sub>1 @ \s\<^sub>2" by(rule wt_instr_appLx[OF wt_instr_appL[where \s' = "\s\<^sub>1"]]) then show ?thesis using False \s_size by(simp add:add.commute) qed } ultimately show ?thesis by(clarsimp simp:wt_instrs_def) qed (*>*) corollary (in TC2) wt_instrs_ext2: "\ \ is\<^sub>2,xt\<^sub>2 [::] \s\<^sub>2; \ is\<^sub>1,xt\<^sub>1 [::] \s\<^sub>1@\s\<^sub>2; size \s\<^sub>1 = size is\<^sub>1 \ \ \ is\<^sub>1@is\<^sub>2, xt\<^sub>1 @ shift (size is\<^sub>1) xt\<^sub>2 [::] \s\<^sub>1@\s\<^sub>2" (*<*)by(rule wt_instrs_ext)(*>*) corollary (in TC2) wt_instrs_ext_prefix [trans]: "\ \ is\<^sub>1,xt\<^sub>1 [::] \s\<^sub>1@\s\<^sub>2; \ is\<^sub>2,xt\<^sub>2 [::] \s\<^sub>3; size \s\<^sub>1 = size is\<^sub>1; prefix \s\<^sub>3 \s\<^sub>2 \ \ \ is\<^sub>1@is\<^sub>2, xt\<^sub>1 @ shift (size is\<^sub>1) xt\<^sub>2 [::] \s\<^sub>1@\s\<^sub>2" (*<*)by(bestsimp simp:prefix_def elim: wt_instrs_ext dest:wt_instrs_extR)(*>*) corollary (in TC2) wt_instrs_app: assumes is\<^sub>1: "\ is\<^sub>1,xt\<^sub>1 [::] \s\<^sub>1@[\]" assumes is\<^sub>2: "\ is\<^sub>2,xt\<^sub>2 [::] \#\s\<^sub>2" assumes s: "size \s\<^sub>1 = size is\<^sub>1" shows "\ is\<^sub>1@is\<^sub>2, xt\<^sub>1@shift (size is\<^sub>1) xt\<^sub>2 [::] \s\<^sub>1@\#\s\<^sub>2" (*<*) proof - from is\<^sub>1 have "\ is\<^sub>1,xt\<^sub>1 [::] (\s\<^sub>1@[\])@\s\<^sub>2" by (rule wt_instrs_extR) hence "\ is\<^sub>1,xt\<^sub>1 [::] \s\<^sub>1@\#\s\<^sub>2" by simp from this is\<^sub>2 s show ?thesis by (rule wt_instrs_ext) qed (*>*) corollary (in TC2) wt_instrs_app_last[trans]: assumes "\ is\<^sub>2,xt\<^sub>2 [::] \#\s\<^sub>2" "\ is\<^sub>1,xt\<^sub>1 [::] \s\<^sub>1" "last \s\<^sub>1 = \" "size \s\<^sub>1 = size is\<^sub>1+1" shows "\ is\<^sub>1@is\<^sub>2, xt\<^sub>1@shift (size is\<^sub>1) xt\<^sub>2 [::] \s\<^sub>1@\s\<^sub>2" (*<*) using assms proof(cases \s\<^sub>1 rule:rev_cases) case (snoc ys y) then show ?thesis using assms by(simp add:wt_instrs_app) qed simp (*>*) corollary (in TC2) wt_instrs_append_last[trans]: assumes wtis: "\ is,xt [::] \s" and wti: "P,T\<^sub>r,mxs,mpc,[] \ i,pc :: \s" and pc: "pc = size is" and mpc: "mpc = size \s" and is_\s: "size is + 1 < size \s" shows "\ is@[i],xt [::] \s" (*<*) proof - have pc_xt: "pc \ pcs xt" using wtis pc by (fastforce simp:wt_instrs_def) have "pcs xt \ {.. pc' < length is" "pc' < Suc (length is)" have "P,T\<^sub>r,mxs,length \s,xt \ i,pc' :: \s" using wt_instr_appLx[where xt = "[]",simplified,OF wti pc_xt] less_antisym[OF pc'] pc mpc by simp } ultimately show ?thesis using wtis is_\s by(clarsimp simp add:wt_instrs_def) qed (*>*) corollary (in TC2) wt_instrs_app2: "\ \ is\<^sub>2,xt\<^sub>2 [::] \'#\s\<^sub>2; \ is\<^sub>1,xt\<^sub>1 [::] \#\s\<^sub>1@[\']; xt' = xt\<^sub>1 @ shift (size is\<^sub>1) xt\<^sub>2; size \s\<^sub>1+1 = size is\<^sub>1 \ \ \ is\<^sub>1@is\<^sub>2,xt' [::] \#\s\<^sub>1@\'#\s\<^sub>2" (*<*)using wt_instrs_app[where ?\s\<^sub>1.0 = "\ # \s\<^sub>1"] by simp (*>*) corollary (in TC2) wt_instrs_app2_simp[trans,simp]: "\ \ is\<^sub>2,xt\<^sub>2 [::] \'#\s\<^sub>2; \ is\<^sub>1,xt\<^sub>1 [::] \#\s\<^sub>1@[\']; size \s\<^sub>1+1 = size is\<^sub>1 \ \ \ is\<^sub>1@is\<^sub>2, xt\<^sub>1@shift (size is\<^sub>1) xt\<^sub>2 [::] \#\s\<^sub>1@\'#\s\<^sub>2" (*<*)using wt_instrs_app[where ?\s\<^sub>1.0 = "\ # \s\<^sub>1"] by simp(*>*) corollary (in TC2) wt_instrs_Cons[simp]: "\ \s \ []; \ [i],[] [::] [\,\']; \ is,xt [::] \'#\s \ \ \ i#is,shift 1 xt [::] \#\'#\s" (*<*) using wt_instrs_app2[where ?is\<^sub>1.0 = "[i]" and ?\s\<^sub>1.0 = "[]" and ?is\<^sub>2.0 = "is" and ?xt\<^sub>1.0 = "[]"] by simp corollary (in TC2) wt_instrs_Cons2[trans]: assumes \s: "\ is,xt [::] \s" assumes i: "P,T\<^sub>r,mxs,mpc,[] \ i,0 :: \#\s" assumes mpc: "mpc = size \s + 1" shows "\ i#is,shift 1 xt [::] \#\s" (*<*) proof - from \s have "\s \ []" by (auto simp: wt_instrs_def) with mpc i have "\ [i],[] [::] [\]@\s" by (simp add: wt_instrs_def) with \s show ?thesis by (fastforce dest: wt_instrs_ext) qed (*>*) lemma (in TC2) wt_instrs_last_incr[trans]: assumes wtis: "\ is,xt [::] \s@[\]" and ss: "P \ \ \' \'" shows "\ is,xt [::] \s@[\']" (*<*) proof - let ?\s = "\s@[\]" and ?\s' = "\s@[\']" { fix pc assume pc: "pc< size is" let ?i = "is!pc" have app_pc: "app (is ! pc) P mxs T\<^sub>r pc (length ?\s) xt (\s ! pc)" using wtis pc by(clarsimp simp add:wt_instrs_def wt_instr_def) then have Apc\': "\pc' \'. (pc',\') \ set (eff ?i P pc xt (?\s!pc)) \ pc' < length ?\s" using wtis pc by(fastforce simp add:wt_instrs_def app_def) have Aepc\': "\pc' \'. (pc',\') \ set (eff ?i P pc xt (?\s!pc)) \ P \ \' \' ?\s!pc'" using wtis pc by(fastforce simp add:wt_instrs_def wt_instr_def) { fix pc1 \1 assume pc\1: "(pc1,\1) \ set (eff ?i P pc xt (?\s'!pc))" then have epc\': "(pc1,\1) \ set (eff ?i P pc xt (?\s!pc))" using wtis pc by(simp add:wt_instrs_def) have "P \ \1 \' ?\s'!pc1" proof(cases "pc1 < length \s") case True then show ?thesis using wtis pc pc\1 by(fastforce simp add:wt_instrs_def wt_instr_def) next case False then have "pc1 < length ?\s" using Apc\'[OF epc\'] by simp then have [simp]: "pc1 = size \s" using False by clarsimp have "P \ \1 \' \" using Aepc\'[OF epc\'] by simp then have "P \ \1 \' \'" by(rule sup_state_opt_trans[OF _ ss]) then show ?thesis by simp qed } then have "P,T\<^sub>r,mxs,size ?\s',xt \ is!pc,pc :: ?\s'" using wtis pc by(clarsimp simp add:wt_instrs_def wt_instr_def) } then show ?thesis using wtis by(simp add:wt_instrs_def) qed (*>*) lemma [iff]: "xcpt_app i P pc mxs [] \" (*<*)by (simp add: xcpt_app_def relevant_entries_def)(*>*) lemma [simp]: "xcpt_eff i P pc \ [] = []" (*<*)by (simp add: xcpt_eff_def relevant_entries_def)(*>*) lemma (in TC2) wt_New: "\ is_class P C; size ST < mxs \ \ \ [New C],[] [::] [ty\<^sub>i' ST E A, ty\<^sub>i' (Class C#ST) E A]" (*<*)by(simp add:wt_defs ty\<^sub>i'_def)(*>*) lemma (in TC2) wt_Cast: "is_class P C \ \ [Checkcast C],[] [::] [ty\<^sub>i' (Class D # ST) E A, ty\<^sub>i' (Class C # ST) E A]" (*<*)by(simp add: ty\<^sub>i'_def wt_defs)(*>*) lemma (in TC2) wt_Push: "\ size ST < mxs; typeof v = Some T \ \ \ [Push v],[] [::] [ty\<^sub>i' ST E A, ty\<^sub>i' (T#ST) E A]" (*<*)by(simp add: ty\<^sub>i'_def wt_defs)(*>*) lemma (in TC2) wt_Pop: "\ [Pop],[] [::] (ty\<^sub>i' (T#ST) E A # ty\<^sub>i' ST E A # \s)" (*<*)by(simp add: ty\<^sub>i'_def wt_defs)(*>*) lemma (in TC2) wt_CmpEq: "\ P \ T\<^sub>1 \ T\<^sub>2 \ P \ T\<^sub>2 \ T\<^sub>1\ \ \ [CmpEq],[] [::] [ty\<^sub>i' (T\<^sub>2 # T\<^sub>1 # ST) E A, ty\<^sub>i' (Boolean # ST) E A]" (*<*) by(auto simp:ty\<^sub>i'_def wt_defs elim!: refTE not_refTE) (*>*) lemma (in TC2) wt_IAdd: "\ [IAdd],[] [::] [ty\<^sub>i' (Integer#Integer#ST) E A, ty\<^sub>i' (Integer#ST) E A]" (*<*)by(simp add:ty\<^sub>i'_def wt_defs)(*>*) lemma (in TC2) wt_Load: "\ size ST < mxs; size E \ mxl; i \\ A; i < size E \ \ \ [Load i],[] [::] [ty\<^sub>i' ST E A, ty\<^sub>i' (E!i # ST) E A]" (*<*)by(auto simp add:ty\<^sub>i'_def wt_defs ty\<^sub>l_def hyperset_defs)(*>*) lemma (in TC2) wt_Store: "\ P \ T \ E!i; i < size E; size E \ mxl \ \ \ [Store i],[] [::] [ty\<^sub>i' (T#ST) E A, ty\<^sub>i' ST E (\{i}\ \ A)]" (*<*) by(auto simp:hyperset_defs nth_list_update ty\<^sub>i'_def wt_defs ty\<^sub>l_def intro:list_all2_all_nthI) (*>*) lemma (in TC2) wt_Get: "\ P \ C sees F:T in D \ \ \ [Getfield F D],[] [::] [ty\<^sub>i' (Class C # ST) E A, ty\<^sub>i' (T # ST) E A]" (*<*)by(auto simp: ty\<^sub>i'_def wt_defs dest: sees_field_idemp sees_field_decl_above)(*>*) lemma (in TC2) wt_Put: "\ P \ C sees F:T in D; P \ T' \ T \ \ \ [Putfield F D],[] [::] [ty\<^sub>i' (T' # Class C # ST) E A, ty\<^sub>i' ST E A]" (*<*)by(auto intro: sees_field_idemp sees_field_decl_above simp: ty\<^sub>i'_def wt_defs)(*>*) lemma (in TC2) wt_Throw: "\ [Throw],[] [::] [ty\<^sub>i' (Class C # ST) E A, \']" (*<*)by(auto simp: ty\<^sub>i'_def wt_defs)(*>*) lemma (in TC2) wt_IfFalse: "\ 2 \ i; nat i < size \s + 2; P \ ty\<^sub>i' ST E A \' \s ! nat(i - 2) \ \ \ [IfFalse i],[] [::] ty\<^sub>i' (Boolean # ST) E A # ty\<^sub>i' ST E A # \s" (*<*) by(simp add: ty\<^sub>i'_def wt_defs eval_nat_numeral nat_diff_distrib) (*>*) lemma wt_Goto: "\ 0 \ int pc + i; nat (int pc + i) < size \s; size \s \ mpc; P \ \s!pc \' \s ! nat (int pc + i) \ \ P,T,mxs,mpc,[] \ Goto i,pc :: \s" (*<*)by(clarsimp simp add: TC2.wt_defs)(*>*) lemma (in TC2) wt_Invoke: "\ size es = size Ts'; P \ C sees M: Ts\T = m in D; P \ Ts' [\] Ts \ \ \ [Invoke M (size es)],[] [::] [ty\<^sub>i' (rev Ts' @ Class C # ST) E A, ty\<^sub>i' (T#ST) E A]" (*<*)by(fastforce simp add: ty\<^sub>i'_def wt_defs)(*>*) corollary (in TC2) wt_instrs_app3[simp]: "\ \ is\<^sub>2,[] [::] (\' # \s\<^sub>2); \ is\<^sub>1,xt\<^sub>1 [::] \ # \s\<^sub>1 @ [\']; size \s\<^sub>1+1 = size is\<^sub>1\ \ \ (is\<^sub>1 @ is\<^sub>2),xt\<^sub>1 [::] \ # \s\<^sub>1 @ \' # \s\<^sub>2" (*<*)using wt_instrs_app2[where ?xt\<^sub>2.0 = "[]"] by (simp add:shift_def)(*>*) corollary (in TC2) wt_instrs_Cons3[simp]: "\ \s \ []; \ [i],[] [::] [\,\']; \ is,[] [::] \'#\s \ \ \ (i # is),[] [::] \ # \' # \s" (*<*) using wt_instrs_Cons[where ?xt = "[]"] by (simp add:shift_def) (*<*) declare nth_append[simp del] declare [[simproc del: list_to_set_comprehension]] (*>*) lemma (in TC2) wt_instrs_xapp[trans]: assumes wtis: "\ is\<^sub>1 @ is\<^sub>2, xt [::] \s\<^sub>1 @ ty\<^sub>i' (Class C # ST) E A # \s\<^sub>2" and P_\s\<^sub>1: "\\ \ set \s\<^sub>1. \ST' LT'. \ = Some(ST',LT') \ size ST \ size ST' \ P \ Some (drop (size ST' - size ST) ST',LT') \' ty\<^sub>i' ST E A" and is_\s: "size is\<^sub>1 = size \s\<^sub>1" and PC: "is_class P C" and ST_mxs: "size ST < mxs" shows "\ is\<^sub>1 @ is\<^sub>2, xt @ [(0,size is\<^sub>1 - 1,C,size is\<^sub>1,size ST)] [::] \s\<^sub>1 @ ty\<^sub>i' (Class C # ST) E A # \s\<^sub>2" (*<*)(is "\ ?is, xt@[?xte] [::] ?\s" is "\ ?is, ?xt' [::] ?\s") proof - let ?is = "is\<^sub>1 @ is\<^sub>2" and ?\s = "\s\<^sub>1 @ ty\<^sub>i' (Class C # ST) E A # \s\<^sub>2" let ?xte = "(0,size is\<^sub>1 - 1,C,size is\<^sub>1,size ST)" let ?xt' = "xt @ [?xte]" have P_\s\<^sub>1': "\\. \ \ set \s\<^sub>1 \ (\ST' LT'. \ = Some(ST',LT') \ size ST \ size ST' \ P \ Some (drop (size ST' - size ST) ST',LT') \' ty\<^sub>i' ST E A)" using P_\s\<^sub>1 by fast have "size ?is < size ?\s" and "pcs ?xt' \ {0..pc. pc< size ?is \ P,T\<^sub>r,mxs,size ?\s,xt \ ?is!pc,pc :: ?\s" using wtis by(simp_all add:wt_instrs_def) moreover { fix pc let ?mpc = "size ?\s" and ?i = "?is!pc" and ?t = "?\s!pc" assume "pc< size ?is" then have wti: "P,T\<^sub>r,mxs,?mpc,xt \ ?i,pc :: ?\s" by(rule P_pc) then have app: "app ?i P mxs T\<^sub>r pc ?mpc xt ?t" and eff_ss: "\pc' \'. (pc',\') \ set (eff ?i P pc xt (?\s!pc)) \ P \ \' \' ?\s!pc'" by(fastforce simp add: wt_instr_def)+ have "app ?i P mxs T\<^sub>r pc ?mpc ?xt' ?t \ (\(pc',\') \ set (eff ?i P pc ?xt' ?t). P \ \' \' ?\s!pc')" proof (cases ?t) case (Some \) obtain ST' LT' where \: "\ = (ST', LT')" by(cases \) simp have app\<^sub>i: "app\<^sub>i (?i,P,pc,mxs,T\<^sub>r,\)" and xcpt_app: "xcpt_app ?i P pc mxs xt \" and eff_pc: "\pc' \'. (pc',\') \ set (eff ?i P pc xt ?t) \ pc' < ?mpc" using Some app by(fastforce simp add: app_def)+ have "xcpt_app ?i P pc mxs ?xt' \" proof(cases "pc < length \s\<^sub>1 - 1") case False then show ?thesis using Some \ is_\s xcpt_app by (clarsimp simp: xcpt_app_def relevant_entries_def is_relevant_entry_def) next case True then have True': "pc < length \s\<^sub>1" by simp then have "\s\<^sub>1 ! pc = ?t" by(fastforce simp: nth_append) moreover have \s\<^sub>1_pc: "\s\<^sub>1 ! pc \ set \s\<^sub>1" by(rule nth_mem[OF True']) ultimately show ?thesis using Some \ PC ST_mxs xcpt_app P_\s\<^sub>1'[OF \s\<^sub>1_pc] by (simp add: xcpt_app_def relevant_entries_def) qed moreover { fix pc' \' assume efft: "(pc',\') \ set (eff ?i P pc ?xt' ?t)" have "pc' < ?mpc \ P \ \' \' ?\s!pc'" (is "?P1 \ ?P2") proof(cases "(pc',\') \ set (eff ?i P pc xt ?t)") case True have ?P1 using True by(rule eff_pc) moreover have ?P2 using True by(rule eff_ss) ultimately show ?thesis by simp next case False then have xte: "(pc',\') \ set (xcpt_eff ?i P pc \ [?xte])" using efft Some by(clarsimp simp: eff_def) then have ?P1 using Some \ is_\s by(clarsimp simp: xcpt_eff_def relevant_entries_def split: if_split_asm) moreover have ?P2 proof(cases "pc < length \s\<^sub>1 - 1") case False': False then show ?thesis using False Some \ xte is_\s by(simp add: xcpt_eff_def relevant_entries_def is_relevant_entry_def) next case True then have True': "pc < length \s\<^sub>1" by simp have \s\<^sub>1_pc: "\s\<^sub>1 ! pc \ set \s\<^sub>1" by(rule nth_mem[OF True']) have "P \ \(Class C # drop (length ST' - length ST) ST', LT')\ \' ty\<^sub>i' (Class C # ST) E A" using True' Some \ P_\s\<^sub>1'[OF \s\<^sub>1_pc] by (fastforce simp: nth_append ty\<^sub>i'_def) then show ?thesis using \ xte is_\s by(simp add: xcpt_eff_def relevant_entries_def split: if_split_asm) qed ultimately show ?thesis by simp qed } ultimately show ?thesis using Some app\<^sub>i by(fastforce simp add: app_def) qed simp then have "P,T\<^sub>r,mxs,size ?\s,?xt' \ ?is!pc,pc :: ?\s" by(simp add: wt_instr_def) } ultimately show ?thesis by(simp add:wt_instrs_def) qed declare [[simproc add: list_to_set_comprehension]] declare nth_append[simp] (*>*) lemma drop_Cons_Suc: "\xs. drop n xs = y#ys \ drop (Suc n) xs = ys" proof(induct n) case (Suc n) then show ?case by(simp add: drop_Suc) qed simp lemma drop_mess: assumes "Suc (length xs\<^sub>0) \ length xs" and "drop (length xs - Suc (length xs\<^sub>0)) xs = x # xs\<^sub>0" shows "drop (length xs - length xs\<^sub>0) xs = xs\<^sub>0" using assms proof(cases xs) case (Cons a list) then show ?thesis using assms proof(cases "length list - length xs\<^sub>0") case Suc then show ?thesis using Cons assms by (simp add: Suc_diff_le drop_Cons_Suc) qed simp qed simp (*<*) declare (in TC0) after_def[simp] pair_eq_ty\<^sub>i'_conv[simp] (*>*) lemma (in TC1) compT_ST_prefix: "\E A ST\<^sub>0. \(ST,LT)\ \ set(compT E A ST\<^sub>0 e) \ size ST\<^sub>0 \ size ST \ drop (size ST - size ST\<^sub>0) ST = ST\<^sub>0" and "\E A ST\<^sub>0. \(ST,LT)\ \ set(compTs E A ST\<^sub>0 es) \ size ST\<^sub>0 \ size ST \ drop (size ST - size ST\<^sub>0) ST = ST\<^sub>0" (*<*) proof(induct e and es rule: compT.induct compTs.induct) case (FAss e\<^sub>1 F D e\<^sub>2) moreover { let ?ST\<^sub>0 = "ty E e\<^sub>1 # ST\<^sub>0" fix A assume "\(ST, LT)\ \ set (compT E A ?ST\<^sub>0 e\<^sub>2)" with FAss have "length ?ST\<^sub>0 \ length ST \ drop (size ST - size ?ST\<^sub>0) ST = ?ST\<^sub>0" by blast hence ?case by (clarsimp simp add: drop_mess) } ultimately show ?case by auto next case TryCatch thus ?case by auto next case Block thus ?case by auto next case Seq thus ?case by auto next case While thus ?case by auto next case Cond thus ?case by auto next case (Call e M es) moreover { let ?ST\<^sub>0 = "ty E e # ST\<^sub>0" fix A assume "\(ST, LT)\ \ set (compTs E A ?ST\<^sub>0 es)" with Call have "length ?ST\<^sub>0 \ length ST \ drop (size ST - size ?ST\<^sub>0) ST = ?ST\<^sub>0" by blast hence ?case by (clarsimp simp add: drop_mess) } ultimately show ?case by auto next case (Cons_exp e es) moreover { let ?ST\<^sub>0 = "ty E e # ST\<^sub>0" fix A assume "\(ST, LT)\ \ set (compTs E A ?ST\<^sub>0 es)" with Cons_exp have "length ?ST\<^sub>0 \ length ST \ drop (size ST - size ?ST\<^sub>0) ST = ?ST\<^sub>0" by blast hence ?case by (clarsimp simp add: drop_mess) } ultimately show ?case by auto next case (BinOp e\<^sub>1 bop e\<^sub>2) moreover { let ?ST\<^sub>0 = "ty E e\<^sub>1 # ST\<^sub>0" fix A assume "\(ST, LT)\ \ set (compT E A ?ST\<^sub>0 e\<^sub>2)" with BinOp have "length ?ST\<^sub>0 \ length ST \ drop (size ST - size ?ST\<^sub>0) ST = ?ST\<^sub>0" by blast hence ?case by (clarsimp simp add: drop_mess) } ultimately show ?case by auto next case new thus ?case by auto next case Val thus ?case by auto next case Cast thus ?case by auto next case Var thus ?case by auto next case LAss thus ?case by auto next case throw thus ?case by auto next case FAcc thus ?case by auto next case Nil_exp thus ?case by auto qed declare (in TC0) after_def[simp del] pair_eq_ty\<^sub>i'_conv[simp del] (*>*) (* FIXME *) lemma fun_of_simp [simp]: "fun_of S x y = ((x,y) \ S)" (*<*) by (simp add: fun_of_def)(*>*) theorem (in TC2) compT_wt_instrs: "\E T A ST. \ P,E \\<^sub>1 e :: T; \ e A; \ e (size E); size ST + max_stack e \ mxs; size E + max_vars e \ mxl \ \ \ compE\<^sub>2 e, compxE\<^sub>2 e 0 (size ST) [::] ty\<^sub>i' ST E A # compT E A ST e @ [after E A ST e]" (*<*)(is "\E T A ST. PROP ?P e E T A ST")(*>*) and "\E Ts A ST. \ P,E \\<^sub>1 es[::]Ts; \s es A; \s es (size E); size ST + max_stacks es \ mxs; size E + max_varss es \ mxl \ \ let \s = ty\<^sub>i' ST E A # compTs E A ST es in \ compEs\<^sub>2 es,compxEs\<^sub>2 es 0 (size ST) [::] \s \ last \s = ty\<^sub>i' (rev Ts @ ST) E (A \ \s es)" (*<*) (is "\E Ts A ST. PROP ?Ps es E Ts A ST") proof(induct e and es rule: compxE\<^sub>2.induct compxEs\<^sub>2.induct) case (TryCatch e\<^sub>1 C i e\<^sub>2) hence [simp]: "i = size E" by simp have wt\<^sub>1: "P,E \\<^sub>1 e\<^sub>1 :: T" and wt\<^sub>2: "P,E@[Class C] \\<^sub>1 e\<^sub>2 :: T" and "class": "is_class P C" using TryCatch by auto let ?A\<^sub>1 = "A \ \ e\<^sub>1" let ?A\<^sub>i = "A \ \{i}\" let ?E\<^sub>i = "E @ [Class C]" let ?\ = "ty\<^sub>i' ST E A" let ?\s\<^sub>1 = "compT E A ST e\<^sub>1" let ?\\<^sub>1 = "ty\<^sub>i' (T#ST) E ?A\<^sub>1" let ?\\<^sub>2 = "ty\<^sub>i' (Class C#ST) E A" let ?\\<^sub>3 = "ty\<^sub>i' ST ?E\<^sub>i ?A\<^sub>i" let ?\s\<^sub>2 = "compT ?E\<^sub>i ?A\<^sub>i ST e\<^sub>2" let ?\\<^sub>2' = "ty\<^sub>i' (T#ST) ?E\<^sub>i (?A\<^sub>i \ \ e\<^sub>2)" let ?\' = "ty\<^sub>i' (T#ST) E (A \ \ e\<^sub>1 \ (\ e\<^sub>2 \ i))" let ?go = "Goto (int(size(compE\<^sub>2 e\<^sub>2)) + 2)" have "PROP ?P e\<^sub>2 ?E\<^sub>i T ?A\<^sub>i ST" by fact hence "\ compE\<^sub>2 e\<^sub>2,compxE\<^sub>2 e\<^sub>2 0 (size ST) [::] (?\\<^sub>3 # ?\s\<^sub>2) @ [?\\<^sub>2']" using TryCatch.prems by(auto simp:after_def) also have "?A\<^sub>i \ \ e\<^sub>2 = (A \ \ e\<^sub>2) \ \{size E}\" by(fastforce simp:hyperset_defs) also have "P \ ty\<^sub>i' (T#ST) ?E\<^sub>i \ \' ty\<^sub>i' (T#ST) E (A \ \ e\<^sub>2)" by(simp add:hyperset_defs ty\<^sub>l_incr ty\<^sub>i'_def) also have "P \ \ \' ty\<^sub>i' (T#ST) E (A \ \ e\<^sub>1 \ (\ e\<^sub>2 \ i))" by(auto intro!: ty\<^sub>l_antimono simp:hyperset_defs ty\<^sub>i'_def) also have "(?\\<^sub>3 # ?\s\<^sub>2) @ [?\'] = ?\\<^sub>3 # ?\s\<^sub>2 @ [?\']" by simp also have "\ [Store i],[] [::] ?\\<^sub>2 # [] @ [?\\<^sub>3]" using TryCatch.prems by(auto simp:nth_list_update wt_defs ty\<^sub>i'_def ty\<^sub>l_def list_all2_conv_all_nth hyperset_defs) also have "[] @ (?\\<^sub>3 # ?\s\<^sub>2 @ [?\']) = (?\\<^sub>3 # ?\s\<^sub>2 @ [?\'])" by simp also have "P,T\<^sub>r,mxs,size(compE\<^sub>2 e\<^sub>2)+3,[] \ ?go,0 :: ?\\<^sub>1#?\\<^sub>2#?\\<^sub>3#?\s\<^sub>2 @ [?\']" by (auto simp: hyperset_defs ty\<^sub>i'_def wt_defs nth_Cons nat_add_distrib fun_of_def intro: ty\<^sub>l_antimono list_all2_refl split:nat.split) also have "\ compE\<^sub>2 e\<^sub>1,compxE\<^sub>2 e\<^sub>1 0 (size ST) [::] ?\ # ?\s\<^sub>1 @ [?\\<^sub>1]" using TryCatch by(auto simp:after_def) also have "?\ # ?\s\<^sub>1 @ ?\\<^sub>1 # ?\\<^sub>2 # ?\\<^sub>3 # ?\s\<^sub>2 @ [?\'] = (?\ # ?\s\<^sub>1 @ [?\\<^sub>1]) @ ?\\<^sub>2 # ?\\<^sub>3 # ?\s\<^sub>2 @ [?\']" by simp also have "compE\<^sub>2 e\<^sub>1 @ ?go # [Store i] @ compE\<^sub>2 e\<^sub>2 = (compE\<^sub>2 e\<^sub>1 @ [?go]) @ (Store i # compE\<^sub>2 e\<^sub>2)" by simp also let "?Q \" = "\ST' LT'. \ = \(ST', LT')\ \ size ST \ size ST' \ P \ Some (drop (size ST' - size ST) ST',LT') \' ty\<^sub>i' ST E A" { have "?Q (ty\<^sub>i' ST E A)" by (clarsimp simp add: ty\<^sub>i'_def) moreover have "?Q (ty\<^sub>i' (T # ST) E ?A\<^sub>1)" by (fastforce simp add: ty\<^sub>i'_def hyperset_defs intro!: ty\<^sub>l_antimono) moreover have "\\. \ \ set (compT E A ST e\<^sub>1) \ ?Q \" using TryCatch.prems by clarsimp (frule compT_ST_prefix, fastforce dest!: compT_LT_prefix simp add: ty\<^sub>i'_def) ultimately have "\\\set (ty\<^sub>i' ST E A # compT E A ST e\<^sub>1 @ [ty\<^sub>i' (T # ST) E ?A\<^sub>1]). ?Q \" by auto } also from TryCatch.prems max_stack1[of e\<^sub>1] have "size ST + 1 \ mxs" by auto ultimately show ?case using wt\<^sub>1 wt\<^sub>2 TryCatch.prems "class" by (simp add:after_def) next case new thus ?case by(auto simp add:after_def wt_New) next case (BinOp e\<^sub>1 bop e\<^sub>2) let ?op = "case bop of Eq \ [CmpEq] | Add \ [IAdd]" have T: "P,E \\<^sub>1 e\<^sub>1 \bop\ e\<^sub>2 :: T" by fact then obtain T\<^sub>1 T\<^sub>2 where T\<^sub>1: "P,E \\<^sub>1 e\<^sub>1 :: T\<^sub>1" and T\<^sub>2: "P,E \\<^sub>1 e\<^sub>2 :: T\<^sub>2" and bopT: "case bop of Eq \ (P \ T\<^sub>1 \ T\<^sub>2 \ P \ T\<^sub>2 \ T\<^sub>1) \ T = Boolean | Add \ T\<^sub>1 = Integer \ T\<^sub>2 = Integer \ T = Integer" by auto let ?A\<^sub>1 = "A \ \ e\<^sub>1" let ?A\<^sub>2 = "?A\<^sub>1 \ \ e\<^sub>2" let ?\ = "ty\<^sub>i' ST E A" let ?\s\<^sub>1 = "compT E A ST e\<^sub>1" let ?\\<^sub>1 = "ty\<^sub>i' (T\<^sub>1#ST) E ?A\<^sub>1" let ?\s\<^sub>2 = "compT E ?A\<^sub>1 (T\<^sub>1#ST) e\<^sub>2" let ?\\<^sub>2 = "ty\<^sub>i' (T\<^sub>2#T\<^sub>1#ST) E ?A\<^sub>2" let ?\' = "ty\<^sub>i' (T#ST) E ?A\<^sub>2" from bopT have "\ ?op,[] [::] [?\\<^sub>2,?\']" by (cases bop) (auto simp add: wt_CmpEq wt_IAdd) also have "PROP ?P e\<^sub>2 E T\<^sub>2 ?A\<^sub>1 (T\<^sub>1#ST)" by fact with BinOp.prems T\<^sub>2 have "\ compE\<^sub>2 e\<^sub>2, compxE\<^sub>2 e\<^sub>2 0 (size (T\<^sub>1#ST)) [::] ?\\<^sub>1#?\s\<^sub>2@[?\\<^sub>2]" by (auto simp: after_def) also from BinOp T\<^sub>1 have "\ compE\<^sub>2 e\<^sub>1, compxE\<^sub>2 e\<^sub>1 0 (size ST) [::] ?\#?\s\<^sub>1@[?\\<^sub>1]" by (auto simp: after_def) finally show ?case using T T\<^sub>1 T\<^sub>2 by (simp add: after_def hyperUn_assoc) next case (Cons_exp e es) have "P,E \\<^sub>1 e # es [::] Ts" by fact then obtain T\<^sub>e Ts' where T\<^sub>e: "P,E \\<^sub>1 e :: T\<^sub>e" and Ts': "P,E \\<^sub>1 es [::] Ts'" and Ts: "Ts = T\<^sub>e#Ts'" by auto let ?A\<^sub>e = "A \ \ e" let ?\ = "ty\<^sub>i' ST E A" let ?\s\<^sub>e = "compT E A ST e" let ?\\<^sub>e = "ty\<^sub>i' (T\<^sub>e#ST) E ?A\<^sub>e" let ?\s' = "compTs E ?A\<^sub>e (T\<^sub>e#ST) es" let ?\s = "?\ # ?\s\<^sub>e @ (?\\<^sub>e # ?\s')" have Ps: "PROP ?Ps es E Ts' ?A\<^sub>e (T\<^sub>e#ST)" by fact with Cons_exp.prems T\<^sub>e Ts' have "\ compEs\<^sub>2 es, compxEs\<^sub>2 es 0 (size (T\<^sub>e#ST)) [::] ?\\<^sub>e#?\s'" by (simp add: after_def) also from Cons_exp T\<^sub>e have "\ compE\<^sub>2 e, compxE\<^sub>2 e 0 (size ST) [::] ?\#?\s\<^sub>e@[?\\<^sub>e]" by (auto simp: after_def) moreover from Ps Cons_exp.prems T\<^sub>e Ts' Ts have "last ?\s = ty\<^sub>i' (rev Ts@ST) E (?A\<^sub>e \ \s es)" by simp ultimately show ?case using T\<^sub>e by (simp add: after_def hyperUn_assoc) next case (FAss e\<^sub>1 F D e\<^sub>2) hence Void: "P,E \\<^sub>1 e\<^sub>1\F{D} := e\<^sub>2 :: Void" by auto then obtain C T T' where C: "P,E \\<^sub>1 e\<^sub>1 :: Class C" and sees: "P \ C sees F:T in D" and T': "P,E \\<^sub>1 e\<^sub>2 :: T'" and T'_T: "P \ T' \ T" by auto let ?A\<^sub>1 = "A \ \ e\<^sub>1" let ?A\<^sub>2 = "?A\<^sub>1 \ \ e\<^sub>2" let ?\ = "ty\<^sub>i' ST E A" let ?\s\<^sub>1 = "compT E A ST e\<^sub>1" let ?\\<^sub>1 = "ty\<^sub>i' (Class C#ST) E ?A\<^sub>1" let ?\s\<^sub>2 = "compT E ?A\<^sub>1 (Class C#ST) e\<^sub>2" let ?\\<^sub>2 = "ty\<^sub>i' (T'#Class C#ST) E ?A\<^sub>2" let ?\\<^sub>3 = "ty\<^sub>i' ST E ?A\<^sub>2" let ?\' = "ty\<^sub>i' (Void#ST) E ?A\<^sub>2" from FAss.prems sees T'_T have "\ [Putfield F D,Push Unit],[] [::] [?\\<^sub>2,?\\<^sub>3,?\']" by (fastforce simp add: wt_Push wt_Put) also have "PROP ?P e\<^sub>2 E T' ?A\<^sub>1 (Class C#ST)" by fact with FAss.prems T' have "\ compE\<^sub>2 e\<^sub>2, compxE\<^sub>2 e\<^sub>2 0 (size ST+1) [::] ?\\<^sub>1#?\s\<^sub>2@[?\\<^sub>2]" by (auto simp add: after_def hyperUn_assoc) also from FAss C have "\ compE\<^sub>2 e\<^sub>1, compxE\<^sub>2 e\<^sub>1 0 (size ST) [::] ?\#?\s\<^sub>1@[?\\<^sub>1]" by (auto simp add: after_def) finally show ?case using Void C T' by (simp add: after_def hyperUn_assoc) next case Val thus ?case by(auto simp:after_def wt_Push) next case Cast thus ?case by (auto simp:after_def wt_Cast) next case (Block i T\<^sub>i e) let ?\s = "ty\<^sub>i' ST E A # compT (E @ [T\<^sub>i]) (A\i) ST e" have IH: "PROP ?P e (E@[T\<^sub>i]) T (A\i) ST" by fact hence "\ compE\<^sub>2 e, compxE\<^sub>2 e 0 (size ST) [::] ?\s @ [ty\<^sub>i' (T#ST) (E@[T\<^sub>i]) (A\(size E) \ \ e)]" using Block.prems by (auto simp add: after_def) also have "P \ ty\<^sub>i' (T # ST) (E@[T\<^sub>i]) (A \ size E \ \ e) \' ty\<^sub>i' (T # ST) (E@[T\<^sub>i]) ((A \ \ e) \ size E)" by(auto simp add:hyperset_defs intro: ty\<^sub>i'_antimono) also have "\ = ty\<^sub>i' (T # ST) E (A \ \ e)" by simp also have "P \ \ \' ty\<^sub>i' (T # ST) E (A \ (\ e \ i))" by(auto simp add:hyperset_defs intro: ty\<^sub>i'_antimono) finally show ?case using Block.prems by(simp add: after_def) next case Var thus ?case by(auto simp:after_def wt_Load) next case FAcc thus ?case by(auto simp:after_def wt_Get) next case (LAss i e) thus ?case using max_stack1[of e] by(auto simp: hyper_insert_comm after_def wt_Store wt_Push) next case Nil_exp thus ?case by auto next case throw thus ?case by(auto simp add: after_def wt_Throw) next case (While e c) obtain Tc where wte: "P,E \\<^sub>1 e :: Boolean" and wtc: "P,E \\<^sub>1 c :: Tc" and [simp]: "T = Void" using While by auto have [simp]: "ty E (while (e) c) = Void" using While by simp let ?A\<^sub>0 = "A \ \ e" let ?A\<^sub>1 = "?A\<^sub>0 \ \ c" let ?\ = "ty\<^sub>i' ST E A" let ?\s\<^sub>e = "compT E A ST e" let ?\\<^sub>e = "ty\<^sub>i' (Boolean#ST) E ?A\<^sub>0" let ?\\<^sub>1 = "ty\<^sub>i' ST E ?A\<^sub>0" let ?\s\<^sub>c = "compT E ?A\<^sub>0 ST c" let ?\\<^sub>c = "ty\<^sub>i' (Tc#ST) E ?A\<^sub>1" let ?\\<^sub>2 = "ty\<^sub>i' ST E ?A\<^sub>1" let ?\' = "ty\<^sub>i' (Void#ST) E ?A\<^sub>0" let ?\s = "(?\ # ?\s\<^sub>e @ [?\\<^sub>e]) @ ?\\<^sub>1 # ?\s\<^sub>c @ [?\\<^sub>c, ?\\<^sub>2, ?\\<^sub>1, ?\']" have "\ [],[] [::] [] @ ?\s" by(simp add:wt_instrs_def) also have "PROP ?P e E Boolean A ST" by fact hence "\ compE\<^sub>2 e,compxE\<^sub>2 e 0 (size ST) [::] ?\ # ?\s\<^sub>e @ [?\\<^sub>e]" using While.prems by (auto simp:after_def) also have "[] @ ?\s = (?\ # ?\s\<^sub>e) @ ?\\<^sub>e # ?\\<^sub>1 # ?\s\<^sub>c @ [?\\<^sub>c,?\\<^sub>2,?\\<^sub>1,?\']" by simp also let ?n\<^sub>e = "size(compE\<^sub>2 e)" let ?n\<^sub>c = "size(compE\<^sub>2 c)" let ?if = "IfFalse (int ?n\<^sub>c + 3)" - have "\ [?if],[] [::] ?\\<^sub>e # ?\\<^sub>1 # ?\s\<^sub>c @ [?\\<^sub>c, ?\\<^sub>2, ?\\<^sub>1, ?\']" - by(simp add: wt_instr_Cons wt_instr_append wt_IfFalse - nat_add_distrib split: nat_diff_split) + have "\ [?if],[] [::] ?\\<^sub>e # ?\\<^sub>1 # ?\s\<^sub>c @ [?\\<^sub>c, ?\\<^sub>2, ?\\<^sub>1, ?\']" + proof-{ + show ?thesis + apply (rule wt_IfFalse) + apply simp + apply simp + apply (subgoal_tac "length (compE\<^sub>2 c) = length (compT E (A \ \ e) ST c) + 1" + "nat (int (length (compE\<^sub>2 c)) + 3 - 2) > length (compT E (A \ \ e) ST c)") + apply simp+ + done + } + qed also have "(?\ # ?\s\<^sub>e) @ (?\\<^sub>e # ?\\<^sub>1 # ?\s\<^sub>c @ [?\\<^sub>c, ?\\<^sub>2, ?\\<^sub>1, ?\']) = ?\s" by simp also have "PROP ?P c E Tc ?A\<^sub>0 ST" by fact hence "\ compE\<^sub>2 c,compxE\<^sub>2 c 0 (size ST) [::] ?\\<^sub>1 # ?\s\<^sub>c @ [?\\<^sub>c]" using While.prems wtc by (auto simp:after_def) also have "?\s = (?\ # ?\s\<^sub>e @ [?\\<^sub>e,?\\<^sub>1] @ ?\s\<^sub>c) @ [?\\<^sub>c,?\\<^sub>2,?\\<^sub>1,?\']" by simp also have "\ [Pop],[] [::] [?\\<^sub>c, ?\\<^sub>2]" by(simp add:wt_Pop) also have "(?\ # ?\s\<^sub>e @ [?\\<^sub>e,?\\<^sub>1] @ ?\s\<^sub>c) @ [?\\<^sub>c,?\\<^sub>2,?\\<^sub>1,?\'] = ?\s" by simp also let ?go = "Goto (-int(?n\<^sub>c+?n\<^sub>e+2))" have "P \ ?\\<^sub>2 \' ?\" by(fastforce intro: ty\<^sub>i'_antimono simp: hyperset_defs) - hence "P,T\<^sub>r,mxs,size ?\s,[] \ ?go,?n\<^sub>e+?n\<^sub>c+2 :: ?\s" - by(simp add: wt_Goto split: nat_diff_split) + hence "P,T\<^sub>r,mxs,size ?\s,[] \ ?go,?n\<^sub>e+?n\<^sub>c+2 :: ?\s" + proof-{ + let ?t1 = "ty\<^sub>i' ST E A" let ?t2 = "ty\<^sub>i' (Boolean # ST) E (A \ \ e)" let ?t3 = "ty\<^sub>i' ST E (A \ \ e)" + let ?t4 = "[ty\<^sub>i' (Tc # ST) E (A \ \ e \ \ c), ty\<^sub>i' ST E (A \ \ e \ \ c), ty\<^sub>i' ST E (A \ \ e), ty\<^sub>i' (Void # ST) E (A \ \ e)]" + let ?c = " compT E (A \ \ e) ST c" let ?e = "compT E A ST e" + assume ass: "P \ ty\<^sub>i' ST E (A \ \ e \ \ c) \' ?t1" + show ?thesis + apply (rule wt_Goto) + apply simp + apply simp + apply simp + proof-{ + let ?s1 = "((?t1 # ?e @ [?t2]) @ ?t3 # ?c)" + have "length (compE\<^sub>2 e) + length (compE\<^sub>2 c) + 2 > length ?s1" by auto + hence "(?s1 @ ?t4) ! (length (compE\<^sub>2 e) + length (compE\<^sub>2 c) + 2) = + ?t4 ! (length (compE\<^sub>2 e) + length (compE\<^sub>2 c) + 2 - length ?s1)" + by (auto simp only:nth_append split: if_splits) + hence "(?s1 @ ?t4) ! (length (compE\<^sub>2 e) + length (compE\<^sub>2 c) + 2) = ty\<^sub>i' ST E (A \ \ e \ \ c)" + using compT_sizes by auto + hence "P \ (?s1 @ ?t4) ! (length (compE\<^sub>2 e) + length (compE\<^sub>2 c) + 2) \' ty\<^sub>i' ST E A" + using ass by simp + thus "P \ ((?t1 # ?e @ [?t2]) @ ?t3 # ?c @ ?t4) ! (length (compE\<^sub>2 e) + length (compE\<^sub>2 c) + 2) \' + ((?t1 # ?e @ [?t2]) @ ?t3 # ?c @ ?t4) ! nat (int (length (compE\<^sub>2 e) + length (compE\<^sub>2 c) + 2) + + - int (length (compE\<^sub>2 c) + length (compE\<^sub>2 e) + 2))" + by simp + }qed + }qed also have "?\s = (?\ # ?\s\<^sub>e @ [?\\<^sub>e,?\\<^sub>1] @ ?\s\<^sub>c @ [?\\<^sub>c, ?\\<^sub>2]) @ [?\\<^sub>1, ?\']" by simp also have "\ [Push Unit],[] [::] [?\\<^sub>1,?\']" using While.prems max_stack1[of c] by(auto simp add:wt_Push) finally show ?case using wtc wte by (simp add:after_def) next case (Cond e e\<^sub>1 e\<^sub>2) obtain T\<^sub>1 T\<^sub>2 where wte: "P,E \\<^sub>1 e :: Boolean" and wt\<^sub>1: "P,E \\<^sub>1 e\<^sub>1 :: T\<^sub>1" and wt\<^sub>2: "P,E \\<^sub>1 e\<^sub>2 :: T\<^sub>2" and sub\<^sub>1: "P \ T\<^sub>1 \ T" and sub\<^sub>2: "P \ T\<^sub>2 \ T" using Cond by auto have [simp]: "ty E (if (e) e\<^sub>1 else e\<^sub>2) = T" using Cond by simp let ?A\<^sub>0 = "A \ \ e" let ?A\<^sub>2 = "?A\<^sub>0 \ \ e\<^sub>2" let ?A\<^sub>1 = "?A\<^sub>0 \ \ e\<^sub>1" let ?A' = "?A\<^sub>0 \ \ e\<^sub>1 \ \ e\<^sub>2" let ?\\<^sub>2 = "ty\<^sub>i' ST E ?A\<^sub>0" let ?\' = "ty\<^sub>i' (T#ST) E ?A'" let ?\s\<^sub>2 = "compT E ?A\<^sub>0 ST e\<^sub>2" have "PROP ?P e\<^sub>2 E T\<^sub>2 ?A\<^sub>0 ST" by fact hence "\ compE\<^sub>2 e\<^sub>2, compxE\<^sub>2 e\<^sub>2 0 (size ST) [::] (?\\<^sub>2#?\s\<^sub>2) @ [ty\<^sub>i' (T\<^sub>2#ST) E ?A\<^sub>2]" using Cond.prems wt\<^sub>2 by(auto simp add:after_def) also have "P \ ty\<^sub>i' (T\<^sub>2#ST) E ?A\<^sub>2 \' ?\'" using sub\<^sub>2 by(auto simp add: hyperset_defs ty\<^sub>i'_def intro!: ty\<^sub>l_antimono) also let ?\\<^sub>3 = "ty\<^sub>i' (T\<^sub>1 # ST) E ?A\<^sub>1" let ?g\<^sub>2 = "Goto(int (size (compE\<^sub>2 e\<^sub>2) + 1))" from sub\<^sub>1 have "P,T\<^sub>r,mxs,size(compE\<^sub>2 e\<^sub>2)+2,[] \ ?g\<^sub>2,0 :: ?\\<^sub>3#(?\\<^sub>2#?\s\<^sub>2)@[?\']" by(auto simp: hyperset_defs wt_defs nth_Cons ty\<^sub>i'_def split:nat.split intro!: ty\<^sub>l_antimono) also let ?\s\<^sub>1 = "compT E ?A\<^sub>0 ST e\<^sub>1" have "PROP ?P e\<^sub>1 E T\<^sub>1 ?A\<^sub>0 ST" by fact hence "\ compE\<^sub>2 e\<^sub>1,compxE\<^sub>2 e\<^sub>1 0 (size ST) [::] ?\\<^sub>2 # ?\s\<^sub>1 @ [?\\<^sub>3]" using Cond.prems wt\<^sub>1 by(auto simp add:after_def) also let ?\s\<^sub>1\<^sub>2 = "?\\<^sub>2 # ?\s\<^sub>1 @ ?\\<^sub>3 # (?\\<^sub>2 # ?\s\<^sub>2) @ [?\']" let ?\\<^sub>1 = "ty\<^sub>i' (Boolean#ST) E ?A\<^sub>0" let ?g\<^sub>1 = "IfFalse(int (size (compE\<^sub>2 e\<^sub>1) + 2))" let ?code = "compE\<^sub>2 e\<^sub>1 @ ?g\<^sub>2 # compE\<^sub>2 e\<^sub>2" - have "\ [?g\<^sub>1],[] [::] [?\\<^sub>1] @ ?\s\<^sub>1\<^sub>2" - by(simp add: wt_IfFalse nat_add_distrib split:nat_diff_split) + have "\ [?g\<^sub>1],[] [::] [?\\<^sub>1] @ ?\s\<^sub>1\<^sub>2" + proof-{ + show ?thesis + apply clarsimp + apply (rule wt_IfFalse) + apply (simp only:nat_add_distrib) + apply simp + apply (auto simp only:nth_append split:if_splits) + apply (simp only:compT_sizes) + apply simp + done + }qed also (wt_instrs_ext2) have "[?\\<^sub>1] @ ?\s\<^sub>1\<^sub>2 = ?\\<^sub>1 # ?\s\<^sub>1\<^sub>2" by simp also let ?\ = "ty\<^sub>i' ST E A" have "PROP ?P e E Boolean A ST" by fact hence "\ compE\<^sub>2 e, compxE\<^sub>2 e 0 (size ST) [::] ?\ # compT E A ST e @ [?\\<^sub>1]" using Cond.prems wte by(auto simp add:after_def) finally show ?case using wte wt\<^sub>1 wt\<^sub>2 by(simp add:after_def hyperUn_assoc) next case (Call e M es) obtain C D Ts m Ts' where C: "P,E \\<^sub>1 e :: Class C" and "method": "P \ C sees M:Ts \ T = m in D" and wtes: "P,E \\<^sub>1 es [::] Ts'" and subs: "P \ Ts' [\] Ts" using Call.prems by auto from wtes have same_size: "size es = size Ts'" by(rule WTs\<^sub>1_same_size) let ?A\<^sub>0 = "A \ \ e" let ?A\<^sub>1 = "?A\<^sub>0 \ \s es" let ?\ = "ty\<^sub>i' ST E A" let ?\s\<^sub>e = "compT E A ST e" let ?\\<^sub>e = "ty\<^sub>i' (Class C # ST) E ?A\<^sub>0" let ?\s\<^sub>e\<^sub>s = "compTs E ?A\<^sub>0 (Class C # ST) es" let ?\\<^sub>1 = "ty\<^sub>i' (rev Ts' @ Class C # ST) E ?A\<^sub>1" let ?\' = "ty\<^sub>i' (T # ST) E ?A\<^sub>1" have "\ [Invoke M (size es)],[] [::] [?\\<^sub>1,?\']" by(rule wt_Invoke[OF same_size "method" subs]) also have "PROP ?Ps es E Ts' ?A\<^sub>0 (Class C # ST)" by fact hence "\ compEs\<^sub>2 es,compxEs\<^sub>2 es 0 (size ST+1) [::] ?\\<^sub>e # ?\s\<^sub>e\<^sub>s" "last (?\\<^sub>e # ?\s\<^sub>e\<^sub>s) = ?\\<^sub>1" using Call.prems wtes by(auto simp add:after_def) also have "(?\\<^sub>e # ?\s\<^sub>e\<^sub>s) @ [?\'] = ?\\<^sub>e # ?\s\<^sub>e\<^sub>s @ [?\']" by simp also have "\ compE\<^sub>2 e,compxE\<^sub>2 e 0 (size ST) [::] ?\ # ?\s\<^sub>e @ [?\\<^sub>e]" using Call C by(auto simp add:after_def) finally show ?case using Call.prems C by(simp add:after_def hyperUn_assoc) next case Seq thus ?case by(auto simp:after_def) (fastforce simp:wt_Push wt_Pop hyperUn_assoc intro:wt_instrs_app2 wt_instrs_Cons) qed (*>*) lemma [simp]: "types (compP f P) = types P" (*<*)by auto(*>*) lemma [simp]: "states (compP f P) mxs mxl = states P mxs mxl" (*<*)by (simp add: JVM_states_unfold)(*>*) lemma [simp]: "app\<^sub>i (i, compP f P, pc, mpc, T, \) = app\<^sub>i (i, P, pc, mpc, T, \)" (*<*)(is "?A = ?B") proof - obtain ST LT where \: "\ = (ST, LT)" by(cases \) simp then show ?thesis proof(cases i) case Invoke show ?thesis proof(rule iffI) assume ?A then show ?B using Invoke \ by auto (fastforce dest!: sees_method_compPD) next assume ?B then show ?A using Invoke \ by auto (force dest: sees_method_compP) qed qed auto qed (*>*) lemma [simp]: "is_relevant_entry (compP f P) i = is_relevant_entry P i" (*<*) proof - { fix pc e have "is_relevant_entry (compP f P) i pc e = is_relevant_entry P i pc e" by(cases i) (auto simp: is_relevant_entry_def) } then show ?thesis by fast qed (*>*) lemma [simp]: "relevant_entries (compP f P) i pc xt = relevant_entries P i pc xt" (*<*) by (simp add: relevant_entries_def)(*>*) lemma [simp]: "app i (compP f P) mpc T pc mxl xt \ = app i P mpc T pc mxl xt \" (*<*) by (simp add: app_def xcpt_app_def eff_def xcpt_eff_def norm_eff_def) (fastforce simp add: image_def) (*>*) lemma [simp]: assumes "app i P mpc T pc mxl xt \" shows "eff i (compP f P) pc xt \ = eff i P pc xt \" (*<*) using assms proof(clarsimp simp: eff_def norm_eff_def xcpt_eff_def app_def, cases i) qed auto (*>*) lemma [simp]: "subtype (compP f P) = subtype P" (*<*)by (rule ext)+ simp(*>*) lemma [simp]: "compP f P \ \ \' \' = P \ \ \' \'" (*<*) by (simp add: sup_state_opt_def sup_state_def sup_ty_opt_def)(*>*) lemma [simp]: "compP f P,T,mpc,mxl,xt \ i,pc :: \s = P,T,mpc,mxl,xt \ i,pc :: \s" (*<*)by (simp add: wt_instr_def cong: conj_cong)(*>*) declare TC1.compT_sizes[simp] TC0.ty_def2[simp] context TC2 begin lemma compT_method: fixes e and A and C and Ts and mxl\<^sub>0 defines [simp]: "E \ Class C # Ts" and [simp]: "A \ \{..size Ts}\" and [simp]: "A' \ A \ \ e" and [simp]: "mxl\<^sub>0 \ max_vars e" assumes mxs: "max_stack e = mxs" and mxl: "Suc (length Ts + max_vars e) = mxl" assumes wf: "wf_prog p P" and wte: "P,E \\<^sub>1 e :: T" and \: "\ e A" and \: "\ e (size E)" and E_P: "set E \ types P" and wid: "P \ T \ T\<^sub>r" shows "wt_method (compP\<^sub>2 P) C Ts T\<^sub>r mxs mxl\<^sub>0 (compE\<^sub>2 e @ [Return]) (compxE\<^sub>2 e 0 0) (ty\<^sub>i' [] E A # compT\<^sub>a E A [] e)" (*<*)(is "wt_method ?P C Ts T\<^sub>r mxs mxl\<^sub>0 ?is ?xt ?\s") proof - let ?n = "length E + mxl\<^sub>0" have wt_compE: "P,T\<^sub>r,mxs \ compE\<^sub>2 e, compxE\<^sub>2 e 0 (length []) [::] TC0.ty\<^sub>i' ?n [] E A # TC1.compT P ?n E A [] e @[TC0.after P ?n E A [] e]" using mxs TC2.compT_wt_instrs [OF wte \ \, of "[]" mxs ?n T\<^sub>r] by simp have "OK (ty\<^sub>i' [T] E A') \ states P mxs mxl" using mxs WT\<^sub>1_is_type[OF wf wte E_P] max_stack1[of e] OK_ty\<^sub>i'_in_statesI[OF E_P] by simp moreover have "OK ` set (compT E A [] e) \ states P mxs mxl" using mxs mxl wid compT_states(1)[OF wf wte E_P] by simp ultimately have "check_types ?P mxs ?n (map OK ?\s)" using mxl wte E_P by (simp add: compT\<^sub>a_def after_def check_types_def) moreover have "wt_start ?P C Ts mxl\<^sub>0 ?\s" using mxl by (auto simp: wt_start_def ty\<^sub>i'_def ty\<^sub>l_def list_all2_conv_all_nth nth_Cons split: nat.split) moreover { fix pc assume pc: "pc < size ?is" then have "?P,T\<^sub>r,mxs,size ?is,?xt \ ?is!pc,pc :: ?\s" proof(cases "pc < length (compE\<^sub>2 e)") case True then show ?thesis using mxs wte wt_compE by (clarsimp simp: compT\<^sub>a_def mxl after_def wt_instrs_def) next case False have "length (compE\<^sub>2 e) = pc" using less_antisym[OF False] pc by simp then show ?thesis using mxl wte E_P wid by (clarsimp simp: compT\<^sub>a_def after_def wt_defs xcpt_app_pcs xcpt_eff_pcs ty\<^sub>i'_def) qed } moreover have "0 < size ?is" and "size ?\s = size ?is" using wte by (simp_all add: compT\<^sub>a_def) ultimately show ?thesis by(simp add: wt_method_def) qed (*>*) end definition compTP :: "J\<^sub>1_prog \ ty\<^sub>P" where "compTP P C M = ( let (D,Ts,T,e) = method P C M; E = Class C # Ts; A = \{..size Ts}\; mxl = 1 + size Ts + max_vars e in (TC0.ty\<^sub>i' mxl [] E A # TC1.compT\<^sub>a P mxl E A [] e))" theorem wt_compP\<^sub>2: assumes wf: "wf_J\<^sub>1_prog P" shows "wf_jvm_prog (compP\<^sub>2 P)" (*<*) proof - let ?\ = "compTP P" and ?f = compMb\<^sub>2 let ?wf\<^sub>2 = "\P C (M, Ts, T\<^sub>r, mxs, mxl\<^sub>0, is, xt). wt_method P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt (?\ C M)" and ?P = "compP ?f P" { fix C M Ts T m assume cM: "P \ C sees M : Ts\T = m in C" and wfm: "wf_mdecl wf_J\<^sub>1_mdecl P C (M,Ts,T,m)" then have Ts_types: "\T'\set Ts. is_type P T'" and T_type: "is_type P T" and wfm\<^sub>1: "wf_J\<^sub>1_mdecl P C (M,Ts,T,m)" by(simp_all add: wf_mdecl_def) then obtain T' where wte: "P,Class C#Ts \\<^sub>1 m :: T'" and wid: "P \ T' \ T" and \: "\ m \{..size Ts}\" and \: "\ m (Suc (size Ts))" by(auto simp: wf_mdecl_def) have CTs_P: "is_class P C \ set Ts \ types P" using sees_wf_mdecl[OF wf cM] sees_method_is_class[OF cM] by (clarsimp simp: wf_mdecl_def) have "?wf\<^sub>2 ?P C (M,Ts,T,?f m)" using cM TC2.compT_method [simplified, OF _ _ wf wte \ \ CTs_P wid] by(simp add: compTP_def) then have "wf_mdecl ?wf\<^sub>2 ?P C (M, Ts, T, ?f m)" using Ts_types T_type by(simp add: wf_mdecl_def) } then have "wf_prog ?wf\<^sub>2 (compP ?f P)" by(rule wf_prog_compPI[OF _ wf]) then show ?thesis by (simp add: wf_jvm_prog_def wf_jvm_prog_phi_def) fast qed (*>*) theorem wt_J2JVM: "wf_J_prog P \ wf_jvm_prog (J2JVM P)" (*<*) by(simp only:o_def J2JVM_def) (blast intro:wt_compP\<^sub>2 compP\<^sub>1_pres_wf) end diff --git a/thys/Jinja/DFA/Abstract_BV.thy b/thys/Jinja/DFA/Abstract_BV.thy --- a/thys/Jinja/DFA/Abstract_BV.thy +++ b/thys/Jinja/DFA/Abstract_BV.thy @@ -1,13 +1,14 @@ (* Title: HOL/MicroJava/BV/Semilat.thy Author: Gerwin Klein Copyright 2003 TUM Abstract Bytecode Verifier. *) (*<*) theory Abstract_BV -imports Typing_Framework_err Kildall LBVCorrect LBVComplete +imports Typing_Framework_err Kildall_2 LBVCorrect LBVComplete begin + end (*>*) diff --git a/thys/Jinja/DFA/Err.thy b/thys/Jinja/DFA/Err.thy --- a/thys/Jinja/DFA/Err.thy +++ b/thys/Jinja/DFA/Err.thy @@ -1,383 +1,410 @@ (* Title: HOL/MicroJava/BV/Err.thy Author: Tobias Nipkow Copyright 2000 TUM The error type. *) section \The Error Type\ theory Err imports Semilat begin datatype 'a err = Err | OK 'a type_synonym 'a ebinop = "'a \ 'a \ 'a err" type_synonym 'a esl = "'a set \ 'a ord \ 'a ebinop" primrec ok_val :: "'a err \ 'a" where "ok_val (OK x) = x" definition lift :: "('a \ 'b err) \ ('a err \ 'b err)" where "lift f e = (case e of Err \ Err | OK x \ f x)" definition lift2 :: "('a \ 'b \ 'c err) \ 'a err \ 'b err \ 'c err" where "lift2 f e\<^sub>1 e\<^sub>2 = (case e\<^sub>1 of Err \ Err | OK x \ (case e\<^sub>2 of Err \ Err | OK y \ f x y))" definition le :: "'a ord \ 'a err ord" where "le r e\<^sub>1 e\<^sub>2 = (case e\<^sub>2 of Err \ True | OK y \ (case e\<^sub>1 of Err \ False | OK x \ x \\<^sub>r y))" definition sup :: "('a \ 'b \ 'c) \ ('a err \ 'b err \ 'c err)" where "sup f = lift2 (\x y. OK (x \\<^sub>f y))" definition err :: "'a set \ 'a err set" where "err A = insert Err {OK x|x. x\A}" definition esl :: "'a sl \ 'a esl" where "esl = (\(A,r,f). (A, r, \x y. OK(f x y)))" definition sl :: "'a esl \ 'a err sl" where "sl = (\(A,r,f). (err A, le r, lift2 f))" abbreviation err_semilat :: "'a esl \ bool" where "err_semilat L == semilat(sl L)" primrec strict :: "('a \ 'b err) \ ('a err \ 'b err)" where "strict f Err = Err" | "strict f (OK x) = f x" lemma err_def': "err A = insert Err {x. \y\A. x = OK y}" (*<*) proof - have eq: "err A = insert Err {x. \y\A. x = OK y}" by (unfold err_def) blast show "err A = insert Err {x. \y\A. x = OK y}" by (simp add: eq) qed (*>*) lemma strict_Some [simp]: "(strict f x = OK y) = (\z. x = OK z \ f z = OK y)" (*<*) by (cases x, auto) (*>*) lemma not_Err_eq: "(x \ Err) = (\a. x = OK a)" (*<*) by (cases x) auto (*>*) lemma not_OK_eq: "(\y. x \ OK y) = (x = Err)" (*<*) by (cases x) auto (*>*) lemma unfold_lesub_err: "e1 \\<^bsub>le r\<^esub> e2 = le r e1 e2" (*<*) by (simp add: lesub_def) (*>*) lemma le_err_refl: "\x. x \\<^sub>r x \ e \\<^bsub>le r\<^esub> e" (*<*) apply (unfold lesub_def le_def) apply (simp split: err.split) done (*>*) +lemma le_err_refl': "(\x\A. x \\<^sub>r x) \ e \ err A \ e \\<^bsub>le r\<^esub> e" +(*<*) +apply (unfold lesub_def le_def err_def) +apply (auto split: err.split) +done + lemma le_err_trans [rule_format]: - "order r \ e1 \\<^bsub>le r\<^esub> e2 \ e2 \\<^bsub>le r\<^esub> e3 \ e1 \\<^bsub>le r\<^esub> e3" + "order r A \ e1 \ err A \ e2 \ err A \ e3 \ err A \ e1 \\<^bsub>le r\<^esub> e2 \ e2 \\<^bsub>le r\<^esub> e3 \ e1 \\<^bsub>le r\<^esub> e3" (*<*) -apply (unfold unfold_lesub_err le_def) +apply (unfold unfold_lesub_err le_def err_def) apply (simp split: err.split) apply (blast intro: order_trans) done (*>*) lemma le_err_antisym [rule_format]: - "order r \ e1 \\<^bsub>le r\<^esub> e2 \ e2 \\<^bsub>le r\<^esub> e1 \ e1=e2" + "order r A \ e1 \ err A \ e2 \ err A \ e3 \ err A \ e1 \\<^bsub>le r\<^esub> e2 \ e2 \\<^bsub>le r\<^esub> e1 \ e1=e2" (*<*) -apply (unfold unfold_lesub_err le_def) +apply (unfold unfold_lesub_err le_def err_def) apply (simp split: err.split) apply (blast intro: order_antisym) done (*>*) lemma OK_le_err_OK: "(OK x \\<^bsub>le r\<^esub> OK y) = (x \\<^sub>r y)" (*<*) by (simp add: unfold_lesub_err le_def) (*>*) -lemma order_le_err [iff]: "order(le r) = order r" +lemma order_le_err [iff]: "order (le r) (err A) = order r A" (*<*) apply (rule iffI) - apply (subst order_def) + apply (subst order_def) + apply (simp only: err_def) apply (blast dest: order_antisym OK_le_err_OK [THEN iffD2] intro: order_trans OK_le_err_OK [THEN iffD1]) -apply (subst order_def) -apply (blast intro: le_err_refl le_err_trans le_err_antisym + apply (subst order_def) +apply (blast intro: le_err_refl' le_err_trans le_err_antisym dest: order_refl) done (*>*) lemma le_Err [iff]: "e \\<^bsub>le r\<^esub> Err" (*<*) by (simp add: unfold_lesub_err le_def) (*>*) lemma Err_le_conv [iff]: "Err \\<^bsub>le r\<^esub> e = (e = Err)" (*<*) by (simp add: unfold_lesub_err le_def split: err.split) (*>*) lemma le_OK_conv [iff]: "e \\<^bsub>le r\<^esub> OK x = (\y. e = OK y \ y \\<^sub>r x)" (*<*) by (simp add: unfold_lesub_err le_def split: err.split) (*>*) lemma OK_le_conv: "OK x \\<^bsub>le r\<^esub> e = (e = Err \ (\y. e = OK y \ x \\<^sub>r y))" (*<*) by (simp add: unfold_lesub_err le_def split: err.split) (*>*) lemma top_Err [iff]: "top (le r) Err" (*<*) by (simp add: top_def) (*>*) lemma OK_less_conv [rule_format, iff]: "OK x \\<^bsub>le r\<^esub> e = (e=Err \ (\y. e = OK y \ x \\<^sub>r y))" (*<*) by (simp add: lesssub_def lesub_def le_def split: err.split) (*>*) lemma not_Err_less [rule_format, iff]: "\(Err \\<^bsub>le r\<^esub> x)" (*<*) by (simp add: lesssub_def lesub_def le_def split: err.split) (*>*) lemma semilat_errI [intro]: assumes "Semilat A r f" shows "semilat(err A, le r, lift2(\x y. OK(f x y)))" (*<*) proof - interpret Semilat A r f by fact show ?thesis apply(insert semilat) apply (simp only: semilat_Def closed_def plussub_def lesub_def lift2_def le_def) + apply(rule conjI) + apply simp apply (simp add: err_def' split: err.split) done qed (*>*) lemma err_semilat_eslI_aux: assumes "Semilat A r f" shows "err_semilat(esl(A,r,f))" (*<*) proof - interpret Semilat A r f by fact show ?thesis apply (unfold sl_def esl_def) apply (simp add: semilat_errI [OF \Semilat A r f\]) done qed (*>*) lemma err_semilat_eslI [intro, simp]: "semilat L \ err_semilat (esl L)" (*<*) apply (cases L) apply simp apply (drule Semilat.intro) apply (simp add: err_semilat_eslI_aux split_tupled_all) done (*>*) lemma acc_err [simp, intro!]: "acc r \ acc(le r)" (*<*) apply (unfold acc_def lesub_def le_def lesssub_def) apply (simp add: wf_eq_minimal split: err.split) apply clarify apply (case_tac "Err : Q") apply blast apply (erule_tac x = "{a . OK a : Q}" in allE) apply (case_tac "x") apply fast apply blast done (*>*) lemma Err_in_err [iff]: "Err : err A" (*<*) by (simp add: err_def') (*>*) lemma Ok_in_err [iff]: "(OK x \ err A) = (x\A)" (*<*) by (auto simp add: err_def') (*>*) subsection \lift\ lemma lift_in_errI: "\ e \ err S; \x\S. e = OK x \ f x \ err S \ \ lift f e \ err S" (*<*) apply (unfold lift_def) apply (simp split: err.split) apply blast done (*>*) lemma Err_lift2 [simp]: "Err \\<^bsub>lift2 f\<^esub> x = Err" (*<*) by (simp add: lift2_def plussub_def) (*>*) lemma lift2_Err [simp]: "x \\<^bsub>lift2 f\<^esub> Err = Err" (*<*) by (simp add: lift2_def plussub_def split: err.split) (*>*) lemma OK_lift2_OK [simp]: "OK x \\<^bsub>lift2 f\<^esub> OK y = x \\<^sub>f y" (*<*) by (simp add: lift2_def plussub_def split: err.split) (*>*) subsection \sup\ lemma Err_sup_Err [simp]: "Err \\<^bsub>sup f\<^esub> x = Err" (*<*) by (simp add: plussub_def sup_def lift2_def) (*>*) lemma Err_sup_Err2 [simp]: "x \\<^bsub>sup f\<^esub> Err = Err" (*<*) by (simp add: plussub_def sup_def lift2_def split: err.split) (*>*) lemma Err_sup_OK [simp]: "OK x \\<^bsub>sup f\<^esub> OK y = OK (x \\<^sub>f y)" (*<*) by (simp add: plussub_def sup_def lift2_def) (*>*) lemma Err_sup_eq_OK_conv [iff]: "(sup f ex ey = OK z) = (\x y. ex = OK x \ ey = OK y \ f x y = z)" (*<*) apply (unfold sup_def lift2_def plussub_def) apply (rule iffI) apply (simp split: err.split_asm) apply clarify apply simp done (*>*) lemma Err_sup_eq_Err [iff]: "(sup f ex ey = Err) = (ex=Err \ ey=Err)" (*<*) apply (unfold sup_def lift2_def plussub_def) apply (simp split: err.split) done (*>*) subsection \semilat (err A) (le r) f\ lemma semilat_le_err_Err_plus [simp]: "\ x\ err A; semilat(err A, le r, f) \ \ Err \\<^sub>f x = Err" (*<*) by (blast intro: Semilat.le_iff_plus_unchanged [THEN iffD1, OF Semilat.intro] Semilat.le_iff_plus_unchanged2 [THEN iffD1, OF Semilat.intro]) (*>*) lemma semilat_le_err_plus_Err [simp]: "\ x\ err A; semilat(err A, le r, f) \ \ x \\<^sub>f Err = Err" (*<*) by (blast intro: Semilat.le_iff_plus_unchanged [THEN iffD1, OF Semilat.intro] Semilat.le_iff_plus_unchanged2 [THEN iffD1, OF Semilat.intro]) (*>*) lemma semilat_le_err_OK1: "\ x\A; y\A; semilat(err A, le r, f); OK x \\<^sub>f OK y = OK z \ \ x \\<^sub>r z" (*<*) apply (rule OK_le_err_OK [THEN iffD1]) apply (erule subst) apply (simp add: Semilat.ub1 [OF Semilat.intro]) done (*>*) lemma semilat_le_err_OK2: "\ x\A; y\A; semilat(err A, le r, f); OK x \\<^sub>f OK y = OK z \ \ y \\<^sub>r z" (*<*) apply (rule OK_le_err_OK [THEN iffD1]) apply (erule subst) apply (simp add: Semilat.ub2 [OF Semilat.intro]) done (*>*) lemma eq_order_le: - "\ x=y; order r \ \ x \\<^sub>r y" + "\ x=y; order r A; x \A ; y \ A \ \ x \\<^sub>r y" (*<*) apply (unfold order_def) apply blast done (*>*) lemma OK_plus_OK_eq_Err_conv [simp]: assumes "x\A" "y\A" "semilat(err A, le r, fe)" shows "(OK x \\<^bsub>fe\<^esub> OK y = Err) = (\(\z\A. x \\<^sub>r z \ y \\<^sub>r z))" (*<*) proof - have plus_le_conv3: "\A x y z f r. \ semilat (A,r,f); x \\<^sub>f y \\<^sub>r z; x\A; y\A; z\A \ \ x \\<^sub>r z \ y \\<^sub>r z" (*<*) by (rule Semilat.plus_le_conv [OF Semilat.intro, THEN iffD1]) (*>*) from assms show ?thesis - apply (rule_tac iffI) - apply clarify - apply (drule OK_le_err_OK [THEN iffD2]) - apply (drule OK_le_err_OK [THEN iffD2]) - apply (drule Semilat.lub[OF Semilat.intro, of _ _ _ "OK x" _ "OK y"]) - apply assumption - apply assumption + apply (rule_tac iffI) + apply clarify + apply (drule OK_le_err_OK [THEN iffD2]) + apply (drule OK_le_err_OK [THEN iffD2]) + apply (drule Semilat.lub[OF Semilat.intro, of _ _ _ "OK x" _ "OK y"]) + apply assumption + apply assumption + apply simp + apply simp apply simp apply simp - apply simp - apply simp - apply (case_tac "OK x \\<^bsub>fe\<^esub> OK y") - apply assumption - apply (rename_tac z) - apply (subgoal_tac "OK z\ err A") - apply (drule eq_order_le) - apply (erule Semilat.orderI [OF Semilat.intro]) - apply (blast dest: plus_le_conv3) - apply (erule subst) - apply (blast intro: Semilat.closedI [OF Semilat.intro] closedD) - done + apply (case_tac "OK x \\<^bsub>fe\<^esub> OK y") + apply assumption + apply (rename_tac z) + apply (subgoal_tac "OK z\ err A") + apply (drule eq_order_le) + apply (erule Semilat.orderI [OF Semilat.intro]) + apply (blast intro: Semilat.closedI [OF Semilat.intro] closedD) + apply simp + apply (blast dest: plus_le_conv3) + apply (erule subst) + apply (blast intro: Semilat.closedI [OF Semilat.intro] closedD) + done qed (*>*) subsection \semilat (err(Union AS))\ (* FIXME? *) lemma all_bex_swap_lemma [iff]: "(\x. (\y\A. x = f y) \ P x) = (\y\A. P(f y))" (*<*) by blast (*>*) lemma closed_err_Union_lift2I: "\ \A\AS. closed (err A) (lift2 f); AS \ {}; \A\AS.\B\AS. A\B \ (\a\A.\b\B. a \\<^sub>f b = Err) \ \ closed (err(Union AS)) (lift2 f)" (*<*) apply (unfold closed_def err_def') apply simp apply clarify apply simp apply fast done (*>*) text \ If @{term "AS = {}"} the thm collapses to - @{prop "order r \ closed {Err} f \ Err \\<^sub>f Err = Err"} + @{prop "order r A \ closed {Err} f \ Err \\<^sub>f Err = Err"} which may not hold \ + +lemma err_semilat_UnionI_auxi: + assumes "\A\AS. order r A " + and "\A\AS. \B\AS. A \ B \ (\a\A. \b\B. \ a \\<^bsub>r\<^esub> b \ a \\<^bsub>f\<^esub> b = Err)" + shows "order r (\ AS)" +proof- + from assms(1) have "\A. A \ AS \ order r A" by auto + then have "\A x. A \ AS \ x \ A \ x \\<^sub>r x" + and g1: "\A x y. A \ AS \ x \ A \ y \ A \x \\<^sub>r y \ y \\<^sub>r x \ x=y" + and g2: "\A x y z. A \ AS \ x \ A \ y \ A \z \ A \x \\<^sub>r y \ y \\<^sub>r z \ x \\<^sub>r z" by (auto dest:order_antisym order_trans) + then have "\x\(\ AS). x \\<^sub>r x" by blast + moreover from g1 have "\x\(\ AS). \y\(\ AS). x \\<^sub>r y \ y \\<^sub>r x \ x=y" using assms(2) by blast + moreover from g2 have "\x\(\ AS). \y\(\ AS). \z\(\ AS). x \\<^sub>r y \ y \\<^sub>r z \ x \\<^sub>r z" using assms(2) by blast + ultimately show "order r (\ AS)" using order_def by blast +qed + lemma err_semilat_UnionI: "\ \A\AS. err_semilat(A, r, f); AS \ {}; \A\AS.\B\AS. A\B \ (\a\A.\b\B. \a \\<^sub>r b \ a \\<^sub>f b = Err) \ \ err_semilat(Union AS, r, f)" (*<*) apply (unfold semilat_def sl_def) apply (simp add: closed_err_Union_lift2I) apply (rule conjI) - apply blast + apply (blast intro: err_semilat_UnionI_auxi) apply (simp add: err_def') apply (rule conjI) apply clarify apply (rename_tac A a u B b) apply (case_tac "A = B") apply simp apply simp apply (rule conjI) apply clarify apply (rename_tac A a u B b) apply (case_tac "A = B") apply simp apply simp apply clarify apply (rename_tac A ya yb B yd z C c a b) apply (case_tac "A = B") apply (case_tac "A = C") apply simp apply simp apply (case_tac "B = C") apply simp apply simp done (*>*) end diff --git a/thys/Jinja/DFA/LBVComplete.thy b/thys/Jinja/DFA/LBVComplete.thy --- a/thys/Jinja/DFA/LBVComplete.thy +++ b/thys/Jinja/DFA/LBVComplete.thy @@ -1,339 +1,385 @@ (* Title: HOL/MicroJava/BV/LBVComplete.thy Author: Gerwin Klein Copyright 2000 Technische Universitaet Muenchen *) section \Completeness of the LBV\ theory LBVComplete -imports LBVSpec Typing_Framework +imports LBVSpec Typing_Framework_1 begin definition is_target :: "'s step_type \ 's list \ nat \ bool" where "is_target step \s pc' \ (\pc s'. pc' \ pc+1 \ pc < size \s \ (pc',s') \ set (step pc (\s!pc)))" definition make_cert :: "'s step_type \ 's list \ 's \ 's certificate" where "make_cert step \s B = map (\pc. if is_target step \s pc then \s!pc else B) [0..s] @ [B]" lemma [code]: "is_target step \s pc' = list_ex (\pc. pc' \ pc+1 \ List.member (map fst (step pc (\s!pc))) pc') [0..s]" (*<*) apply (simp add: list_ex_iff is_target_def member_def) apply force done (*>*) locale lbvc = lbv + fixes \s :: "'a list" fixes c :: "'a list" defines cert_def: "c \ make_cert step \s \" assumes mono: "mono r step (size \s) A" assumes pres: "pres_type step (size \s) A" assumes \s: "\pc < size \s. \s!pc \ A \ \s!pc \ \" assumes bounded: "bounded step (size \s)" assumes B_neq_T: "\ \ \" lemma (in lbvc) cert: "cert_ok c (size \s) \ \ A" (*<*) proof (unfold cert_ok_def, intro strip conjI) note [simp] = make_cert_def cert_def nth_append show "c!size \s = \" by simp fix pc assume pc: "pc < size \s" from pc \s B_A show "c!pc \ A" by simp from pc \s B_neq_T show "c!pc \ \" by simp qed (*>*) lemmas [simp del] = split_paired_Ex lemma (in lbvc) cert_target [intro?]: "\ (pc',s') \ set (step pc (\s!pc)); pc' \ pc+1; pc < size \s; pc' < size \s \ \ c!pc' = \s!pc'" (*<*) by (auto simp add: cert_def make_cert_def nth_append is_target_def) (*>*) lemma (in lbvc) cert_approx [intro?]: "\ pc < size \s; c!pc \ \ \ \ c!pc = \s!pc" (*<*) by (auto simp add: cert_def make_cert_def nth_append) (*>*) lemma (in lbv) le_top [simp, intro]: "x <=_r \" (*<*) by (insert top) simp (*>*) lemma (in lbv) merge_mono: assumes less: "set ss\<^sub>2 {\\<^bsub>r\<^esub>} set ss\<^sub>1" assumes x: "x \ A" assumes ss\<^sub>1: "snd`set ss\<^sub>1 \ A" assumes ss\<^sub>2: "snd`set ss\<^sub>2 \ A" + assumes boun: "\x\(fst`set ss\<^sub>1). x < size \s" + assumes cert: "cert_ok c (size \s) T B A" shows "merge c pc ss\<^sub>2 x \\<^sub>r merge c pc ss\<^sub>1 x" (is "?s\<^sub>2 \\<^sub>r ?s\<^sub>1") (*<*) proof- have "?s\<^sub>1 = \ \ ?thesis" by simp moreover { assume merge: "?s\<^sub>1 \ T" from x ss\<^sub>1 have "?s\<^sub>1 = (if \(pc',s')\set ss\<^sub>1. pc' \ pc + 1 \ s' \\<^sub>r c!pc' then (map snd [(p', t') \ ss\<^sub>1 . p'=pc+1]) \\<^bsub>f\<^esub> x else \)" by (rule merge_def) with merge obtain app: "\(pc',s')\set ss\<^sub>1. pc' \ pc+1 \ s' \\<^sub>r c!pc'" (is "?app ss\<^sub>1") and sum: "(map snd [(p',t') \ ss\<^sub>1 . p' = pc+1] \\<^bsub>f\<^esub> x) = ?s\<^sub>1" (is "?map ss\<^sub>1 \\<^bsub>f\<^esub> x = _" is "?sum ss\<^sub>1 = _") by (simp split: if_split_asm) - from app less have "?app ss\<^sub>2" by (blast dest: trans_r lesub_step_typeD) + + have "?app ss\<^sub>2" + proof(intro strip, clarsimp ) + fix a b + assume a_b: "(a, b) \ set ss\<^sub>2" and neq: "a \ Suc pc" + from less a_b obtain y where y: "(a, y) \ set ss\<^sub>1 " and b_lt_y: "b \\<^sub>r y" by (blast dest:lesub_step_typeD) + with app have "a \ pc + 1 \ y \\<^bsub>r\<^esub> c!a" by auto + with neq have "y \\<^bsub>r\<^esub> c!a" by auto + moreover from y ss\<^sub>1 have "y \ A" by auto + moreover from a_b ss\<^sub>2 have "b \ A" by auto + moreover from y boun cert have "c!a \ A" by (auto dest:cert_okD1) + ultimately show "b \\<^bsub>r\<^esub> c!a" using b_lt_y by (auto intro:trans_r) + qed + moreover { from ss\<^sub>1 have map1: "set (?map ss\<^sub>1) \ A" by auto with x and semilat Semilat_axioms have "?sum ss\<^sub>1 \ A" by (auto intro!: plusplus_closed) with sum have "?s\<^sub>1 \ A" by simp moreover have mapD: "\x ss. x \ set (?map ss) \ \p. (p,x) \ set ss \ p=pc+1" by auto from x map1 have "\x \ set (?map ss\<^sub>1). x \\<^sub>r ?sum ss\<^sub>1" by clarify (rule pp_ub1) with sum have "\x \ set (?map ss\<^sub>1). x \\<^sub>r ?s\<^sub>1" by simp - with less have "\x \ set (?map ss\<^sub>2). x \\<^sub>r ?s\<^sub>1" - by (fastforce dest!: mapD lesub_step_typeD intro: trans_r) + then have "\x \ set (?map ss\<^sub>2). x \\<^sub>r ?s\<^sub>1" + proof(intro strip, clarsimp) + fix b + assume xa: "\xa. (Suc pc, xa) \set ss\<^sub>1 \ xa \\<^bsub>r\<^esub> merge c pc ss\<^sub>1 x" + and b: "(Suc pc, b) \ set ss\<^sub>2" + from less b obtain y where y: "(Suc pc, y) \ set ss\<^sub>1 " and b_lt_y: "b \\<^sub>r y" by (blast dest:lesub_step_typeD) + from y xa have "y \\<^bsub>r\<^esub> merge c pc ss\<^sub>1 x" by auto + moreover from y ss\<^sub>1 have "y \ A" by auto + moreover from b ss\<^sub>2 have "b \ A" by auto + moreover from ss\<^sub>1 x have "merge c pc ss\<^sub>1 x \ A" by (auto intro:merge_pres) + ultimately show "b \\<^bsub>r\<^esub> merge c pc ss\<^sub>1 x" using b_lt_y by (auto intro:trans_r) + qed moreover from map1 x have "x \\<^sub>r (?sum ss\<^sub>1)" by (rule pp_ub2) with sum have "x \\<^sub>r ?s\<^sub>1" by simp moreover from ss\<^sub>2 have "set (?map ss\<^sub>2) \ A" by auto ultimately have "?sum ss\<^sub>2 \\<^sub>r ?s\<^sub>1" using x by - (rule pp_lub) } moreover from x ss\<^sub>2 have "?s\<^sub>2 = (if \(pc', s')\set ss\<^sub>2. pc' \ pc + 1 \ s' \\<^sub>r c!pc' then map snd [(p', t') \ ss\<^sub>2 . p' = pc + 1] \\<^bsub>f\<^esub> x else \)" by (rule merge_def) ultimately have ?thesis by simp } ultimately show ?thesis by (cases "?s\<^sub>1 = \") auto qed (*>*) lemma (in lbvc) wti_mono: assumes less: "s\<^sub>2 \\<^sub>r s\<^sub>1" assumes pc: "pc < size \s" and s\<^sub>1: "s\<^sub>1 \ A" and s\<^sub>2: "s\<^sub>2 \ A" shows "wti c pc s\<^sub>2 \\<^sub>r wti c pc s\<^sub>1" (is "?s\<^sub>2' \\<^sub>r ?s\<^sub>1'") (*<*) proof - + from bounded pc have "\(q,\) \ set(step pc s\<^sub>1). q < size \s" by (simp add:bounded_def) + then have u: "\q \ fst ` set (step pc s\<^sub>1). q < size \s" by auto from mono pc s\<^sub>2 less have "set (step pc s\<^sub>2) {\\<^bsub>r\<^esub>} set (step pc s\<^sub>1)" by (rule monoD) moreover from cert B_A pc have "c!Suc pc \ A" by (rule cert_okD3) moreover from pres s\<^sub>1 pc have "snd`set (step pc s\<^sub>1) \ A" by (rule pres_typeD2) moreover from pres s\<^sub>2 pc have "snd`set (step pc s\<^sub>2) \ A" by (rule pres_typeD2) - ultimately show ?thesis by (simp add: wti merge_mono) + ultimately show ?thesis using cert u by (simp add: wti merge_mono) qed (*>*) lemma (in lbvc) wtc_mono: assumes less: "s\<^sub>2 \\<^sub>r s\<^sub>1" assumes pc: "pc < size \s" and s\<^sub>1: "s\<^sub>1 \ A" and s\<^sub>2: "s\<^sub>2 \ A" shows "wtc c pc s\<^sub>2 \\<^sub>r wtc c pc s\<^sub>1" (is "?s\<^sub>2' \\<^sub>r ?s\<^sub>1'") (*<*) proof (cases "c!pc = \") case True moreover from less pc s\<^sub>1 s\<^sub>2 have "wti c pc s\<^sub>2 \\<^sub>r wti c pc s\<^sub>1" by (rule wti_mono) ultimately show ?thesis by (simp add: wtc) next case False + with pc have "c!pc = \s!pc" using cert_approx by auto + with \s pc have c_pc_inA: "c!pc \ A" by auto + moreover from cert B_A pc have "c ! (pc + 1) \ A" by (auto dest: cert_okD3) + ultimately have inA: "wtc c pc s\<^sub>2 \ A" using pres wtc_pres pc s\<^sub>2 by auto have "?s\<^sub>1' = \ \ ?thesis" by simp moreover { assume "?s\<^sub>1' \ \" with False have c: "s\<^sub>1 \\<^sub>r c!pc" by (simp add: wtc split: if_split_asm) - with less have "s\<^sub>2 \\<^sub>r c!pc" .. - with False c have ?thesis by (simp add: wtc) + from semilat have "order r A" by auto + from this less c s\<^sub>2 s\<^sub>1 c_pc_inA have "s\<^sub>2 \\<^sub>r c!pc" by (rule order_trans) + with False c have ?thesis using inA by (simp add: wtc) } ultimately show ?thesis by (cases "?s\<^sub>1' = \") auto qed (*>*) -lemma (in lbv) top_le_conv [simp]: "\ \\<^sub>r x = (x = \)" -(*<*) by (insert semilat) (simp add: top top_le_conv) (*>*) +lemma (in lbv) top_le_conv [simp]: "x \ A \ \ \\<^sub>r x = (x = \)" +(*<*) + apply(subgoal_tac "order r A") + apply (insert semilat T_A top) + apply (rule top_le_conv) + apply assumption+ + apply (simp add:semilat_def) + done + (*>*) -lemma (in lbv) neq_top [simp, elim]: "\ x \\<^sub>r y; y \ \ \ \ x \ \" +lemma (in lbv) neq_top [simp, elim]: "\ x \\<^sub>r y; y \ \; y \ A \ \ x \ \" (*<*) by (cases "x = T") auto (*>*) lemma (in lbvc) stable_wti: assumes stable: "stable r step \s pc" and pc: "pc < size \s" shows "wti c pc (\s!pc) \ \" (*<*) proof - let ?step = "step pc (\s!pc)" from stable have less: "\(q,s')\set ?step. s' \\<^sub>r \s!q" by (simp add: stable_def) from cert B_A pc have cert_suc: "c!Suc pc \ A" by (rule cert_okD3) moreover from \s pc have "\s!pc \ A" by simp with pres pc have stepA: "snd`set ?step \ A" by - (rule pres_typeD2) ultimately have "merge c pc ?step (c!Suc pc) = (if \(pc',s')\set ?step. pc'\pc+1 \ s' \\<^sub>r c!pc' then map snd [(p',t') \ ?step.p'=pc+1] \\<^bsub>f\<^esub> c!Suc pc else \)" unfolding mrg_def by (rule lbv.merge_def [OF lbvc.axioms(1), OF lbvc_axioms]) moreover { fix pc' s' assume s': "(pc',s') \ set ?step" and suc_pc: "pc' \ pc+1" with less have "s' \\<^sub>r \s!pc'" by auto also from bounded pc s' have "pc' < size \s" by (rule boundedD) with s' suc_pc pc have "c!pc' = \s!pc'" .. hence "\s!pc' = c!pc'" .. finally have "s' \\<^sub>r c!pc'" . } hence "\(pc',s')\set ?step. pc'\pc+1 \ s' \\<^sub>r c!pc'" by auto moreover from pc have "Suc pc = size \s \ Suc pc < size \s" by auto hence "map snd [(p',t') \ ?step.p'=pc+1] \\<^bsub>f\<^esub> c!Suc pc \ \" (is "?map \\<^bsub>f\<^esub> _ \ _") proof (rule disjE) assume pc': "Suc pc = size \s" with cert have "c!Suc pc = \" by (simp add: cert_okD2) moreover from pc' bounded pc have "\(p',t')\set ?step. p'\pc+1" by clarify (drule boundedD, auto) hence "[(p',t') \ ?step. p'=pc+1] = []" by (blast intro: filter_False) hence "?map = []" by simp ultimately show ?thesis by (simp add: B_neq_T) next assume pc': "Suc pc < size \s" - from pc' \s have "\s!Suc pc \ A" by simp + from pc' \s have \s_inA: "\s!Suc pc \ A" by simp moreover note cert_suc moreover from stepA have "set ?map \ A" by auto moreover have "\s. s \ set ?map \ \t. (Suc pc, t) \ set ?step" by auto with less have "\s' \ set ?map. s' \\<^sub>r \s!Suc pc" by auto - moreover from pc' have "c!Suc pc \\<^sub>r \s!Suc pc" + moreover from pc' \s_inA have "c!Suc pc \\<^sub>r \s!Suc pc" by (cases "c!Suc pc = \") (auto dest: cert_approx) ultimately have "?map \\<^bsub>f\<^esub> c!Suc pc \\<^sub>r \s!Suc pc" by (rule pp_lub) moreover from pc' \s have "\s!Suc pc \ \" by simp - ultimately show ?thesis by auto + ultimately show ?thesis using \s_inA by auto qed ultimately have "merge c pc ?step (c!Suc pc) \ \" by simp thus ?thesis by (simp add: wti) qed (*>*) lemma (in lbvc) wti_less: assumes stable: "stable r step \s pc" and suc_pc: "Suc pc < size \s" shows "wti c pc (\s!pc) \\<^sub>r \s!Suc pc" (is "?wti \\<^sub>r _") (*<*) proof - let ?step = "step pc (\s!pc)" from stable have less: "\(q,s')\set ?step. s' \\<^sub>r \s!q" by (simp add: stable_def) from suc_pc have pc: "pc < size \s" by simp with cert B_A have cert_suc: "c!Suc pc \ A" by (rule cert_okD3) moreover from \s pc have "\s!pc \ A" by simp with pres pc have stepA: "snd`set ?step \ A" by - (rule pres_typeD2) moreover from stable pc have "?wti \ \" by (rule stable_wti) hence "merge c pc ?step (c!Suc pc) \ \" by (simp add: wti) ultimately have "merge c pc ?step (c!Suc pc) = map snd [(p',t') \ ?step.p'=pc+1] \\<^bsub>f\<^esub> c!Suc pc" by (rule merge_not_top_s) hence "?wti = \" (is "_ = (?map \\<^bsub>f\<^esub> _)" is "_ = ?sum") by (simp add: wti) also { - from suc_pc \s have "\s!Suc pc \ A" by simp + from suc_pc \s have \s_inA: "\s!Suc pc \ A" by simp moreover note cert_suc moreover from stepA have "set ?map \ A" by auto moreover have "\s. s \ set ?map \ \t. (Suc pc, t) \ set ?step" by auto with less have "\s' \ set ?map. s' \\<^sub>r \s!Suc pc" by auto - moreover from suc_pc have "c!Suc pc \\<^sub>r \s!Suc pc" + moreover from suc_pc \s_inA have "c!Suc pc \\<^sub>r \s!Suc pc" by (cases "c!Suc pc = \") (auto dest: cert_approx) ultimately have "?sum \\<^sub>r \s!Suc pc" by (rule pp_lub) } finally show ?thesis . qed (*>*) lemma (in lbvc) stable_wtc: assumes stable: "stable r step \s pc" and pc: "pc < size \s" shows "wtc c pc (\s!pc) \ \" (*<*) proof - from stable pc have wti: "wti c pc (\s!pc) \ \" by (rule stable_wti) show ?thesis proof (cases "c!pc = \") case True with wti show ?thesis by (simp add: wtc) next case False with pc have "c!pc = \s!pc" .. - with False wti show ?thesis by (simp add: wtc) + with False wti \s pc show ?thesis by (simp add: wtc) qed qed (*>*) lemma (in lbvc) wtc_less: assumes stable: "stable r step \s pc" and suc_pc: "Suc pc < size \s" shows "wtc c pc (\s!pc) \\<^sub>r \s!Suc pc" (is "?wtc \\<^sub>r _") (*<*) proof (cases "c!pc = \") case True moreover from stable suc_pc have "wti c pc (\s!pc) \\<^sub>r \s!Suc pc" by (rule wti_less) ultimately show ?thesis by (simp add: wtc) next case False from suc_pc have pc: "pc < size \s" by simp with stable have "?wtc \ \" by (rule stable_wtc) with False have "?wtc = wti c pc (c!pc)" by (unfold wtc) (simp split: if_split_asm) also from pc False have "c!pc = \s!pc" .. finally have "?wtc = wti c pc (\s!pc)" . also from stable suc_pc have "wti c pc (\s!pc) \\<^sub>r \s!Suc pc" by (rule wti_less) finally show ?thesis . qed (*>*) lemma (in lbvc) wt_step_wtl_lemma: assumes wt_step: "wt_step r \ step \s" shows "\pc s. pc+size ls = size \s \ s \\<^sub>r \s!pc \ s \ A \ s\\ \ wtl ls c pc s \ \" (is "\pc s. _ \ _ \ _ \ _ \ ?wtl ls pc s \ _") (*<*) proof (induct ls) fix pc s assume "s\\" thus "?wtl [] pc s \ \" by simp next fix pc s i ls assume "\pc s. pc+size ls=size \s \ s \\<^sub>r \s!pc \ s \ A \ s\\ \ ?wtl ls pc s \ \" moreover assume pc_l: "pc + size (i#ls) = size \s" hence suc_pc_l: "Suc pc + size ls = size \s" by simp ultimately have IH: "\s. s \\<^sub>r \s!Suc pc \ s \ A \ s \ \ \ ?wtl ls (Suc pc) s \ \" . from pc_l obtain pc: "pc < size \s" by simp with wt_step have stable: "stable r step \s pc" by (simp add: wt_step_def) moreover note pc ultimately have wt_\s: "wtc c pc (\s!pc) \ \" by (rule stable_wtc) assume s_\s: "s \\<^sub>r \s!pc" assume sA: "s \ A" - from \s pc have \s_pc: "\s!pc \ A" by simp + + from \s pc have \s_pc:"\s!pc \ A" by auto + moreover from cert pc have c1: "c!pc \ A" by (rule cert_okD1) + moreover from cert B_A pc have c2: "c!Suc pc \ A" by (rule cert_okD3) + ultimately have wtc1: "wtc c pc (\s!pc) \ A" + and wtc2: "wtc c pc s \ A" using pres sA pc by (auto intro:wtc_pres) + from s_\s pc \s_pc sA have wt_s_\s: "wtc c pc s \\<^sub>r wtc c pc (\s!pc)" by (rule wtc_mono) - with wt_\s have wt_s: "wtc c pc s \ \" by simp + with wt_\s have wt_s: "wtc c pc s \ \" using wtc1 by simp moreover assume s: "s \ \" ultimately have "ls = [] \ ?wtl (i#ls) pc s \ \" by simp moreover { assume "ls \ []" with pc_l have suc_pc: "Suc pc < size \s" by (auto simp add: neq_Nil_conv) - with stable have "wtc c pc (\s!pc) \\<^sub>r \s!Suc pc" by (rule wtc_less) - with wt_s_\s have "wtc c pc s \\<^sub>r \s!Suc pc" by (rule trans_r) + with stable have t: "wtc c pc (\s!pc) \\<^sub>r \s!Suc pc" by (rule wtc_less) + from suc_pc \s have "\s!Suc pc \ A" by auto + with t wt_s_\s wtc1 wtc2 have "wtc c pc s \\<^sub>r \s!Suc pc" by (auto intro: trans_r) moreover from cert suc_pc have "c!pc \ A" "c!(pc+1) \ A" by (auto simp add: cert_ok_def) from pres this sA pc have "wtc c pc s \ A" by (rule wtc_pres) ultimately have "?wtl ls (Suc pc) (wtc c pc s) \ \" using IH wt_s by blast with s wt_s have "?wtl (i#ls) pc s \ \" by simp } ultimately show "?wtl (i#ls) pc s \ \" by (cases ls) blast+ qed (*>*) theorem (in lbvc) wtl_complete: assumes wt: "wt_step r \ step \s" assumes s: "s \\<^sub>r \s!0" "s \ A" "s \ \" and eq: "size ins = size \s" shows "wtl ins c 0 s \ \" (*<*) proof - from eq have "0+size ins = size \s" by simp from wt this s show ?thesis by (rule wt_step_wtl_lemma) qed (*>*) end diff --git a/thys/Jinja/DFA/LBVCorrect.thy b/thys/Jinja/DFA/LBVCorrect.thy --- a/thys/Jinja/DFA/LBVCorrect.thy +++ b/thys/Jinja/DFA/LBVCorrect.thy @@ -1,212 +1,221 @@ (* Author: Gerwin Klein Copyright 1999 Technische Universitaet Muenchen *) section \Correctness of the LBV\ theory LBVCorrect -imports LBVSpec Typing_Framework +imports LBVSpec Typing_Framework_1 begin locale lbvs = lbv + fixes s\<^sub>0 :: 'a fixes c :: "'a list" fixes ins :: "'b list" fixes \s :: "'a list" defines phi_def: "\s \ map (\pc. if c!pc = \ then wtl (take pc ins) c 0 s\<^sub>0 else c!pc) [0.. \ A" assumes pres: "pres_type step (size ins) A" lemma (in lbvs) phi_None [intro?]: "\ pc < size ins; c!pc = \ \ \ \s!pc = wtl (take pc ins) c 0 s\<^sub>0" (*<*) by (simp add: phi_def) (*>*) lemma (in lbvs) phi_Some [intro?]: "\ pc < size ins; c!pc \ \ \ \ \s!pc = c!pc" (*<*) by (simp add: phi_def) (*>*) lemma (in lbvs) phi_len [simp]: "size \s = size ins" (*<*) by (simp add: phi_def) (*>*) lemma (in lbvs) wtl_suc_pc: assumes all: "wtl ins c 0 s\<^sub>0 \ \" assumes pc: "pc+1 < size ins" + assumes sA: "s\<^sub>0 \ A" shows "wtl (take (pc+1) ins) c 0 s\<^sub>0 \\<^sub>r \s!(pc+1)" (*<*) proof - from all pc have "wtc c (pc+1) (wtl (take (pc+1) ins) c 0 s\<^sub>0) \ T" by (rule wtl_all) - with pc show ?thesis by (simp add: phi_def wtc split: if_split_asm) + with pc show ?thesis using sA pres cert all wtl_pres by (simp add: phi_def wtc split: if_split_asm) qed (*>*) lemma (in lbvs) wtl_stable: assumes wtl: "wtl ins c 0 s\<^sub>0 \ \" assumes s\<^sub>0: "s\<^sub>0 \ A" and pc: "pc < size ins" shows "stable r step \s pc" (*<*) proof (unfold stable_def, clarify) fix pc' s' assume step: "(pc',s') \ set (step pc (\s ! pc))" (is "(pc',s') \ set (?step pc)") from bounded pc step have pc': "pc' < size ins" by (rule boundedD) have tkpc: "wtl (take pc ins) c 0 s\<^sub>0 \ \" (is "?s\<^sub>1 \ _") using wtl by (rule wtl_take) have s\<^sub>2: "wtl (take (pc+1) ins) c 0 s\<^sub>0 \ \" (is "?s\<^sub>2 \ _") using wtl by (rule wtl_take) from wtl pc have wt_s\<^sub>1: "wtc c pc ?s\<^sub>1 \ \" by (rule wtl_all) have c_Some: "\pc t. pc < size ins \ c!pc \ \ \ \s!pc = c!pc" by (simp add: phi_def) have c_None: "c!pc = \ \ \s!pc = ?s\<^sub>1" using pc .. from wt_s\<^sub>1 pc c_None c_Some have inst: "wtc c pc ?s\<^sub>1 = wti c pc (\s!pc)" by (simp add: wtc split: if_split_asm) have "?s\<^sub>1 \ A" using pres cert s\<^sub>0 wtl pc by (rule wtl_pres) with pc c_Some cert c_None have "\s!pc \ A" by (cases "c!pc = \") (auto dest: cert_okD1) with pc pres have step_in_A: "snd`set (?step pc) \ A" by (auto dest: pres_typeD2) + then have inA1: "s' \ A" using step by auto show "s' \\<^sub>r \s!pc'" proof (cases "pc' = pc+1") case True with pc' cert have cert_in_A: "c!(pc+1) \ A" by (auto dest: cert_okD1) from True pc' have pc1: "pc+1 < size ins" by simp - with tkpc have "?s\<^sub>2 = wtc c pc ?s\<^sub>1" by - (rule wtl_Suc) + with pres cert s\<^sub>0 wtl have inA2: "wtl (take (pc + 1) ins) c 0 s\<^sub>0 \ A" by (auto dest:wtl_pres) + + have c_None': "c!(pc +1)= \ \ \s!(pc + 1)= ?s\<^sub>2" using pc1 .. + have "?s\<^sub>2 \ A" using pres cert s\<^sub>0 wtl pc1 by (rule wtl_pres) + with pc1 c_Some cert c_None' + have inA3: "\s!(pc+1) \ A" by (cases "c!(pc+1) = \") (auto dest: cert_okD1) + + from pc1 tkpc have "?s\<^sub>2 = wtc c pc ?s\<^sub>1" by - (rule wtl_Suc) with inst have merge: "?s\<^sub>2 = merge c pc (?step pc) (c!(pc+1))" by (simp add: wti) also from s\<^sub>2 merge have "\ \ \" (is "?merge \ _") by simp with cert_in_A step_in_A have "?merge = (map snd [(p',t') \ ?step pc. p'=pc+1] \\<^bsub>f\<^esub> c!(pc+1))" by (rule merge_not_top_s) finally have "s' \\<^sub>r ?s\<^sub>2" using step_in_A cert_in_A True step by (auto intro: pp_ub1') - also from wtl pc1 have "?s\<^sub>2 \\<^sub>r \s!(pc+1)" by (rule wtl_suc_pc) + also from wtl pc1 have "?s\<^sub>2 \\<^sub>r \s!(pc+1)" using s\<^sub>0 by (auto dest: wtl_suc_pc) also note True [symmetric] - finally show ?thesis by simp + finally show ?thesis using inA1 inA2 inA3 by simp next case False from wt_s\<^sub>1 inst have "merge c pc (?step pc) (c!(pc+1)) \ \" by (simp add: wti) with step_in_A have "\(pc', s')\set (?step pc). pc'\pc+1 \ s' \\<^sub>r c!pc'" by - (rule merge_not_top) with step False have ok: "s' \\<^sub>r c!pc'" by blast - moreover from ok have "c!pc' = \ \ s' = \" by simp + moreover from ok have "c!pc' = \ \ s' = \" using inA1 by simp moreover from c_Some pc' have "c!pc' \ \ \ \s!pc' = c!pc'" by auto ultimately show ?thesis by (cases "c!pc' = \") auto qed qed (*>*) lemma (in lbvs) phi_not_top: assumes wtl: "wtl ins c 0 s\<^sub>0 \ \" and pc: "pc < size ins" shows "\s!pc \ \" (*<*) proof (cases "c!pc = \") case False with pc have "\s!pc = c!pc" .. also from cert pc have "\ \ \" by (rule cert_okD4) finally show ?thesis . next case True with pc have "\s!pc = wtl (take pc ins) c 0 s\<^sub>0" .. also from wtl have "\ \ \" by (rule wtl_take) finally show ?thesis . qed (*>*) lemma (in lbvs) phi_in_A: assumes wtl: "wtl ins c 0 s\<^sub>0 \ \" and s\<^sub>0: "s\<^sub>0 \ A" shows "\s \ list (size ins) A" (*<*) proof - { fix x assume "x \ set \s" then obtain xs ys where "\s = xs @ x # ys" by (auto simp add: in_set_conv_decomp) then obtain pc where pc: "pc < size \s" and x: "\s!pc = x" by (simp add: that [of "size xs"] nth_append) from pres cert wtl s\<^sub>0 pc have "wtl (take pc ins) c 0 s\<^sub>0 \ A" by (auto intro!: wtl_pres) moreover from pc have "pc < size ins" by simp with cert have "c!pc \ A" .. ultimately have "\s!pc \ A" using pc by (simp add: phi_def) hence "x \ A" using x by simp } hence "set \s \ A" .. thus ?thesis by (unfold list_def) simp qed (*>*) lemma (in lbvs) phi0: - assumes wtl: "wtl ins c 0 s\<^sub>0 \ \" and 0: "0 < size ins" + assumes wtl: "wtl ins c 0 s\<^sub>0 \ \" and 0: "0 < size ins" and s\<^sub>0: "s\<^sub>0 \ A" shows "s\<^sub>0 \\<^sub>r \s!0" (*<*) proof (cases "c!0 = \") case True with 0 have "\s!0 = wtl (take 0 ins) c 0 s\<^sub>0" .. moreover have "wtl (take 0 ins) c 0 s\<^sub>0 = s\<^sub>0" by simp ultimately have "\s!0 = s\<^sub>0" by simp - thus ?thesis by simp + thus ?thesis using s\<^sub>0 by simp next case False with 0 have "\s!0 = c!0" .. moreover have "wtl (take 1 ins) c 0 s\<^sub>0 \ \" using wtl by (rule wtl_take) with 0 False have "s\<^sub>0 \\<^sub>r c!0" by (auto simp add: neq_Nil_conv wtc split: if_split_asm) ultimately show ?thesis by simp qed (*>*) theorem (in lbvs) wtl_sound: assumes wtl: "wtl ins c 0 s\<^sub>0 \ \" and s\<^sub>0: "s\<^sub>0 \ A" shows "\\s. wt_step r \ step \s" (*<*) proof - have "wt_step r \ step \s" proof (unfold wt_step_def, intro strip conjI) fix pc assume "pc < size \s" then obtain pc: "pc < size ins" by simp with wtl show "\s!pc \ \" by (rule phi_not_top) from wtl s\<^sub>0 pc show "stable r step \s pc" by (rule wtl_stable) qed thus ?thesis .. qed (*>*) theorem (in lbvs) wtl_sound_strong: assumes wtl: "wtl ins c 0 s\<^sub>0 \ \" assumes s\<^sub>0: "s\<^sub>0 \ A" and ins: "0 < size ins" shows "\\s \ list (size ins) A. wt_step r \ step \s \ s\<^sub>0 \\<^sub>r \s!0" (*<*) proof - have "\s \ list (size ins) A" using wtl s\<^sub>0 by (rule phi_in_A) moreover have "wt_step r \ step \s" proof (unfold wt_step_def, intro strip conjI) fix pc assume "pc < size \s" then obtain pc: "pc < size ins" by simp with wtl show "\s!pc \ \" by (rule phi_not_top) from wtl s\<^sub>0 and pc show "stable r step \s pc" by (rule wtl_stable) qed - moreover from wtl ins have "s\<^sub>0 \\<^sub>r \s!0" by (rule phi0) + moreover from wtl ins have "s\<^sub>0 \\<^sub>r \s!0" using s\<^sub>0 by (rule phi0) ultimately show ?thesis by fast qed (*>*) end diff --git a/thys/Jinja/DFA/LBVSpec.thy b/thys/Jinja/DFA/LBVSpec.thy --- a/thys/Jinja/DFA/LBVSpec.thy +++ b/thys/Jinja/DFA/LBVSpec.thy @@ -1,396 +1,398 @@ (* Title: HOL/MicroJava/BV/LBVSpec.thy Author: Gerwin Klein Copyright 1999 Technische Universitaet Muenchen *) section \The Lightweight Bytecode Verifier\ theory LBVSpec imports SemilatAlg Opt begin type_synonym 's certificate = "'s list" primrec merge :: "'s certificate \ 's binop \ 's ord \ 's \ nat \ (nat \ 's) list \ 's \ 's" where "merge cert f r T pc [] x = x" | "merge cert f r T pc (s#ss) x = merge cert f r T pc ss (let (pc',s') = s in if pc'=pc+1 then s' \\<^sub>f x else if s' \\<^sub>r cert!pc' then x else T)" definition wtl_inst :: "'s certificate \ 's binop \ 's ord \ 's \ 's step_type \ nat \ 's \ 's" where "wtl_inst cert f r T step pc s = merge cert f r T pc (step pc s) (cert!(pc+1))" definition wtl_cert :: "'s certificate \ 's binop \ 's ord \ 's \ 's \ 's step_type \ nat \ 's \ 's" where "wtl_cert cert f r T B step pc s = (if cert!pc = B then wtl_inst cert f r T step pc s else if s \\<^sub>r cert!pc then wtl_inst cert f r T step pc (cert!pc) else T)" primrec wtl_inst_list :: "'a list \ 's certificate \ 's binop \ 's ord \ 's \ 's \ 's step_type \ nat \ 's \ 's" where "wtl_inst_list [] cert f r T B step pc s = s" | "wtl_inst_list (i#is) cert f r T B step pc s = (let s' = wtl_cert cert f r T B step pc s in if s' = T \ s = T then T else wtl_inst_list is cert f r T B step (pc+1) s')" definition cert_ok :: "'s certificate \ nat \ 's \ 's \ 's set \ bool" where "cert_ok cert n T B A \ (\i < n. cert!i \ A \ cert!i \ T) \ (cert!n = B)" definition bottom :: "'a ord \ 'a \ bool" where "bottom r B \ (\x. B \\<^sub>r x)" locale lbv = Semilat + fixes T :: "'a" ("\") fixes B :: "'a" ("\") fixes step :: "'a step_type" assumes top: "top r \" assumes T_A: "\ \ A" assumes bot: "bottom r \" assumes B_A: "\ \ A" fixes merge :: "'a certificate \ nat \ (nat \ 'a) list \ 'a \ 'a" defines mrg_def: "merge cert \ LBVSpec.merge cert f r \" fixes wti :: "'a certificate \ nat \ 'a \ 'a" defines wti_def: "wti cert \ wtl_inst cert f r \ step" fixes wtc :: "'a certificate \ nat \ 'a \ 'a" defines wtc_def: "wtc cert \ wtl_cert cert f r \ \ step" fixes wtl :: "'b list \ 'a certificate \ nat \ 'a \ 'a" defines wtl_def: "wtl ins cert \ wtl_inst_list ins cert f r \ \ step" lemma (in lbv) wti: "wti c pc s = merge c pc (step pc s) (c!(pc+1))" (*<*) by (simp add: wti_def mrg_def wtl_inst_def) (*>*) lemma (in lbv) wtc: "wtc c pc s = (if c!pc = \ then wti c pc s else if s \\<^sub>r c!pc then wti c pc (c!pc) else \)" (*<*) by (unfold wtc_def wti_def wtl_cert_def) rule (*>*) lemma cert_okD1 [intro?]: "cert_ok c n T B A \ pc < n \ c!pc \ A" (*<*) by (unfold cert_ok_def) fast (*>*) lemma cert_okD2 [intro?]: "cert_ok c n T B A \ c!n = B" (*<*) by (simp add: cert_ok_def) (*>*) lemma cert_okD3 [intro?]: "cert_ok c n T B A \ B \ A \ pc < n \ c!Suc pc \ A" (*<*) by (drule Suc_leI) (auto simp add: le_eq_less_or_eq dest: cert_okD1 cert_okD2) (*>*) lemma cert_okD4 [intro?]: "cert_ok c n T B A \ pc < n \ c!pc \ T" (*<*) by (simp add: cert_ok_def) (*>*) declare Let_def [simp] subsection "more semilattice lemmas" lemma (in lbv) sup_top [simp, elim]: assumes x: "x \ A" shows "x \\<^sub>f \ = \" (*<*) proof - - from top have "x \\<^sub>f \ \\<^sub>r \" .. - moreover from x T_A have "\ \\<^sub>r x \\<^sub>f \" .. - ultimately show ?thesis .. + from semilat have "order r A" by auto + moreover from x T_A have xfT_inA: "x \\<^sub>f \ \ A" by auto + with T_A top have xfT: "x \\<^sub>f \ \\<^sub>r \" by (simp add:top_def) + moreover from x T_A have "\ \\<^sub>r x \\<^sub>f \" .. + ultimately show ?thesis using xfT_inA T_A by (rule order_antisym) qed (*>*) lemma (in lbv) plusplussup_top [simp, elim]: "set xs \ A \ xs \\<^bsub>f\<^esub> \ = \" by (induct xs) auto lemma (in Semilat) pp_ub1': assumes S: "snd`set S \ A" assumes y: "y \ A" and ab: "(a, b) \ set S" shows "b \\<^sub>r map snd [(p', t') \ S . p' = a] \\<^bsub>f\<^esub> y" (*<*) proof - from S have "\(x,y) \ set S. y \ A" by auto with Semilat_axioms show ?thesis using y ab by (rule ub1') qed (*>*) lemma (in lbv) bottom_le [simp, intro!]: "\ \\<^sub>r x" by (insert bot) (simp add: bottom_def) -lemma (in lbv) le_bottom [simp]: "x \\<^sub>r \ = (x = \)" - by (blast intro: antisym_r) +lemma (in lbv) le_bottom [simp]: "x \ A \ x \\<^sub>r \ = (x = \)" + using B_A by (blast intro: antisym_r) subsection "merge" lemma (in lbv) merge_Nil [simp]: "merge c pc [] x = x" by (simp add: mrg_def) lemma (in lbv) merge_Cons [simp]: "merge c pc (l#ls) x = merge c pc ls (if fst l=pc+1 then snd l +_f x else if snd l \\<^sub>r c!fst l then x else \)" by (simp add: mrg_def split_beta) lemma (in lbv) merge_Err [simp]: "snd`set ss \ A \ merge c pc ss \ = \" by (induct ss) auto lemma (in lbv) merge_not_top: "\x. snd`set ss \ A \ merge c pc ss x \ \ \ \(pc',s') \ set ss. (pc' \ pc+1 \ s' \\<^sub>r c!pc')" (is "\x. ?set ss \ ?merge ss x \ ?P ss") (*<*) proof (induct ss) show "?P []" by simp next fix x ls l assume "?set (l#ls)" then obtain set: "snd`set ls \ A" by simp assume merge: "?merge (l#ls) x" moreover obtain pc' s' where [simp]: "l = (pc',s')" by (cases l) ultimately obtain x' where merge': "?merge ls x'" by simp assume "\x. ?set ls \ ?merge ls x \ ?P ls" hence "?P ls" using set merge' . moreover from merge set have "pc' \ pc+1 \ s' \\<^sub>r c!pc'" by (simp split: if_split_asm) ultimately show "?P (l#ls)" by simp qed (*>*) lemma (in lbv) merge_def: shows "\x. x \ A \ snd`set ss \ A \ merge c pc ss x = (if \(pc',s') \ set ss. pc'\pc+1 \ s' \\<^sub>r c!pc' then map snd [(p',t') \ ss. p'=pc+1] \\<^bsub>f\<^esub> x else \)" (is "\x. _ \ _ \ ?merge ss x = ?if ss x" is "\x. _ \ _ \ ?P ss x") (*<*) proof (induct ss) fix x show "?P [] x" by simp next fix x assume x: "x \ A" fix l::"nat \ 'a" and ls assume "snd`set (l#ls) \ A" then obtain l: "snd l \ A" and ls: "snd`set ls \ A" by auto assume "\x. x \ A \ snd`set ls \ A \ ?P ls x" hence IH: "\x. x \ A \ ?P ls x" using ls by iprover obtain pc' s' where [simp]: "l = (pc',s')" by (cases l) hence "?merge (l#ls) x = ?merge ls (if pc'=pc+1 then s' \\<^sub>f x else if s' \\<^sub>r c!pc' then x else \)" (is "?merge (l#ls) x = ?merge ls ?if'") by simp also have "\ = ?if ls ?if'" proof - from l have "s' \ A" by simp with x have "s' \\<^sub>f x \ A" by simp with x T_A have "?if' \ A" by auto hence "?P ls ?if'" by (rule IH) thus ?thesis by simp qed also have "\ = ?if (l#ls) x" proof (cases "\(pc', s')\set (l#ls). pc'\pc+1 \ s' \\<^sub>r c!pc'") case True hence "\(pc', s')\set ls. pc'\pc+1 \ s' \\<^sub>r c!pc'" by auto moreover from True have "map snd [(p',t') \ ls . p'=pc+1] \\<^bsub>f\<^esub> ?if' = (map snd [(p',t') \ l#ls . p'=pc+1] \\<^bsub>f\<^esub> x)" by simp ultimately show ?thesis using True by simp next case False moreover from ls have "set (map snd [(p', t') \ ls . p' = Suc pc]) \ A" by auto ultimately show ?thesis by auto qed finally show "?P (l#ls) x" . qed (*>*) lemma (in lbv) merge_not_top_s: assumes x: "x \ A" and ss: "snd`set ss \ A" assumes m: "merge c pc ss x \ \" shows "merge c pc ss x = (map snd [(p',t') \ ss. p'=pc+1] \\<^bsub>f\<^esub> x)" (*<*) proof - from ss m have "\(pc',s') \ set ss. (pc' \ pc+1 \ s' <=_r c!pc')" by (rule merge_not_top) with x ss m show ?thesis by - (drule merge_def, auto split: if_split_asm) qed (*>*) subsection "wtl-inst-list" lemmas [iff] = not_Err_eq lemma (in lbv) wtl_Nil [simp]: "wtl [] c pc s = s" by (simp add: wtl_def) lemma (in lbv) wtl_Cons [simp]: "wtl (i#is) c pc s = (let s' = wtc c pc s in if s' = \ \ s = \ then \ else wtl is c (pc+1) s')" by (simp add: wtl_def wtc_def) lemma (in lbv) wtl_Cons_not_top: "wtl (i#is) c pc s \ \ = (wtc c pc s \ \ \ s \ T \ wtl is c (pc+1) (wtc c pc s) \ \)" by (auto simp del: split_paired_Ex) lemma (in lbv) wtl_top [simp]: "wtl ls c pc \ = \" by (cases ls) auto lemma (in lbv) wtl_not_top: "wtl ls c pc s \ \ \ s \ \" by (cases "s=\") auto lemma (in lbv) wtl_append [simp]: "\pc s. wtl (a@b) c pc s = wtl b c (pc+length a) (wtl a c pc s)" by (induct a) auto lemma (in lbv) wtl_take: "wtl is c pc s \ \ \ wtl (take pc' is) c pc s \ \" (is "?wtl is \ _ \ _") (*<*) proof - assume "?wtl is \ \" hence "?wtl (take pc' is @ drop pc' is) \ \" by simp thus ?thesis by (auto dest!: wtl_not_top simp del: append_take_drop_id) qed (*>*) lemma take_Suc: "\n. n < length l \ take (Suc n) l = (take n l)@[l!n]" (is "?P l") (*<*) proof (induct l) show "?P []" by simp next fix x xs assume IH: "?P xs" show "?P (x#xs)" proof (intro strip) fix n assume "n < length (x#xs)" with IH show "take (Suc n) (x # xs) = take n (x # xs) @ [(x # xs) ! n]" by (cases n, auto) qed qed (*>*) lemma (in lbv) wtl_Suc: assumes suc: "pc+1 < length is" assumes wtl: "wtl (take pc is) c 0 s \ \" shows "wtl (take (pc+1) is) c 0 s = wtc c pc (wtl (take pc is) c 0 s)" (*<*) proof - from suc have "take (pc+1) is=(take pc is)@[is!pc]" by (simp add: take_Suc) with suc wtl show ?thesis by (simp add: min_def) qed (*>*) lemma (in lbv) wtl_all: assumes all: "wtl is c 0 s \ \" (is "?wtl is \ _") assumes pc: "pc < length is" shows "wtc c pc (wtl (take pc is) c 0 s) \ \" (*<*) proof - from pc have "0 < length (drop pc is)" by simp then obtain i r where Cons: "drop pc is = i#r" by (auto simp add: neq_Nil_conv simp del: length_drop drop_eq_Nil) hence "i#r = drop pc is" .. with all have take: "?wtl (take pc is@i#r) \ \" by simp from pc have "is!pc = drop pc is ! 0" by simp with Cons have "is!pc = i" by simp with take pc show ?thesis by (auto simp add: min_def split: if_split_asm) qed (*>*) subsection "preserves-type" lemma (in lbv) merge_pres: assumes s0: "snd`set ss \ A" and x: "x \ A" shows "merge c pc ss x \ A" (*<*) proof - from s0 have "set (map snd [(p', t') \ ss . p'=pc+1]) \ A" by auto with x semilat Semilat_axioms have "(map snd [(p', t') \ ss . p'=pc+1] \\<^bsub>f\<^esub> x) \ A" by (auto intro!: plusplus_closed) with s0 x show ?thesis by (simp add: merge_def T_A) qed (*>*) lemma pres_typeD2: "pres_type step n A \ s \ A \ p < n \ snd`set (step p s) \ A" by auto (drule pres_typeD) lemma (in lbv) wti_pres [intro?]: assumes pres: "pres_type step n A" assumes cert: "c!(pc+1) \ A" assumes s_pc: "s \ A" "pc < n" shows "wti c pc s \ A" (*<*) proof - from pres s_pc have "snd`set (step pc s) \ A" by (rule pres_typeD2) with cert show ?thesis by (simp add: wti merge_pres) qed (*>*) lemma (in lbv) wtc_pres: assumes "pres_type step n A" assumes "c!pc \ A" and "c!(pc+1) \ A" assumes "s \ A" and "pc < n" shows "wtc c pc s \ A" (*<*) proof - have "wti c pc s \ A" using assms(1,3-5) .. moreover have "wti c pc (c!pc) \ A" using assms(1,3,2,5) .. ultimately show ?thesis using T_A by (simp add: wtc) qed (*>*) lemma (in lbv) wtl_pres: assumes pres: "pres_type step (length is) A" assumes cert: "cert_ok c (length is) \ \ A" assumes s: "s \ A" assumes all: "wtl is c 0 s \ \" shows "pc < length is \ wtl (take pc is) c 0 s \ A" (is "?len pc \ ?wtl pc \ A") (*<*) proof (induct pc) from s show "?wtl 0 \ A" by simp next fix n assume Suc_n: "Suc n < length is" hence n1: "n+1 < length is" by simp then obtain n: "n < length is" by simp assume "n < length is \ ?wtl n \ A" hence "?wtl n \ A" using n . from pres _ _ this n have "wtc c n (?wtl n) \ A" proof (rule wtc_pres) from cert n show "c!n \ A" by (rule cert_okD1) from cert n1 show "c!(n+1) \ A" by (rule cert_okD1) qed also from all n have "?wtl n \ \" by - (rule wtl_take) with n1 have "wtc c n (?wtl n) = ?wtl (n+1)" by (rule wtl_Suc [symmetric]) finally show "?wtl (Suc n) \ A" by simp qed (*>*) end diff --git a/thys/Jinja/DFA/Listn.thy b/thys/Jinja/DFA/Listn.thy --- a/thys/Jinja/DFA/Listn.thy +++ b/thys/Jinja/DFA/Listn.thy @@ -1,632 +1,710 @@ (* Title: HOL/MicroJava/BV/Listn.thy Author: Tobias Nipkow Copyright 2000 TUM Lists of a fixed length. *) section \Fixed Length Lists\ theory Listn imports Err begin definition list :: "nat \ 'a set \ 'a list set" where "list n A = {xs. size xs = n \ set xs \ A}" definition le :: "'a ord \ ('a list)ord" where "le r = list_all2 (\x y. x \\<^sub>r y)" abbreviation lesublist :: "'a list \ 'a ord \ 'a list \ bool" ("(_ /[\\<^bsub>_\<^esub>] _)" [50, 0, 51] 50) where "x [\\<^bsub>r\<^esub>] y == x <=_(Listn.le r) y" abbreviation lesssublist :: "'a list \ 'a ord \ 'a list \ bool" ("(_ /[\\<^bsub>_\<^esub>] _)" [50, 0, 51] 50) where "x [\\<^bsub>r\<^esub>] y == x <_(Listn.le r) y" (*<*) notation (ASCII) lesublist ("(_ /[<=_] _)" [50, 0, 51] 50) and lesssublist ("(_ /[<_] _)" [50, 0, 51] 50) abbreviation (input) lesublist2 :: "'a list \ 'a ord \ 'a list \ bool" ("(_ /[\\<^sub>_] _)" [50, 0, 51] 50) where "x [\\<^sub>r] y == x [\\<^bsub>r\<^esub>] y" abbreviation (input) lesssublist2 :: "'a list \ 'a ord \ 'a list \ bool" ("(_ /[\\<^sub>_] _)" [50, 0, 51] 50) where "x [\\<^sub>r] y == x [\\<^bsub>r\<^esub>] y" (*>*) abbreviation plussublist :: "'a list \ ('a \ 'b \ 'c) \ 'b list \ 'c list" ("(_ /[\\<^bsub>_\<^esub>] _)" [65, 0, 66] 65) where "x [\\<^bsub>f\<^esub>] y == x \\<^bsub>map2 f\<^esub> y" (*<*) notation (ASCII) plussublist ("(_ /[+_] _)" [65, 0, 66] 65) abbreviation (input) plussublist2 :: "'a list \ ('a \ 'b \ 'c) \ 'b list \ 'c list" ("(_ /[\\<^sub>_] _)" [65, 0, 66] 65) where "x [\\<^sub>f] y == x [\\<^bsub>f\<^esub>] y" (*>*) primrec coalesce :: "'a err list \ 'a list err" where "coalesce [] = OK[]" | "coalesce (ex#exs) = Err.sup (#) ex (coalesce exs)" definition sl :: "nat \ 'a sl \ 'a list sl" where "sl n = (\(A,r,f). (list n A, le r, map2 f))" definition sup :: "('a \ 'b \ 'c err) \ 'a list \ 'b list \ 'c list err" where "sup f = (\xs ys. if size xs = size ys then coalesce(xs [\\<^bsub>f\<^esub>] ys) else Err)" definition upto_esl :: "nat \ 'a esl \ 'a list esl" where "upto_esl m = (\(A,r,f). (Union{list n A |n. n \ m}, le r, sup f))" lemmas [simp] = set_update_subsetI lemma unfold_lesub_list: "xs [\\<^bsub>r\<^esub>] ys = Listn.le r xs ys" (*<*) by (simp add: lesub_def) (*>*) lemma Nil_le_conv [iff]: "([] [\\<^bsub>r\<^esub>] ys) = (ys = [])" (*<*) apply (unfold lesub_def Listn.le_def) apply simp done (*>*) lemma Cons_notle_Nil [iff]: "\ x#xs [\\<^bsub>r\<^esub>] []" (*<*) apply (unfold lesub_def Listn.le_def) apply simp done (*>*) lemma Cons_le_Cons [iff]: "x#xs [\\<^bsub>r\<^esub>] y#ys = (x \\<^sub>r y \ xs [\\<^bsub>r\<^esub>] ys)" (*<*) by (simp add: lesub_def Listn.le_def) (*>*) -lemma Cons_less_Conss [simp]: - "order r \ x#xs [\\<^sub>r] y#ys = (x \\<^sub>r y \ xs [\\<^bsub>r\<^esub>] ys \ x = y \ xs [\\<^sub>r] ys)" -(*<*) -apply (unfold lesssub_def) -apply blast -done -(*>*) + + + + + + + lemma list_update_le_cong: "\ i\<^bsub>r\<^esub>] ys; x \\<^sub>r y \ \ xs[i:=x] [\\<^bsub>r\<^esub>] ys[i:=y]" (*<*) apply (unfold unfold_lesub_list) apply (unfold Listn.le_def) apply (simp add: list_all2_update_cong) done (*>*) lemma le_listD: "\ xs [\\<^bsub>r\<^esub>] ys; p < size xs \ \ xs!p \\<^sub>r ys!p" (*<*) by (simp add: Listn.le_def lesub_def list_all2_nthD) (*>*) lemma le_list_refl: "\x. x \\<^sub>r x \ xs [\\<^bsub>r\<^esub>] xs" (*<*) apply (simp add: unfold_lesub_list lesub_def Listn.le_def list_all2_refl) done (*>*) -lemma le_list_trans: "\ order r; xs [\\<^bsub>r\<^esub>] ys; ys [\\<^bsub>r\<^esub>] zs \ \ xs [\\<^bsub>r\<^esub>] zs" + +lemma le_list_trans: + assumes ord: "order r A" + and xs: "xs \ list n A" and ys: "ys \ list n A" and zs: "zs \ list n A" + and "xs [\\<^bsub>r\<^esub>] ys" and "ys [\\<^bsub>r\<^esub>] zs" + shows "xs [\\<^bsub>r\<^esub>] zs" (*<*) + using assms +proof (unfold le_def lesssub_def lesub_def) + assume "list_all2 r xs ys" and "list_all2 r ys zs" + hence xs_ys_zs: "\i < length xs. r (xs!i) (ys!i) \ r (ys!i) (zs!i)" + and len_xs_zs: "length xs = length zs" by (auto simp add: list_all2_conv_all_nth) + from xs ys zs have inA: "\i A \ ys!i \ A \ zs!i \ A" by (unfold list_def) auto + + from ord have "\x\A. \y\A. \z\A. x \\<^sub>r y \ y \\<^sub>r z \ x \\<^sub>r z" by (unfold order_def) blast + hence "\x\A. \y\A. \z\A. r x y \ r y z \ r x z" by (unfold lesssub_def lesub_def) + with xs_ys_zs inA have "\i list n A" and ys: "ys \ list n A" + and "xs [\\<^bsub>r\<^esub>] ys" and "ys [\\<^bsub>r\<^esub>] xs" + shows "xs = ys" +(*<*) + using assms +proof (simp add: le_def lesssub_def lesub_def) + assume "list_all2 r xs ys" and "list_all2 r ys xs" + hence xs_ys: "\i < length xs. r (xs!i) (ys!i) \ r (ys!i) (xs!i)" + and len_xs_ys: "length xs = length ys" by (auto simp add: list_all2_conv_all_nth) + from xs ys have inA: "\i A \ ys!i \ A" by (unfold list_def) auto + + from ord have "\x\A. \y\A. x \\<^sub>r y \ y \\<^sub>r x \ x = y" by (unfold order_def) blast + hence "\x\A. \y\A. r x y \ r y x \ x = y" by (unfold lesssub_def lesub_def) + with xs_ys inA have "\i*) + +lemma listE_set [simp]: "xs \ list n A \ set xs \ A" +(*<*) +apply (unfold list_def) +apply blast done (*>*) -lemma le_list_antisym: "\ order r; xs [\\<^bsub>r\<^esub>] ys; ys [\\<^bsub>r\<^esub>] xs \ \ xs = ys" + +lemma nth_in [rule_format, simp]: + "\i n. size xs = n \ set xs \ A \ i < n \ (xs!i) \ A" (*<*) -apply (unfold unfold_lesub_list) -apply (unfold Listn.le_def) -apply (rule list_all2_antisym) -apply (rule order_antisym) -apply assumption+ +apply (induct "xs") + apply simp +apply (simp add: nth_Cons split: nat.split) done (*>*) -lemma order_listI [simp, intro!]: "order r \ order(Listn.le r)" +(* FIXME: remove simp *) +lemma listE_length [simp]: "xs \ list n A \ size xs = n" (*<*) -apply (subst order_def) -apply (blast intro: le_list_refl le_list_trans le_list_antisym - dest: order_refl) +apply (unfold list_def) +apply blast done (*>*) +lemma listE_nth_in: "\ xs \ list n A; i < n \ \ xs!i \ A" +(*<*) by auto (*>*) + + +lemma le_list_refl': "order r A \ xs \ list n A \ xs \\<^bsub>Listn.le r\<^esub> xs" + apply (unfold le_def lesssub_def lesub_def list_all2_conv_all_nth) + apply auto + apply (subgoal_tac "xs ! i \ A") + apply (subgoal_tac "(xs ! i) \\<^sub>r (xs ! i)") + apply (simp only: lesssub_def lesub_def) + apply (fastforce dest: listE_nth_in) + apply (fastforce dest: order_refl) + done + +lemma order_listI [simp, intro!]: "order r A \ order(Listn.le r) (list n A)" +(*<*) +proof- + assume ord: "order r A" + let ?r = "Listn.le r" + let ?A = "list n A" + have "\x\?A. x \\<^bsub>?r\<^esub> x" using ord le_list_refl' by auto + moreover have "\x\?A. \y\?A. x \\<^bsub>?r\<^esub> y \ y \\<^bsub>?r\<^esub> x \ x=y" using ord le_list_antisym by auto + moreover have "\x\?A. \y\?A. \z\?A. x \\<^bsub>?r\<^esub>y \ y \\<^bsub>?r\<^esub> z \ x \\<^bsub>?r\<^esub> z" using ord le_list_trans by auto + ultimately show ?thesis by (auto simp only: order_def) +qed +(*>*) + +lemma le_list_refl2': "order r A \ xs \ (\{list n A |n. n \ mxs})\ xs \\<^bsub>Listn.le r\<^esub> xs" + by (auto simp add:le_list_refl') +lemma le_list_trans2: + assumes "order r A" + and "xs \ (\{list n A |n. n \ mxs})" and "ys \(\{list n A |n. n \ mxs})" and "zs \(\{list n A |n. n \ mxs})" + and "xs [\\<^bsub>r\<^esub>] ys" and "ys [\\<^bsub>r\<^esub>] zs" + shows "xs [\\<^bsub>r\<^esub>] zs" +(*<*)using assms +proof(auto simp only:list_def le_def lesssub_def lesub_def) + assume xA: "set xs \ A" and "length xs \ mxs " and " length ys \ mxs " + and yA: "set ys \ A" and len_zs: "length zs \ mxs" + and zA: "set zs \ A" and xy: "list_all2 r xs ys" and yz: "list_all2 r ys zs" + and ord: "order r A" + from xy yz have xs_ys: "length xs = length ys" and ys_zs: "length ys = length zs" by (auto simp add:list_all2_conv_all_nth) + from ord have "\x\A. \y\A. \z\A. x \\<^sub>r y \ y \\<^sub>r z \ x \\<^sub>r z" by (unfold order_def) blast + hence trans: "\x\A. \y\A. \z\A. r x y \ r y z \ r x z" by (simp only:lesssub_def lesub_def) + + from xA yA zA have x_inA: "\i A" + and y_inA: "\i A" + and z_inA: "\i A" using xs_ys ys_zs by auto + + from xy yz have "\i < length xs. r (xs!i) (ys!i)" + and "\i < length ys. r (ys!i) (zs!i)" by (auto simp add:list_all2_conv_all_nth) + hence "\i < length xs. r (xs!i) (ys!i) \ r (ys!i) (zs!i)" using xs_ys ys_zs by auto + + with x_inA y_inA z_inA + have "\i(\{list n A |n. n \ mxs})" and "ys \(\{list n A |n. n \ mxs})" + and "xs [\\<^bsub>r\<^esub>] ys" and "ys [\\<^bsub>r\<^esub>] xs" + shows " xs = ys" +(*<*) + using assms +proof(auto simp only:list_def le_def lesssub_def lesub_def) + assume xA: "set xs \ A" and len_ys: "length ys \ mxs" and len_xs: "length xs \ mxs" + and yA: "set ys \ A" and xy: "list_all2 r xs ys" and yx: "list_all2 r ys xs" + and ord: "order r A" + from xy have len_eq_xs_ys: "length xs = length ys" by (simp add:list_all2_conv_all_nth) + + from ord have antisymm:"\x\A. \y\A. r x y \ r y x\ x=y " by (auto simp only:lesssub_def lesub_def order_def) + from xA yA have x_inA: "\i A" and y_inA: "\i A" by auto + from xy yx have "\i < length xs. r (xs!i) (ys!i)" and "\i < length ys. r (ys!i) (xs!i)" by (auto simp add:list_all2_conv_all_nth) + with x_inA y_inA + have "\i order(Listn.le r) (\{list n A |n. n \ mxs})" +(*<*) +proof- + assume ord: "order r A" + let ?r = "Listn.le r" + let ?A = "(\{list n A |n. n \ mxs})" + have "\x\?A. x \\<^bsub>?r\<^esub> x" using ord le_list_refl2' by auto \\ use order_listI \ + moreover have "\x\?A. \y\?A. x \\<^bsub>?r\<^esub> y \ y \\<^bsub>?r\<^esub> x \ x=y" using ord le_list_antisym2 by blast + moreover have "\x\?A. \y\?A. \z\?A. x \\<^bsub>?r\<^esub>y \ y \\<^bsub>?r\<^esub> z \ x \\<^bsub>?r\<^esub> z" using ord le_list_trans2 by blast + ultimately show ?thesis by (auto simp only: order_def) +qed + + lemma lesub_list_impl_same_size [simp]: "xs [\\<^bsub>r\<^esub>] ys \ size ys = size xs" (*<*) apply (unfold Listn.le_def lesub_def) apply (simp add: list_all2_lengthD) done (*>*) lemma lesssub_lengthD: "xs [\\<^sub>r] ys \ size ys = size xs" (*<*) apply (unfold lesssub_def) apply auto done (*>*) + lemma le_list_appendI: "a [\\<^bsub>r\<^esub>] b \ c [\\<^bsub>r\<^esub>] d \ a@c [\\<^bsub>r\<^esub>] b@d" (*<*) apply (unfold Listn.le_def lesub_def) apply (rule list_all2_appendI, assumption+) done (*>*) lemma le_listI: assumes "length a = length b" assumes "\n. n < length a \ a!n \\<^sub>r b!n" shows "a [\\<^bsub>r\<^esub>] b" (*<*) proof - from assms have "list_all2 r a b" by (simp add: list_all2_all_nthI lesub_def) then show ?thesis by (simp add: Listn.le_def lesub_def) qed (*>*) lemma listI: "\ size xs = n; set xs \ A \ \ xs \ list n A" (*<*) apply (unfold list_def) apply blast done (*>*) -(* FIXME: remove simp *) -lemma listE_length [simp]: "xs \ list n A \ size xs = n" -(*<*) -apply (unfold list_def) -apply blast -done -(*>*) - lemma less_lengthI: "\ xs \ list n A; p < n \ \ p < size xs" (*<*) by simp (*>*) -lemma listE_set [simp]: "xs \ list n A \ set xs \ A" -(*<*) -apply (unfold list_def) -apply blast -done -(*>*) - lemma list_0 [simp]: "list 0 A = {[]}" (*<*) apply (unfold list_def) apply auto done (*>*) lemma in_list_Suc_iff: "(xs \ list (Suc n) A) = (\y\A. \ys \ list n A. xs = y#ys)" (*<*) apply (unfold list_def) apply (case_tac "xs") apply auto done (*>*) + lemma Cons_in_list_Suc [iff]: "(x#xs \ list (Suc n) A) = (x\A \ xs \ list n A)" (*<*) apply (simp add: in_list_Suc_iff) done (*>*) lemma list_not_empty: "\a. a\A \ \xs. xs \ list n A" (*<*) apply (induct "n") apply simp apply (simp add: in_list_Suc_iff) apply blast done (*>*) - -lemma nth_in [rule_format, simp]: - "\i n. size xs = n \ set xs \ A \ i < n \ (xs!i) \ A" -(*<*) -apply (induct "xs") - apply simp -apply (simp add: nth_Cons split: nat.split) -done -(*>*) - -lemma listE_nth_in: "\ xs \ list n A; i < n \ \ xs!i \ A" -(*<*) by auto (*>*) - lemma listn_Cons_Suc [elim!]: "l#xs \ list n A \ (\n'. n = Suc n' \ l \ A \ xs \ list n' A \ P) \ P" (*<*) by (cases n) auto (*>*) lemma listn_appendE [elim!]: "a@b \ list n A \ (\n1 n2. n=n1+n2 \ a \ list n1 A \ b \ list n2 A \ P) \ P" (*<*) proof - have "\n. a@b \ list n A \ \n1 n2. n=n1+n2 \ a \ list n1 A \ b \ list n2 A" (is "\n. ?list a n \ \n1 n2. ?P a n n1 n2") proof (induct a) fix n assume "?list [] n" hence "?P [] n 0 n" by simp thus "\n1 n2. ?P [] n n1 n2" by fast next fix n l ls assume "?list (l#ls) n" then obtain n' where n: "n = Suc n'" "l \ A" and n': "ls@b \ list n' A" by fastforce assume "\n. ls @ b \ list n A \ \n1 n2. n = n1 + n2 \ ls \ list n1 A \ b \ list n2 A" from this and n' have "\n1 n2. n' = n1 + n2 \ ls \ list n1 A \ b \ list n2 A" . then obtain n1 n2 where "n' = n1 + n2" "ls \ list n1 A" "b \ list n2 A" by fast with n have "?P (l#ls) n (n1+1) n2" by simp thus "\n1 n2. ?P (l#ls) n n1 n2" by fastforce qed moreover assume "a@b \ list n A" "\n1 n2. n=n1+n2 \ a \ list n1 A \ b \ list n2 A \ P" ultimately show ?thesis by blast qed (*>*) lemma listt_update_in_list [simp, intro!]: "\ xs \ list n A; x\A \ \ xs[i := x] \ list n A" (*<*) apply (unfold list_def) apply simp done (*>*) lemma list_appendI [intro?]: "\ a \ list n A; b \ list m A \ \ a @ b \ list (n+m) A" (*<*) by (unfold list_def) auto (*>*) lemma list_map [simp]: "(map f xs \ list (size xs) A) = (f ` set xs \ A)" (*<*) by (unfold list_def) simp (*>*) lemma list_replicateI [intro]: "x \ A \ replicate n x \ list n A" (*<*) by (induct n) auto (*>*) lemma plus_list_Nil [simp]: "[] [\\<^bsub>f\<^esub>] xs = []" (*<*) apply (unfold plussub_def) apply simp done (*>*) lemma plus_list_Cons [simp]: "(x#xs) [\\<^bsub>f\<^esub>] ys = (case ys of [] \ [] | y#ys \ (x \\<^sub>f y)#(xs [\\<^bsub>f\<^esub>] ys))" (*<*) by (simp add: plussub_def split: list.split) (*>*) lemma length_plus_list [rule_format, simp]: "\ys. size(xs [\\<^bsub>f\<^esub>] ys) = min(size xs) (size ys)" (*<*) apply (induct xs) apply simp apply clarify apply (simp (no_asm_simp) split: list.split) done (*>*) lemma nth_plus_list [rule_format, simp]: "\xs ys i. size xs = n \ size ys = n \ i (xs [\\<^bsub>f\<^esub>] ys)!i = (xs!i) \\<^sub>f (ys!i)" (*<*) apply (induct n) apply simp apply clarify apply (case_tac xs) apply simp apply (force simp add: nth_Cons split: list.split nat.split) done (*>*) lemma (in Semilat) plus_list_ub1 [rule_format]: "\ set xs \ A; set ys \ A; size xs = size ys \ \ xs [\\<^bsub>r\<^esub>] xs [\\<^bsub>f\<^esub>] ys" (*<*) apply (unfold unfold_lesub_list) apply (simp add: Listn.le_def list_all2_conv_all_nth) done (*>*) lemma (in Semilat) plus_list_ub2: "\set xs \ A; set ys \ A; size xs = size ys \ \ ys [\\<^bsub>r\<^esub>] xs [\\<^bsub>f\<^esub>] ys" (*<*) apply (unfold unfold_lesub_list) apply (simp add: Listn.le_def list_all2_conv_all_nth) done (*>*) lemma (in Semilat) plus_list_lub [rule_format]: shows "\xs ys zs. set xs \ A \ set ys \ A \ set zs \ A \ size xs = n \ size ys = n \ xs [\\<^bsub>r\<^esub>] zs \ ys [\\<^bsub>r\<^esub>] zs \ xs [\\<^bsub>f\<^esub>] ys [\\<^bsub>r\<^esub>] zs" (*<*) apply (unfold unfold_lesub_list) apply (simp add: Listn.le_def list_all2_conv_all_nth) done (*>*) lemma (in Semilat) list_update_incr [rule_format]: "x\A \ set xs \ A \ (\i. i xs [\\<^bsub>r\<^esub>] xs[i := x \\<^sub>f xs!i])" (*<*) apply (unfold unfold_lesub_list) apply (simp add: Listn.le_def list_all2_conv_all_nth) apply (induct xs) apply simp apply (simp add: in_list_Suc_iff) apply clarify apply (simp add: nth_Cons split: nat.split) done (*>*) -lemma acc_le_listI [intro!]: - "\ order r; acc r \ \ acc(Listn.le r)" -(*<*) -apply (unfold acc_def) -apply (subgoal_tac - "wf(UN n. {(ys,xs). size xs = n \ size ys = n \ xs <_(Listn.le r) ys})") - apply (erule wf_subset) - apply (blast intro: lesssub_lengthD) -apply (rule wf_UN) - prefer 2 - apply (rename_tac m n) - apply (case_tac "m=n") - apply simp - apply (fast intro!: equals0I dest: not_sym) -apply (rename_tac n) -apply (induct_tac n) - apply (simp add: lesssub_def cong: conj_cong) -apply (rename_tac k) -apply (simp add: wf_eq_minimal) -apply (simp (no_asm) add: length_Suc_conv cong: conj_cong) -apply clarify -apply (rename_tac M m) -apply (case_tac "\x xs. size xs = k \ x#xs \ M") - prefer 2 - apply (erule thin_rl) - apply (erule thin_rl) - apply blast -apply (erule_tac x = "{a. \xs. size xs = k \ a#xs:M}" in allE) -apply (erule impE) - apply blast -apply (thin_tac "\x xs. P x xs" for P) -apply clarify -apply (rename_tac maxA xs) -apply (erule_tac x = "{ys. size ys = size xs \ maxA#ys \ M}" in allE) -apply (erule impE) - apply blast -apply clarify -apply (thin_tac "m \ M") -apply (thin_tac "maxA#xs \ M") -apply (rule bexI) - prefer 2 - apply assumption -apply clarify -apply simp -apply blast -done -(*>*) lemma closed_listI: "closed S f \ closed (list n S) (map2 f)" (*<*) apply (unfold closed_def) apply (induct n) apply simp apply clarify apply (simp add: in_list_Suc_iff) apply clarify apply simp done (*>*) lemma Listn_sl_aux: assumes "Semilat A r f" shows "semilat (Listn.sl n (A,r,f))" (*<*) proof - interpret Semilat A r f by fact show ?thesis apply (unfold Listn.sl_def) apply (simp (no_asm) only: semilat_Def split_conv) apply (rule conjI) apply simp apply (rule conjI) apply (simp only: closedI closed_listI) apply (simp (no_asm) only: list_def) apply (simp (no_asm_simp) add: plus_list_ub1 plus_list_ub2 plus_list_lub) done qed (*>*) lemma Listn_sl: "semilat L \ semilat (Listn.sl n L)" (*<*) apply (cases L) apply simp apply (drule Semilat.intro) by (simp add: Listn_sl_aux split_tupled_all) (*>*) lemma coalesce_in_err_list [rule_format]: "\xes. xes \ list n (err A) \ coalesce xes \ err(list n A)" (*<*) apply (induct n) apply simp apply clarify apply (simp add: in_list_Suc_iff) apply clarify apply (simp (no_asm) add: plussub_def Err.sup_def lift2_def split: err.split) apply force done (*>*) lemma lem: "\x xs. x \\<^bsub>(#)\<^esub> xs = x#xs" (*<*) by (simp add: plussub_def) (*>*) lemma coalesce_eq_OK1_D [rule_format]: "semilat(err A, Err.le r, lift2 f) \ \xs. xs \ list n A \ (\ys. ys \ list n A \ (\zs. coalesce (xs [\\<^bsub>f\<^esub>] ys) = OK zs \ xs [\\<^bsub>r\<^esub>] zs))" (*<*) apply (induct n) apply simp apply clarify apply (simp add: in_list_Suc_iff) apply clarify apply (simp split: err.split_asm add: lem Err.sup_def lift2_def) apply (force simp add: semilat_le_err_OK1) done (*>*) lemma coalesce_eq_OK2_D [rule_format]: "semilat(err A, Err.le r, lift2 f) \ \xs. xs \ list n A \ (\ys. ys \ list n A \ (\zs. coalesce (xs [\\<^bsub>f\<^esub>] ys) = OK zs \ ys [\\<^bsub>r\<^esub>] zs))" (*<*) apply (induct n) apply simp apply clarify apply (simp add: in_list_Suc_iff) apply clarify apply (simp split: err.split_asm add: lem Err.sup_def lift2_def) apply (force simp add: semilat_le_err_OK2) done (*>*) lemma lift2_le_ub: "\ semilat(err A, Err.le r, lift2 f); x\A; y\A; x \\<^sub>f y = OK z; u\A; x \\<^sub>r u; y \\<^sub>r u \ \ z \\<^sub>r u" (*<*) apply (unfold semilat_Def plussub_def err_def') apply (simp add: lift2_def) apply clarify apply (rotate_tac -3) apply (erule thin_rl) apply (erule thin_rl) apply force done (*>*) lemma coalesce_eq_OK_ub_D [rule_format]: "semilat(err A, Err.le r, lift2 f) \ \xs. xs \ list n A \ (\ys. ys \ list n A \ (\zs us. coalesce (xs [\\<^bsub>f\<^esub>] ys) = OK zs \ xs [\\<^bsub>r\<^esub>] us \ ys [\\<^bsub>r\<^esub>] us \ us \ list n A \ zs [\\<^bsub>r\<^esub>] us))" (*<*) apply (induct n) apply simp apply clarify apply (simp add: in_list_Suc_iff) apply clarify apply (simp (no_asm_use) split: err.split_asm add: lem Err.sup_def lift2_def) apply clarify apply (rule conjI) apply (blast intro: lift2_le_ub) apply blast done (*>*) lemma lift2_eq_ErrD: "\ x \\<^sub>f y = Err; semilat(err A, Err.le r, lift2 f); x\A; y\A \ \ \(\u\A. x \\<^sub>r u \ y \\<^sub>r u)" (*<*) by (simp add: OK_plus_OK_eq_Err_conv [THEN iffD1]) (*>*) lemma coalesce_eq_Err_D [rule_format]: "\ semilat(err A, Err.le r, lift2 f) \ \ \xs. xs \ list n A \ (\ys. ys \ list n A \ coalesce (xs [\\<^bsub>f\<^esub>] ys) = Err \ \(\zs \ list n A. xs [\\<^bsub>r\<^esub>] zs \ ys [\\<^bsub>r\<^esub>] zs))" (*<*) apply (induct n) apply simp apply clarify apply (simp add: in_list_Suc_iff) apply clarify apply (simp split: err.split_asm add: lem Err.sup_def lift2_def) apply (blast dest: lift2_eq_ErrD) done (*>*) lemma closed_err_lift2_conv: "closed (err A) (lift2 f) = (\x\A. \y\A. x \\<^sub>f y \ err A)" (*<*) apply (unfold closed_def) apply (simp add: err_def') done (*>*) lemma closed_map2_list [rule_format]: "closed (err A) (lift2 f) \ \xs. xs \ list n A \ (\ys. ys \ list n A \ map2 f xs ys \ list n (err A))" (*<*) apply (induct n) apply simp apply clarify apply (simp add: in_list_Suc_iff) apply clarify apply (simp add: plussub_def closed_err_lift2_conv) done (*>*) lemma closed_lift2_sup: "closed (err A) (lift2 f) \ closed (err (list n A)) (lift2 (sup f))" (*<*) by (fastforce simp add: closed_def plussub_def sup_def lift2_def coalesce_in_err_list closed_map2_list split: err.split) (*>*) lemma err_semilat_sup: "err_semilat (A,r,f) \ err_semilat (list n A, Listn.le r, sup f)" (*<*) apply (unfold Err.sl_def) apply (simp only: split_conv) apply (simp (no_asm) only: semilat_Def plussub_def) apply (simp (no_asm_simp) only: Semilat.closedI [OF Semilat.intro] closed_lift2_sup) apply (rule conjI) apply (drule Semilat.orderI [OF Semilat.intro]) apply simp apply (simp (no_asm) only: unfold_lesub_err Err.le_def err_def' sup_def lift2_def) apply (simp (no_asm_simp) add: coalesce_eq_OK1_D coalesce_eq_OK2_D split: err.split) apply (blast intro: coalesce_eq_OK_ub_D dest: coalesce_eq_Err_D) done (*>*) lemma err_semilat_upto_esl: "\L. err_semilat L \ err_semilat(upto_esl m L)" (*<*) apply (unfold Listn.upto_esl_def) apply (simp (no_asm_simp) only: split_tupled_all) apply simp apply (fastforce intro!: err_semilat_UnionI err_semilat_sup dest: lesub_list_impl_same_size simp add: plussub_def Listn.sup_def) done (*>*) end diff --git a/thys/Jinja/DFA/Opt.thy b/thys/Jinja/DFA/Opt.thy --- a/thys/Jinja/DFA/Opt.thy +++ b/thys/Jinja/DFA/Opt.thy @@ -1,239 +1,239 @@ (* Title: HOL/MicroJava/BV/Opt.thy Author: Tobias Nipkow Copyright 2000 TUM More about options. *) section \More about Options\ theory Opt imports Err begin definition le :: "'a ord \ 'a option ord" where "le r o\<^sub>1 o\<^sub>2 = (case o\<^sub>2 of None \ o\<^sub>1=None | Some y \ (case o\<^sub>1 of None \ True | Some x \ x \\<^sub>r y))" definition opt :: "'a set \ 'a option set" where "opt A = insert None {Some y |y. y \ A}" definition sup :: "'a ebinop \ 'a option ebinop" where "sup f o\<^sub>1 o\<^sub>2 = (case o\<^sub>1 of None \ OK o\<^sub>2 | Some x \ (case o\<^sub>2 of None \ OK o\<^sub>1 | Some y \ (case f x y of Err \ Err | OK z \ OK (Some z))))" definition esl :: "'a esl \ 'a option esl" where "esl = (\(A,r,f). (opt A, le r, sup f))" lemma unfold_le_opt: "o\<^sub>1 \\<^bsub>le r\<^esub> o\<^sub>2 = (case o\<^sub>2 of None \ o\<^sub>1=None | Some y \ (case o\<^sub>1 of None \ True | Some x \ x \\<^sub>r y))" (*<*) apply (unfold lesub_def le_def) apply (rule refl) done (*>*) -lemma le_opt_refl: "order r \ x \\<^bsub>le r\<^esub> x" -(*<*) by (simp add: unfold_le_opt split: option.split) (*<*) +lemma le_opt_refl: "order r A \ x \ opt A \ x \\<^bsub>le r\<^esub> x" +(*<*) by (auto simp add: unfold_le_opt opt_def split: option.split) (*<*) lemma le_opt_trans [rule_format]: - "order r \ x \\<^bsub>le r\<^esub> y \ y \\<^bsub>le r\<^esub> z \ x \\<^bsub>le r\<^esub> z" + "order r A \ x \ opt A \ y \ opt A \ z \ opt A \ x \\<^bsub>le r\<^esub> y \ y \\<^bsub>le r\<^esub> z \ x \\<^bsub>le r\<^esub> z" (*<*) -apply (simp add: unfold_le_opt split: option.split) +apply (simp add: unfold_le_opt opt_def split: option.split) apply (blast intro: order_trans) done (*>*) lemma le_opt_antisym [rule_format]: - "order r \ x \\<^bsub>le r\<^esub> y \ y \\<^bsub>le r\<^esub> x \ x=y" + "order r A \ x \ opt A \ y \ opt A \ z \ opt A \ x \\<^bsub>le r\<^esub> y \ y \\<^bsub>le r\<^esub> x \ x=y" (*<*) -apply (simp add: unfold_le_opt split: option.split) +apply (simp add: unfold_le_opt opt_def split: option.split) apply (blast intro: order_antisym) done (*>*) -lemma order_le_opt [intro!,simp]: "order r \ order(le r)" +lemma order_le_opt [intro!,simp]: "order r A \ order(le r) (opt A)" (*<*) apply (subst order_def) apply (blast intro: le_opt_refl le_opt_trans le_opt_antisym) done (*>*) lemma None_bot [iff]: "None \\<^bsub>le r\<^esub> ox" (*<*) apply (unfold lesub_def le_def) apply (simp split: option.split) done (*>*) lemma Some_le [iff]: "(Some x \\<^bsub>le r\<^esub> z) = (\y. z = Some y \ x \\<^sub>r y)" (*<*) apply (unfold lesub_def le_def) apply (simp split: option.split) done (*>*) lemma le_None [iff]: "(x \\<^bsub>le r\<^esub> None) = (x = None)" (*<*) apply (unfold lesub_def le_def) apply (simp split: option.split) done (*>*) lemma OK_None_bot [iff]: "OK None \\<^bsub>Err.le (le r)\<^esub> x" (*<*) by (simp add: lesub_def Err.le_def le_def split: option.split err.split) (*>*) lemma sup_None1 [iff]: "x \\<^bsub>sup f\<^esub> None = OK x" (*<*) by (simp add: plussub_def sup_def split: option.split) (*>*) lemma sup_None2 [iff]: "None \\<^bsub>sup f\<^esub> x = OK x" (*<*) by (simp add: plussub_def sup_def split: option.split) (*>*) lemma None_in_opt [iff]: "None \ opt A" (*<*) by (simp add: opt_def) (*>*) lemma Some_in_opt [iff]: "(Some x \ opt A) = (x \ A)" (*<*) by (unfold opt_def) auto (*>*) lemma semilat_opt [intro, simp]: "err_semilat L \ err_semilat (Opt.esl L)" (*<*) proof - assume s: "err_semilat L" obtain A r f where [simp]: "L = (A,r,f)" by (cases L) let ?A0 = "err A" and ?r0 = "Err.le r" and ?f0 = "lift2 f" from s obtain - ord: "order ?r0" and + ord: "order ?r0 ?A" and clo: "closed ?A0 ?f0" and ub1: "\x\?A0. \y\?A0. x \\<^bsub>?r0\<^esub> x \\<^bsub>?f0\<^esub> y" and ub2: "\x\?A0. \y\?A0. y \\<^bsub>?r0\<^esub> x \\<^bsub>?f0\<^esub> y" and lub: "\x\?A0. \y\?A0. \z\?A0. x \\<^bsub>?r0\<^esub> z \ y \\<^bsub>?r0\<^esub> z \ x \\<^bsub>?f0\<^esub> y \\<^bsub>?r0\<^esub> z" by (unfold semilat_def sl_def) simp let ?A = "err (opt A)" and ?r = "Err.le (Opt.le r)" and ?f = "lift2 (Opt.sup f)" - from ord have "order ?r" by simp + from ord have "order ?r ?A" by simp moreover have "closed ?A ?f" proof (unfold closed_def, intro strip) fix x y assume x: "x \ ?A" and y: "y \ ?A" { fix a b assume ab: "x = OK a" "y = OK b" with x have a: "\c. a = Some c \ c \ A" by (clarsimp simp add: opt_def) from ab y have b: "\d. b = Some d \ d \ A" by (clarsimp simp add: opt_def) { fix c d assume "a = Some c" "b = Some d" with ab x y have "c \ A & d \ A" by (simp add: err_def opt_def Bex_def) with clo have "f c d \ err A" by (simp add: closed_def plussub_def err_def' lift2_def) moreover fix z assume "f c d = OK z" ultimately have "z \ A" by simp } note f_closed = this have "sup f a b \ ?A" proof (cases a) case None thus ?thesis by (simp add: sup_def opt_def) (cases b, simp, simp add: b Bex_def) next case Some thus ?thesis by (auto simp add: sup_def opt_def Bex_def a b f_closed split: err.split option.split) qed } thus "x \\<^bsub>?f\<^esub> y \ ?A" by (simp add: plussub_def lift2_def split: err.split) qed moreover { fix a b c assume "a \ opt A" and "b \ opt A" and "a \\<^bsub>sup f\<^esub> b = OK c" - moreover from ord have "order r" by simp + moreover from ord have "order r A" by simp moreover { fix x y z assume "x \ A" and "y \ A" hence "OK x \ err A \ OK y \ err A" by simp with ub1 ub2 have "(OK x) \\<^bsub>Err.le r\<^esub> (OK x) \\<^bsub>lift2 f\<^esub> (OK y) \ (OK y) \\<^bsub>Err.le r\<^esub> (OK x) \\<^bsub>lift2 f\<^esub> (OK y)" by blast moreover assume "x \\<^sub>f y = OK z" ultimately have "x \\<^sub>r z \ y \\<^sub>r z" by (auto simp add: plussub_def lift2_def Err.le_def lesub_def) } ultimately have "a \\<^bsub>le r\<^esub> c \ b \\<^bsub>le r\<^esub> c" by (auto simp add: sup_def le_def lesub_def plussub_def dest: order_refl split: option.splits err.splits) } hence "(\x\?A. \y\?A. x \\<^bsub>?r\<^esub> x \\<^bsub>?f\<^esub> y) \ (\x\?A. \y\?A. y \\<^bsub>?r\<^esub> x \\<^bsub>?f\<^esub> y)" by (auto simp add: lesub_def plussub_def Err.le_def lift2_def split: err.split) moreover have "\x\?A. \y\?A. \z\?A. x \\<^bsub>?r\<^esub> z \ y \\<^bsub>?r\<^esub> z \ x \\<^bsub>?f\<^esub> y \\<^bsub>?r\<^esub> z" proof (intro strip, elim conjE) fix x y z assume xyz: "x \ ?A" "y \ ?A" "z \ ?A" assume xz: "x \\<^bsub>?r\<^esub> z" and yz: "y \\<^bsub>?r\<^esub> z" { fix a b c assume ok: "x = OK a" "y = OK b" "z = OK c" { fix d e g assume some: "a = Some d" "b = Some e" "c = Some g" with ok xyz obtain "OK d:err A" "OK e:err A" "OK g:err A" by simp with lub have "\ OK d \\<^bsub>Err.le r\<^esub> OK g; OK e \\<^bsub>Err.le r\<^esub> OK g \ \ OK d \\<^bsub>lift2 f\<^esub> OK e \\<^bsub>Err.le r\<^esub> OK g" by blast hence "\ d \\<^sub>r g; e \\<^sub>r g \ \ \y. d \\<^sub>f e = OK y \ y \\<^sub>r g" by simp with ok some xyz xz yz have "x \\<^bsub>?f\<^esub> y \\<^bsub>?r\<^esub> z" by (auto simp add: sup_def le_def lesub_def lift2_def plussub_def Err.le_def) } note this [intro!] from ok xyz xz yz have "x \\<^bsub>?f\<^esub> y \\<^bsub>?r\<^esub> z" by - (cases a, simp, cases b, simp, cases c, simp, blast) } with xyz xz yz show "x \\<^bsub>?f\<^esub> y \\<^bsub>?r\<^esub> z" by - (cases x, simp, cases y, simp, cases z, simp+) qed ultimately show "err_semilat (Opt.esl L)" by (unfold semilat_def esl_def sl_def) simp qed (*>*) lemma top_le_opt_Some [iff]: "top (le r) (Some T) = top r T" (*<*) apply (unfold top_def) apply (rule iffI) apply blast apply (rule allI) apply (case_tac "x") apply simp+ done (*>*) -lemma Top_le_conv: "\ order r; top r T \ \ (T \\<^sub>r x) = (x = T)" +lemma Top_le_conv: "\ order r A; top r T; x \ A; T \ A \ \ (T \\<^sub>r x) = (x = T)" (*<*) apply (unfold top_def) apply (blast intro: order_antisym) done (*>*) lemma acc_le_optI [intro!]: "acc r \ acc(le r)" (*<*) apply (unfold acc_def lesub_def le_def lesssub_def) apply (simp add: wf_eq_minimal split: option.split) apply clarify apply (case_tac "\a. Some a \ Q") apply (erule_tac x = "{a . Some a \ Q}" in allE) apply blast apply (case_tac "x") apply blast apply blast done (*>*) lemma map_option_in_optionI: "\ ox \ opt S; \x\S. ox = Some x \ f x \ S \ \ map_option f ox \ opt S" (*<*) apply (unfold map_option_case) apply (simp split: option.split) apply blast done (*>*) end diff --git a/thys/Jinja/DFA/Product.thy b/thys/Jinja/DFA/Product.thy --- a/thys/Jinja/DFA/Product.thy +++ b/thys/Jinja/DFA/Product.thy @@ -1,161 +1,176 @@ (* Title: HOL/MicroJava/BV/Product.thy Author: Tobias Nipkow Copyright 2000 TUM Products as semilattices. *) section \Products as Semilattices\ theory Product imports Err begin definition le :: "'a ord \ 'b ord \ ('a \ 'b) ord" where "le r\<^sub>A r\<^sub>B = (\(a\<^sub>1,b\<^sub>1) (a\<^sub>2,b\<^sub>2). a\<^sub>1 \\<^bsub>r\<^sub>A\<^esub> a\<^sub>2 \ b\<^sub>1 \\<^bsub>r\<^sub>B\<^esub> b\<^sub>2)" definition sup :: "'a ebinop \ 'b ebinop \ ('a \ 'b) ebinop" where "sup f g = (\(a\<^sub>1,b\<^sub>1)(a\<^sub>2,b\<^sub>2). Err.sup Pair (a\<^sub>1 \\<^sub>f a\<^sub>2) (b\<^sub>1 \\<^sub>g b\<^sub>2))" definition esl :: "'a esl \ 'b esl \ ('a \ 'b ) esl" where "esl = (\(A,r\<^sub>A,f\<^sub>A) (B,r\<^sub>B,f\<^sub>B). (A \ B, le r\<^sub>A r\<^sub>B, sup f\<^sub>A f\<^sub>B))" abbreviation lesubprod :: "'a \ 'b \ ('a \ 'a \ bool) \ ('b \ 'b \ bool) \ 'a \ 'b \ bool" ("(_ /\'(_,_') _)" [50, 0, 0, 51] 50) where "p \(rA,rB) q == p \\<^bsub>Product.le rA rB\<^esub> q" (*<*) notation lesubprod ("(_ /<='(_,_') _)" [50, 0, 0, 51] 50) (*>*) lemma unfold_lesub_prod: "x \(r\<^sub>A,r\<^sub>B) y = le r\<^sub>A r\<^sub>B x y" (*<*) by (simp add: lesub_def) (*>*) lemma le_prod_Pair_conv [iff]: "((a\<^sub>1,b\<^sub>1) \(r\<^sub>A,r\<^sub>B) (a\<^sub>2,b\<^sub>2)) = (a\<^sub>1 \\<^bsub>r\<^sub>A\<^esub> a\<^sub>2 & b\<^sub>1 \\<^bsub>r\<^sub>B\<^esub> b\<^sub>2)" (*<*) by (simp add: lesub_def le_def) (*>*) lemma less_prod_Pair_conv: "((a\<^sub>1,b\<^sub>1) \\<^bsub>Product.le r\<^sub>A r\<^sub>B\<^esub> (a\<^sub>2,b\<^sub>2)) = (a\<^sub>1 \\<^bsub>r\<^sub>A\<^esub> a\<^sub>2 & b\<^sub>1 \\<^bsub>r\<^sub>B\<^esub> b\<^sub>2 | a\<^sub>1 \\<^bsub>r\<^sub>A\<^esub> a\<^sub>2 & b\<^sub>1 \\<^bsub>r\<^sub>B\<^esub> b\<^sub>2)" (*<*) -apply (unfold lesssub_def) -apply simp -apply blast -done + apply (unfold lesssub_def) + apply simp + apply blast + done (*>*) -lemma order_le_prod [iff]: "order(Product.le r\<^sub>A r\<^sub>B) = (order r\<^sub>A & order r\<^sub>B)" +lemma order_le_prodI [iff]: "(order r\<^sub>A A & order r\<^sub>B B) \ order (Product.le r\<^sub>A r\<^sub>B) (A \ B)" + apply (unfold order_def) + apply safe + apply blast+ +done + +lemma order_le_prodE: "A \ {} \ B \ {} \ order (Product.le r\<^sub>A r\<^sub>B) (A \ B) \ (order r\<^sub>A A & order r\<^sub>B B)" + apply (unfold order_def) + apply simp + apply safe + apply blast+ + done + +lemma order_le_prod [iff]: "A \ {} \ B \ {} \ order(Product.le r\<^sub>A r\<^sub>B) (A \ B) = (order r\<^sub>A A & order r\<^sub>B B)" (*<*) -apply (unfold order_def) -apply simp -apply safe -apply blast+ -done + apply (unfold order_def) + apply simp + apply safe + apply blast+ + done + (*>*) - lemma acc_le_prodI [intro!]: "\ acc r\<^sub>A; acc r\<^sub>B \ \ acc(Product.le r\<^sub>A r\<^sub>B)" (*<*) apply (unfold acc_def) apply (rule wf_subset) apply (erule wf_lex_prod) apply assumption apply (auto simp add: lesssub_def less_prod_Pair_conv lex_prod_def) done (*>*) lemma closed_lift2_sup: "\ closed (err A) (lift2 f); closed (err B) (lift2 g) \ \ closed (err(A\B)) (lift2(sup f g))" (*<*) apply (unfold closed_def plussub_def lift2_def err_def' sup_def) apply (simp split: err.split) apply blast done (*>*) lemma unfold_plussub_lift2: "e\<^sub>1 \\<^bsub>lift2 f\<^esub> e\<^sub>2 = lift2 f e\<^sub>1 e\<^sub>2" (*<*) by (simp add: plussub_def) (*>*) lemma plus_eq_Err_conv [simp]: assumes "x\A" "y\A" "semilat(err A, Err.le r, lift2 f)" shows "(x \\<^sub>f y = Err) = (\(\z\A. x \\<^sub>r z \ y \\<^sub>r z))" (*<*) proof - have plus_le_conv2: "\r f z. \ z \ err A; semilat (err A, r, f); OK x \ err A; OK y \ err A; OK x \\<^sub>f OK y \\<^sub>r z\ \ OK x \\<^sub>r z \ OK y \\<^sub>r z" (*<*) by (rule Semilat.plus_le_conv [OF Semilat.intro, THEN iffD1]) (*>*) from assms show ?thesis apply (rule_tac iffI) apply clarify apply (drule OK_le_err_OK [THEN iffD2]) apply (drule OK_le_err_OK [THEN iffD2]) apply (drule Semilat.lub[OF Semilat.intro, of _ _ _ "OK x" _ "OK y"]) apply assumption apply assumption apply simp apply simp apply simp apply simp apply (case_tac "x \\<^sub>f y") apply assumption apply (rename_tac "z") apply (subgoal_tac "OK z: err A") apply (frule plus_le_conv2) apply assumption apply simp apply blast apply simp apply (blast dest: Semilat.orderI [OF Semilat.intro] order_refl) apply blast apply (erule subst) apply (unfold semilat_def err_def' closed_def) apply simp done qed (*>*) lemma err_semilat_Product_esl: "\L\<^sub>1 L\<^sub>2. \ err_semilat L\<^sub>1; err_semilat L\<^sub>2 \ \ err_semilat(Product.esl L\<^sub>1 L\<^sub>2)" (*<*) apply (unfold esl_def Err.sl_def) apply (simp (no_asm_simp) only: split_tupled_all) apply simp apply (simp (no_asm) only: semilat_Def) apply (simp (no_asm_simp) only: Semilat.closedI [OF Semilat.intro] closed_lift2_sup) apply (simp (no_asm) only: unfold_lesub_err Err.le_def unfold_plussub_lift2 sup_def) apply (auto elim: semilat_le_err_OK1 semilat_le_err_OK2 simp add: lift2_def split: err.split) + apply(rule order_le_prodI) + apply (rule conjI) apply (blast dest: Semilat.orderI [OF Semilat.intro]) apply (blast dest: Semilat.orderI [OF Semilat.intro]) apply (rule OK_le_err_OK [THEN iffD1]) apply (erule subst, subst OK_lift2_OK [symmetric], rule Semilat.lub [OF Semilat.intro]) apply simp apply simp apply simp apply simp apply simp apply simp apply (rule OK_le_err_OK [THEN iffD1]) apply (erule subst, subst OK_lift2_OK [symmetric], rule Semilat.lub [OF Semilat.intro]) apply simp apply simp apply simp apply simp apply simp apply simp done (*>*) end diff --git a/thys/Jinja/DFA/Semilat.thy b/thys/Jinja/DFA/Semilat.thy --- a/thys/Jinja/DFA/Semilat.thy +++ b/thys/Jinja/DFA/Semilat.thy @@ -1,393 +1,395 @@ (* Title: HOL/MicroJava/BV/Semilat.thy Author: Tobias Nipkow Copyright 2000 TUM Semilattices. *) chapter \Bytecode Verifier \label{cha:bv}\ section \Semilattices\ theory Semilat imports Main "HOL-Library.While_Combinator" begin type_synonym 'a ord = "'a \ 'a \ bool" type_synonym 'a binop = "'a \ 'a \ 'a" type_synonym 'a sl = "'a set \ 'a ord \ 'a binop" definition lesub :: "'a \ 'a ord \ 'a \ bool" where "lesub x r y \ r x y" definition lesssub :: "'a \ 'a ord \ 'a \ bool" where "lesssub x r y \ lesub x r y \ x \ y" definition plussub :: "'a \ ('a \ 'b \ 'c) \ 'b \ 'c" where "plussub x f y = f x y" notation (ASCII) "lesub" ("(_ /<='__ _)" [50, 1000, 51] 50) and "lesssub" ("(_ /<'__ _)" [50, 1000, 51] 50) and "plussub" ("(_ /+'__ _)" [65, 1000, 66] 65) notation "lesub" ("(_ /\\<^bsub>_\<^esub> _)" [50, 0, 51] 50) and "lesssub" ("(_ /\\<^bsub>_\<^esub> _)" [50, 0, 51] 50) and "plussub" ("(_ /\\<^bsub>_\<^esub> _)" [65, 0, 66] 65) (* allow \ instead of \..\ *) abbreviation (input) lesub1 :: "'a \ 'a ord \ 'a \ bool" ("(_ /\\<^sub>_ _)" [50, 1000, 51] 50) where "x \\<^sub>r y == x \\<^bsub>r\<^esub> y" abbreviation (input) lesssub1 :: "'a \ 'a ord \ 'a \ bool" ("(_ /\\<^sub>_ _)" [50, 1000, 51] 50) where "x \\<^sub>r y == x \\<^bsub>r\<^esub> y" abbreviation (input) plussub1 :: "'a \ ('a \ 'b \ 'c) \ 'b \ 'c" ("(_ /\\<^sub>_ _)" [65, 1000, 66] 65) where "x \\<^sub>f y == x \\<^bsub>f\<^esub> y" definition ord :: "('a \ 'a) set \ 'a ord" where "ord r = (\x y. (x,y) \ r)" -definition order :: "'a ord \ bool" +definition order :: "'a ord \ 'a set \ bool" where - "order r \ (\x. x \\<^sub>r x) \ (\x y. x \\<^sub>r y \ y \\<^sub>r x \ x=y) \ (\x y z. x \\<^sub>r y \ y \\<^sub>r z \ x \\<^sub>r z)" + "order r A \ (\x\A. x \\<^sub>r x) \ (\x\A. \y\A. x \\<^sub>r y \ y \\<^sub>r x \ x=y) \ (\x\A. \y\A. \z\A. x \\<^sub>r y \ y \\<^sub>r z \ x \\<^sub>r z)" definition top :: "'a ord \ 'a \ bool" where "top r T \ (\x. x \\<^sub>r T)" definition acc :: "'a ord \ bool" where "acc r \ wf {(y,x). x \\<^sub>r y}" definition closed :: "'a set \ 'a binop \ bool" where "closed A f \ (\x\A. \y\A. x \\<^sub>f y \ A)" definition semilat :: "'a sl \ bool" where - "semilat = (\(A,r,f). order r \ closed A f \ + "semilat = (\(A,r,f). order r A \ closed A f \ (\x\A. \y\A. x \\<^sub>r x \\<^sub>f y) \ (\x\A. \y\A. y \\<^sub>r x \\<^sub>f y) \ (\x\A. \y\A. \z\A. x \\<^sub>r z \ y \\<^sub>r z \ x \\<^sub>f y \\<^sub>r z))" definition is_ub :: "('a \ 'a) set \ 'a \ 'a \ 'a \ bool" where "is_ub r x y u \ (x,u)\r \ (y,u)\r" definition is_lub :: "('a \ 'a) set \ 'a \ 'a \ 'a \ bool" where "is_lub r x y u \ is_ub r x y u \ (\z. is_ub r x y z \ (u,z)\r)" definition some_lub :: "('a \ 'a) set \ 'a \ 'a \ 'a" where "some_lub r x y = (SOME z. is_lub r x y z)" locale Semilat = fixes A :: "'a set" fixes r :: "'a ord" fixes f :: "'a binop" assumes semilat: "semilat (A, r, f)" -lemma order_refl [simp, intro]: "order r \ x \\<^sub>r x" - (*<*) by (unfold order_def) (simp (no_asm_simp)) (*>*) - -lemma order_antisym: "\ order r; x \\<^sub>r y; y \\<^sub>r x \ \ x = y" +lemma order_refl [simp, intro]: "order r A \ x \ A \ x \\<^sub>r x" (*<*) by (unfold order_def) (simp (no_asm_simp)) (*>*) -lemma order_trans: "\ order r; x \\<^sub>r y; y \\<^sub>r z \ \ x \\<^sub>r z" +lemma order_antisym: "\ order r A; x \\<^sub>r y; y \\<^sub>r x; x \ A; y \ A \ \ x = y" + (*<*) by (unfold order_def) ( simp (no_asm_simp)) (*>*) + +lemma order_trans: "\ order r A; x \\<^sub>r y; y \\<^sub>r z; x \ A; y \ A; z \ A \ \ x \\<^sub>r z" (*<*) by (unfold order_def) blast (*>*) -lemma order_less_irrefl [intro, simp]: "order r \ \ x \\<^sub>r x" +lemma order_less_irrefl [intro, simp]: "order r A \ x \ A \ \ x \\<^sub>r x" (*<*) by (unfold order_def lesssub_def) blast (*>*) -lemma order_less_trans: "\ order r; x \\<^sub>r y; y \\<^sub>r z \ \ x \\<^sub>r z" +lemma order_less_trans: "\ order r A; x \\<^sub>r y; y \\<^sub>r z; x \ A; y \ A; z \ A \ \ x \\<^sub>r z" (*<*) by (unfold order_def lesssub_def) blast (*>*) lemma topD [simp, intro]: "top r T \ x \\<^sub>r T" (*<*) by (simp add: top_def) (*>*) -lemma top_le_conv [simp]: "\ order r; top r T \ \ (T \\<^sub>r x) = (x = T)" +lemma top_le_conv [simp]: "\ order r A; top r T; x \ A; T \ A\ \ (T \\<^sub>r x) = (x = T)" (*<*) by (blast intro: order_antisym) (*>*) lemma semilat_Def: -"semilat(A,r,f) \ order r \ closed A f \ +"semilat(A,r,f) \ order r A \ closed A f \ (\x\A. \y\A. x \\<^sub>r x \\<^sub>f y) \ (\x\A. \y\A. y \\<^sub>r x \\<^sub>f y) \ (\x\A. \y\A. \z\A. x \\<^sub>r z \ y \\<^sub>r z \ x \\<^sub>f y \\<^sub>r z)" (*<*) by (unfold semilat_def) clarsimp (*>*) -lemma (in Semilat) orderI [simp, intro]: "order r" +lemma (in Semilat) orderI [simp, intro]: "order r A" (*<*) using semilat by (simp add: semilat_Def) (*>*) lemma (in Semilat) closedI [simp, intro]: "closed A f" (*<*) using semilat by (simp add: semilat_Def) (*>*) lemma closedD: "\ closed A f; x\A; y\A \ \ x \\<^sub>f y \ A" (*<*) by (unfold closed_def) blast (*>*) lemma closed_UNIV [simp]: "closed UNIV f" (*<*) by (simp add: closed_def) (*>*) lemma (in Semilat) closed_f [simp, intro]: "\x \ A; y \ A\ \ x \\<^sub>f y \ A" (*<*) by (simp add: closedD [OF closedI]) (*>*) -lemma (in Semilat) refl_r [intro, simp]: "x \\<^sub>r x" by simp +lemma (in Semilat) refl_r [intro, simp]: "x \ A \ x \\<^sub>r x" by auto -lemma (in Semilat) antisym_r [intro?]: "\ x \\<^sub>r y; y \\<^sub>r x \ \ x = y" +lemma (in Semilat) antisym_r [intro?]: "\ x \\<^sub>r y; y \\<^sub>r x; x \ A; y \ A \ \ x = y" (*<*) by (rule order_antisym) auto (*>*) - -lemma (in Semilat) trans_r [trans, intro?]: "\x \\<^sub>r y; y \\<^sub>r z\ \ x \\<^sub>r z" + +lemma (in Semilat) trans_r [trans, intro?]: "\x \\<^sub>r y; y \\<^sub>r z; x \ A; y \ A; z \ A \ \ x \\<^sub>r z" (*<*) by (auto intro: order_trans) (*>*) lemma (in Semilat) ub1 [simp, intro?]: "\ x \ A; y \ A \ \ x \\<^sub>r x \\<^sub>f y" (*<*) by (insert semilat) (unfold semilat_Def, simp) (*>*) lemma (in Semilat) ub2 [simp, intro?]: "\ x \ A; y \ A \ \ y \\<^sub>r x \\<^sub>f y" (*<*) by (insert semilat) (unfold semilat_Def, simp) (*>*) lemma (in Semilat) lub [simp, intro?]: "\ x \\<^sub>r z; y \\<^sub>r z; x \ A; y \ A; z \ A \ \ x \\<^sub>f y \\<^sub>r z" (*<*) by (insert semilat) (unfold semilat_Def, simp) (*>*) lemma (in Semilat) plus_le_conv [simp]: "\ x \ A; y \ A; z \ A \ \ (x \\<^sub>f y \\<^sub>r z) = (x \\<^sub>r z \ y \\<^sub>r z)" (*<*) by (blast intro: ub1 ub2 lub order_trans) (*>*) lemma (in Semilat) le_iff_plus_unchanged: assumes "x \ A" and "y \ A" shows "x \\<^sub>r y \ x \\<^sub>f y = y" (is "?P \ ?Q") (*<*) proof assume ?P with assms show ?Q by (blast intro: antisym_r lub ub2) next assume ?Q then have "y = x \\<^bsub>f\<^esub> y" by simp moreover from assms have "x \\<^bsub>r\<^esub> x \\<^bsub>f\<^esub> y" by simp ultimately show ?P by simp qed (*>*) lemma (in Semilat) le_iff_plus_unchanged2: assumes "x \ A" and "y \ A" shows "x \\<^sub>r y \ y \\<^sub>f x = y" (is "?P \ ?Q") (*<*) proof assume ?P with assms show ?Q by (blast intro: antisym_r lub ub1) next assume ?Q then have "y = y \\<^bsub>f\<^esub> x" by simp moreover from assms have "x \\<^bsub>r\<^esub> y \\<^bsub>f\<^esub> x" by simp ultimately show ?P by simp qed (*>*) lemma (in Semilat) plus_assoc [simp]: assumes a: "a \ A" and b: "b \ A" and c: "c \ A" shows "a \\<^sub>f (b \\<^sub>f c) = a \\<^sub>f b \\<^sub>f c" (*<*) proof - from a b have ab: "a \\<^sub>f b \ A" .. from this c have abc: "(a \\<^sub>f b) \\<^sub>f c \ A" .. from b c have bc: "b \\<^sub>f c \ A" .. from a this have abc': "a \\<^sub>f (b \\<^sub>f c) \ A" .. show ?thesis proof show "a \\<^sub>f (b \\<^sub>f c) \\<^sub>r (a \\<^sub>f b) \\<^sub>f c" proof - - from a b have "a \\<^sub>r a \\<^sub>f b" .. - also from ab c have "\ \\<^sub>r \ \\<^sub>f c" .. - finally have "a<": "a \\<^sub>r (a \\<^sub>f b) \\<^sub>f c" . - from a b have "b \\<^sub>r a \\<^sub>f b" .. - also from ab c have "\ \\<^sub>r \ \\<^sub>f c" .. - finally have "b<": "b \\<^sub>r (a \\<^sub>f b) \\<^sub>f c" . + from a b have 1: "a \\<^sub>r a \\<^sub>f b" .. + from ab c have 2: "\ \\<^sub>r \ \\<^sub>f c" .. + with 1 have "a<": "a \\<^sub>r (a \\<^sub>f b) \\<^sub>f c" using a ab abc by (rule trans_r) + from a b have 11: "b \\<^sub>r a \\<^sub>f b" .. + hence "b<": "b \\<^sub>r (a \\<^sub>f b) \\<^sub>f c" using 2 b ab abc by (rule trans_r) from ab c have "c<": "c \\<^sub>r (a \\<^sub>f b) \\<^sub>f c" .. from "b<" "c<" b c abc have "b \\<^sub>f c \\<^sub>r (a \\<^sub>f b) \\<^sub>f c" .. from "a<" this a bc abc show ?thesis .. qed show "(a \\<^sub>f b) \\<^sub>f c \\<^sub>r a \\<^sub>f (b \\<^sub>f c)" proof - - from b c have "b \\<^sub>r b \\<^sub>f c" .. - also from a bc have "\ \\<^sub>r a \\<^sub>f \" .. - finally have "b<": "b \\<^sub>r a \\<^sub>f (b \\<^sub>f c)" . - from b c have "c \\<^sub>r b \\<^sub>f c" .. - also from a bc have "\ \\<^sub>r a \\<^sub>f \" .. - finally have "c<": "c \\<^sub>r a \\<^sub>f (b \\<^sub>f c)" . + from b c have 3:"b \\<^sub>r b \\<^sub>f c" .. + from a bc have bc_abc: "\ \\<^sub>r a \\<^sub>f \" .. + with 3 have "b<": "b \\<^sub>r a \\<^sub>f (b \\<^sub>f c)"using b bc abc' by (rule trans_r) + from b c have 4: "c \\<^sub>r b \\<^sub>f c" .. + hence "c<": "c \\<^sub>r a \\<^sub>f (b \\<^sub>f c)"using bc_abc c bc abc' by (rule trans_r) from a bc have "a<": "a \\<^sub>r a \\<^sub>f (b \\<^sub>f c)" .. from "a<" "b<" a b abc' have "a \\<^sub>f b \\<^sub>r a \\<^sub>f (b \\<^sub>f c)" .. from this "c<" ab c abc' show ?thesis .. qed + next + show "a \\<^bsub>f\<^esub> (b \\<^bsub>f\<^esub> c) \ A" using abc' by auto + next + show "a \\<^bsub>f\<^esub> b \\<^bsub>f\<^esub> c \ A" using abc by auto qed qed (*>*) lemma (in Semilat) plus_com_lemma: "\a \ A; b \ A\ \ a \\<^sub>f b \\<^sub>r b \\<^sub>f a" (*<*) proof - assume a: "a \ A" and b: "b \ A" from b a have "a \\<^sub>r b \\<^sub>f a" .. moreover from b a have "b \\<^sub>r b \\<^sub>f a" .. moreover note a b moreover from b a have "b \\<^sub>f a \ A" .. ultimately show ?thesis .. qed (*>*) lemma (in Semilat) plus_commutative: "\a \ A; b \ A\ \ a \\<^sub>f b = b \\<^sub>f a" (*<*) by(blast intro: order_antisym plus_com_lemma) (*>*) lemma is_lubD: "is_lub r x y u \ is_ub r x y u \ (\z. is_ub r x y z \ (u,z) \ r)" (*<*) by (simp add: is_lub_def) (*>*) lemma is_ubI: "\ (x,u) \ r; (y,u) \ r \ \ is_ub r x y u" (*<*) by (simp add: is_ub_def) (*>*) lemma is_ubD: "is_ub r x y u \ (x,u) \ r \ (y,u) \ r" (*<*) by (simp add: is_ub_def) (*>*) lemma is_lub_bigger1 [iff]: "is_lub (r^* ) x y y = ((x,y)\r^* )" (*<*) apply (unfold is_lub_def is_ub_def) apply blast done (*>*) lemma is_lub_bigger2 [iff]: "is_lub (r^* ) x y x = ((y,x)\r^* )" (*<*) apply (unfold is_lub_def is_ub_def) apply blast done (*>*) lemma extend_lub: "\ single_valued r; is_lub (r^* ) x y u; (x',x) \ r \ \ \v. is_lub (r^* ) x' y v" (*<*) apply (unfold is_lub_def is_ub_def) apply (case_tac "(y,x) \ r^*") apply (case_tac "(y,x') \ r^*") apply blast apply (blast elim: converse_rtranclE dest: single_valuedD) apply (rule exI) apply (rule conjI) apply (blast intro: converse_rtrancl_into_rtrancl dest: single_valuedD) apply (blast intro: rtrancl_into_rtrancl converse_rtrancl_into_rtrancl elim: converse_rtranclE dest: single_valuedD) done (*>*) lemma single_valued_has_lubs [rule_format]: "\ single_valued r; (x,u) \ r^* \ \ (\y. (y,u) \ r^* \ (\z. is_lub (r^* ) x y z))" (*<*) apply (erule converse_rtrancl_induct) apply clarify apply (erule converse_rtrancl_induct) apply blast apply (blast intro: converse_rtrancl_into_rtrancl) apply (blast intro: extend_lub) done (*>*) lemma some_lub_conv: "\ acyclic r; is_lub (r^* ) x y u \ \ some_lub (r^* ) x y = u" (*<*) apply (simp only: some_lub_def is_lub_def) apply (rule someI2) apply (simp only: is_lub_def) apply (blast intro: antisymD dest!: acyclic_impl_antisym_rtrancl) done (*>*) lemma is_lub_some_lub: "\ single_valued r; acyclic r; (x,u)\r^*; (y,u)\r^* \ \ is_lub (r^* ) x y (some_lub (r^* ) x y)" (*<*) by (fastforce dest: single_valued_has_lubs simp add: some_lub_conv) (*>*) subsection\An executable lub-finder\ definition exec_lub :: "('a * 'a) set \ ('a \ 'a) \ 'a binop" where "exec_lub r f x y = while (\z. (x,z) \ r\<^sup>*) f y" lemma exec_lub_refl: "exec_lub r f T T = T" by (simp add: exec_lub_def while_unfold) lemma acyclic_single_valued_finite: "\acyclic r; single_valued r; (x,y) \ r\<^sup>*\ \ finite (r \ {a. (x, a) \ r\<^sup>*} \ {b. (b, y) \ r\<^sup>*})" (*<*) apply(erule converse_rtrancl_induct) apply(rule_tac B = "{}" in finite_subset) apply(simp only:acyclic_def) apply(blast intro:rtrancl_into_trancl2 rtrancl_trancl_trancl) apply simp apply(rename_tac x x') apply(subgoal_tac "r \ {a. (x,a) \ r\<^sup>*} \ {b. (b,y) \ r\<^sup>*} = insert (x,x') (r \ {a. (x', a) \ r\<^sup>*} \ {b. (b, y) \ r\<^sup>*})") apply simp apply(blast intro:converse_rtrancl_into_rtrancl elim:converse_rtranclE dest:single_valuedD) done (*>*) lemma exec_lub_conv: "\ acyclic r; \x y. (x,y) \ r \ f x = y; is_lub (r\<^sup>*) x y u \ \ exec_lub r f x y = u" (*<*) apply(unfold exec_lub_def) apply(rule_tac P = "\z. (y,z) \ r\<^sup>* \ (z,u) \ r\<^sup>*" and r = "(r \ {(a,b). (y,a) \ r\<^sup>* \ (b,u) \ r\<^sup>*})^-1" in while_rule) apply(blast dest: is_lubD is_ubD) apply(erule conjE) apply(erule_tac z = u in converse_rtranclE) apply(blast dest: is_lubD is_ubD) apply(blast dest:rtrancl_into_rtrancl) apply(rename_tac s) apply(subgoal_tac "is_ub (r\<^sup>*) x y s") prefer 2 apply(simp add:is_ub_def) apply(subgoal_tac "(u, s) \ r\<^sup>*") prefer 2 apply(blast dest:is_lubD) apply(erule converse_rtranclE) apply blast apply(simp only:acyclic_def) apply(blast intro:rtrancl_into_trancl2 rtrancl_trancl_trancl) apply(rule finite_acyclic_wf) apply simp apply(erule acyclic_single_valued_finite) apply(blast intro:single_valuedI) apply(simp add:is_lub_def is_ub_def) apply simp apply(erule acyclic_subset) apply blast apply simp apply(erule conjE) apply(erule_tac z = u in converse_rtranclE) apply(blast dest: is_lubD is_ubD) apply(blast dest:rtrancl_into_rtrancl) done (*>*) lemma is_lub_exec_lub: "\ single_valued r; acyclic r; (x,u):r^*; (y,u):r^*; \x y. (x,y) \ r \ f x = y \ \ is_lub (r^* ) x y (exec_lub r f x y)" (*<*) by (fastforce dest: single_valued_has_lubs simp add: exec_lub_conv) (*>*) end diff --git a/thys/Jinja/DFA/SemilatAlg.thy b/thys/Jinja/DFA/SemilatAlg.thy --- a/thys/Jinja/DFA/SemilatAlg.thy +++ b/thys/Jinja/DFA/SemilatAlg.thy @@ -1,211 +1,224 @@ (* Title: HOL/MicroJava/BV/SemilatAlg.thy Author: Gerwin Klein Copyright 2002 Technische Universitaet Muenchen *) section \More on Semilattices\ theory SemilatAlg -imports Typing_Framework +imports Typing_Framework_1 begin definition lesubstep_type :: "(nat \ 's) set \ 's ord \ (nat \ 's) set \ bool" ("(_ /{\\<^bsub>_\<^esub>} _)" [50, 0, 51] 50) where "A {\\<^bsub>r\<^esub>} B \ \(p,\) \ A. \\'. (p,\') \ B \ \ \\<^sub>r \'" notation (ASCII) lesubstep_type ("(_ /{<='__} _)" [50, 0, 51] 50) primrec pluslussub :: "'a list \ ('a \ 'a \ 'a) \ 'a \ 'a" ("(_ /\\<^bsub>_\<^esub> _)" [65, 0, 66] 65) where "pluslussub [] f y = y" | "pluslussub (x#xs) f y = pluslussub xs f (x \\<^sub>f y)" (*<*) notation (ASCII) pluslussub ("(_ /++'__ _)" [65, 1000, 66] 65) (*>*) definition bounded :: "'s step_type \ nat \ bool" where "bounded step n \ (\p\. \(q,\') \ set (step p \). q nat \ 's set \ bool" where "pres_type step n A \ (\\\A. \p(q,\') \ set (step p \). \' \ A)" definition mono :: "'s ord \ 's step_type \ nat \ 's set \ bool" where "mono r step n A \ (\\ p \'. \ \ A \ p \ \\<^sub>r \' \ set (step p \) {\\<^bsub>r\<^esub>} set (step p \'))" lemma [iff]: "{} {\\<^bsub>r\<^esub>} B" (*<*) by (simp add: lesubstep_type_def) (*>*) lemma [iff]: "(A {\\<^bsub>r\<^esub>} {}) = (A = {})" (*<*) by (cases "A={}") (auto simp add: lesubstep_type_def) (*>*) lemma lesubstep_union: "\ A\<^sub>1 {\\<^bsub>r\<^esub>} B\<^sub>1; A\<^sub>2 {\\<^bsub>r\<^esub>} B\<^sub>2 \ \ A\<^sub>1 \ A\<^sub>2 {\\<^bsub>r\<^esub>} B\<^sub>1 \ B\<^sub>2" (*<*) by (auto simp add: lesubstep_type_def) (*>*) lemma pres_typeD: "\ pres_type step n A; s\A; pset (step p s) \ \ s' \ A" (*<*) by (unfold pres_type_def, blast) (*>*) lemma monoD: "\ mono r step n A; p < n; s\A; s \\<^sub>r t \ \ set (step p s) {\\<^bsub>r\<^esub>} set (step p t)" (*<*) by (unfold mono_def, blast) (*>*) lemma boundedD: "\ bounded step n; p < n; (q,t) \ set (step p xs) \ \ q < n" (*<*) by (unfold bounded_def, blast) (*>*) lemma lesubstep_type_refl [simp, intro]: "(\x. x \\<^sub>r x) \ A {\\<^bsub>r\<^esub>} A" (*<*) by (unfold lesubstep_type_def) auto (*>*) lemma lesub_step_typeD: "A {\\<^bsub>r\<^esub>} B \ (x,y) \ A \ \y'. (x, y') \ B \ y \\<^sub>r y'" (*<*) by (unfold lesubstep_type_def) blast (*>*) lemma list_update_le_listI [rule_format]: "set xs \ A \ set ys \ A \ xs [\\<^sub>r] ys \ p < size xs \ x \\<^sub>r ys!p \ semilat(A,r,f) \ x\A \ xs[p := x \\<^sub>f xs!p] [\\<^sub>r] ys" (*<*) apply (simp only: Listn.le_def lesub_def semilat_def) apply (simp add: list_all2_conv_all_nth nth_list_update) done (*>*) lemma plusplus_closed: assumes "Semilat A r f" shows "\y. \ set x \ A; y \ A\ \ x \\<^bsub>f\<^esub> y \ A" (*<*) proof (induct x) interpret Semilat A r f by fact show "\y. y \ A \ [] \\<^bsub>f\<^esub> y \ A" by simp fix y x xs assume y: "y \ A" and xxs: "set (x#xs) \ A" assume IH: "\y. \ set xs \ A; y \ A\ \ xs \\<^bsub>f\<^esub> y \ A" from xxs obtain x: "x \ A" and xs: "set xs \ A" by simp from x y have "x \\<^bsub>f\<^esub> y \ A" .. with xs have "xs \\<^bsub>f\<^esub> (x \\<^bsub>f\<^esub> y) \ A" by (rule IH) thus "x#xs \\<^bsub>f\<^esub> y \ A" by simp qed (*>*) + lemma (in Semilat) pp_ub2: "\y. \ set x \ A; y \ A\ \ y \\<^bsub>r\<^esub> x \\<^bsub>f\<^esub> y" (*<*) + proof (induct x) - from semilat show "\y. y \\<^bsub>r\<^esub> [] \\<^bsub>f\<^esub> y" by simp + thm plusplus_closed[OF Semilat.intro, OF semilat] + thm semilat_Def + thm Semilat.intro + fix y assume "y \ A" + with semilat show "y\\<^bsub>r\<^esub> [] \\<^bsub>f\<^esub> y" by simp fix y a l assume y: "y \ A" and "set (a#l) \ A" then obtain a: "a \ A" and x: "set l \ A" by simp + + from semilat x y have l_y_inA: "l \\<^bsub>f\<^esub> y \ A" by (auto dest: plusplus_closed[OF Semilat.intro, OF semilat]) assume "\y. \set l \ A; y \ A\ \ y \\<^bsub>r\<^esub> l \\<^bsub>f\<^esub> y" from this and x have IH: "\y. y \ A \ y \\<^bsub>r\<^esub> l \\<^bsub>f\<^esub> y" . + from a y have "a \\<^bsub>f\<^esub> y \ A" .. + then have l_ay_inA: "l \\<^bsub>f\<^esub> ( a \\<^bsub>f\<^esub> y) \ A" using plusplus_closed[OF Semilat.intro, OF semilat] x by auto + from a y have "y \\<^bsub>r\<^esub> a \\<^bsub>f\<^esub> y" .. - also from a y have "a \\<^bsub>f\<^esub> y \ A" .. + also from a y have a_y_inA: "a \\<^bsub>f\<^esub> y \ A" .. hence "(a \\<^bsub>f\<^esub> y) \\<^bsub>r\<^esub> l \\<^bsub>f\<^esub> (a \\<^bsub>f\<^esub> y)" by (rule IH) - finally have "y \\<^bsub>r\<^esub> l \\<^bsub>f\<^esub> (a \\<^bsub>f\<^esub> y)" . + finally have "y \\<^bsub>r\<^esub> l \\<^bsub>f\<^esub> (a \\<^bsub>f\<^esub> y)" using y a_y_inA l_ay_inA . thus "y \\<^bsub>r\<^esub> (a#l) \\<^bsub>f\<^esub> y" by simp qed (*>*) lemma (in Semilat) pp_ub1: shows "\y. \set ls \ A; y \ A; x \ set ls\ \ x \\<^bsub>r\<^esub> ls \\<^bsub>f\<^esub> y" (*<*) proof (induct ls) show "\y. x \ set [] \ x \\<^bsub>r\<^esub> [] \\<^bsub>f\<^esub> y" by simp fix y s ls assume "set (s#ls) \ A" then obtain s: "s \ A" and ls: "set ls \ A" by simp assume y: "y \ A" assume "\y. \set ls \ A; y \ A; x \ set ls\ \ x \\<^bsub>r\<^esub> ls \\<^bsub>f\<^esub> y" from this and ls have IH: "\y. x \ set ls \ y \ A \ x \\<^bsub>r\<^esub> ls \\<^bsub>f\<^esub> y" . + from s y have s_y_inA: "s \\<^bsub>f\<^esub> y \ A" .. + then have ls_sy_inA: "ls \\<^bsub>f\<^esub> (s \\<^bsub>f\<^esub> y) \ A"using plusplus_closed[OF Semilat.intro, OF semilat] ls by auto + assume "x \ set (s#ls)" then obtain xls: "x = s \ x \ set ls" by simp moreover { assume xs: "x = s" from s y have "s \\<^bsub>r\<^esub> s \\<^bsub>f\<^esub> y" .. - also from s y have "s \\<^bsub>f\<^esub> y \ A" .. - with ls have "(s \\<^bsub>f\<^esub> y) \\<^bsub>r\<^esub> ls \\<^bsub>f\<^esub> (s \\<^bsub>f\<^esub> y)" by (rule pp_ub2) - finally have "s \\<^bsub>r\<^esub> ls \\<^bsub>f\<^esub> (s \\<^bsub>f\<^esub> y)" . + also from ls s_y_inA have "(s \\<^bsub>f\<^esub> y) \\<^bsub>r\<^esub> ls \\<^bsub>f\<^esub> (s \\<^bsub>f\<^esub> y)" by (rule pp_ub2) + finally have "s \\<^bsub>r\<^esub> ls \\<^bsub>f\<^esub> (s \\<^bsub>f\<^esub> y)" using s s_y_inA ls_sy_inA . with xs have "x \\<^bsub>r\<^esub> ls \\<^bsub>f\<^esub> (s \\<^bsub>f\<^esub> y)" by simp } moreover { assume "x \ set ls" hence "\y. y \ A \ x \\<^bsub>r\<^esub> ls \\<^bsub>f\<^esub> y" by (rule IH) moreover from s y have "s \\<^bsub>f\<^esub> y \ A" .. ultimately have "x \\<^bsub>r\<^esub> ls \\<^bsub>f\<^esub> (s \\<^bsub>f\<^esub> y)" . } ultimately have "x \\<^bsub>r\<^esub> ls \\<^bsub>f\<^esub> (s \\<^bsub>f\<^esub> y)" by blast thus "x \\<^bsub>r\<^esub> (s#ls) \\<^bsub>f\<^esub> y" by simp qed (*>*) lemma (in Semilat) pp_lub: assumes z: "z \ A" shows "\y. y \ A \ set xs \ A \ \x \ set xs. x \\<^bsub>r\<^esub> z \ y \\<^bsub>r\<^esub> z \ xs \\<^bsub>f\<^esub> y \\<^bsub>r\<^esub> z" (*<*) proof (induct xs) fix y assume "y \\<^bsub>r\<^esub> z" thus "[] \\<^bsub>f\<^esub> y \\<^bsub>r\<^esub> z" by simp next fix y l ls assume y: "y \ A" and "set (l#ls) \ A" then obtain l: "l \ A" and ls: "set ls \ A" by auto assume "\x \ set (l#ls). x \\<^bsub>r\<^esub> z" then obtain lz: "l \\<^bsub>r\<^esub> z" and lsz: "\x \ set ls. x \\<^bsub>r\<^esub> z" by auto assume "y \\<^bsub>r\<^esub> z" with lz have "l \\<^bsub>f\<^esub> y \\<^bsub>r\<^esub> z" using l y z .. moreover from l y have "l \\<^bsub>f\<^esub> y \ A" .. moreover assume "\y. y \ A \ set ls \ A \ \x \ set ls. x \\<^bsub>r\<^esub> z \ y \\<^bsub>r\<^esub> z \ ls \\<^bsub>f\<^esub> y \\<^bsub>r\<^esub> z" ultimately have "ls \\<^bsub>f\<^esub> (l \\<^bsub>f\<^esub> y) \\<^bsub>r\<^esub> z" using ls lsz by - thus "(l#ls) \\<^bsub>f\<^esub> y \\<^bsub>r\<^esub> z" by simp qed (*>*) lemma ub1': assumes "Semilat A r f" shows "\\(p,s) \ set S. s \ A; y \ A; (a,b) \ set S\ \ b \\<^bsub>r\<^esub> map snd [(p', t') \ S. p' = a] \\<^bsub>f\<^esub> y" (*<*) proof - interpret Semilat A r f by fact let "b \\<^bsub>r\<^esub> ?map \\<^bsub>f\<^esub> y" = ?thesis assume "y \ A" moreover assume "\(p,s) \ set S. s \ A" hence "set ?map \ A" by auto moreover assume "(a,b) \ set S" hence "b \ set ?map" by (induct S, auto) ultimately show ?thesis by - (rule pp_ub1) qed (*>*) lemma plusplus_empty: "\s'. (q, s') \ set S \ s' \\<^bsub>f\<^esub> ss ! q = ss ! q \ (map snd [(p', t') \ S. p' = q] \\<^bsub>f\<^esub> ss ! q) = ss ! q" (*<*) apply (induct S) apply auto done (*>*) end diff --git a/thys/Jinja/DFA/Typing_Framework_err.thy b/thys/Jinja/DFA/Typing_Framework_err.thy --- a/thys/Jinja/DFA/Typing_Framework_err.thy +++ b/thys/Jinja/DFA/Typing_Framework_err.thy @@ -1,267 +1,268 @@ (* Title: HOL/MicroJava/BV/Typing_Framework_err.thy Author: Gerwin Klein Copyright 2000 TUM *) section \Lifting the Typing Framework to err, app, and eff\ -theory Typing_Framework_err imports Typing_Framework SemilatAlg begin +theory Typing_Framework_err imports SemilatAlg begin definition wt_err_step :: "'s ord \ 's err step_type \ 's err list \ bool" where "wt_err_step r step \s \ wt_step (Err.le r) Err step \s" definition wt_app_eff :: "'s ord \ (nat \ 's \ bool) \ 's step_type \ 's list \ bool" where "wt_app_eff r app step \s \ (\p < size \s. app p (\s!p) \ (\(q,\) \ set (step p (\s!p)). \ <=_r \s!q))" definition map_snd :: "('b \ 'c) \ ('a \ 'b) list \ ('a \ 'c) list" where "map_snd f = map (\(x,y). (x, f y))" definition error :: "nat \ (nat \ 'a err) list" where "error n = map (\x. (x,Err)) [0.. (nat \ 's \ bool) \ 's step_type \ 's err step_type" where "err_step n app step p t = (case t of Err \ error n | OK \ \ if app p \ then map_snd OK (step p \) else error n)" definition app_mono :: "'s ord \ (nat \ 's \ bool) \ nat \ 's set \ bool" where "app_mono r app n A \ (\s p t. s \ A \ p < n \ s \\<^sub>r t \ app p t \ app p s)" lemmas err_step_defs = err_step_def map_snd_def error_def lemma bounded_err_stepD: "\ bounded (err_step n app step) n; p < n; app p a; (q,b) \ set (step p a) \ \ q < n" (*<*) apply (simp add: bounded_def err_step_def) apply (erule allE, erule impE, assumption) apply (erule_tac x = "OK a" in allE, drule bspec) apply (simp add: map_snd_def) apply fast apply simp done (*>*) lemma in_map_sndD: "(a,b) \ set (map_snd f xs) \ \b'. (a,b') \ set xs" (*<*) apply (induct xs) apply (auto simp add: map_snd_def) done (*>*) lemma bounded_err_stepI: "\p. p < n \ (\s. ap p s \ (\(q,s') \ set (step p s). q < n)) \ bounded (err_step n ap step) n" (*<*) apply (clarsimp simp: bounded_def err_step_def split: err.splits) apply (simp add: error_def image_def) apply (blast dest: in_map_sndD) done (*>*) lemma bounded_lift: "bounded step n \ bounded (err_step n app step) n" (*<*) apply (unfold bounded_def err_step_def error_def) apply clarify apply (erule allE, erule impE, assumption) apply (case_tac \) apply (auto simp add: map_snd_def split: if_split_asm) done (*>*) lemma le_list_map_OK [simp]: "\b. (map OK a [\\<^bsub>Err.le r\<^esub>] map OK b) = (a [\\<^sub>r] b)" (*<*) apply (induct a) apply simp apply simp apply (case_tac b) apply simp apply simp done (*>*) lemma map_snd_lessI: "set xs {\\<^bsub>r\<^esub>} set ys \ set (map_snd OK xs) {\\<^bsub>Err.le r\<^esub>} set (map_snd OK ys)" (*<*) apply (induct xs) apply (unfold lesubstep_type_def map_snd_def) apply auto done (*>*) lemma mono_lift: - "\ order r; app_mono r app n A; bounded (err_step n app step) n; + "\ order r A; app_mono r app n A; bounded (err_step n app step) n; \s p t. s \ A \ p < n \ s \\<^sub>r t \ app p t \ set (step p s) {\\<^bsub>r\<^esub>} set (step p t) \ \ mono (Err.le r) (err_step n app step) n (err A)" (*<*) apply (simp only: app_mono_def SemilatAlg.mono_def err_step_def) apply clarify apply (case_tac \) - apply simp + defer apply simp apply (case_tac \') apply simp apply clarify apply (simp add: lesubstep_type_def error_def) apply clarify apply (drule in_map_sndD) apply clarify apply (drule bounded_err_stepD, assumption+) apply (rule exI [of _ Err]) apply simp apply simp apply (erule allE, erule allE, erule allE, erule impE) apply (rule conjI, assumption) apply (rule conjI, assumption) apply assumption apply (rule conjI) apply clarify apply (erule allE, erule allE, erule allE, erule impE) apply (rule conjI, assumption) apply (rule conjI, assumption) apply assumption apply (erule impE, assumption) apply (rule map_snd_lessI, assumption) apply clarify apply (simp add: lesubstep_type_def error_def) apply clarify apply (drule in_map_sndD) apply clarify apply (drule bounded_err_stepD, assumption+) apply (rule exI [of _ Err]) apply simp + apply (auto simp: lesubstep_type_def error_def) done (*>*) lemma in_errorD: "(x,y) \ set (error n) \ y = Err" (*<*) by (auto simp add: error_def) (*>*) lemma pres_type_lift: "\s\A. \p. p < n \ app p s \ (\(q, s')\set (step p s). s' \ A) \ pres_type (err_step n app step) n (err A)" (*<*) apply (unfold pres_type_def err_step_def) apply clarify apply (case_tac b) apply simp apply (case_tac \) apply simp apply (drule in_errorD) apply simp apply (simp add: map_snd_def split: if_split_asm) apply fast apply (drule in_errorD) apply simp done (*>*) lemma wt_err_imp_wt_app_eff: assumes wt: "wt_err_step r (err_step (size ts) app step) ts" assumes b: "bounded (err_step (size ts) app step) (size ts)" shows "wt_app_eff r app step (map ok_val ts)" (*<*) proof (unfold wt_app_eff_def, intro strip, rule conjI) fix p assume "p < size (map ok_val ts)" hence lp: "p < size ts" by simp hence ts: "0 < size ts" by (cases p) auto hence err: "(0,Err) \ set (error (size ts))" by (simp add: error_def) from wt lp have [intro?]: "\p. p < size ts \ ts ! p \ Err" by (unfold wt_err_step_def wt_step_def) simp show app: "app p (map ok_val ts ! p)" proof (rule ccontr) from wt lp obtain s where OKp: "ts ! p = OK s" and less: "\(q,t) \ set (err_step (size ts) app step p (ts!p)). t <=_(Err.le r) ts!q" by (unfold wt_err_step_def wt_step_def stable_def) (auto iff: not_Err_eq) assume "\ app p (map ok_val ts ! p)" with OKp lp have "\ app p s" by simp with OKp have "err_step (size ts) app step p (ts!p) = error (size ts)" by (simp add: err_step_def) with err ts obtain q where "(q,Err) \ set (err_step (size ts) app step p (ts!p))" and q: "q < size ts" by auto with less have "ts!q = Err" by auto moreover from q have "ts!q \ Err" .. ultimately show False by blast qed show "\(q,t)\set(step p (map ok_val ts ! p)). t \\<^sub>r map ok_val ts ! q" proof clarify fix q t assume q: "(q,t) \ set (step p (map ok_val ts ! p))" from wt lp q obtain s where OKp: "ts ! p = OK s" and less: "\(q,t) \ set (err_step (size ts) app step p (ts!p)). t <=_(Err.le r) ts!q" by (unfold wt_err_step_def wt_step_def stable_def) (auto iff: not_Err_eq) from b lp app q have lq: "q < size ts" by (rule bounded_err_stepD) hence "ts!q \ Err" .. then obtain s' where OKq: "ts ! q = OK s'" by (auto iff: not_Err_eq) from lp lq OKp OKq app less q show "t \\<^sub>r map ok_val ts ! q" by (auto simp add: err_step_def map_snd_def) qed qed (*>*) lemma wt_app_eff_imp_wt_err: assumes app_eff: "wt_app_eff r app step ts" assumes bounded: "bounded (err_step (size ts) app step) (size ts)" shows "wt_err_step r (err_step (size ts) app step) (map OK ts)" (*<*) proof (unfold wt_err_step_def wt_step_def, intro strip, rule conjI) fix p assume "p < size (map OK ts)" hence p: "p < size ts" by simp thus "map OK ts ! p \ Err" by simp { fix q t assume q: "(q,t) \ set (err_step (size ts) app step p (map OK ts ! p))" with p app_eff obtain "app p (ts ! p)" "\(q,t) \ set (step p (ts!p)). t \\<^sub>r ts!q" by (unfold wt_app_eff_def) blast moreover from q p bounded have "q < size ts" by - (rule boundedD) hence "map OK ts ! q = OK (ts!q)" by simp moreover have "p < size ts" by (rule p) moreover note q ultimately have "t \\<^bsub>Err.le r\<^esub> map OK ts ! q" by (auto simp add: err_step_def map_snd_def) } thus "stable (Err.le r) (err_step (size ts) app step) (map OK ts) p" by (unfold stable_def) blast qed (*>*) end