diff --git a/thys/AI_Planning_Languages_Semantics/SASP_Checker.thy b/thys/AI_Planning_Languages_Semantics/SASP_Checker.thy --- a/thys/AI_Planning_Languages_Semantics/SASP_Checker.thy +++ b/thys/AI_Planning_Languages_Semantics/SASP_Checker.thy @@ -1,347 +1,347 @@ theory SASP_Checker imports SASP_Semantics "HOL-Library.Code_Target_Nat" begin section \An Executable Checker for Multi-Valued Planning Problem Solutions\ subsection \Auxiliary Lemmas\ lemma map_of_leI: assumes "distinct (map fst l)" assumes "\k v. (k,v)\set l \ m k = Some v" shows "map_of l \\<^sub>m m" using assms by (metis (no_types, opaque_lifting) domIff map_le_def map_of_SomeD not_Some_eq) lemma [simp]: "fst \ (\(a, b, c, d). (f a b c d, g a b c d)) = (\(a,b,c,d). f a b c d)" by auto lemma map_mp: "m\\<^sub>mm' \ m k = Some v \ m' k = Some v" by (auto simp: map_le_def dom_def) lemma map_add_map_of_fold: fixes ps and m :: "'a \ 'b" assumes "distinct (map fst ps)" shows "m ++ map_of ps = fold (\(k, v) m. m(k \ v)) ps m" proof - - have X1: "fold (\(k, v) m. m(k \ v)) ps m(a \ b) + have X1: "(fold (\(k, v) m. m(k \ v)) ps m)(a \ b) = fold (\(k, v) m. m(k \ v)) ps (m(a \ b))" if "a \ fst ` set ps" for a b ps and m :: "'a \ 'b" using that by (induction ps arbitrary: m) (auto simp: fun_upd_twist) show ?thesis using assms by (induction ps arbitrary: m) (auto simp: X1) qed subsection \Well-formedness Check\ lemmas wf_code_thms = ast_problem.astDom_def ast_problem.astI_def ast_problem.astG_def ast_problem.ast\_def ast_problem.numVars_def ast_problem.numVals_def ast_problem.wf_partial_state_def ast_problem.wf_operator_def ast_problem.well_formed_def declare wf_code_thms[code] export_code ast_problem.well_formed in SML subsection \Execution\ definition match_pre :: "ast_precond \ state \ bool" where "match_pre \ \(x,v) s. s x = Some v" definition match_pres :: "ast_precond list \ state \ bool" where "match_pres pres s \ \pre\set pres. match_pre pre s" definition match_implicit_pres :: "ast_effect list \ state \ bool" where "match_implicit_pres effs s \ \(_,x,vp,_)\set effs. (case vp of None \ True | Some v \ s x = Some v)" definition enabled_opr' :: "ast_operator \ state \ bool" where "enabled_opr' \ \(name,pres,effs,cost) s. match_pres pres s \ match_implicit_pres effs s" definition eff_enabled' :: "state \ ast_effect \ bool" where "eff_enabled' s \ \(pres,_,_,_). match_pres pres s" definition "execute_opr' \ \(name,_,effs,_) s. let effs = filter (eff_enabled' s) effs in fold (\(_,x,_,v) s. s(x\v)) effs s " definition lookup_operator' :: "ast_problem \ name \ ast_operator" where "lookup_operator' \ \(D,I,G,\) name. find (\(n,_,_,_). n=name) \" definition enabled' :: "ast_problem \ name \ state \ bool" where "enabled' problem name s \ case lookup_operator' problem name of Some \ \ enabled_opr' \ s | None \ False" definition execute' :: "ast_problem \ name \ state \ state" where "execute' problem name s \ case lookup_operator' problem name of Some \ \ execute_opr' \ s | None \ undefined" context wf_ast_problem begin lemma match_pres_correct: assumes D: "distinct (map fst pres)" assumes "s\valid_states" shows "match_pres pres s \ s\subsuming_states (map_of pres)" proof - have "match_pres pres s \ map_of pres \\<^sub>m s" unfolding match_pres_def match_pre_def apply (auto split: prod.splits) using map_le_def map_of_SomeD apply fastforce by (metis (full_types) D domIff map_le_def map_of_eq_Some_iff option.simps(3)) with assms show ?thesis unfolding subsuming_states_def by auto qed lemma match_implicit_pres_correct: assumes D: "distinct (map (\(_, v, _, _). v) effs)" assumes "s\valid_states" shows "match_implicit_pres effs s \ s\subsuming_states (map_of (implicit_pres effs))" proof - from assms show ?thesis unfolding subsuming_states_def unfolding match_implicit_pres_def implicit_pres_def apply (auto split: prod.splits option.splits simp: distinct_map_filter intro!: map_of_leI) apply (force simp: distinct_map_filter split: prod.split elim: map_mp) done qed lemma enabled_opr'_correct: assumes V: "s\valid_states" assumes "lookup_operator name = Some \" shows "enabled_opr' \ s \ enabled name s" using lookup_operator_wf[OF assms(2)] assms unfolding enabled_opr'_def enabled_def wf_operator_def by (auto simp: match_pres_correct[OF _ V] match_implicit_pres_correct[OF _ V] simp: wf_partial_state_def split: option.split ) lemma eff_enabled'_correct: assumes V: "s\valid_states" assumes "case eff of (pres,_,_,_) \ wf_partial_state pres" shows "eff_enabled' s eff \ eff_enabled s eff" using assms unfolding eff_enabled'_def eff_enabled_def wf_partial_state_def by (auto simp: match_pres_correct) lemma execute_opr'_correct: assumes V: "s\valid_states" assumes LO: "lookup_operator name = Some \" shows "execute_opr' \ s = execute name s" proof (cases \) case [simp]: (fields name pres effs) have [simp]: "filter (eff_enabled' s) effs = filter (eff_enabled s) effs" apply (rule filter_cong[OF refl]) apply (rule eff_enabled'_correct[OF V]) using lookup_operator_wf[OF LO] unfolding wf_operator_def by auto have X1: "distinct (map fst (map (\(_, x, _, y). (x, y)) (filter (eff_enabled s) effs)))" using lookup_operator_wf[OF LO] unfolding wf_operator_def by (auto simp: distinct_map_filter) term "filter (eff_enabled s) effs" have [simp]: "fold (\(_, x, _, v) s. s(x \ v)) l s = fold (\(k, v) m. m(k \ v)) (map (\(_, x, _, y). (x, y)) l) s" for l :: "ast_effect list" by (induction l arbitrary: s) auto show ?thesis unfolding execute_opr'_def execute_def using LO by (auto simp: map_add_map_of_fold[OF X1]) qed lemma lookup_operator'_correct: "lookup_operator' problem name = lookup_operator name" unfolding lookup_operator'_def lookup_operator_def unfolding ast\_def by (auto split: prod.split) lemma enabled'_correct: assumes V: "s\valid_states" shows "enabled' problem name s = enabled name s" unfolding enabled'_def using enabled_opr'_correct[OF V] by (auto split: option.splits simp: enabled_def lookup_operator'_correct) lemma execute'_correct: assumes V: "s\valid_states" assumes "enabled name s" (* Intentionally put this here, also we could resolve non-enabled case by reflexivity (undefined=undefined) *) shows "execute' problem name s = execute name s" unfolding execute'_def using execute_opr'_correct[OF V] \enabled name s\ by (auto split: option.splits simp: enabled_def lookup_operator'_correct) end context ast_problem begin fun simulate_plan :: "plan \ state \ state" where "simulate_plan [] s = Some s" | "simulate_plan (\#\s) s = ( if enabled \ s then let s' = execute \ s in simulate_plan \s s' else None )" lemma simulate_plan_correct: "simulate_plan \s s = Some s' \ path_to s \s s'" by (induction s \s s' rule: path_to.induct) auto definition check_plan :: "plan \ bool" where "check_plan \s = ( case simulate_plan \s I of None \ False | Some s' \ s' \ G)" lemma check_plan_correct: "check_plan \s \ valid_plan \s" unfolding check_plan_def valid_plan_def by (auto split: option.split simp: simulate_plan_correct[symmetric]) end fun simulate_plan' :: "ast_problem \ plan \ state \ state" where "simulate_plan' problem [] s = Some s" | "simulate_plan' problem (\#\s) s = ( if enabled' problem \ s then let s = execute' problem \ s in simulate_plan' problem \s s else None )" text \Avoiding duplicate lookup.\ (*[code] *) lemma simulate_plan'_code[code]: "simulate_plan' problem [] s = Some s" "simulate_plan' problem (\#\s) s = ( case lookup_operator' problem \ of None \ None | Some \ \ if enabled_opr' \ s then simulate_plan' problem \s (execute_opr' \ s) else None )" by (auto simp: enabled'_def execute'_def split: option.split) definition initial_state' :: "ast_problem \ state" where "initial_state' problem \ let astI = ast_problem.astI problem in ( \v. if v plan \ bool" where "check_plan' problem \s = ( case simulate_plan' problem \s (initial_state' problem) of None \ False | Some s' \ match_pres (ast_problem.astG problem) s')" context wf_ast_problem begin lemma simulate_plan'_correct: assumes "s\valid_states" shows "simulate_plan' problem \s s = simulate_plan \s s" using assms apply (induction \s s rule: simulate_plan.induct) apply (auto simp: enabled'_correct execute'_correct execute_preserves_valid) done lemma simulate_plan'_correct_paper: (* For presentation in paper. Summarizing intermediate refinement step. *) assumes "s\valid_states" shows "simulate_plan' problem \s s = Some s' \ path_to s \s s'" using simulate_plan'_correct[OF assms] simulate_plan_correct by simp lemma initial_state'_correct: "initial_state' problem = I" unfolding initial_state'_def I_def by (auto simp: Let_def) lemma check_plan'_correct: "check_plan' problem \s = check_plan \s" proof - have D: "distinct (map fst astG)" using wf_goal unfolding wf_partial_state_def by auto have S'V: "s'\valid_states" if "simulate_plan \s I = Some s'" for s' using that by (auto simp: simulate_plan_correct path_to_pres_valid[OF I_valid]) show ?thesis unfolding check_plan'_def check_plan_def by (auto split: option.splits simp: initial_state'_correct simulate_plan'_correct[OF I_valid] simp: match_pres_correct[OF D S'V] G_def ) qed end (* Overall checker *) definition verify_plan :: "ast_problem \ plan \ String.literal + unit" where "verify_plan problem \s = ( if ast_problem.well_formed problem then if check_plan' problem \s then Inr () else Inl (STR ''Invalid plan'') else Inl (STR ''Problem not well formed'') )" lemma verify_plan_correct: "verify_plan problem \s = Inr () \ ast_problem.well_formed problem \ ast_problem.valid_plan problem \s" proof - { assume "ast_problem.well_formed problem" then interpret wf_ast_problem by unfold_locales from check_plan'_correct check_plan_correct have "check_plan' problem \s = valid_plan \s" by simp } then show ?thesis unfolding verify_plan_def by auto qed definition nat_opt_of_integer :: "integer \ nat option" where "nat_opt_of_integer i = (if (i \ 0) then Some (nat_of_integer i) else None)" (*Export functions, which includes constructors*) export_code verify_plan nat_of_integer integer_of_nat nat_opt_of_integer Inl Inr String.explode String.implode in SML module_name SASP_Checker_Exported end diff --git a/thys/Automatic_Refinement/Lib/Misc.thy b/thys/Automatic_Refinement/Lib/Misc.thy --- a/thys/Automatic_Refinement/Lib/Misc.thy +++ b/thys/Automatic_Refinement/Lib/Misc.thy @@ -1,4507 +1,4507 @@ (* Title: Miscellaneous Definitions and Lemmas Author: Peter Lammich Maintainer: Peter Lammich Thomas Tuerk *) (* CHANGELOG: 2010-05-09: Removed AC, AI locales, they are superseeded by concepts from OrderedGroups 2010-09-22: Merges with ext/Aux 2017-06-12: Added a bunch of lemmas from various other projects *) section \Miscellaneous Definitions and Lemmas\ theory Misc imports Main "HOL-Library.Multiset" "HOL-ex.Quicksort" "HOL-Library.Option_ord" "HOL-Library.Infinite_Set" "HOL-Eisbach.Eisbach" begin text_raw \\label{thy:Misc}\ subsection \Isabelle Distribution Move Proposals\ subsubsection \Pure\ lemma TERMI: "TERM x" unfolding Pure.term_def . subsubsection \HOL\ (* Stronger disjunction elimination rules. *) lemma disjE1: "\ P \ Q; P \ R; \\P;Q\ \ R \ \ R" by metis lemma disjE2: "\ P \ Q; \P; \Q\ \ R; Q \ R \ \ R" by blast lemma imp_mp_iff[simp]: "a \ (a \ b) \ a \ b" "(a \ b) \ a \ a \ b" (* is Inductive.imp_conj_iff, but not in simpset by default *) by blast+ lemma atomize_Trueprop_eq[atomize]: "(Trueprop x \ Trueprop y) \ Trueprop (x=y)" apply rule apply (rule) apply (erule equal_elim_rule1) apply assumption apply (erule equal_elim_rule2) apply assumption apply simp done subsubsection \Set\ lemma inter_compl_diff_conv[simp]: "A \ -B = A - B" by auto lemma pair_set_inverse[simp]: "{(a,b). P a b}\ = {(b,a). P a b}" by auto lemma card_doubleton_eq_2_iff[simp]: "card {a,b} = 2 \ a\b" by auto subsubsection \List\ (* TODO: Move, analogous to List.length_greater_0_conv *) thm List.length_greater_0_conv lemma length_ge_1_conv[iff]: "Suc 0 \ length l \ l\[]" by (cases l) auto \ \Obtains a list from the pointwise characterization of its elements\ lemma obtain_list_from_elements: assumes A: "\ili. P li i)" obtains l where "length l = n" "\il. length l=n \ (\iii l!j" by (simp add: sorted_iff_nth_mono) also from nth_eq_iff_index_eq[OF D] B have "l!i \ l!j" by auto finally show ?thesis . qed lemma distinct_sorted_strict_mono_iff: assumes "distinct l" "sorted l" assumes "i il!j \ i\j" by (metis assms distinct_sorted_strict_mono_iff leD le_less_linear) (* List.thy has: declare map_eq_Cons_D [dest!] Cons_eq_map_D [dest!] We could, analogously, declare rules for "map _ _ = _@_" as dest!, or use elim!, or declare the _conv-rule as simp *) lemma map_eq_appendE: assumes "map f ls = fl@fl'" obtains l l' where "ls=l@l'" and "map f l=fl" and "map f l' = fl'" using assms proof (induction fl arbitrary: ls thesis) case (Cons x xs) then obtain l ls' where [simp]: "ls = l#ls'" "f l = x" by force with Cons.prems(2) have "map f ls' = xs @ fl'" by simp from Cons.IH[OF _ this] obtain ll ll' where "ls' = ll @ ll'" "map f ll = xs" "map f ll' = fl'" . with Cons.prems(1)[of "l#ll" ll'] show thesis by simp qed simp lemma map_eq_append_conv: "map f ls = fl@fl' \ (\l l'. ls = l@l' \ map f l = fl \ map f l' = fl')" by (auto elim!: map_eq_appendE) lemmas append_eq_mapE = map_eq_appendE[OF sym] lemma append_eq_map_conv: "fl@fl' = map f ls \ (\l l'. ls = l@l' \ map f l = fl \ map f l' = fl')" by (auto elim!: append_eq_mapE) lemma distinct_mapI: "distinct (map f l) \ distinct l" by (induct l) auto lemma map_distinct_upd_conv: "\i \ (map f l)[i := x] = map (f(l!i := x)) l" \ \Updating a mapped distinct list is equal to updating the mapping function\ by (auto simp: nth_eq_iff_index_eq intro: nth_equalityI) lemma zip_inj: "\length a = length b; length a' = length b'; zip a b = zip a' b'\ \ a=a' \ b=b'" proof (induct a b arbitrary: a' b' rule: list_induct2) case Nil then show ?case by (cases a'; cases b'; auto) next case (Cons x xs y ys) then show ?case by (cases a'; cases b'; auto) qed lemma zip_eq_zip_same_len[simp]: "\ length a = length b; length a' = length b' \ \ zip a b = zip a' b' \ a=a' \ b=b'" by (auto dest: zip_inj) lemma upt_merge[simp]: "i\j \ j\k \ [i.. (ys \ [] \ butlast ys = xs \ last ys = x)" by auto (* Case distinction how two elements of a list can be related to each other *) lemma list_match_lel_lel: assumes "c1 @ qs # c2 = c1' @ qs' # c2'" obtains (left) c21' where "c1 = c1' @ qs' # c21'" "c2' = c21' @ qs # c2" | (center) "c1' = c1" "qs' = qs" "c2' = c2" | (right) c21 where "c1' = c1 @ qs # c21" "c2 = c21 @ qs' # c2'" using assms apply (clarsimp simp: append_eq_append_conv2) subgoal for us by (cases us) auto done lemma xy_in_set_cases[consumes 2, case_names EQ XY YX]: assumes A: "x\set l" "y\set l" and C: "!!l1 l2. \ x=y; l=l1@y#l2 \ \ P" "!!l1 l2 l3. \ x\y; l=l1@x#l2@y#l3 \ \ P" "!!l1 l2 l3. \ x\y; l=l1@y#l2@x#l3 \ \ P" shows P proof (cases "x=y") case True with A(1) obtain l1 l2 where "l=l1@y#l2" by (blast dest: split_list) with C(1) True show ?thesis by blast next case False from A(1) obtain l1 l2 where S1: "l=l1@x#l2" by (blast dest: split_list) from A(2) obtain l1' l2' where S2: "l=l1'@y#l2'" by (blast dest: split_list) from S1 S2 have M: "l1@x#l2 = l1'@y#l2'" by simp thus P proof (cases rule: list_match_lel_lel[consumes 1, case_names 1 2 3]) case (1 c) with S1 have "l=l1'@y#c@x#l2" by simp with C(3) False show ?thesis by blast next case 2 with False have False by blast thus ?thesis .. next case (3 c) with S1 have "l=l1@x#c@y#l2'" by simp with C(2) False show ?thesis by blast qed qed lemma list_e_eq_lel[simp]: "[e] = l1@e'#l2 \ l1=[] \ e'=e \ l2=[]" "l1@e'#l2 = [e] \ l1=[] \ e'=e \ l2=[]" apply (cases l1, auto) [] apply (cases l1, auto) [] done lemma list_ee_eq_leel[simp]: "([e1,e2] = l1@e1'#e2'#l2) \ (l1=[] \ e1=e1' \ e2=e2' \ l2=[])" "(l1@e1'#e2'#l2 = [e1,e2]) \ (l1=[] \ e1=e1' \ e2=e2' \ l2=[])" apply (cases l1, auto) [] apply (cases l1, auto) [] done subsubsection \Transitive Closure\ text \A point-free induction rule for elements reachable from an initial set\ lemma rtrancl_reachable_induct[consumes 0, case_names base step]: assumes I0: "I \ INV" assumes IS: "E``INV \ INV" shows "E\<^sup>*``I \ INV" by (metis I0 IS Image_closed_trancl Image_mono subset_refl) lemma acyclic_empty[simp, intro!]: "acyclic {}" by (unfold acyclic_def) auto lemma acyclic_union: "acyclic (A\B) \ acyclic A" "acyclic (A\B) \ acyclic B" by (metis Un_upper1 Un_upper2 acyclic_subset)+ text \Here we provide a collection of miscellaneous definitions and helper lemmas\ subsection "Miscellaneous (1)" text \This stuff is used in this theory itself, and thus occurs in first place or is simply not sorted into any other section of this theory.\ lemma IdD: "(a,b)\Id \ a=b" by simp text \Conversion Tag\ definition [simp]: "CNV x y \ x=y" lemma CNV_I: "CNV x x" by simp lemma CNV_eqD: "CNV x y \ x=y" by simp lemma CNV_meqD: "CNV x y \ x\y" by simp (* TODO: Move. Shouldn't this be detected by simproc? *) lemma ex_b_b_and_simp[simp]: "(\b. b \ Q b) \ Q True" by auto lemma ex_b_not_b_and_simp[simp]: "(\b. \b \ Q b) \ Q False" by auto method repeat_all_new methods m = m;(repeat_all_new \m\)? subsubsection "AC-operators" text \Locale to declare AC-laws as simplification rules\ locale Assoc = fixes f assumes assoc[simp]: "f (f x y) z = f x (f y z)" locale AC = Assoc + assumes commute[simp]: "f x y = f y x" lemma (in AC) left_commute[simp]: "f x (f y z) = f y (f x z)" by (simp only: assoc[symmetric]) simp lemmas (in AC) AC_simps = commute assoc left_commute text \Locale to define functions from surjective, unique relations\ locale su_rel_fun = fixes F and f assumes unique: "\(A,B)\F; (A,B')\F\ \ B=B'" assumes surjective: "\!!B. (A,B)\F \ P\ \ P" assumes f_def: "f A == THE B. (A,B)\F" lemma (in su_rel_fun) repr1: "(A,f A)\F" proof (unfold f_def) obtain B where "(A,B)\F" by (rule surjective) with theI[where P="\B. (A,B)\F", OF this] show "(A, THE x. (A, x) \ F) \ F" by (blast intro: unique) qed lemma (in su_rel_fun) repr2: "(A,B)\F \ B=f A" using repr1 by (blast intro: unique) lemma (in su_rel_fun) repr: "(f A = B) = ((A,B)\F)" using repr1 repr2 by (blast) \ \Contract quantification over two variables to pair\ lemma Ex_prod_contract: "(\a b. P a b) \ (\z. P (fst z) (snd z))" by auto lemma All_prod_contract: "(\a b. P a b) \ (\z. P (fst z) (snd z))" by auto lemma nat_geq_1_eq_neqz: "x\1 \ x\(0::nat)" by auto lemma nat_in_between_eq: "(a b\Suc a) \ b = Suc a" "(a\b \ b b = a" by auto lemma Suc_n_minus_m_eq: "\ n\m; m>1 \ \ Suc (n - m) = n - (m - 1)" by simp lemma Suc_to_right: "Suc n = m \ n = m - Suc 0" by simp lemma Suc_diff[simp]: "\n m. n\m \ m\1 \ Suc (n - m) = n - (m - 1)" by simp lemma if_not_swap[simp]: "(if \c then a else b) = (if c then b else a)" by auto lemma all_to_meta: "Trueprop (\a. P a) \ (\a. P a)" apply rule by auto lemma imp_to_meta: "Trueprop (P\Q) \ (P\Q)" apply rule by auto (* for some reason, there is no such rule in HOL *) lemma iffI2: "\P \ Q; \ P \ \ Q\ \ P \ Q" by metis lemma iffExI: "\ \x. P x \ Q x; \x. Q x \ P x \ \ (\x. P x) \ (\x. Q x)" by metis lemma bex2I[intro?]: "\ (a,b)\S; (a,b)\S \ P a b \ \ \a b. (a,b)\S \ P a b" by blast (* TODO: Move lemma to HOL! *) lemma cnv_conj_to_meta: "(P \ Q \ PROP X) \ (\P;Q\ \ PROP X)" by (rule BNF_Fixpoint_Base.conj_imp_eq_imp_imp) subsection \Sets\ lemma remove_subset: "x\S \ S-{x} \ S" by auto lemma subset_minus_empty: "A\B \ A-B = {}" by auto lemma insert_minus_eq: "x\y \ insert x A - {y} = insert x (A - {y})" by auto lemma set_notEmptyE: "\S\{}; !!x. x\S \ P\ \ P" by (metis equals0I) lemma subset_Collect_conv: "S \ Collect P \ (\x\S. P x)" by auto lemma memb_imp_not_empty: "x\S \ S\{}" by auto lemma disjoint_mono: "\ a\a'; b\b'; a'\b'={} \ \ a\b={}" by auto lemma disjoint_alt_simp1: "A-B = A \ A\B = {}" by auto lemma disjoint_alt_simp2: "A-B \ A \ A\B \ {}" by auto lemma disjoint_alt_simp3: "A-B \ A \ A\B \ {}" by auto lemma disjointI[intro?]: "\ \x. \x\a; x\b\ \ False \ \ a\b={}" by auto lemmas set_simps = subset_minus_empty disjoint_alt_simp1 disjoint_alt_simp2 disjoint_alt_simp3 Un_absorb1 Un_absorb2 lemma set_minus_singleton_eq: "x\X \ X-{x} = X" by auto lemma set_diff_diff_left: "A-B-C = A-(B\C)" by auto lemma image_update[simp]: "x\A \ f(x:=n)`A = f`A" by auto lemma eq_or_mem_image_simp[simp]: "{f l |l. l = a \ l\B} = insert (f a) {f l|l. l\B}" by blast lemma set_union_code [code_unfold]: "set xs \ set ys = set (xs @ ys)" by auto lemma in_fst_imageE: assumes "x \ fst`S" obtains y where "(x,y)\S" using assms by auto lemma in_snd_imageE: assumes "y \ snd`S" obtains x where "(x,y)\S" using assms by auto lemma fst_image_mp: "\fst`A \ B; (x,y)\A \ \ x\B" by (metis Domain.DomainI fst_eq_Domain in_mono) lemma snd_image_mp: "\snd`A \ B; (x,y)\A \ \ y\B" by (metis Range.intros rev_subsetD snd_eq_Range) lemma inter_eq_subsetI: "\ S\S'; A\S' = B\S' \ \ A\S = B\S" by auto text \ Decompose general union over sum types. \ lemma Union_plus: "(\ x \ A <+> B. f x) = (\ a \ A. f (Inl a)) \ (\b \ B. f (Inr b))" by auto lemma Union_sum: "(\x. f (x::'a+'b)) = (\l. f (Inl l)) \ (\r. f (Inr r))" (is "?lhs = ?rhs") proof - have "?lhs = (\x \ UNIV <+> UNIV. f x)" by simp thus ?thesis by (simp only: Union_plus) qed subsubsection \Finite Sets\ lemma card_1_singletonI: "\finite S; card S = 1; x\S\ \ S={x}" proof (safe, rule ccontr, goal_cases) case prems: (1 x') hence "finite (S-{x})" "S-{x} \ {}" by auto hence "card (S-{x}) \ 0" by auto moreover from prems(1-3) have "card (S-{x}) = 0" by auto ultimately have False by simp thus ?case .. qed lemma card_insert_disjoint': "\finite A; x \ A\ \ card (insert x A) - Suc 0 = card A" by (drule (1) card_insert_disjoint) auto lemma card_eq_UNIV[simp]: "card (S::'a::finite set) = card (UNIV::'a set) \ S=UNIV" proof (auto) fix x assume A: "card S = card (UNIV::'a set)" show "x\S" proof (rule ccontr) assume "x\S" hence "S\UNIV" by auto with psubset_card_mono[of UNIV S] have "card S < card (UNIV::'a set)" by auto with A show False by simp qed qed lemma card_eq_UNIV2[simp]: "card (UNIV::'a set) = card (S::'a::finite set) \ S=UNIV" using card_eq_UNIV[of S] by metis lemma card_ge_UNIV[simp]: "card (UNIV::'a::finite set) \ card (S::'a set) \ S=UNIV" using card_mono[of "UNIV::'a::finite set" S, simplified] by auto lemmas length_remdups_card = length_remdups_concat[of "[l]", simplified] for l lemma fs_contract: "fst ` { p | p. f (fst p) (snd p) \ S } = { a . \b. f a b \ S }" by (simp add: image_Collect) lemma finite_Collect: "finite S \ inj f \ finite {a. f a : S}" by(simp add: finite_vimageI vimage_def[symmetric]) \ \Finite sets have an injective mapping to an initial segments of the natural numbers\ (* This lemma is also in the standard library (from Isabelle2009-1 on) as @{thm [source] Finite_Set.finite_imp_inj_to_nat_seg}. However, it is formulated with HOL's \ there rather then with the meta-logic obtain *) lemma finite_imp_inj_to_nat_seg': fixes A :: "'a set" assumes A: "finite A" obtains f::"'a \ nat" and n::"nat" where "f`A = {i. i finite (lists P \ { l. length l = n })" proof (induct n) case 0 thus ?case by auto next case (Suc n) have "lists P \ { l. length l = Suc n } = (\(a,l). a#l) ` (P \ (lists P \ {l. length l = n}))" apply auto apply (case_tac x) apply auto done moreover from Suc have "finite \" by auto ultimately show ?case by simp qed lemma lists_of_len_fin2: "finite P \ finite (lists P \ { l. n = length l })" proof - assume A: "finite P" have S: "{ l. n = length l } = { l. length l = n }" by auto have "finite (lists P \ { l. n = length l }) \ finite (lists P \ { l. length l = n })" by (subst S) simp thus ?thesis using lists_of_len_fin1[OF A] by auto qed lemmas lists_of_len_fin = lists_of_len_fin1 lists_of_len_fin2 (* Try (simp only: cset_fin_simps, fastforce intro: cset_fin_intros) when reasoning about finiteness of collected sets *) lemmas cset_fin_simps = Ex_prod_contract fs_contract[symmetric] image_Collect[symmetric] lemmas cset_fin_intros = finite_imageI finite_Collect inj_onI lemma Un_interval: fixes b1 :: "'a::linorder" assumes "b1\b2" and "b2\b3" shows "{ f i | i. b1\i \ i { f i | i. b2\i \ ii \ i The standard library proves that a generalized union is finite if the index set is finite and if for every index the component set is itself finite. Conversely, we show that every component set must be finite when the union is finite. \ lemma finite_UNION_then_finite: "finite (\(B ` A)) \ a \ A \ finite (B a)" by (metis Set.set_insert UN_insert Un_infinite) lemma finite_if_eq_beyond_finite: "finite S \ finite {s. s - S = s' - S}" proof (rule finite_subset[where B="(\s. s \ (s' - S)) ` Pow S"], clarsimp) fix s have "s = (s \ S) \ (s - S)" by auto also assume "s - S = s' - S" finally show "s \ (\s. s \ (s' - S)) ` Pow S" by blast qed blast lemma distinct_finite_subset: assumes "finite x" shows "finite {ys. set ys \ x \ distinct ys}" (is "finite ?S") proof (rule finite_subset) from assms show "?S \ {ys. set ys \ x \ length ys \ card x}" by clarsimp (metis distinct_card card_mono) from assms show "finite ..." by (rule finite_lists_length_le) qed lemma distinct_finite_set: shows "finite {ys. set ys = x \ distinct ys}" (is "finite ?S") proof (cases "finite x") case False hence "{ys. set ys = x} = {}" by auto thus ?thesis by simp next case True show ?thesis proof (rule finite_subset) show "?S \ {ys. set ys \ x \ length ys \ card x}" using distinct_card by force from True show "finite ..." by (rule finite_lists_length_le) qed qed lemma finite_set_image: assumes f: "finite (set ` A)" and dist: "\xs. xs \ A \ distinct xs" shows "finite A" proof (rule finite_subset) from f show "finite (set -` (set ` A) \ {xs. distinct xs})" proof (induct rule: finite_induct) case (insert x F) have "finite (set -` {x} \ {xs. distinct xs})" apply (simp add: vimage_def) by (metis Collect_conj_eq distinct_finite_set) with insert show ?case apply (subst vimage_insert) apply (subst Int_Un_distrib2) apply (rule finite_UnI) apply simp_all done qed simp moreover from dist show "A \ ..." by (auto simp add: vimage_image_eq) qed subsubsection \Infinite Set\ lemma INFM_nat_inductI: assumes P0: "P (0::nat)" assumes PS: "\i. P i \ \j>i. P j \ Q j" shows "\\<^sub>\i. Q i" proof - have "\i. \j>i. P j \ Q j" proof fix i show "\j>i. P j \ Q j" apply (induction i) using PS[OF P0] apply auto [] by (metis PS Suc_lessI) qed thus ?thesis unfolding INFM_nat by blast qed subsection \Functions\ lemma fun_neq_ext_iff: "m\m' \ (\x. m x \ m' x)" by auto definition "inv_on f A x == SOME y. y\A \ f y = x" lemma inv_on_f_f[simp]: "\inj_on f A; x\A\ \ inv_on f A (f x) = x" by (auto simp add: inv_on_def inj_on_def) lemma f_inv_on_f: "\ y\f`A \ \ f (inv_on f A y) = y" by (auto simp add: inv_on_def intro: someI2) lemma inv_on_f_range: "\ y \ f`A \ \ inv_on f A y \ A" by (auto simp add: inv_on_def intro: someI2) lemma inj_on_map_inv_f [simp]: "\set l \ A; inj_on f A\ \ map (inv_on f A) (map f l) = l" apply (simp) apply (induct l) apply auto done lemma comp_cong_right: "x = y \ f o x = f o y" by (simp) lemma comp_cong_left: "x = y \ x o f = y o f" by (simp) lemma fun_comp_eq_conv: "f o g = fg \ (\x. f (g x) = fg x)" by auto abbreviation comp2 (infixl "oo" 55) where "f oo g \ \x. f o (g x)" abbreviation comp3 (infixl "ooo" 55) where "f ooo g \ \x. f oo (g x)" notation comp2 (infixl "\\" 55) and comp3 (infixl "\\\" 55) definition [code_unfold, simp]: "swap_args2 f x y \ f y x" subsection \Multisets\ (* The following is a syntax extension for multisets. Unfortunately, it depends on a change in the Library/Multiset.thy, so it is commented out here, until it will be incorporated into Library/Multiset.thy by its maintainers. The required change in Library/Multiset.thy is removing the syntax for single: - single :: "'a => 'a multiset" ("{#_#}") + single :: "'a => 'a multiset" And adding the following translations instead: + syntax + "_multiset" :: "args \ 'a multiset" ("{#(_)#}") + translations + "{#x, xs#}" == "{#x#} + {#xs#}" + "{# x #}" == "single x" This translates "{# \ #}" into a sum of singletons, that is parenthesized to the right. ?? Can we also achieve left-parenthesizing ?? *) (* Let's try what happens if declaring AC-rules for multiset union as simp-rules *) (*declare union_ac[simp] -- don't do it !*) lemma count_mset_set_finite_iff: "finite S \ count (mset_set S) a = (if a \ S then 1 else 0)" by simp lemma ex_Melem_conv: "(\x. x \# A) = (A \ {#})" by (simp add: ex_in_conv) subsubsection \Count\ lemma count_ne_remove: "\ x ~= t\ \ count S x = count (S-{#t#}) x" by (auto) lemma mset_empty_count[simp]: "(\p. count M p = 0) = (M={#})" by (auto simp add: multiset_eq_iff) subsubsection \Union, difference and intersection\ lemma size_diff_se: "t \# S \ size S = size (S - {#t#}) + 1" proof (unfold size_multiset_overloaded_eq) let ?SIZE = "sum (count S) (set_mset S)" assume A: "t \# S" from A have SPLITPRE: "finite (set_mset S) & {t}\(set_mset S)" by auto hence "?SIZE = sum (count S) (set_mset S - {t}) + sum (count S) {t}" by (blast dest: sum.subset_diff) hence "?SIZE = sum (count S) (set_mset S - {t}) + count (S) t" by auto moreover with A have "count S t = count (S-{#t#}) t + 1" by auto ultimately have D: "?SIZE = sum (count S) (set_mset S - {t}) + count (S-{#t#}) t + 1" by (arith) moreover have "sum (count S) (set_mset S - {t}) = sum (count (S-{#t#})) (set_mset S - {t})" proof - have "\x\(set_mset S - {t}). count S x = count (S-{#t#}) x" by (auto iff add: count_ne_remove) thus ?thesis by simp qed ultimately have D: "?SIZE = sum (count (S-{#t#})) (set_mset S - {t}) + count (S-{#t#}) t + 1" by (simp) moreover { assume CASE: "t \# S - {#t#}" from CASE have "set_mset S - {t} = set_mset (S - {#t#})" by (auto simp add: in_diff_count split: if_splits) with CASE D have "?SIZE = sum (count (S-{#t#})) (set_mset (S - {#t#})) + 1" by (simp add: not_in_iff) } moreover { assume CASE: "t \# S - {#t#}" from CASE have "t \# S" by (rule in_diffD) with CASE have 1: "set_mset S = set_mset (S-{#t#})" by (auto simp add: in_diff_count split: if_splits) moreover from D have "?SIZE = sum (count (S-{#t#})) (set_mset S - {t}) + sum (count (S-{#t#})) {t} + 1" by simp moreover from SPLITPRE sum.subset_diff have "sum (count (S-{#t#})) (set_mset S) = sum (count (S-{#t#})) (set_mset S - {t}) + sum (count (S-{#t#})) {t}" by (blast) ultimately have "?SIZE = sum (count (S-{#t#})) (set_mset (S-{#t#})) + 1" by simp } ultimately show "?SIZE = sum (count (S-{#t#})) (set_mset (S - {#t#})) + 1" by blast qed (* TODO: Check whether this proof can be done simpler *) lemma mset_union_diff_comm: "t \# S \ T + (S - {#t#}) = (T + S) - {#t#}" proof - assume "t \# S" then obtain S' where S: "S = add_mset t S'" by (metis insert_DiffM) then show ?thesis by auto qed (* lemma mset_diff_diff_left: "A-B-C = A-((B::'a multiset)+C)" proof - have "\e . count (A-B-C) e = count (A-(B+C)) e" by auto thus ?thesis by (simp add: multiset_eq_conv_count_eq) qed lemma mset_diff_commute: "A-B-C = A-C-(B::'a multiset)" proof - have "A-B-C = A-(B+C)" by (simp add: mset_diff_diff_left) also have "\ = A-(C+B)" by (simp add: union_commute) thus ?thesis by (simp add: mset_diff_diff_left) qed lemma mset_diff_same_empty[simp]: "(S::'a multiset) - S = {#}" proof - have "\e . count (S-S) e = 0" by auto hence "\e . ~ (e : set_mset (S-S))" by auto hence "set_mset (S-S) = {}" by blast thus ?thesis by (auto) qed *) lemma mset_right_cancel_union: "\a \# A+B; ~(a \# B)\ \ a\#A" by (simp) lemma mset_left_cancel_union: "\a \# A+B; ~(a \# A)\ \ a\#B" by (simp) lemmas mset_cancel_union = mset_right_cancel_union mset_left_cancel_union lemma mset_right_cancel_elem: "\a \# A+{#b#}; a~=b\ \ a\#A" by simp lemma mset_left_cancel_elem: "\a \# {#b#}+A; a~=b\ \ a\#A" by simp lemmas mset_cancel_elem = mset_right_cancel_elem mset_left_cancel_elem lemma mset_diff_cancel1elem[simp]: "~(a \# B) \ {#a#}-B = {#a#}" by (auto simp add: not_in_iff intro!: multiset_eqI) (* lemma diff_union_inverse[simp]: "A + B - B = (A::'a multiset)" by (auto iff add: multiset_eq_conv_count_eq) lemma diff_union_inverse2[simp]: "B + A - B = (A::'a multiset)" by (auto iff add: multiset_eq_conv_count_eq) *) (*lemma union_diff_assoc_se2: "t \# A \ (A+B)-{#t#} = (A-{#t#}) + B" by (auto iff add: multiset_eq_conv_count_eq) lemmas union_diff_assoc_se = union_diff_assoc_se1 union_diff_assoc_se2*) lemma union_diff_assoc: "C-B={#} \ (A+B)-C = A + (B-C)" by (simp add: multiset_eq_iff) lemmas mset_neutral_cancel1 = union_left_cancel[where N="{#}", simplified] union_right_cancel[where N="{#}", simplified] declare mset_neutral_cancel1[simp] lemma mset_union_2_elem: "{#a, b#} = add_mset c M \ {#a#}=M & b=c | a=c & {#b#}=M" by (auto simp: add_eq_conv_diff) lemma mset_un_cases[cases set, case_names left right]: "\a \# A + B; a \# A \ P; a \# B \ P\ \ P" by (auto) lemma mset_unplusm_dist_cases[cases set, case_names left right]: assumes A: "{#s#}+A = B+C" assumes L: "\B={#s#}+(B-{#s#}); A=(B-{#s#})+C\ \ P" assumes R: "\C={#s#}+(C-{#s#}); A=B+(C-{#s#})\ \ P" shows P proof - from A[symmetric] have "s \# B+C" by simp thus ?thesis proof (cases rule: mset_un_cases) case left hence 1: "B={#s#}+(B-{#s#})" by simp with A have "{#s#}+A = {#s#}+((B-{#s#})+C)" by (simp add: union_ac) hence 2: "A = (B-{#s#})+C" by (simp) from L[OF 1 2] show ?thesis . next case right hence 1: "C={#s#}+(C-{#s#})" by simp with A have "{#s#}+A = {#s#}+(B+(C-{#s#}))" by (simp add: union_ac) hence 2: "A = B+(C-{#s#})" by (simp) from R[OF 1 2] show ?thesis . qed qed lemma mset_unplusm_dist_cases2[cases set, case_names left right]: assumes A: "B+C = {#s#}+A" assumes L: "\B={#s#}+(B-{#s#}); A=(B-{#s#})+C\ \ P" assumes R: "\C={#s#}+(C-{#s#}); A=B+(C-{#s#})\ \ P" shows P using mset_unplusm_dist_cases[OF A[symmetric]] L R by blast lemma mset_single_cases[cases set, case_names loc env]: assumes A: "add_mset s c = add_mset r' c'" assumes CASES: "\s=r'; c=c'\ \ P" "\c'={#s#}+(c'-{#s#}); c={#r'#}+(c-{#r'#}); c-{#r'#} = c'-{#s#} \ \ P" shows "P" proof - { assume CASE: "s=r'" with A have "c=c'" by simp with CASE CASES have ?thesis by auto } moreover { assume CASE: "s\r'" have "s \# {#s#}+c" by simp with A have "s \# {#r'#}+c'" by simp with CASE have "s \# c'" by simp hence 1: "c' = {#s#} + (c' - {#s#})" by simp with A have "{#s#}+c = {#s#}+({#r'#}+(c' - {#s#}))" by (auto simp add: union_ac) hence 2: "c={#r'#}+(c' - {#s#})" by (auto) hence 3: "c-{#r'#} = (c' - {#s#})" by auto from 1 2 3 CASES have ?thesis by auto } ultimately show ?thesis by blast qed lemma mset_single_cases'[cases set, case_names loc env]: assumes A: "add_mset s c = add_mset r' c'" assumes CASES: "\s=r'; c=c'\ \ P" "!!cc. \c'={#s#}+cc; c={#r'#}+cc; c'-{#s#}=cc; c-{#r'#}=cc\ \ P" shows "P" using A CASES by (auto elim!: mset_single_cases) lemma mset_single_cases2[cases set, case_names loc env]: assumes A: "add_mset s c = add_mset r' c'" assumes CASES: "\s=r'; c=c'\ \ P" "\c'=(c'-{#s#})+{#s#}; c=(c-{#r'#})+{#r'#}; c-{#r'#} = c'-{#s#} \ \ P" shows "P" proof - from A have "add_mset s c = add_mset r' c'" by (simp add: union_ac) thus ?thesis using CASES by (cases rule: mset_single_cases) simp_all qed lemma mset_single_cases2'[cases set, case_names loc env]: assumes A: "add_mset s c = add_mset r' c'" assumes CASES: "\s=r'; c=c'\ \ P" "!!cc. \c'=cc+{#s#}; c=cc+{#r'#}; c'-{#s#}=cc; c-{#r'#}=cc\ \ P" shows "P" using A CASES by (auto elim!: mset_single_cases2) lemma mset_un_single_un_cases [consumes 1, case_names left right]: assumes A: "add_mset a A = B + C" and CASES: "a \# B \ A = (B - {#a#}) + C \ P" "a \# C \ A = B + (C - {#a#}) \ P" shows "P" proof - have "a \# A+{#a#}" by simp with A have "a \# B+C" by auto thus ?thesis proof (cases rule: mset_un_cases) case left hence "B=B-{#a#}+{#a#}" by auto with A have "A+{#a#} = (B-{#a#})+C+{#a#}" by (auto simp add: union_ac) hence "A=(B-{#a#})+C" by simp with CASES(1)[OF left] show ?thesis by blast next case right hence "C=C-{#a#}+{#a#}" by auto with A have "A+{#a#} = B+(C-{#a#})+{#a#}" by (auto simp add: union_ac) hence "A=B+(C-{#a#})" by simp with CASES(2)[OF right] show ?thesis by blast qed qed (* TODO: Can this proof be done more automatically ? *) lemma mset_distrib[consumes 1, case_names dist]: assumes A: "(A::'a multiset)+B = M+N" "!!Am An Bm Bn. \A=Am+An; B=Bm+Bn; M=Am+Bm; N=An+Bn\ \ P" shows "P" proof - have BN_MA: "B - N = M - A" by (metis (no_types) add_diff_cancel_right assms(1) union_commute) have H: "A = A\# C + (A - C) \# D" if "A + B = C + D" for A B C D :: "'a multiset" by (metis add.commute diff_intersect_left_idem mset_subset_eq_add_left subset_eq_diff_conv subset_mset.add_diff_inverse subset_mset.inf_absorb1 subset_mset.inf_le1 that) have A': "A = A\# M + (A - M) \# N" using A(1) H by blast moreover have B': "B = (B - N) \# M + B\# N" using A(1) H[of B A N M] by (auto simp: ac_simps) moreover have "M = A \# M + (B - N) \# M" using H[of M N A B] BN_MA[symmetric] A(1) by (metis (no_types) diff_intersect_left_idem diff_union_cancelR multiset_inter_commute subset_mset.diff_add subset_mset.inf.cobounded1 union_commute) moreover have "N = (A - M) \# N + B \# N" by (metis A' assms(1) diff_union_cancelL inter_union_distrib_left inter_union_distrib_right mset_subset_eq_multiset_union_diff_commute subset_mset.inf.cobounded1 subset_mset.inf.commute) ultimately show P using A(2) by blast qed subsubsection \Singleton multisets\ lemma mset_size_le1_cases[case_names empty singleton,consumes 1]: "\ size M \ Suc 0; M={#} \ P; !!m. M={#m#} \ P \ \ P" by (cases M) auto lemma diff_union_single_conv2: "a \# J \ J + I - {#a#} = (J - {#a#}) + I" by simp lemmas diff_union_single_convs = diff_union_single_conv diff_union_single_conv2 lemma mset_contains_eq: "(m \# M) = ({#m#}+(M-{#m#})=M)" using diff_single_trivial by fastforce subsubsection \Pointwise ordering\ (*declare mset_le_trans[trans] Seems to be in there now. Why is this not done in Multiset.thy or order-class ? *) lemma mset_le_incr_right1: "a\#(b::'a multiset) \ a\#b+c" using mset_subset_eq_mono_add[of a b "{#}" c, simplified] . lemma mset_le_incr_right2: "a\#(b::'a multiset) \ a\#c+b" using mset_le_incr_right1 by (auto simp add: union_commute) lemmas mset_le_incr_right = mset_le_incr_right1 mset_le_incr_right2 lemma mset_le_decr_left1: "a+c\#(b::'a multiset) \ a\#b" using mset_le_incr_right1 mset_subset_eq_mono_add_right_cancel by blast lemma mset_le_decr_left2: "c+a\#(b::'a multiset) \ a\#b" using mset_le_decr_left1 by (auto simp add: union_ac) lemma mset_le_add_mset_decr_left1: "add_mset c a\#(b::'a multiset) \ a\#b" by (simp add: mset_subset_eq_insertD subset_mset.dual_order.strict_implies_order) lemma mset_le_add_mset_decr_left2: "add_mset c a\#(b::'a multiset) \ {#c#}\#b" by (simp add: mset_subset_eq_insertD subset_mset.dual_order.strict_implies_order) lemmas mset_le_decr_left = mset_le_decr_left1 mset_le_decr_left2 mset_le_add_mset_decr_left1 mset_le_add_mset_decr_left2 lemma mset_le_subtract: "A\#B \ A-C \# B-(C::'a multiset)" apply (unfold subseteq_mset_def count_diff) apply clarify apply (subgoal_tac "count A a \ count B a") apply arith apply simp done lemma mset_union_subset: "A+B \# C \ A\#C \ B\#(C::'a multiset)" by (auto dest: mset_le_decr_left) lemma mset_le_add_mset: "add_mset x B \# C \ {#x#}\#C \ B\#(C::'a multiset)" by (auto dest: mset_le_decr_left) lemma mset_le_subtract_left: "A+B \# (X::'a multiset) \ B \# X-A \ A\#X" by (auto dest: mset_le_subtract[of "A+B" "X" "A"] mset_union_subset) lemma mset_le_subtract_right: "A+B \# (X::'a multiset) \ A \# X-B \ B\#X" by (auto dest: mset_le_subtract[of "A+B" "X" "B"] mset_union_subset) lemma mset_le_subtract_add_mset_left: "add_mset x B \# (X::'a multiset) \ B \# X-{#x#} \ {#x#}\#X" by (auto dest: mset_le_subtract[of "add_mset x B" "X" "{#x#}"] mset_le_add_mset) lemma mset_le_subtract_add_mset_right: "add_mset x B \# (X::'a multiset) \ {#x#} \# X-B \ B\#X" by (auto dest: mset_le_subtract[of "add_mset x B" "X" "B"] mset_le_add_mset) lemma mset_le_addE: "\ xs \# (ys::'a multiset); !!zs. ys=xs+zs \ P \ \ P" using mset_subset_eq_exists_conv by blast declare subset_mset.diff_add[simp, intro] lemma mset_2dist2_cases: assumes A: "{#a#}+{#b#} \# A+B" assumes CASES: "{#a#}+{#b#} \# A \ P" "{#a#}+{#b#} \# B \ P" "\a \# A; b \# B\ \ P" "\a \# B; b \# A\ \ P" shows "P" proof - { assume C: "a \# A" "b \# A-{#a#}" with mset_subset_eq_mono_add[of "{#a#}" "{#a#}" "{#b#}" "A-{#a#}"] have "{#a#}+{#b#} \# A" by auto } moreover { assume C: "a \# A" "\ (b \# A-{#a#})" with A have "b \# B" by (metis diff_union_single_conv2 mset_le_subtract_left mset_subset_eq_insertD mset_un_cases) } moreover { assume C: "\ (a \# A)" "b \# B-{#a#}" with A have "a \# B" by (auto dest: mset_subset_eqD) with C mset_subset_eq_mono_add[of "{#a#}" "{#a#}" "{#b#}" "B-{#a#}"] have "{#a#}+{#b#} \# B" by auto } moreover { assume C: "\ (a \# A)" "\ (b \# B-{#a#})" with A have "a \# B \ b \# A" apply (intro conjI) apply (auto dest!: mset_subset_eq_insertD simp: insert_union_subset_iff; fail)[] by (metis mset_diff_cancel1elem mset_le_subtract_left multiset_diff_union_assoc single_subset_iff subset_eq_diff_conv) } ultimately show P using CASES by blast qed lemma mset_union_subset_s: "{#a#}+B \# C \ a \# C \ B \# C" by (drule mset_union_subset) simp (* TODO: Check which of these lemmas are already introduced by order-classes ! *) lemma mset_le_single_cases[consumes 1, case_names empty singleton]: "\M\#{#a#}; M={#} \ P; M={#a#} \ P\ \ P" by (induct M) auto lemma mset_le_distrib[consumes 1, case_names dist]: "\(X::'a multiset)\#A+B; !!Xa Xb. \X=Xa+Xb; Xa\#A; Xb\#B\ \ P \ \ P" by (auto elim!: mset_le_addE mset_distrib) lemma mset_le_mono_add_single: "\a \# ys; b \# ws\ \ {#a#} + {#b#} \# ys + ws" by (meson mset_subset_eq_mono_add single_subset_iff) lemma mset_size1elem: "\size P \ 1; q \# P\ \ P={#q#}" by (auto elim: mset_size_le1_cases) lemma mset_size2elem: "\size P \ 2; {#q#}+{#q'#} \# P\ \ P={#q#}+{#q'#}" by (auto elim: mset_le_addE) subsubsection \Image under function\ notation image_mset (infixr "`#" 90) lemma mset_map_split_orig: "!!M1 M2. \f `# P = M1+M2; !!P1 P2. \P=P1+P2; f `# P1 = M1; f `# P2 = M2\ \ Q \ \ Q" by (induct P) (force elim!: mset_un_single_un_cases)+ lemma mset_map_id: "\!!x. f (g x) = x\ \ f `# g `# X = X" by (induct X) auto text \The following is a very specialized lemma. Intuitively, it splits the original multiset by a splitting of some pointwise supermultiset of its image. Application: This lemma came in handy when proving the correctness of a constraint system that collects at most k sized submultisets of the sets of spawned threads. \ lemma mset_map_split_orig_le: assumes A: "f `# P \# M1+M2" and EX: "!!P1 P2. \P=P1+P2; f `# P1 \# M1; f `# P2 \# M2\ \ Q" shows "Q" using A EX by (auto elim: mset_le_distrib mset_map_split_orig) subsection \Lists\ lemma len_greater_imp_nonempty[simp]: "length l > x \ l\[]" by auto lemma list_take_induct_tl2: "\length xs = length ys; \n \ \n < length (tl xs). P ((tl ys) ! n) ((tl xs) ! n)" by (induct xs ys rule: list_induct2) auto lemma not_distinct_split_distinct: assumes "\ distinct xs" obtains y ys zs where "distinct ys" "y \ set ys" "xs = ys@[y]@zs" using assms proof (induct xs rule: rev_induct) case Nil thus ?case by simp next case (snoc x xs) thus ?case by (cases "distinct xs") auto qed lemma distinct_length_le: assumes d: "distinct ys" and eq: "set ys = set xs" shows "length ys \ length xs" proof - from d have "length ys = card (set ys)" by (simp add: distinct_card) also from eq List.card_set have "card (set ys) = length (remdups xs)" by simp also have "... \ length xs" by simp finally show ?thesis . qed lemma find_SomeD: "List.find P xs = Some x \ P x" "List.find P xs = Some x \ x\set xs" by (auto simp add: find_Some_iff) lemma in_hd_or_tl_conv[simp]: "l\[] \ x=hd l \ x\set (tl l) \ x\set l" by (cases l) auto lemma length_dropWhile_takeWhile: assumes "x < length (dropWhile P xs)" shows "x + length (takeWhile P xs) < length xs" using assms by (induct xs) auto text \Elim-version of @{thm neq_Nil_conv}.\ lemma neq_NilE: assumes "l\[]" obtains x xs where "l=x#xs" using assms by (metis list.exhaust) lemma length_Suc_rev_conv: "length xs = Suc n \ (\ys y. xs=ys@[y] \ length ys = n)" by (cases xs rule: rev_cases) auto subsubsection \List Destructors\ lemma not_hd_in_tl: "x \ hd xs \ x \ set xs \ x \ set (tl xs)" by (induct xs) simp_all lemma distinct_hd_tl: "distinct xs \ x = hd xs \ x \ set (tl (xs))" by (induct xs) simp_all lemma in_set_tlD: "x \ set (tl xs) \ x \ set xs" by (induct xs) simp_all lemma nth_tl: "xs \ [] \ tl xs ! n = xs ! Suc n" by (induct xs) simp_all lemma tl_subset: "xs \ [] \ set xs \ A \ set (tl xs) \ A" by (metis in_set_tlD rev_subsetD subsetI) lemma tl_last: "tl xs \ [] \ last xs = last (tl xs)" by (induct xs) simp_all lemma tl_obtain_elem: assumes "xs \ []" "tl xs = []" obtains e where "xs = [e]" using assms by (induct xs rule: list_nonempty_induct) simp_all lemma butlast_subset: "xs \ [] \ set xs \ A \ set (butlast xs) \ A" by (metis in_set_butlastD rev_subsetD subsetI) lemma butlast_rev_tl: "xs \ [] \ butlast (rev xs) = rev (tl xs)" by (induct xs rule: rev_induct) simp_all lemma hd_butlast: "length xs > 1 \ hd (butlast xs) = hd xs" by (induct xs) simp_all lemma butlast_upd_last_eq[simp]: "length l \ 2 \ (butlast l) [ length l - 2 := x ] = take (length l - 2) l @ [x]" apply (case_tac l rule: rev_cases) apply simp apply simp apply (case_tac ys rule: rev_cases) apply simp apply simp done lemma distinct_butlast_swap[simp]: "distinct pq \ distinct (butlast (pq[i := last pq]))" apply (cases pq rule: rev_cases) apply (auto simp: list_update_append distinct_list_update split: nat.split) done subsubsection \Splitting list according to structure of other list\ context begin private definition "SPLIT_ACCORDING m l \ length l = length m" private lemma SPLIT_ACCORDINGE: assumes "length m = length l" obtains "SPLIT_ACCORDING m l" unfolding SPLIT_ACCORDING_def using assms by auto private lemma SPLIT_ACCORDING_simp: "SPLIT_ACCORDING m (l1@l2) \ (\m1 m2. m=m1@m2 \ SPLIT_ACCORDING m1 l1 \ SPLIT_ACCORDING m2 l2)" "SPLIT_ACCORDING m (x#l') \ (\y m'. m=y#m' \ SPLIT_ACCORDING m' l')" apply (fastforce simp: SPLIT_ACCORDING_def intro: exI[where x = "take (length l1) m"] exI[where x = "drop (length l1) m"]) apply (cases m;auto simp: SPLIT_ACCORDING_def) done text \Split structure of list @{term m} according to structure of list @{term l}.\ method split_list_according for m :: "'a list" and l :: "'b list" = (rule SPLIT_ACCORDINGE[of m l], (simp; fail), ( simp only: SPLIT_ACCORDING_simp, elim exE conjE, simp only: SPLIT_ACCORDING_def ) ) end subsubsection \\list_all2\\ lemma list_all2_induct[consumes 1, case_names Nil Cons]: assumes "list_all2 P l l'" assumes "Q [] []" assumes "\x x' ls ls'. \ P x x'; list_all2 P ls ls'; Q ls ls' \ \ Q (x#ls) (x'#ls')" shows "Q l l'" using list_all2_lengthD[OF assms(1)] assms apply (induct rule: list_induct2) apply auto done subsubsection \Indexing\ lemma ran_nth_set_encoding_conv[simp]: "ran (\i. if ii l[j:=x]!i = l!i" apply (induction l arbitrary: i j) apply (auto split: nat.splits) done lemma nth_list_update': "l[i:=x]!j = (if i=j \ i length l \ n\0 \ last (take n l) = l!(n - 1)" apply (induction l arbitrary: n) apply (auto simp: take_Cons split: nat.split) done lemma nth_append_first[simp]: "i (l@l')!i = l!i" by (simp add: nth_append) lemma in_set_image_conv_nth: "f x \ f`set l \ (\ii. i f (l!i) = f (l'!i)" shows "f`set l = f`set l'" using assms by (fastforce simp: in_set_conv_nth in_set_image_conv_nth) lemma insert_swap_set_eq: "i insert (l!i) (set (l[i:=x])) = insert x (set l)" by (auto simp: in_set_conv_nth nth_list_update split: if_split_asm) subsubsection \Reverse lists\ lemma neq_Nil_revE: assumes "l\[]" obtains ll e where "l = ll@[e]" using assms by (cases l rule: rev_cases) auto lemma neq_Nil_rev_conv: "l\[] \ (\xs x. l = xs@[x])" by (cases l rule: rev_cases) auto text \Caution: Same order of case variables in snoc-case as @{thm [source] rev_exhaust}, the other way round than @{thm [source] rev_induct} !\ lemma length_compl_rev_induct[case_names Nil snoc]: "\P []; !! l e . \!! ll . length ll <= length l \ P ll\ \ P (l@[e])\ \ P l" apply(induct_tac l rule: length_induct) apply(case_tac "xs" rule: rev_cases) apply(auto) done lemma list_append_eq_Cons_cases[consumes 1]: "\ys@zs = x#xs; \ys=[]; zs=x#xs\ \ P; !!ys'. \ ys=x#ys'; ys'@zs=xs \ \ P \ \ P" by (auto iff add: append_eq_Cons_conv) lemma list_Cons_eq_append_cases[consumes 1]: "\x#xs = ys@zs; \ys=[]; zs=x#xs\ \ P; !!ys'. \ ys=x#ys'; ys'@zs=xs \ \ P \ \ P" by (auto iff add: Cons_eq_append_conv) lemma map_of_rev_distinct[simp]: "distinct (map fst m) \ map_of (rev m) = map_of m" apply (induct m) apply simp apply simp apply (subst map_add_comm) apply force apply simp done \ \Tail-recursive, generalized @{const rev}. May also be used for tail-recursively getting a list with all elements of the two operands, if the order does not matter, e.g. when implementing sets by lists.\ fun revg where "revg [] b = b" | "revg (a#as) b = revg as (a#b)" lemma revg_fun[simp]: "revg a b = rev a @ b" by (induct a arbitrary: b) auto lemma rev_split_conv[simp]: "l \ [] \ rev (tl l) @ [hd l] = rev l" by (induct l) simp_all lemma rev_butlast_is_tl_rev: "rev (butlast l) = tl (rev l)" by (induct l) auto lemma hd_last_singletonI: "\xs \ []; hd xs = last xs; distinct xs\ \ xs = [hd xs]" by (induct xs rule: list_nonempty_induct) auto lemma last_filter: "\xs \ []; P (last xs)\ \ last (filter P xs) = last xs" by (induct xs rule: rev_nonempty_induct) simp_all (* As the following two rules are similar in nature to list_induct2', they are named accordingly. *) lemma rev_induct2' [case_names empty snocl snocr snoclr]: assumes empty: "P [] []" and snocl: "\x xs. P (xs@[x]) []" and snocr: "\y ys. P [] (ys@[y])" and snoclr: "\x xs y ys. P xs ys \ P (xs@[x]) (ys@[y])" shows "P xs ys" proof (induct xs arbitrary: ys rule: rev_induct) case Nil thus ?case using empty snocr by (cases ys rule: rev_exhaust) simp_all next case snoc thus ?case using snocl snoclr by (cases ys rule: rev_exhaust) simp_all qed lemma rev_nonempty_induct2' [case_names single snocl snocr snoclr, consumes 2]: assumes "xs \ []" "ys \ []" assumes single': "\x y. P [x] [y]" and snocl: "\x xs y. xs \ [] \ P (xs@[x]) [y]" and snocr: "\x y ys. ys \ [] \ P [x] (ys@[y])" and snoclr: "\x xs y ys. \P xs ys; xs \ []; ys\[]\ \ P (xs@[x]) (ys@[y])" shows "P xs ys" using assms(1,2) proof (induct xs arbitrary: ys rule: rev_nonempty_induct) case single then obtain z zs where "ys = zs@[z]" by (metis rev_exhaust) thus ?case using single' snocr by (cases "zs = []") simp_all next case (snoc x xs) then obtain z zs where zs: "ys = zs@[z]" by (metis rev_exhaust) thus ?case using snocl snoclr snoc by (cases "zs = []") simp_all qed subsubsection "Folding" text "Ugly lemma about foldl over associative operator with left and right neutral element" lemma foldl_A1_eq: "!!i. \ !! e. f n e = e; !! e. f e n = e; !!a b c. f a (f b c) = f (f a b) c \ \ foldl f i ww = f i (foldl f n ww)" proof (induct ww) case Nil thus ?case by simp next case (Cons a ww i) note IHP[simplified]=this have "foldl f i (a # ww) = foldl f (f i a) ww" by simp also from IHP have "\ = f (f i a) (foldl f n ww)" by blast also from IHP(4) have "\ = f i (f a (foldl f n ww))" by simp also from IHP(1)[OF IHP(2,3,4), where i=a] have "\ = f i (foldl f a ww)" by simp also from IHP(2)[of a] have "\ = f i (foldl f (f n a) ww)" by simp also have "\ = f i (foldl f n (a#ww))" by simp finally show ?case . qed lemmas foldl_conc_empty_eq = foldl_A1_eq[of "(@)" "[]", simplified] lemmas foldl_un_empty_eq = foldl_A1_eq[of "(\)" "{}", simplified, OF Un_assoc[symmetric]] lemma foldl_set: "foldl (\) {} l = \{x. x\set l}" apply (induct l) apply simp_all apply (subst foldl_un_empty_eq) apply auto done lemma (in monoid_mult) foldl_absorb1: "x*foldl (*) 1 zs = foldl (*) x zs" apply (rule sym) apply (rule foldl_A1_eq) apply (auto simp add: mult.assoc) done text \Towards an invariant rule for foldl\ lemma foldl_rule_aux: fixes I :: "'\ \ 'a list \ bool" assumes initial: "I \0 l0" assumes step: "!!l1 l2 x \. \ l0=l1@x#l2; I \ (x#l2) \ \ I (f \ x) l2" shows "I (foldl f \0 l0) []" using initial step apply (induct l0 arbitrary: \0) apply auto done lemma foldl_rule_aux_P: fixes I :: "'\ \ 'a list \ bool" assumes initial: "I \0 l0" assumes step: "!!l1 l2 x \. \ l0=l1@x#l2; I \ (x#l2) \ \ I (f \ x) l2" assumes final: "!!\. I \ [] \ P \" shows "P (foldl f \0 l0)" using foldl_rule_aux[of I \0 l0, OF initial, OF step] final by simp lemma foldl_rule: fixes I :: "'\ \ 'a list \ 'a list \ bool" assumes initial: "I \0 [] l0" assumes step: "!!l1 l2 x \. \ l0=l1@x#l2; I \ l1 (x#l2) \ \ I (f \ x) (l1@[x]) l2" shows "I (foldl f \0 l0) l0 []" using initial step apply (rule_tac I="\\ lr. \ll. l0=ll@lr \ I \ ll lr" in foldl_rule_aux_P) apply auto done text \ Invariant rule for foldl. The invariant is parameterized with the state, the list of items that have already been processed and the list of items that still have to be processed. \ lemma foldl_rule_P: fixes I :: "'\ \ 'a list \ 'a list \ bool" \ \The invariant holds for the initial state, no items processed yet and all items to be processed:\ assumes initial: "I \0 [] l0" \ \The invariant remains valid if one item from the list is processed\ assumes step: "!!l1 l2 x \. \ l0=l1@x#l2; I \ l1 (x#l2) \ \ I (f \ x) (l1@[x]) l2" \ \The proposition follows from the invariant in the final state, i.e. all items processed and nothing to be processed\ assumes final: "!!\. I \ l0 [] \ P \" shows "P (foldl f \0 l0)" using foldl_rule[of I, OF initial step] by (simp add: final) text \Invariant reasoning over @{const foldl} for distinct lists. Invariant rule makes no assumptions about ordering.\ lemma distinct_foldl_invar: "\ distinct S; I (set S) \0; \x it \. \x \ it; it \ set S; I it \\ \ I (it - {x}) (f \ x) \ \ I {} (foldl f \0 S)" proof (induct S arbitrary: \0) case Nil thus ?case by auto next case (Cons x S) note [simp] = Cons.prems(1)[simplified] show ?case apply simp apply (rule Cons.hyps) proof - from Cons.prems(1) show "distinct S" by simp from Cons.prems(3)[of x "set (x#S)", simplified, OF Cons.prems(2)[simplified]] show "I (set S) (f \0 x)" . fix xx it \ assume A: "xx\it" "it \ set S" "I it \" show "I (it - {xx}) (f \ xx)" using A(2) apply (rule_tac Cons.prems(3)) apply (simp_all add: A(1,3)) apply blast done qed qed lemma foldl_length_aux: "foldl (\i x. Suc i) a l = a + length l" by (induct l arbitrary: a) auto lemmas foldl_length[simp] = foldl_length_aux[where a=0, simplified] lemma foldr_length_aux: "foldr (\x i. Suc i) l a = a + length l" by (induct l arbitrary: a rule: rev_induct) auto lemmas foldr_length[simp] = foldr_length_aux[where a=0, simplified] context comp_fun_commute begin lemma foldl_f_commute: "f a (foldl (\a b. f b a) b xs) = foldl (\a b. f b a) (f a b) xs" by(induct xs arbitrary: b)(simp_all add: fun_left_comm) lemma foldr_conv_foldl: "foldr f xs a = foldl (\a b. f b a) a xs" by(induct xs arbitrary: a)(simp_all add: foldl_f_commute) end lemma filter_conv_foldr: "filter P xs = foldr (\x xs. if P x then x # xs else xs) xs []" by(induct xs) simp_all lemma foldr_Cons: "foldr Cons xs [] = xs" by(induct xs) simp_all lemma foldr_snd_zip: "length xs \ length ys \ foldr (\(x, y). f y) (zip xs ys) b = foldr f ys b" proof(induct ys arbitrary: xs) case (Cons y ys) thus ?case by(cases xs) simp_all qed simp lemma foldl_snd_zip: "length xs \ length ys \ foldl (\b (x, y). f b y) b (zip xs ys) = foldl f b ys" proof(induct ys arbitrary: xs b) case (Cons y ys) thus ?case by(cases xs) simp_all qed simp lemma fst_foldl: "fst (foldl (\(a, b) x. (f a x, g a b x)) (a, b) xs) = foldl f a xs" by(induct xs arbitrary: a b) simp_all lemma foldl_foldl_conv_concat: "foldl (foldl f) a xs = foldl f a (concat xs)" by(induct xs arbitrary: a) simp_all lemma foldl_list_update: "n < length xs \ foldl f a (xs[n := x]) = foldl f (f (foldl f a (take n xs)) x) (drop (Suc n) xs)" by(simp add: upd_conv_take_nth_drop) lemma map_by_foldl: fixes l :: "'a list" and f :: "'a \ 'b" shows "foldl (\l x. l@[f x]) [] l = map f l" proof - { fix l' have "foldl (\l x. l@[f x]) l' l = l'@map f l" by (induct l arbitrary: l') auto } thus ?thesis by simp qed subsubsection \Sorting\ lemma sorted_in_between: assumes A: "0\i" "i x" "xk" and "kx" and "xk. i\k \ k l!k\x \ x x") case True from True Suc.hyps have "d = j - (i + 1)" by simp moreover from True have "i+1 < j" by (metis Suc.prems Suc_eq_plus1 Suc_lessI not_less) moreover from True have "0\i+1" by simp ultimately obtain k where "i+1\k" "k x" "xk" "k x" "xsorted l; l\[]\ \ hd l \ last l" by (metis eq_iff hd_Cons_tl last_in_set not_hd_in_tl sorted_wrt.simps(2)) lemma (in linorder) sorted_hd_min: "\xs \ []; sorted xs\ \ \x \ set xs. hd xs \ x" by (induct xs, auto) lemma sorted_append_bigger: "\sorted xs; \x \ set xs. x \ y\ \ sorted (xs @ [y])" proof (induct xs) case Nil then show ?case by simp next case (Cons x xs) then have s: "sorted xs" by (cases xs) simp_all from Cons have a: "\x\set xs. x \ y" by simp from Cons(1)[OF s a] Cons(2-) show ?case by (cases xs) simp_all qed lemma sorted_filter': "sorted l \ sorted (filter P l)" using sorted_filter[where f=id, simplified] . subsubsection \Map\ (* List.thy has: declare map_eq_Cons_D [dest!] Cons_eq_map_D [dest!] *) lemma map_eq_consE: "\map f ls = fa#fl; !!a l. \ ls=a#l; f a=fa; map f l = fl \ \ P\ \ P" by auto lemma map_fst_mk_snd[simp]: "map fst (map (\x. (x,k)) l) = l" by (induct l) auto lemma map_snd_mk_fst[simp]: "map snd (map (\x. (k,x)) l) = l" by (induct l) auto lemma map_fst_mk_fst[simp]: "map fst (map (\x. (k,x)) l) = replicate (length l) k" by (induct l) auto lemma map_snd_mk_snd[simp]: "map snd (map (\x. (x,k)) l) = replicate (length l) k" by (induct l) auto lemma map_zip1: "map (\x. (x,k)) l = zip l (replicate (length l) k)" by (induct l) auto lemma map_zip2: "map (\x. (k,x)) l = zip (replicate (length l) k) l" by (induct l) auto lemmas map_zip=map_zip1 map_zip2 (* TODO/FIXME: hope nobody changes nth to be underdefined! *) lemma map_eq_nth_eq: assumes A: "map f l = map f l'" shows "f (l!i) = f (l'!i)" proof - from A have "length l = length l'" by (metis length_map) thus ?thesis using A apply (induct arbitrary: i rule: list_induct2) apply simp apply (simp add: nth_def split: nat.split) done qed lemma map_upd_eq: "\i f (l!i) = f x\ \ map f (l[i:=x]) = map f l" by (metis list_update_beyond list_update_id map_update not_le_imp_less) lemma inj_map_inv_f [simp]: "inj f \ map (inv f) (map f l) = l" by (simp) lemma inj_on_map_the: "\D \ dom m; inj_on m D\ \ inj_on (the\m) D" apply (rule inj_onI) apply simp apply (case_tac "m x") apply (case_tac "m y") apply (auto intro: inj_onD) [1] apply (auto intro: inj_onD) [1] apply (case_tac "m y") apply (auto intro: inj_onD) [1] apply simp apply (rule inj_onD) apply assumption apply auto done lemma map_consI: "w=map f ww \ f a#w = map f (a#ww)" "w@l=map f ww@l \ f a#w@l = map f (a#ww)@l" by auto lemma restrict_map_subset_eq: fixes R shows "\m |` R = m'; R'\R\ \ m|` R' = m' |` R'" by (auto simp add: Int_absorb1) lemma restrict_map_self[simp]: "m |` dom m = m" apply (rule ext) apply (case_tac "m x") apply (auto simp add: restrict_map_def) done lemma restrict_map_UNIV[simp]: "f |` UNIV = f" by (auto simp add: restrict_map_def) lemma restrict_map_inv[simp]: "f |` (- dom f) = Map.empty" by (auto simp add: restrict_map_def intro: ext) lemma restrict_map_upd: "(f |` S)(k \ v) = f(k\v) |` (insert k S)" by (auto simp add: restrict_map_def intro: ext) lemma map_upd_eq_restrict: "m (x:=None) = m |` (-{x})" by (auto intro: ext) declare Map.finite_dom_map_of [simp, intro!] lemma dom_const'[simp]: "dom (\x. Some (f x)) = UNIV" by auto lemma restrict_map_eq : "((m |` A) k = None) \ (k \ dom m \ A)" "((m |` A) k = Some v) \ (m k = Some v \ k \ A)" unfolding restrict_map_def by (simp_all add: dom_def) definition "rel_of m P == {(k,v). m k = Some v \ P (k, v)}" lemma rel_of_empty[simp]: "rel_of Map.empty P = {}" by (auto simp add: rel_of_def) lemma remove1_tl: "xs \ [] \ remove1 (hd xs) xs = tl xs" by (cases xs) auto lemma set_oo_map_alt: "(set \\ map) f = (\l. f ` set l)" by auto subsubsection "Filter and Revert" primrec filter_rev_aux where "filter_rev_aux a P [] = a" | "filter_rev_aux a P (x#xs) = ( if P x then filter_rev_aux (x#a) P xs else filter_rev_aux a P xs)" lemma filter_rev_aux_alt: "filter_rev_aux a P l = filter P (rev l) @ a" by (induct l arbitrary: a) auto definition "filter_rev == filter_rev_aux []" lemma filter_rev_alt: "filter_rev P l = filter P (rev l)" unfolding filter_rev_def by (simp add: filter_rev_aux_alt) definition "remove_rev x == filter_rev (Not o (=) x)" lemma remove_rev_alt_def : "remove_rev x xs = (filter (\y. y \ x) (rev xs))" unfolding remove_rev_def apply (simp add: filter_rev_alt comp_def) by metis subsubsection "zip" declare zip_map_fst_snd[simp] lemma pair_list_split: "\ !!l1 l2. \ l = zip l1 l2; length l1=length l2; length l=length l2 \ \ P \ \ P" proof (induct l arbitrary: P) case Nil thus ?case by auto next case (Cons a l) from Cons.hyps obtain l1 l2 where IHAPP: "l=zip l1 l2" "length l1 = length l2" "length l=length l2" . obtain a1 a2 where [simp]: "a=(a1,a2)" by (cases a) auto from IHAPP have "a#l = zip (a1#l1) (a2#l2)" "length (a1#l1) = length (a2#l2)" "length (a#l) = length (a2#l2)" by (simp_all only:) (simp_all (no_asm_use)) with Cons.prems show ?case by blast qed lemma set_zip_cart: "x\set (zip l l') \ x\set l \ set l'" by (auto simp add: set_zip) lemma map_prod_fun_zip: "map (\(x, y). (f x, g y)) (zip xs ys) = zip (map f xs) (map g ys)" proof(induct xs arbitrary: ys) case Nil thus ?case by simp next case (Cons x xs) thus ?case by(cases ys) simp_all qed subsubsection \Generalized Zip\ text \Zip two lists element-wise, where the combination of two elements is specified by a function. Note that this function is underdefined for lists of different length.\ fun zipf :: "('a\'b\'c) \ 'a list \ 'b list \ 'c list" where "zipf f [] [] = []" | "zipf f (a#as) (b#bs) = f a b # zipf f as bs" lemma zipf_zip: "\length l1 = length l2\ \ zipf Pair l1 l2 = zip l1 l2" apply (induct l1 arbitrary: l2) apply auto apply (case_tac l2) apply auto done \ \All quantification over zipped lists\ fun list_all_zip where "list_all_zip P [] [] \ True" | "list_all_zip P (a#as) (b#bs) \ P a b \ list_all_zip P as bs" | "list_all_zip P _ _ \ False" lemma list_all_zip_alt: "list_all_zip P as bs \ length as = length bs \ (\iP as bs rule: list_all_zip.induct) apply auto apply (case_tac i) apply auto done lemma list_all_zip_map1: "list_all_zip P (List.map f as) bs \ list_all_zip (\a b. P (f a) b) as bs" apply (induct as arbitrary: bs) apply (case_tac bs) apply auto [2] apply (case_tac bs) apply auto [2] done lemma list_all_zip_map2: "list_all_zip P as (List.map f bs) \ list_all_zip (\a b. P a (f b)) as bs" apply (induct as arbitrary: bs) apply (case_tac bs) apply auto [2] apply (case_tac bs) apply auto [2] done declare list_all_zip_alt[mono] lemma lazI[intro?]: "\ length a = length b; !!i. i P (a!i) (b!i) \ \ list_all_zip P a b" by (auto simp add: list_all_zip_alt) lemma laz_conj[simp]: "list_all_zip (\x y. P x y \ Q x y) a b \ list_all_zip P a b \ list_all_zip Q a b" by (auto simp add: list_all_zip_alt) lemma laz_len: "list_all_zip P a b \ length a = length b" by (simp add: list_all_zip_alt) lemma laz_eq: "list_all_zip (=) a b \ a=b" apply (induct a arbitrary: b) apply (case_tac b) apply simp apply simp apply (case_tac b) apply simp apply simp done lemma laz_swap_ex: assumes A: "list_all_zip (\a b. \c. P a b c) A B" obtains C where "list_all_zip (\a c. \b. P a b c) A C" "list_all_zip (\b c. \a. P a b c) B C" proof - from A have [simp]: "length A = length B" and IC: "\ici. P (A!i) (B!i) ci" by (auto simp add: list_all_zip_alt) from obtain_list_from_elements[OF IC] obtain C where "length C = length B" "\ia b. P a) A B \ (length A = length B) \ (\a\set A. P a)" by (auto simp add: list_all_zip_alt set_conv_nth) lemma laz_weak_Pb[simp]: "list_all_zip (\a b. P b) A B \ (length A = length B) \ (\b\set B. P b)" by (force simp add: list_all_zip_alt set_conv_nth) subsubsection "Collecting Sets over Lists" definition "list_collect_set f l == \{ f a | a. a\set l }" lemma list_collect_set_simps[simp]: "list_collect_set f [] = {}" "list_collect_set f [a] = f a" "list_collect_set f (a#l) = f a \ list_collect_set f l" "list_collect_set f (l@l') = list_collect_set f l \ list_collect_set f l'" by (unfold list_collect_set_def) auto lemma list_collect_set_map_simps[simp]: "list_collect_set f (map x []) = {}" "list_collect_set f (map x [a]) = f (x a)" "list_collect_set f (map x (a#l)) = f (x a) \ list_collect_set f (map x l)" "list_collect_set f (map x (l@l')) = list_collect_set f (map x l) \ list_collect_set f (map x l')" by simp_all lemma list_collect_set_alt: "list_collect_set f l = \{ f (l!i) | i. i(set (map f l))" by (unfold list_collect_set_def) auto subsubsection \Sorted List with arbitrary Relations\ lemma (in linorder) sorted_wrt_rev_linord [simp] : "sorted_wrt (\) l \ sorted (rev l)" by (simp add: sorted_wrt_rev) lemma (in linorder) sorted_wrt_map_linord [simp] : "sorted_wrt (\(x::'a \ 'b) y. fst x \ fst y) l \ sorted (map fst l)" by (simp add: sorted_wrt_map) lemma (in linorder) sorted_wrt_map_rev_linord [simp] : "sorted_wrt (\(x::'a \ 'b) y. fst x \ fst y) l \ sorted (rev (map fst l))" by (induct l) (auto simp add: sorted_append) subsubsection \Take and Drop\ lemma take_update[simp]: "take n (l[i:=x]) = (take n l)[i:=x]" apply (induct l arbitrary: n i) apply (auto split: nat.split) apply (case_tac n) apply simp_all apply (case_tac n) apply simp_all done lemma take_update_last: "length list>n \ (take (Suc n) list) [n:=x] = take n list @ [x]" by (induct list arbitrary: n) (auto split: nat.split) lemma drop_upd_irrelevant: "m < n \ drop n (l[m:=x]) = drop n l" apply (induct n arbitrary: l m) apply simp apply (case_tac l) apply (auto split: nat.split) done lemma set_drop_conv: "set (drop n l) = { l!i | i. n\i \ i < length l }" (is "?L=?R") proof (intro equalityI subsetI) fix x assume "x\?L" then obtain i where L: "i = l!(n+i)" using L by simp finally show "x\?R" using L by auto next fix x assume "x\?R" then obtain i where L: "n\i" "i?L" by (auto simp add: in_set_conv_nth) qed lemma filter_upt_take_conv: "[i\[n..[n..set (drop n l) \ (\i. n\i \ i x = l!i)" apply (clarsimp simp: in_set_conv_nth) apply safe apply simp apply (metis le_add2 less_diff_conv add.commute) apply (rule_tac x="i-n" in exI) apply auto [] done lemma Union_take_drop_id: "\(set (drop n l)) \ \(set (take n l)) = \(set l)" by (metis Union_Un_distrib append_take_drop_id set_union_code sup_commute) lemma Un_set_drop_extend: "\j\Suc 0; j < length l\ \ l ! (j - Suc 0) \ \(set (drop j l)) = \(set (drop (j - Suc 0) l))" apply safe apply simp_all apply (metis diff_Suc_Suc diff_zero gr0_implies_Suc in_set_drop_conv_nth le_refl less_eq_Suc_le order.strict_iff_order) apply (metis Nat.diff_le_self set_drop_subset_set_drop subset_code(1)) by (metis diff_Suc_Suc gr0_implies_Suc in_set_drop_conv_nth less_eq_Suc_le order.strict_iff_order minus_nat.diff_0) lemma drop_take_drop_unsplit: "i\j \ drop i (take j l) @ drop j l = drop i l" proof - assume "i \ j" then obtain skf where "i + skf = j" by (metis le_iff_add) thus "drop i (take j l) @ drop j l = drop i l" by (metis append_take_drop_id diff_add_inverse drop_drop drop_take add.commute) qed lemma drop_last_conv[simp]: "l\[] \ drop (length l - Suc 0) l = [last l]" by (cases l rule: rev_cases) auto lemma take_butlast_conv[simp]: "take (length l - Suc 0) l = butlast l" by (cases l rule: rev_cases) auto lemma drop_takeWhile: assumes "i\length (takeWhile P l)" shows "drop i (takeWhile P l) = takeWhile P (drop i l)" using assms proof (induction l arbitrary: i) case Nil thus ?case by auto next case (Cons x l) thus ?case by (cases i) auto qed lemma less_length_takeWhile_conv: "i < length (takeWhile P l) \ (i (\j\i. P (l!j)))" apply safe subgoal using length_takeWhile_le less_le_trans by blast subgoal by (metis dual_order.strict_trans2 nth_mem set_takeWhileD takeWhile_nth) subgoal by (meson less_le_trans not_le_imp_less nth_length_takeWhile) done lemma eq_len_takeWhile_conv: "i=length (takeWhile P l) \ i\length l \ (\j (i \P (l!i))" apply safe subgoal using length_takeWhile_le less_le_trans by blast subgoal by (auto simp: less_length_takeWhile_conv) subgoal using nth_length_takeWhile by blast subgoal by (metis length_takeWhile_le nth_length_takeWhile order.order_iff_strict) subgoal by (metis dual_order.strict_trans2 leI less_length_takeWhile_conv linorder_neqE_nat nth_length_takeWhile) done subsubsection \Up-to\ lemma upt_eq_append_conv: "i\j \ [i.. (\k. i\k \ k\j \ [i.. [k..j" thus "\k\i. k \ j \ [i.. [k.. length l \ map (nth l) [M.. is1 = [l.. is2 = [Suc i.. l\i \ ii. i + ofs) [a..Last and butlast\ lemma butlast_upt: "butlast [m.. butlast [n..length l \ take (n - Suc 0) l = butlast (take n l)" by (simp add: butlast_take) lemma butlast_eq_cons_conv: "butlast l = x#xs \ (\xl. l=x#xs@[xl])" by (metis Cons_eq_appendI append_butlast_last_id butlast.simps butlast_snoc eq_Nil_appendI) lemma butlast_eq_consE: assumes "butlast l = x#xs" obtains xl where "l=x#xs@[xl]" using assms by (auto simp: butlast_eq_cons_conv) lemma drop_eq_ConsD: "drop n xs = x # xs' \ drop (Suc n) xs = xs'" by(induct xs arbitrary: n)(simp_all add: drop_Cons split: nat.split_asm) subsubsection \List Slices\ text \Based on Lars Hupel's code.\ definition slice :: "nat \ nat \ 'a list \ 'a list" where "slice from to list = take (to - from) (drop from list)" lemma slice_len[simp]: "\ from \ to; to \ length xs \ \ length (slice from to xs) = to - from" unfolding slice_def by simp lemma slice_head: "\ from < to; to \ length xs \ \ hd (slice from to xs) = xs ! from" unfolding slice_def proof - assume a1: "from < to" assume "to \ length xs" then have "\n. to - (to - n) \ length (take to xs)" by (metis (no_types) slice_def diff_le_self drop_take length_drop slice_len) then show "hd (take (to - from) (drop from xs)) = xs ! from" using a1 by (metis diff_diff_cancel drop_take hd_drop_conv_nth leI le_antisym less_or_eq_imp_le nth_take) qed lemma slice_eq_bounds_empty[simp]: "slice i i xs = []" unfolding slice_def by auto lemma slice_nth: "\ from < to; to \ length xs; i < to - from \ \ slice from to xs ! i = xs ! (from + i)" unfolding slice_def by (induction "to - from" arbitrary: "from" to i) simp+ lemma slice_prepend: "\ i \ k; k \ length xs \ \ let p = length ys in slice i k xs = slice (i + p) (k + p) (ys @ xs)" unfolding slice_def Let_def by force lemma slice_Nil[simp]: "slice begin end [] = []" unfolding slice_def by auto lemma slice_Cons: "slice begin end (x#xs) = (if begin=0 \ end>0 then x#slice begin (end-1) xs else slice (begin - 1) (end - 1) xs)" unfolding slice_def by (auto simp: take_Cons' drop_Cons') lemma slice_complete[simp]: "slice 0 (length xs) xs = xs" unfolding slice_def by simp subsubsection \Miscellaneous\ lemma length_compl_induct[case_names Nil Cons]: "\P []; !! e l . \!! ll . length ll <= length l \ P ll\ \ P (e#l)\ \ P l" apply(induct_tac l rule: length_induct) apply(case_tac "xs") apply(auto) done lemma in_set_list_format: "\ e\set l; !!l1 l2. l=l1@e#l2 \ P \ \ P" proof (induct l arbitrary: P) case Nil thus ?case by auto next case (Cons a l) show ?case proof (cases "a=e") case True with Cons show ?thesis by force next case False with Cons.prems(1) have "e\set l" by auto with Cons.hyps obtain l1 l2 where "l=l1@e#l2" by blast hence "a#l = (a#l1)@e#l2" by simp with Cons.prems(2) show P by blast qed qed lemma in_set_upd_cases: assumes "x\set (l[i:=y])" obtains "iset l" by (metis assms in_set_conv_nth length_list_update nth_list_update_eq nth_list_update_neq) lemma in_set_upd_eq_aux: assumes "iset (l[i:=y]) \ x=y \ (\y. x\set (l[i:=y]))" by (metis in_set_upd_cases assms list_update_overwrite set_update_memI) lemma in_set_upd_eq: assumes "iset (l[i:=y]) \ x=y \ (x\set l \ (\y. x\set (l[i:=y])))" by (metis in_set_upd_cases in_set_upd_eq_aux assms) text \Simultaneous induction over two lists, prepending an element to one of the lists in each step\ lemma list_2pre_induct[case_names base left right]: assumes BASE: "P [] []" and LEFT: "!!e w1' w2. P w1' w2 \ P (e#w1') w2" and RIGHT: "!!e w1 w2'. P w1 w2' \ P w1 (e#w2')" shows "P w1 w2" proof - { \ \The proof is done by induction over the sum of the lengths of the lists\ fix n have "!!w1 w2. \length w1 + length w2 = n; P [] []; !!e w1' w2. P w1' w2 \ P (e#w1') w2; !!e w1 w2'. P w1 w2' \ P w1 (e#w2') \ \ P w1 w2 " apply (induct n) apply simp apply (case_tac w1) apply auto apply (case_tac w2) apply auto done } from this[OF _ BASE LEFT RIGHT] show ?thesis by blast qed lemma list_decomp_1: "length l=1 \ \a. l=[a]" by (case_tac l, auto) lemma list_decomp_2: "length l=2 \ \a b. l=[a,b]" by (case_tac l, auto simp add: list_decomp_1) lemma list_rest_coinc: "\length s2 \ length s1; s1@r1 = s2@r2\ \ \r1p. r2=r1p@r1" by (metis append_eq_append_conv_if) lemma list_tail_coinc: "n1#r1 = n2#r2 \ n1=n2 & r1=r2" by (auto) lemma last_in_set[intro]: "\l\[]\ \ last l \ set l" by (induct l) auto lemma empty_append_eq_id[simp]: "(@) [] = (\x. x)" by auto lemma op_conc_empty_img_id[simp]: "((@) [] ` L) = L" by auto lemma distinct_match: "\ distinct (al@e#bl) \ \ (al@e#bl = al'@e#bl') \ (al=al' \ bl=bl')" proof (rule iffI, induct al arbitrary: al') case Nil thus ?case by (cases al') auto next case (Cons a al) note Cprems=Cons.prems note Chyps=Cons.hyps show ?case proof (cases al') case Nil with Cprems have False by auto thus ?thesis .. next case [simp]: (Cons a' all') with Cprems have [simp]: "a=a'" and P: "al@e#bl = all'@e#bl'" by auto from Cprems(1) have D: "distinct (al@e#bl)" by auto from Chyps[OF D P] have [simp]: "al=all'" "bl=bl'" by auto show ?thesis by simp qed qed simp lemma prop_match: "\ list_all P al; \P e; \P e'; list_all P bl \ \ (al@e#bl = al'@e'#bl') \ (al=al' \ e=e' \ bl=bl')" apply (rule iffI, induct al arbitrary: al') apply (case_tac al', fastforce, fastforce)+ done lemmas prop_matchD = rev_iffD1[OF _ prop_match[where P=P]] for P declare distinct_tl[simp] lemma list_se_match[simp]: "l1 \ [] \ l1@l2 = [a] \ l1 = [a] \ l2 = []" "l2 \ [] \ l1@l2 = [a] \ l1 = [] \ l2 = [a]" "l1 \ [] \ [a] = l1@l2 \ l1 = [a] \ l2 = []" "l2 \ [] \ [a] = l1@l2 \ l1 = [] \ l2 = [a]" apply (cases l1, simp_all) apply (cases l1, simp_all) apply (cases l1, auto) [] apply (cases l1, auto) [] done (* Placed here because it depends on xy_in_set_cases *) lemma distinct_map_eq: "\ distinct (List.map f l); f x = f y; x\set l; y\set l \ \ x=y" by (erule (2) xy_in_set_cases) auto lemma upt_append: assumes "iu'" assumes NP: "\i. u\i \ i \P i" shows "[i\[0..[0..[0..[0..[u..[u..[0.. P (l!i)" proof assume A: "P (l!i)" have "[0..i by (simp add: upt_append) also have "[i..i by (auto simp: upt_conv_Cons) finally have "[k\[0..[Suc i..P (l!i)\ by simp hence "j = last (i#[k\[Suc i.. \ i" proof - have "sorted (i#[k\[Suc i.. last ?l" by (rule sorted_hd_last) simp thus ?thesis by simp qed finally have "i\j" . thus False using \j by simp qed lemma all_set_conv_nth: "(\x\set l. P x) \ (\i *) lemma upt_0_eq_Nil_conv[simp]: "[0.. j=0" by auto lemma filter_eq_snocD: "filter P l = l'@[x] \ x\set l \ P x" proof - assume A: "filter P l = l'@[x]" hence "x\set (filter P l)" by simp thus ?thesis by simp qed lemma lists_image_witness: assumes A: "x\lists (f`Q)" obtains xo where "xo\lists Q" "x=map f xo" proof - have "\ x\lists (f`Q) \ \ \xo\lists Q. x=map f xo" proof (induct x) case Nil thus ?case by auto next case (Cons x xs) then obtain xos where "xos\lists Q" "xs=map f xos" by force moreover from Cons.prems have "x\f`Q" by auto then obtain xo where "xo\Q" "x=f xo" by auto ultimately show ?case by (rule_tac x="xo#xos" in bexI) auto qed thus ?thesis apply (simp_all add: A) apply (erule_tac bexE) apply (rule_tac that) apply assumption+ done qed lemma map_of_None_filterD: "map_of xs x = None \ map_of (filter P xs) x = None" by(induct xs) auto lemma map_of_concat: "map_of (concat xss) = foldr (\xs f. f ++ map_of xs) xss Map.empty" by(induct xss) simp_all lemma map_of_Some_split: "map_of xs k = Some v \ \ys zs. xs = ys @ (k, v) # zs \ map_of ys k = None" proof(induct xs) case (Cons x xs) obtain k' v' where x: "x = (k', v')" by(cases x) show ?case proof(cases "k' = k") case True with \map_of (x # xs) k = Some v\ x have "x # xs = [] @ (k, v) # xs" "map_of [] k = None" by simp_all thus ?thesis by blast next case False with \map_of (x # xs) k = Some v\ x have "map_of xs k = Some v" by simp from \map_of xs k = Some v \ \ys zs. xs = ys @ (k, v) # zs \ map_of ys k = None\[OF this] obtain ys zs where "xs = ys @ (k, v) # zs" "map_of ys k = None" by blast with False x have "x # xs = (x # ys) @ (k, v) # zs" "map_of (x # ys) k = None" by simp_all thus ?thesis by blast qed qed simp lemma map_add_find_left: "g k = None \ (f ++ g) k = f k" by(simp add: map_add_def) lemma map_add_left_None: "f k = None \ (f ++ g) k = g k" by(simp add: map_add_def split: option.split) lemma map_of_Some_filter_not_in: "\ map_of xs k = Some v; \ P (k, v); distinct (map fst xs) \ \ map_of (filter P xs) k = None" apply(induct xs) apply(auto) apply(auto simp add: map_of_eq_None_iff) done lemma distinct_map_fst_filterI: "distinct (map fst xs) \ distinct (map fst (filter P xs))" by(induct xs) auto lemma distinct_map_fstD: "\ distinct (map fst xs); (x, y) \ set xs; (x, z) \ set xs \ \ y = z" by(induct xs)(fastforce elim: notE rev_image_eqI)+ lemma concat_filter_neq_Nil: "concat [ys\xs. ys \ Nil] = concat xs" by(induct xs) simp_all lemma distinct_concat': "\distinct [ys\xs. ys \ Nil]; \ys. ys \ set xs \ distinct ys; \ys zs. \ys \ set xs; zs \ set xs; ys \ zs\ \ set ys \ set zs = {}\ \ distinct (concat xs)" by(erule distinct_concat[of "[ys\xs. ys \ Nil]", unfolded concat_filter_neq_Nil]) auto lemma distinct_idx: assumes "distinct (map f l)" assumes "im. n \ m \ m < length xs \ filter P xs ! n = xs ! m \ filter P (take m xs) = take n (filter P xs)" using assms proof(induct xs rule: rev_induct) case Nil thus ?case by simp next case (snoc x xs) show ?case proof(cases "P x") case [simp]: False from \n < length (filter P (xs @ [x]))\ have "n < length (filter P xs)" by simp hence "\m\n. m < length xs \ filter P xs ! n = xs ! m \ filter P (take m xs) = take n (filter P xs)" by(rule snoc) thus ?thesis by(auto simp add: nth_append) next case [simp]: True show ?thesis proof(cases "n = length (filter P xs)") case False with \n < length (filter P (xs @ [x]))\ have "n < length (filter P xs)" by simp moreover hence "\m\n. m < length xs \ filter P xs ! n = xs ! m \ filter P (take m xs) = take n (filter P xs)" by(rule snoc) ultimately show ?thesis by(auto simp add: nth_append) next case [simp]: True hence "filter P (xs @ [x]) ! n = (xs @ [x]) ! length xs" by simp moreover have "length xs < length (xs @ [x])" by simp moreover have "length xs \ n" by simp moreover have "filter P (take (length xs) (xs @ [x])) = take n (filter P (xs @ [x]))" by simp ultimately show ?thesis by blast qed qed qed lemma set_map_filter: "set (List.map_filter g xs) = {y. \x. x \ set xs \ g x = Some y}" by (induct xs) (auto simp add: List.map_filter_def set_eq_iff) subsection \Quicksort by Relation\ text \A functional implementation of quicksort on lists. It it similar to the one in Isabelle/HOL's example directory. However, it uses tail-recursion for append and arbitrary relations.\ fun partition_rev :: "('a \ bool) \ ('a list \ 'a list) \ 'a list \ ('a list \ 'a list)" where "partition_rev P (yes, no) [] = (yes, no)" | "partition_rev P (yes, no) (x # xs) = partition_rev P (if P x then (x # yes, no) else (yes, x # no)) xs" lemma partition_rev_filter_conv : "partition_rev P (yes, no) xs = (rev (filter P xs) @ yes, rev (filter (Not \ P) xs) @ no)" by (induct xs arbitrary: yes no) (simp_all) function quicksort_by_rel :: "('a \ 'a \ bool) \ 'a list \ 'a list \ 'a list" where "quicksort_by_rel R sl [] = sl" | "quicksort_by_rel R sl (x#xs) = (let (xs_s, xs_b) = partition_rev (\y. R y x) ([],[]) xs in quicksort_by_rel R (x # (quicksort_by_rel R sl xs_b)) xs_s)" by pat_completeness simp_all termination by (relation "measure (\(_, _, xs). length xs)") (simp_all add: partition_rev_filter_conv less_Suc_eq_le) lemma quicksort_by_rel_remove_acc : "quicksort_by_rel R sl xs = (quicksort_by_rel R [] xs @ sl)" proof (induct xs arbitrary: sl rule: measure_induct_rule[of "length"]) case (less xs) note ind_hyp = this show ?case proof (cases xs) case Nil thus ?thesis by simp next case (Cons x xs') note xs_eq[simp] = this obtain xs1 xs2 where part_rev_eq[simp]: "partition_rev (\y. R y x) ([], []) xs' = (xs1, xs2)" by (rule prod.exhaust) from part_rev_eq[symmetric] have length_le: "length xs1 < length xs" "length xs2 < length xs" unfolding partition_rev_filter_conv by (simp_all add: less_Suc_eq_le) note ind_hyp1a = ind_hyp[OF length_le(1), of "x # quicksort_by_rel R [] xs2"] note ind_hyp1b = ind_hyp[OF length_le(1), of "x # quicksort_by_rel R [] xs2 @ sl"] note ind_hyp2 = ind_hyp[OF length_le(2), of sl] show ?thesis by (simp add: ind_hyp1a ind_hyp1b ind_hyp2) qed qed lemma quicksort_by_rel_remove_acc_guared : "sl \ [] \ quicksort_by_rel R sl xs = (quicksort_by_rel R [] xs @ sl)" by (metis quicksort_by_rel_remove_acc) lemma quicksort_by_rel_permutes [simp]: "mset (quicksort_by_rel R sl xs) = mset (xs @ sl)" proof (induct xs arbitrary: sl rule: measure_induct_rule[of "length"]) case (less xs) note ind_hyp = this show ?case proof (cases xs) case Nil thus ?thesis by simp next case (Cons x xs') note xs_eq[simp] = this obtain xs1 xs2 where part_rev_eq[simp]: "partition_rev (\y. R y x) ([], []) xs' = (xs1, xs2)" by (rule prod.exhaust) from part_rev_eq[symmetric] have xs'_multi_eq : "mset xs' = mset xs1 + mset xs2" unfolding partition_rev_filter_conv by (simp add: mset_filter) from part_rev_eq[symmetric] have length_le: "length xs1 < length xs" "length xs2 < length xs" unfolding partition_rev_filter_conv by (simp_all add: less_Suc_eq_le) note ind_hyp[OF length_le(1)] ind_hyp[OF length_le(2)] thus ?thesis by (simp add: xs'_multi_eq union_assoc) qed qed lemma set_quicksort_by_rel [simp]: "set (quicksort_by_rel R sl xs) = set (xs @ sl)" unfolding set_mset_comp_mset [symmetric] o_apply by simp lemma sorted_wrt_quicksort_by_rel: fixes R:: "'x \ 'x \ bool" assumes lin : "\x y. (R x y) \ (R y x)" and trans_R: "\x y z. R x y \ R y z \ R x z" shows "sorted_wrt R (quicksort_by_rel R [] xs)" proof (induct xs rule: measure_induct_rule[of "length"]) case (less xs) note ind_hyp = this show ?case proof (cases xs) case Nil thus ?thesis by simp next case (Cons x xs') note xs_eq[simp] = this obtain xs1 xs2 where part_rev_eq[simp]: "partition_rev (\y. R y x) ([], []) xs' = (xs1, xs2)" by (rule prod.exhaust) from part_rev_eq[symmetric] have xs1_props: "\y. y \ set xs1 \ (R y x)" and xs2_props: "\y. y \ set xs2 \ \(R y x)" unfolding partition_rev_filter_conv by simp_all from xs2_props lin have xs2_props': "\y. y \ set xs2 \ (R x y)" by blast from xs2_props' xs1_props trans_R have xs1_props': "\y1 y2. y1 \ set xs1 \ y2 \ set xs2 \ (R y1 y2)" by metis from part_rev_eq[symmetric] have length_le: "length xs1 < length xs" "length xs2 < length xs" unfolding partition_rev_filter_conv by (simp_all add: less_Suc_eq_le) note ind_hyps = ind_hyp[OF length_le(1)] ind_hyp[OF length_le(2)] thus ?thesis by (simp add: quicksort_by_rel_remove_acc_guared sorted_wrt_append Ball_def xs1_props xs2_props' xs1_props') qed qed lemma sorted_quicksort_by_rel: "sorted (quicksort_by_rel (\) [] xs)" by (rule sorted_wrt_quicksort_by_rel) auto lemma sort_quicksort_by_rel: "sort = quicksort_by_rel (\) []" apply (rule ext, rule properties_for_sort) apply(simp_all add: sorted_quicksort_by_rel) done lemma [code]: "quicksort = quicksort_by_rel (\) []" apply (subst sort_quicksort[symmetric]) by (rule sort_quicksort_by_rel) subsection \Mergesort by Relation\ text \A functional implementation of mergesort on lists. It it similar to the one in Isabelle/HOL's example directory. However, it uses tail-recursion for append and arbitrary relations.\ fun mergesort_by_rel_split :: "('a list \ 'a list) \ 'a list \ ('a list \ 'a list)" where "mergesort_by_rel_split (xs1, xs2) [] = (xs1, xs2)" | "mergesort_by_rel_split (xs1, xs2) [x] = (x # xs1, xs2)" | "mergesort_by_rel_split (xs1, xs2) (x1 # x2 # xs) = mergesort_by_rel_split (x1 # xs1, x2 # xs2) xs" lemma list_induct_first2 [consumes 0, case_names Nil Sing Cons2]: assumes "P []" "\x. P [x]" "\x1 x2 xs. P xs \ P (x1 # x2 #xs)" shows "P xs" proof (induct xs rule: length_induct) case (1 xs) note ind_hyp = this show ?case proof (cases xs) case Nil thus ?thesis using assms(1) by simp next case (Cons x1 xs') note xs_eq[simp] = this thus ?thesis proof (cases xs') case Nil thus ?thesis using assms(2) by simp next case (Cons x2 xs'') note xs'_eq[simp] = this show ?thesis by (simp add: ind_hyp assms(3)) qed qed qed lemma mergesort_by_rel_split_length : "length (fst (mergesort_by_rel_split (xs1, xs2) xs)) = length xs1 + (length xs div 2) + (length xs mod 2) \ length (snd (mergesort_by_rel_split (xs1, xs2) xs)) = length xs2 + (length xs div 2)" by (induct xs arbitrary: xs1 xs2 rule: list_induct_first2) (simp_all) lemma mset_mergesort_by_rel_split [simp]: "mset (fst (mergesort_by_rel_split (xs1, xs2) xs)) + mset (snd (mergesort_by_rel_split (xs1, xs2) xs)) = mset xs + mset xs1 + mset xs2" apply (induct xs arbitrary: xs1 xs2 rule: list_induct_first2) apply (simp_all add: ac_simps) done fun mergesort_by_rel_merge :: "('a \ 'a \ bool) \ 'a list \ 'a list \ 'a list" where "mergesort_by_rel_merge R (x#xs) (y#ys) = (if R x y then x # mergesort_by_rel_merge R xs (y#ys) else y # mergesort_by_rel_merge R (x#xs) ys)" | "mergesort_by_rel_merge R xs [] = xs" | "mergesort_by_rel_merge R [] ys = ys" declare mergesort_by_rel_merge.simps [simp del] lemma mergesort_by_rel_merge_simps[simp] : "mergesort_by_rel_merge R (x#xs) (y#ys) = (if R x y then x # mergesort_by_rel_merge R xs (y#ys) else y # mergesort_by_rel_merge R (x#xs) ys)" "mergesort_by_rel_merge R xs [] = xs" "mergesort_by_rel_merge R [] ys = ys" apply (simp_all add: mergesort_by_rel_merge.simps) apply (cases ys) apply (simp_all add: mergesort_by_rel_merge.simps) done lemma mergesort_by_rel_merge_induct [consumes 0, case_names Nil1 Nil2 Cons1 Cons2]: assumes "\xs::'a list. P xs []" "\ys::'b list. P [] ys" "\x xs y ys. R x y \ P xs (y # ys) \ P (x # xs) (y # ys)" "\x xs y ys. \(R x y) \ P (x # xs) ys \ P (x # xs) (y # ys)" shows "P xs ys" proof (induct xs arbitrary: ys) case Nil thus ?case using assms(2) by simp next case (Cons x xs) note P_xs = this show ?case proof (induct ys) case Nil thus ?case using assms(1) by simp next case (Cons y ys) note P_x_xs_ys = this show ?case using assms(3,4)[of x y xs ys] P_x_xs_ys P_xs by metis qed qed lemma mset_mergesort_by_rel_merge [simp]: "mset (mergesort_by_rel_merge R xs ys) = mset xs + mset ys" by (induct R xs ys rule: mergesort_by_rel_merge.induct) (simp_all add: ac_simps) lemma set_mergesort_by_rel_merge [simp]: "set (mergesort_by_rel_merge R xs ys) = set xs \ set ys" by (induct R xs ys rule: mergesort_by_rel_merge.induct) auto lemma sorted_wrt_mergesort_by_rel_merge [simp]: assumes lin : "\x y. (R x y) \ (R y x)" and trans_R: "\x y z. R x y \ R y z \ R x z" shows "sorted_wrt R (mergesort_by_rel_merge R xs ys) \ sorted_wrt R xs \ sorted_wrt R ys" proof (induct xs ys rule: mergesort_by_rel_merge_induct[where R = R]) case Nil1 thus ?case by simp next case Nil2 thus ?case by simp next case (Cons1 x xs y ys) thus ?case by (simp add: Ball_def) (metis trans_R) next case (Cons2 x xs y ys) thus ?case apply (auto simp add: Ball_def) apply (metis lin) apply (metis lin trans_R) done qed function mergesort_by_rel :: "('a \ 'a \ bool) \ 'a list \ 'a list" where "mergesort_by_rel R xs = (if length xs < 2 then xs else (mergesort_by_rel_merge R (mergesort_by_rel R (fst (mergesort_by_rel_split ([], []) xs))) (mergesort_by_rel R (snd (mergesort_by_rel_split ([], []) xs)))))" by auto termination by (relation "measure (\(_, xs). length xs)") (simp_all add: mergesort_by_rel_split_length not_less minus_div_mult_eq_mod [symmetric]) declare mergesort_by_rel.simps [simp del] lemma mergesort_by_rel_simps [simp, code] : "mergesort_by_rel R [] = []" "mergesort_by_rel R [x] = [x]" "mergesort_by_rel R (x1 # x2 # xs) = (let (xs1, xs2) = (mergesort_by_rel_split ([x1], [x2]) xs) in mergesort_by_rel_merge R (mergesort_by_rel R xs1) (mergesort_by_rel R xs2))" apply (simp add: mergesort_by_rel.simps) apply (simp add: mergesort_by_rel.simps) apply (simp add: mergesort_by_rel.simps[of _ "x1 # x2 # xs"] split: prod.splits) done lemma mergesort_by_rel_permutes [simp]: "mset (mergesort_by_rel R xs) = mset xs" proof (induct xs rule: length_induct) case (1 xs) note ind_hyp = this show ?case proof (cases xs) case Nil thus ?thesis by simp next case (Cons x1 xs') note xs_eq[simp] = this show ?thesis proof (cases xs') case Nil thus ?thesis by simp next case (Cons x2 xs'') note xs'_eq[simp] = this have "length (fst (mergesort_by_rel_split ([], []) xs)) < length xs" "length (snd (mergesort_by_rel_split ([], []) xs)) < length xs" by (simp_all add: mergesort_by_rel_split_length) with ind_hyp show ?thesis unfolding mergesort_by_rel.simps[of _ xs] by (simp add: ac_simps) qed qed qed lemma set_mergesort_by_rel [simp]: "set (mergesort_by_rel R xs) = set xs" unfolding set_mset_comp_mset [symmetric] o_apply by simp lemma sorted_wrt_mergesort_by_rel: fixes R:: "'x \ 'x \ bool" assumes lin : "\x y. (R x y) \ (R y x)" and trans_R: "\x y z. R x y \ R y z \ R x z" shows "sorted_wrt R (mergesort_by_rel R xs)" proof (induct xs rule: measure_induct_rule[of "length"]) case (less xs) note ind_hyp = this show ?case proof (cases xs) case Nil thus ?thesis by simp next case (Cons x xs') note xs_eq[simp] = this thus ?thesis proof (cases xs') case Nil thus ?thesis by simp next case (Cons x2 xs'') note xs'_eq[simp] = this have "length (fst (mergesort_by_rel_split ([], []) xs)) < length xs" "length (snd (mergesort_by_rel_split ([], []) xs)) < length xs" by (simp_all add: mergesort_by_rel_split_length) with ind_hyp show ?thesis unfolding mergesort_by_rel.simps[of _ xs] by (simp add: sorted_wrt_mergesort_by_rel_merge[OF lin trans_R]) qed qed qed lemma sorted_mergesort_by_rel: "sorted (mergesort_by_rel (\) xs)" by (rule sorted_wrt_mergesort_by_rel) auto lemma sort_mergesort_by_rel: "sort = mergesort_by_rel (\)" apply (rule ext, rule properties_for_sort) apply(simp_all add: sorted_mergesort_by_rel) done definition "mergesort = mergesort_by_rel (\)" lemma sort_mergesort: "sort = mergesort" unfolding mergesort_def by (rule sort_mergesort_by_rel) subsubsection \Mergesort with Remdup\ term merge fun merge :: "'a::{linorder} list \ 'a list \ 'a list" where "merge [] l2 = l2" | "merge l1 [] = l1" | "merge (x1 # l1) (x2 # l2) = (if (x1 < x2) then x1 # (merge l1 (x2 # l2)) else (if (x1 = x2) then x1 # (merge l1 l2) else x2 # (merge (x1 # l1) l2)))" lemma merge_correct : assumes l1_OK: "distinct l1 \ sorted l1" assumes l2_OK: "distinct l2 \ sorted l2" shows "distinct (merge l1 l2) \ sorted (merge l1 l2) \ set (merge l1 l2) = set l1 \ set l2" using assms proof (induct l1 arbitrary: l2) case Nil thus ?case by simp next case (Cons x1 l1 l2) note x1_l1_props = Cons(2) note l2_props = Cons(3) from x1_l1_props have l1_props: "distinct l1 \ sorted l1" and x1_nin_l1: "x1 \ set l1" and x1_le: "\x. x \ set l1 \ x1 \ x" by (simp_all add: Ball_def) note ind_hyp_l1 = Cons(1)[OF l1_props] show ?case using l2_props proof (induct l2) case Nil with x1_l1_props show ?case by simp next case (Cons x2 l2) note x2_l2_props = Cons(2) from x2_l2_props have l2_props: "distinct l2 \ sorted l2" and x2_nin_l2: "x2 \ set l2" and x2_le: "\x. x \ set l2 \ x2 \ x" by (simp_all add: Ball_def) note ind_hyp_l2 = Cons(1)[OF l2_props] show ?case proof (cases "x1 < x2") case True note x1_less_x2 = this from ind_hyp_l1[OF x2_l2_props] x1_less_x2 x1_nin_l1 x1_le x2_le show ?thesis apply (auto simp add: Ball_def) apply (metis linorder_not_le) apply (metis linorder_not_less xt1(6) xt1(9)) done next case False note x2_le_x1 = this show ?thesis proof (cases "x1 = x2") case True note x1_eq_x2 = this from ind_hyp_l1[OF l2_props] x1_le x2_le x2_nin_l2 x1_eq_x2 x1_nin_l1 show ?thesis by (simp add: x1_eq_x2 Ball_def) next case False note x1_neq_x2 = this with x2_le_x1 have x2_less_x1 : "x2 < x1" by auto from ind_hyp_l2 x2_le_x1 x1_neq_x2 x2_le x2_nin_l2 x1_le show ?thesis apply (simp add: x2_less_x1 Ball_def) apply (metis linorder_not_le x2_less_x1 xt1(7)) done qed qed qed qed function merge_list :: "'a::{linorder} list list \ 'a list list \ 'a list" where "merge_list [] [] = []" | "merge_list [] [l] = l" | "merge_list (la # acc2) [] = merge_list [] (la # acc2)" | "merge_list (la # acc2) [l] = merge_list [] (l # la # acc2)" | "merge_list acc2 (l1 # l2 # ls) = merge_list ((merge l1 l2) # acc2) ls" by pat_completeness simp_all termination by (relation "measure (\(acc, ls). 3 * length acc + 2 * length ls)") (simp_all) lemma merge_list_correct : assumes ls_OK: "\l. l \ set ls \ distinct l \ sorted l" assumes as_OK: "\l. l \ set as \ distinct l \ sorted l" shows "distinct (merge_list as ls) \ sorted (merge_list as ls) \ set (merge_list as ls) = set (concat (as @ ls))" using assms proof (induct as ls rule: merge_list.induct) case 1 thus ?case by simp next case 2 thus ?case by simp next case 3 thus ?case by simp next case 4 thus ?case by auto next case (5 acc l1 l2 ls) note ind_hyp = 5(1) note l12_l_OK = 5(2) note acc_OK = 5(3) from l12_l_OK acc_OK merge_correct[of l1 l2] have set_merge_eq: "set (merge l1 l2) = set l1 \ set l2" by auto from l12_l_OK acc_OK merge_correct[of l1 l2] have "distinct (merge_list (merge l1 l2 # acc) ls) \ sorted (merge_list (merge l1 l2 # acc) ls) \ set (merge_list (merge l1 l2 # acc) ls) = set (concat ((merge l1 l2 # acc) @ ls))" by (rule_tac ind_hyp) auto with set_merge_eq show ?case by auto qed definition mergesort_remdups where "mergesort_remdups xs = merge_list [] (map (\x. [x]) xs)" lemma mergesort_remdups_correct : "distinct (mergesort_remdups l) \ sorted (mergesort_remdups l) \ (set (mergesort_remdups l) = set l)" proof - let ?l' = "map (\x. [x]) l" { fix xs assume "xs \ set ?l'" then obtain x where xs_eq: "xs = [x]" by auto hence "distinct xs \ sorted xs" by simp } note l'_OK = this from merge_list_correct[of "?l'" "[]", OF l'_OK] show ?thesis unfolding mergesort_remdups_def by simp qed (* TODO: Move *) lemma ex1_eqI: "\\!x. P x; P a; P b\ \ a=b" by blast lemma remdup_sort_mergesort_remdups: "remdups o sort = mergesort_remdups" (is "?lhs=?rhs") proof fix l have "set (?lhs l) = set l" and "sorted (?lhs l)" and "distinct (?lhs l)" by simp_all moreover note mergesort_remdups_correct ultimately show "?lhs l = ?rhs l" by (auto intro!: ex1_eqI[OF finite_sorted_distinct_unique[OF finite_set]]) qed subsection \Native Integers\ lemma int_of_integer_less_iff: "int_of_integer x < int_of_integer y \ x0 \ y\0 \ nat_of_integer x < nat_of_integer y \ xn' < n. \ P n') \ P (n::nat)" shows "\n' \ n. P n'" proof (rule classical) assume contra: "\ (\n'\n. P n')" hence "\n' < n. \ P n'" by auto hence "P n" by (rule hyp) thus "\n'\n. P n'" by auto qed subsubsection \Induction on nat\ lemma nat_compl_induct[case_names 0 Suc]: "\P 0; \n . \nn. nn \ n \ P nn \ P (Suc n)\ \ P n" apply(induct_tac n rule: nat_less_induct) apply(case_tac n) apply(auto) done lemma nat_compl_induct'[case_names 0 Suc]: "\P 0; !! n . \!! nn . nn \ n \ P nn\ \ P (Suc n)\ \ P n" apply(induct_tac n rule: nat_less_induct) apply(case_tac n) apply(auto) done lemma nz_le_conv_less: "0 k \ m \ k - Suc 0 < m" by auto lemma min_Suc_gt[simp]: "a min (Suc a) b = Suc a" "a min b (Suc a) = Suc a" by auto subsection \Integer\ text \Some setup from \int\ transferred to \integer\\ lemma atLeastLessThanPlusOne_atLeastAtMost_integer: "{l.. u \ {(0::integer).. u") apply (subst image_atLeastZeroLessThan_integer, assumption) apply (rule finite_imageI) apply auto done lemma finite_atLeastLessThan_integer [iff]: "finite {l..Functions of type @{typ "bool\bool"}\ lemma boolfun_cases_helper: "g=(\x. False) | g=(\x. x) | g=(\x. True) | g= (\x. \x)" proof - { assume "g False" "g True" hence "g = (\x. True)" by (rule_tac ext, case_tac x, auto) } moreover { assume "g False" "\g True" hence "g = (\x. \x)" by (rule_tac ext, case_tac x, auto) } moreover { assume "\g False" "g True" hence "g = (\x. x)" by (rule_tac ext, case_tac x, auto) } moreover { assume "\g False" "\g True" hence "g = (\x. False)" by (rule_tac ext, case_tac x, auto) } ultimately show ?thesis by fast qed lemma boolfun_cases[case_names False Id True Neg]: "\g=(\x. False) \ P; g=(\x. x) \ P; g=(\x. True) \ P; g=(\x. \x) \ P\ \ P" proof - note boolfun_cases_helper[of g] moreover assume "g=(\x. False) \ P" "g=(\x. x) \ P" "g=(\x. True) \ P" "g=(\x. \x) \ P" ultimately show ?thesis by fast qed subsection \Definite and indefinite description\ text "Combined definite and indefinite description for binary predicate" lemma some_theI: assumes EX: "\a b . P a b" and BUN: "!! b1 b2 . \\a . P a b1; \a . P a b2\ \ b1=b2" shows "P (SOME a . \b . P a b) (THE b . \a . P a b)" proof - from EX have "\b. P (SOME a. \b. P a b) b" by (rule someI_ex) moreover from EX have "\b. \a. P a b" by blast with BUN theI'[of "\b. \a. P a b"] have "\a. P a (THE b. \a. P a b)" by (unfold Ex1_def, blast) moreover note BUN ultimately show ?thesis by (fast) qed lemma some_insert_self[simp]: "S\{} \ insert (SOME x. x\S) S = S" by (auto intro: someI) lemma some_elem[simp]: "S\{} \ (SOME x. x\S) \ S" by (auto intro: someI) subsubsection\Hilbert Choice with option\ definition Eps_Opt where "Eps_Opt P = (if (\x. P x) then Some (SOME x. P x) else None)" lemma some_opt_eq_trivial[simp] : "Eps_Opt (\y. y = x) = Some x" unfolding Eps_Opt_def by simp lemma some_opt_sym_eq_trivial[simp] : "Eps_Opt ((=) x) = Some x" unfolding Eps_Opt_def by simp lemma some_opt_false_trivial[simp] : "Eps_Opt (\_. False) = None" unfolding Eps_Opt_def by simp lemma Eps_Opt_eq_None[simp] : "Eps_Opt P = None \ \(Ex P)" unfolding Eps_Opt_def by simp lemma Eps_Opt_eq_Some_implies : "Eps_Opt P = Some x \ P x" unfolding Eps_Opt_def by (metis option.inject option.simps(2) someI_ex) lemma Eps_Opt_eq_Some : assumes P_prop: "\x'. P x \ P x' \ x' = x" shows "Eps_Opt P = Some x \ P x" using P_prop unfolding Eps_Opt_def by (metis option.inject option.simps(2) someI_ex) subsection \Product Type\ lemma nested_case_prod_simp: "(\(a,b) c. f a b c) x y = (case_prod (\a b. f a b y) x)" by (auto split: prod.split) lemma fn_fst_conv: "(\x. (f (fst x))) = (\(a,_). f a)" by auto lemma fn_snd_conv: "(\x. (f (snd x))) = (\(_,b). f b)" by auto fun pairself where "pairself f (a,b) = (f a, f b)" lemma pairself_image_eq[simp]: "pairself f ` {(a,b). P a b} = {(f a, f b)| a b. P a b}" by force lemma pairself_image_cart[simp]: "pairself f ` (A\B) = f`A \ f`B" by (auto simp: image_def) lemma in_prod_fst_sndI: "fst x \ A \ snd x \ B \ x\A\B" by (cases x) auto lemma inj_Pair[simp]: "inj_on (\x. (x,c x)) S" "inj_on (\x. (c x,x)) S" by (auto intro!: inj_onI) declare Product_Type.swap_inj_on[simp] lemma img_fst [intro]: assumes "(a,b) \ S" shows "a \ fst ` S" by (rule image_eqI[OF _ assms]) simp lemma img_snd [intro]: assumes "(a,b) \ S" shows "b \ snd ` S" by (rule image_eqI[OF _ assms]) simp lemma range_prod: "range f \ (range (fst \ f)) \ (range (snd \ f))" proof fix y assume "y \ range f" then obtain x where y: "y = f x" by auto hence "y = (fst(f x), snd(f x))" by simp thus "y \ (range (fst \ f)) \ (range (snd \ f))" by (fastforce simp add: image_def) qed lemma finite_range_prod: assumes fst: "finite (range (fst \ f))" and snd: "finite (range (snd \ f))" shows "finite (range f)" proof - from fst snd have "finite (range (fst \ f) \ range (snd \ f))" by (rule finite_SigmaI) thus ?thesis by (rule finite_subset[OF range_prod]) qed lemma fstE: "x = (a,b) \ P (fst x) \ P a" by (metis fst_conv) lemma sndE: "x = (a,b) \ P (snd x) \ P b" by (metis snd_conv) subsubsection \Uncurrying\ (* TODO: Move to HOL/Product_Type? Lars H: "It's equal to case_prod, should use an abbreviation"*) definition uncurry :: "('a \ 'b \ 'c) \ 'a \ 'b \ 'c" where "uncurry f \ \(a,b). f a b" lemma uncurry_apply[simp]: "uncurry f (a,b) = f a b" unfolding uncurry_def by simp lemma curry_uncurry_id[simp]: "curry (uncurry f) = f" unfolding uncurry_def by simp lemma uncurry_curry_id[simp]: "uncurry (curry f) = f" unfolding uncurry_def by simp lemma do_curry: "f (a,b) = curry f a b" by simp lemma do_uncurry: "f a b = uncurry f (a,b)" by simp subsection \Sum Type\ lemma map_sum_Inr_conv: "map_sum fl fr s = Inr y \ (\x. s=Inr x \ y = fr x)" by (cases s) auto lemma map_sum_Inl_conv: "map_sum fl fr s = Inl y \ (\x. s=Inl x \ y = fl x)" by (cases s) auto subsection \Directed Graphs and Relations\ subsubsection "Reflexive-Transitive Closure" lemma r_le_rtrancl[simp]: "S\S\<^sup>*" by auto lemma rtrancl_mono_rightI: "S\S' \ S\S'\<^sup>*" by auto lemma trancl_sub: "R \ R\<^sup>+" by auto lemma trancl_single[simp]: "{(a,b)}\<^sup>+ = {(a,b)}" by (auto simp: trancl_insert) text \Pick first non-reflexive step\ lemma converse_rtranclE'[consumes 1, case_names base step]: assumes "(u,v)\R\<^sup>*" obtains "u=v" | vh where "u\vh" and "(u,vh)\R" and "(vh,v)\R\<^sup>*" using assms apply (induct rule: converse_rtrancl_induct) apply auto [] apply (case_tac "y=z") apply auto done lemma in_rtrancl_insert: "x\R\<^sup>* \ x\(insert r R)\<^sup>*" by (metis in_mono rtrancl_mono subset_insertI) lemma rtrancl_apply_insert: "R\<^sup>*``(insert x S) = insert x (R\<^sup>*``(S\R``{x}))" apply (auto) apply (erule converse_rtranclE) apply auto [2] apply (erule converse_rtranclE) apply (auto intro: converse_rtrancl_into_rtrancl) [2] done text \A path in a graph either does not use nodes from S at all, or it has a prefix leading to a node in S and a suffix that does not use nodes in S\ lemma rtrancl_last_visit[cases set, case_names no_visit last_visit_point]: shows "\ (q,q')\R\<^sup>*; (q,q')\(R-UNIV\S)\<^sup>* \ P; !!qt. \ qt\S; (q,qt)\R\<^sup>+; (qt,q')\(R-UNIV\S)\<^sup>* \ \ P \ \ P" proof (induct rule: converse_rtrancl_induct[case_names refl step]) case refl thus ?case by auto next case (step q qh) show P proof (rule step.hyps(3)) assume A: "(qh,q')\(R-UNIV\S)\<^sup>*" show P proof (cases "qh\S") case False with step.hyps(1) A have "(q,q')\(R-UNIV\S)\<^sup>*" by (auto intro: converse_rtrancl_into_rtrancl) with step.prems(1) show P . next case True from step.hyps(1) have "(q,qh)\R\<^sup>+" by auto with step.prems(2) True A show P by blast qed next fix qt assume A: "qt\S" "(qh,qt)\R\<^sup>+" "(qt,q')\(R-UNIV\S)\<^sup>*" with step.hyps(1) have "(q,qt)\R\<^sup>+" by auto with step.prems(2) A(1,3) show P by blast qed qed text \Less general version of \rtrancl_last_visit\, but there's a short automatic proof\ lemma rtrancl_last_visit': "\ (q,q')\R\<^sup>*; (q,q')\(R-UNIV\S)\<^sup>* \ P; !!qt. \ qt\S; (q,qt)\R\<^sup>*; (qt,q')\(R-UNIV\S)\<^sup>* \ \ P \ \ P" by (induct rule: converse_rtrancl_induct) (auto intro: converse_rtrancl_into_rtrancl) lemma rtrancl_last_visit_node: assumes "(s,s')\R\<^sup>*" shows "s\sh \ (s,s')\(R \ (UNIV \ (-{sh})))\<^sup>* \ (s,sh)\R\<^sup>* \ (sh,s')\(R \ (UNIV \ (-{sh})))\<^sup>*" using assms proof (induct rule: converse_rtrancl_induct) case base thus ?case by auto next case (step s st) moreover { assume P: "(st,s')\ (R \ UNIV \ - {sh})\<^sup>*" { assume "st=sh" with step have ?case by auto } moreover { assume "st\sh" with \(s,st)\R\ have "(s,st)\(R \ UNIV \ - {sh})\<^sup>*" by auto also note P finally have ?case by blast } ultimately have ?case by blast } moreover { assume P: "(st, sh) \ R\<^sup>* \ (sh, s') \ (R \ UNIV \ - {sh})\<^sup>*" with step(1) have ?case by (auto dest: converse_rtrancl_into_rtrancl) } ultimately show ?case by blast qed text \Find last point where a path touches a set\ lemma rtrancl_last_touch: "\ (q,q')\R\<^sup>*; q\S; !!qt. \ qt\S; (q,qt)\R\<^sup>*; (qt,q')\(R-UNIV\S)\<^sup>* \ \ P \ \ P" by (erule rtrancl_last_visit') auto text \A path either goes over edge once, or not at all\ lemma trancl_over_edgeE: assumes "(u,w)\(insert (v1,v2) E)\<^sup>+" obtains "(u,w)\E\<^sup>+" | "(u,v1)\E\<^sup>*" and "(v2,w)\E\<^sup>*" using assms proof induct case (base z) thus ?thesis by (metis insertE prod.inject r_into_trancl' rtrancl_eq_or_trancl) next case (step y z) thus ?thesis by (metis (opaque_lifting, no_types) Pair_inject insertE rtrancl.simps trancl.simps trancl_into_rtrancl) qed lemma rtrancl_image_advance: "\q\R\<^sup>* `` Q0; (q,x)\R\ \ x\R\<^sup>* `` Q0" by (auto intro: rtrancl_into_rtrancl) lemma trancl_image_by_rtrancl: "(E\<^sup>+)``Vi \ Vi = (E\<^sup>*)``Vi" by (metis Image_Id Un_Image rtrancl_trancl_reflcl) lemma reachable_mono: "\R\R'; X\X'\ \ R\<^sup>*``X \ R'\<^sup>*``X'" by (metis Image_mono rtrancl_mono) lemma finite_reachable_advance: "\ finite (E\<^sup>*``{v0}); (v0,v)\E\<^sup>* \ \ finite (E\<^sup>*``{v})" by (erule finite_subset[rotated]) auto lemma rtrancl_mono_mp: "U\V \ x\U\<^sup>* \ x\V\<^sup>*" by (metis in_mono rtrancl_mono) lemma trancl_mono_mp: "U\V \ x\U\<^sup>+ \ x\V\<^sup>+" by (metis trancl_mono) lemma rtrancl_mapI: "(a,b)\E\<^sup>* \ (f a, f b)\(pairself f `E)\<^sup>*" apply (induction rule: rtrancl_induct) apply (force intro: rtrancl.intros)+ done lemma rtrancl_image_advance_rtrancl: assumes "q \ R\<^sup>*``Q0" assumes "(q,x) \ R\<^sup>*" shows "x \ R\<^sup>*``Q0" using assms by (metis rtrancl_idemp rtrancl_image_advance) lemma nth_step_trancl: "\n m. \ \ n. n < length xs - 1 \ (xs ! Suc n, xs ! n) \ R \ \ n < length xs \ m < n \ (xs ! n, xs ! m) \ R\<^sup>+" proof (induction xs) case (Cons x xs) hence "\n. n < length xs - 1 \ (xs ! Suc n, xs ! n) \ R" apply clarsimp by (metis One_nat_def diff_Suc_eq_diff_pred nth_Cons_Suc zero_less_diff) note IH = this[THEN Cons.IH] from Cons obtain n' where n': "Suc n' = n" by (cases n) blast+ show ?case proof (cases m) case "0" with Cons have "xs \ []" by auto with "0" Cons.prems(1)[of m] have "(xs ! 0, x) \ R" by simp moreover from IH[where m = 0] have "\n. n < length xs \ n > 0 \ (xs ! n, xs ! 0) \ R\<^sup>+" by simp ultimately have "\n. n < length xs \ (xs ! n, x) \ R\<^sup>+" by (metis trancl_into_trancl gr0I r_into_trancl') with Cons "0" show ?thesis by auto next case (Suc m') with Cons.prems n' have "n' < length xs" "m' < n'" by auto with IH have "(xs ! n', xs ! m') \ R\<^sup>+" by simp with Suc n' show ?thesis by auto qed qed simp lemma Image_empty_trancl_Image_empty: "R `` {v} = {} \ R\<^sup>+ `` {v} = {}" unfolding Image_def by (auto elim: converse_tranclE) lemma Image_empty_rtrancl_Image_id: "R `` {v} = {} \ R\<^sup>* `` {v} = {v}" unfolding Image_def by (auto elim: converse_rtranclE) lemma trans_rtrancl_eq_reflcl: "trans A \ A^* = A^=" by (simp add: rtrancl_trancl_reflcl) lemma refl_on_reflcl_Image: "refl_on B A \ C \ B \ A^= `` C = A `` C" by (auto simp add: Image_def dest: refl_onD) lemma Image_absorb_rtrancl: "\ trans A; refl_on B A; C \ B \ \ A^* `` C = A `` C" by (simp add: trans_rtrancl_eq_reflcl refl_on_reflcl_Image) lemma trancl_Image_unfold_left: "E\<^sup>+``S = E\<^sup>*``E``S" by (auto simp: trancl_unfold_left) lemma trancl_Image_unfold_right: "E\<^sup>+``S = E``E\<^sup>*``S" by (auto simp: trancl_unfold_right) lemma trancl_Image_advance_ss: "(u,v)\E \ E\<^sup>+``{v} \ E\<^sup>+``{u}" by auto lemma rtrancl_Image_advance_ss: "(u,v)\E \ E\<^sup>*``{v} \ E\<^sup>*``{u}" by auto (* FIXME: nicer name *) lemma trancl_union_outside: assumes "(v,w) \ (E\U)\<^sup>+" and "(v,w) \ E\<^sup>+" shows "\x y. (v,x) \ (E\U)\<^sup>* \ (x,y) \ U \ (y,w) \ (E\U)\<^sup>*" using assms proof (induction) case base thus ?case by auto next case (step w x) show ?case proof (cases "(v,w)\E\<^sup>+") case True from step have "(v,w)\(E\U)\<^sup>*" by simp moreover from True step have "(w,x) \ U" by (metis Un_iff trancl.simps) moreover have "(x,x) \ (E\U)\<^sup>*" by simp ultimately show ?thesis by blast next case False with step.IH obtain a b where "(v,a) \ (E\U)\<^sup>*" "(a,b) \ U" "(b,w) \ (E\U)\<^sup>*" by blast moreover with step have "(b,x) \ (E\U)\<^sup>*" by (metis rtrancl_into_rtrancl) ultimately show ?thesis by blast qed qed lemma trancl_restrict_reachable: assumes "(u,v) \ E\<^sup>+" assumes "E``S \ S" assumes "u\S" shows "(u,v) \ (E\S\S)\<^sup>+" using assms by (induction rule: converse_trancl_induct) (auto intro: trancl_into_trancl2) lemma rtrancl_image_unfold_right: "E``E\<^sup>*``V \ E\<^sup>*``V" by (auto intro: rtrancl_into_rtrancl) lemma trancl_Image_in_Range: "R\<^sup>+ `` V \ Range R" by (auto elim: trancl.induct) lemma rtrancl_Image_in_Field: "R\<^sup>* `` V \ Field R \ V" proof - from trancl_Image_in_Range have "R\<^sup>+ `` V \ Field R" unfolding Field_def by fast hence "R\<^sup>+ `` V \ V \ Field R \ V" by blast with trancl_image_by_rtrancl show ?thesis by metis qed lemma rtrancl_sub_insert_rtrancl: "R\<^sup>* \ (insert x R)\<^sup>*" by (auto elim: rtrancl.induct rtrancl_into_rtrancl) lemma trancl_sub_insert_trancl: "R\<^sup>+ \ (insert x R)\<^sup>+" by (auto elim: trancl.induct trancl_into_trancl) lemma Restr_rtrancl_mono: "(v,w) \ (Restr E U)\<^sup>* \ (v,w) \ E\<^sup>*" by (metis inf_le1 rtrancl_mono subsetCE) lemma Restr_trancl_mono: "(v,w) \ (Restr E U)\<^sup>+ \ (v,w) \ E\<^sup>+" by (metis inf_le1 trancl_mono) subsubsection "Converse Relation" lemmas converse_add_simps = converse_Times trancl_converse[symmetric] converse_Un converse_Int lemma dom_ran_disj_comp[simp]: "Domain R \ Range R = {} \ R O R = {}" by auto lemma below_Id_inv[simp]: "R\\Id \ R\Id" by (auto) subsubsection "Cyclicity" lemma cyclicE: "\\acyclic g; !!x. (x,x)\g\<^sup>+ \ P\ \ P" by (unfold acyclic_def) blast lemma acyclic_insert_cyclic: "\acyclic g; \acyclic (insert (x,y) g)\ \ (y,x)\g\<^sup>*" by (unfold acyclic_def) (auto simp add: trancl_insert) text \ This lemma makes a case distinction about a path in a graph where a couple of edges with the same endpoint have been inserted: If there is a path from a to b, then there's such a path in the original graph, or there's a path that uses an inserted edge only once. Originally, this lemma was used to reason about the graph of an updated acquisition history. Any path in this graph is either already contained in the original graph, or passes via an inserted edge. Because all the inserted edges point to the same target node, in the second case, the path can be short-circuited to use exactly one inserted edge. \ lemma trancl_multi_insert[cases set, case_names orig via]: "\ (a,b)\(r\X\{m})\<^sup>+; (a,b)\r\<^sup>+ \ P; !!x. \ x\X; (a,x)\r\<^sup>*; (m,b)\r\<^sup>* \ \ P \ \ P" proof (induct arbitrary: P rule: trancl_induct) case (base b) thus ?case by auto next case (step b c) show ?case proof (rule step.hyps(3)) assume A: "(a,b)\r\<^sup>+" note step.hyps(2) moreover { assume "(b,c)\r" with A have "(a,c)\r\<^sup>+" by auto with step.prems have P by blast } moreover { assume "b\X" "c=m" with A have P by (rule_tac step.prems(2)) simp+ } ultimately show P by auto next fix x assume A: "x \ X" "(a, x) \ r\<^sup>*" "(m, b) \ r\<^sup>*" note step.hyps(2) moreover { assume "(b,c)\r" with A(3) have "(m,c)\r\<^sup>*" by auto with step.prems(2)[OF A(1,2)] have P by blast } moreover { assume "b\X" "c=m" with A have P by (rule_tac step.prems(2)) simp+ } ultimately show P by auto qed qed text \ Version of @{thm [source] trancl_multi_insert} for inserted edges with the same startpoint. \ lemma trancl_multi_insert2[cases set, case_names orig via]: "\(a,b)\(r\{m}\X)\<^sup>+; (a,b)\r\<^sup>+ \ P; !!x. \ x\X; (a,m)\r\<^sup>*; (x,b)\r\<^sup>* \ \ P \ \ P" proof goal_cases case prems: 1 from prems(1) have "(b,a)\((r\{m}\X)\<^sup>+)\" by simp also have "((r\{m}\X)\<^sup>+)\ = (r\\X\{m})\<^sup>+" by (simp add: converse_add_simps) finally have "(b, a) \ (r\ \ X \ {m})\<^sup>+" . thus ?case by (auto elim!: trancl_multi_insert intro: prems(2,3) simp add: trancl_converse rtrancl_converse ) qed lemma cyclic_subset: "\ \ acyclic R; R \ S \ \ \ acyclic S" unfolding acyclic_def by (blast intro: trancl_mono) subsubsection \Wellfoundedness\ lemma wf_min: assumes A: "wf R" "R\{}" "!!m. m\Domain R - Range R \ P" shows P proof - have H: "!!x. wf R \ \y. (x,y)\R \ x\Domain R - Range R \ (\m. m\Domain R - Range R)" by (erule_tac wf_induct_rule[where P="\x. \y. (x,y)\R \ x\Domain R - Range R \ (\m. m\Domain R - Range R)"]) auto from A(2) obtain x y where "(x,y)\R" by auto with A(1,3) H show ?thesis by blast qed lemma finite_wf_eq_wf_converse: "finite R \ wf (R\) \ wf R" by (metis acyclic_converse finite_acyclic_wf finite_acyclic_wf_converse wf_acyclic) lemma wf_max: assumes A: "wf (R\)" "R\{}" and C: "!!m. m\Range R - Domain R \ P" shows "P" proof - from A(2) have NE: "R\\{}" by auto from wf_min[OF A(1) NE] obtain m where "m\Range R - Domain R" by auto thus P by (blast intro: C) qed \ \Useful lemma to show well-foundedness of some process approaching a finite upper bound\ lemma wf_bounded_supset: "finite S \ wf {(Q',Q). Q'\Q \ Q'\ S}" proof - assume [simp]: "finite S" hence [simp]: "!!x. finite (S-x)" by auto have "{(Q',Q). Q\Q' \ Q'\ S} \ inv_image ({(s'::nat,s). s'Q. card (S-Q))" proof (intro subsetI, case_tac x, simp) fix a b assume A: "b\a \ a\S" hence "S-a \ S-b" by blast thus "card (S-a) < card (S-b)" by (auto simp add: psubset_card_mono) qed moreover have "wf ({(s'::nat,s). s' Range R = {} \ wf R" apply (rule wf_no_loop) by simp text \Extend a wf-relation by a break-condition\ definition "brk_rel R \ {((False,x),(False,y)) | x y. (x,y)\R} \ {((True,x),(False,y)) | x y. True}" lemma brk_rel_wf[simp,intro!]: assumes WF[simp]: "wf R" shows "wf (brk_rel R)" proof - have "wf {((False,x),(False,y)) | x y. (x,y)\R}" proof - have "{((False,x),(False,y)) | x y. (x,y)\R} \ inv_image R snd" by auto from wf_subset[OF wf_inv_image[OF WF] this] show ?thesis . qed moreover have "wf {((True,x),(False,y)) | x y. True}" by (rule wf_no_path) auto ultimately show ?thesis unfolding brk_rel_def apply (subst Un_commute) by (blast intro: wf_Un) qed subsubsection \Restrict Relation\ definition rel_restrict :: "('a \ 'a) set \ 'a set \ ('a \ 'a) set" where "rel_restrict R A \ {(v,w). (v,w) \ R \ v \ A \ w \ A}" lemma rel_restrict_alt_def: "rel_restrict R A = R \ (-A) \ (-A)" unfolding rel_restrict_def by auto lemma rel_restrict_empty[simp]: "rel_restrict R {} = R" by (simp add: rel_restrict_def) lemma rel_restrict_notR: assumes "(x,y) \ rel_restrict A R" shows "x \ R" and "y \ R" using assms unfolding rel_restrict_def by auto lemma rel_restrict_sub: "rel_restrict R A \ R" unfolding rel_restrict_def by auto lemma rel_restrict_Int_empty: "A \ Field R = {} \ rel_restrict R A = R" unfolding rel_restrict_def Field_def by auto lemma Domain_rel_restrict: "Domain (rel_restrict R A) \ Domain R - A" unfolding rel_restrict_def by auto lemma Range_rel_restrict: "Range (rel_restrict R A) \ Range R - A" unfolding rel_restrict_def by auto lemma Field_rel_restrict: "Field (rel_restrict R A) \ Field R - A" unfolding rel_restrict_def Field_def by auto lemma rel_restrict_compl: "rel_restrict R A \ rel_restrict R (-A) = {}" unfolding rel_restrict_def by auto lemma finite_rel_restrict: "finite R \ finite (rel_restrict R A)" by (metis finite_subset rel_restrict_sub) lemma R_subset_Field: "R \ Field R \ Field R" unfolding Field_def by auto lemma homo_rel_restrict_mono: "R \ B \ B \ rel_restrict R A \ (B - A) \ (B - A)" proof - assume A: "R \ B \ B" hence "Field R \ B" unfolding Field_def by auto with Field_rel_restrict have "Field (rel_restrict R A) \ B - A" by (metis Diff_mono order_refl order_trans) with R_subset_Field show ?thesis by blast qed lemma rel_restrict_union: "rel_restrict R (A \ B) = rel_restrict (rel_restrict R A) B" unfolding rel_restrict_def by auto lemma rel_restrictI: "x \ R \ y \ R \ (x,y) \ E \ (x,y) \ rel_restrict E R" unfolding rel_restrict_def by auto lemma rel_restrict_lift: "(x,y) \ rel_restrict E R \ (x,y) \ E" unfolding rel_restrict_def by simp lemma rel_restrict_trancl_mem: "(a,b) \ (rel_restrict A R)\<^sup>+ \ (a,b) \ rel_restrict (A\<^sup>+) R" by (induction rule: trancl_induct) (auto simp add: rel_restrict_def) lemma rel_restrict_trancl_sub: "(rel_restrict A R)\<^sup>+ \ rel_restrict (A\<^sup>+) R" by (metis subrelI rel_restrict_trancl_mem) lemma rel_restrict_mono: "A \ B \ rel_restrict A R \ rel_restrict B R" unfolding rel_restrict_def by auto lemma rel_restrict_mono2: "R \ S \ rel_restrict A S \ rel_restrict A R" unfolding rel_restrict_def by auto lemma rel_restrict_Sigma_sub: "rel_restrict ((A\A)\<^sup>+) R \ ((A - R) \ (A - R))\<^sup>+" unfolding rel_restrict_def by auto (metis DiffI converse_tranclE mem_Sigma_iff r_into_trancl tranclE) lemma finite_reachable_restrictedI: assumes F: "finite Q" assumes I: "I\Q" assumes R: "Range E \ Q" shows "finite (E\<^sup>*``I)" proof - from I R have "E\<^sup>*``I \ Q" by (force elim: rtrancl.cases) also note F finally (finite_subset) show ?thesis . qed context begin private lemma rtrancl_restrictI_aux: assumes "(u,v)\(E-UNIV\R)\<^sup>*" assumes "u\R" shows "(u,v)\(rel_restrict E R)\<^sup>* \ v\R" using assms by (induction) (auto simp: rel_restrict_def intro: rtrancl.intros) corollary rtrancl_restrictI: assumes "(u,v)\(E-UNIV\R)\<^sup>*" assumes "u\R" shows "(u,v)\(rel_restrict E R)\<^sup>*" using rtrancl_restrictI_aux[OF assms] .. end lemma E_closed_restr_reach_cases: assumes P: "(u,v)\E\<^sup>*" assumes CL: "E``R \ R" obtains "v\R" | "u\R" "(u,v)\(rel_restrict E R)\<^sup>*" using P proof (cases rule: rtrancl_last_visit[where S=R]) case no_visit show ?thesis proof (cases "u\R") case True with P have "v\R" using rtrancl_reachable_induct[OF _ CL, where I="{u}"] by auto thus ?thesis .. next case False with no_visit have "(u,v)\(rel_restrict E R)\<^sup>*" by (rule rtrancl_restrictI) with False show ?thesis .. qed next case (last_visit_point x) from \(x, v) \ (E - UNIV \ R)\<^sup>*\ have "(x,v)\E\<^sup>*" by (rule rtrancl_mono_mp[rotated]) auto with \x\R\ have "v\R" using rtrancl_reachable_induct[OF _ CL, where I="{x}"] by auto thus ?thesis .. qed lemma rel_restrict_trancl_notR: assumes "(v,w) \ (rel_restrict E R)\<^sup>+" shows "v \ R" and "w \ R" using assms by (metis rel_restrict_trancl_mem rel_restrict_notR)+ lemma rel_restrict_tranclI: assumes "(x,y) \ E\<^sup>+" and "x \ R" "y \ R" and "E `` R \ R" shows "(x,y) \ (rel_restrict E R)\<^sup>+" using assms proof (induct) case base thus ?case by (metis r_into_trancl rel_restrictI) next case (step y z) hence "y \ R" by auto with step show ?case by (metis trancl_into_trancl rel_restrictI) qed subsubsection \Single-Valued Relations\ lemma single_valued_inter1: "single_valued R \ single_valued (R\S)" by (auto intro: single_valuedI dest: single_valuedD) lemma single_valued_inter2: "single_valued R \ single_valued (S\R)" by (auto intro: single_valuedI dest: single_valuedD) lemma single_valued_below_Id: "R\Id \ single_valued R" by (auto intro: single_valuedI) subsubsection \Bijective Relations\ definition "bijective R \ (\x y z. (x,y)\R \ (x,z)\R \ y=z) \ (\x y z. (x,z)\R \ (y,z)\R \ x=y)" lemma bijective_alt: "bijective R \ single_valued R \ single_valued (R\)" unfolding bijective_def single_valued_def by blast lemma bijective_Id[simp, intro!]: "bijective Id" by (auto simp: bijective_def) lemma bijective_Empty[simp, intro!]: "bijective {}" by (auto simp: bijective_def) subsubsection \Miscellaneous\ lemma pair_vimage_is_Image[simp]: "(Pair u -` E) = E``{u}" by auto lemma fst_in_Field: "fst ` R \ Field R" by (simp add: Field_def fst_eq_Domain) lemma snd_in_Field: "snd ` R \ Field R" by (simp add: Field_def snd_eq_Range) lemma ran_map_of: "ran (map_of xs) \ snd ` set (xs)" by (induct xs) (auto simp add: ran_def) lemma Image_subset_snd_image: "A `` B \ snd ` A" unfolding Image_def image_def by force lemma finite_Image_subset: "finite (A `` B) \ C \ A \ finite (C `` B)" by (metis Image_mono order_refl rev_finite_subset) lemma finite_Field_eq_finite[simp]: "finite (Field R) \ finite R" by (metis finite_cartesian_product finite_subset R_subset_Field finite_Field) definition "fun_of_rel R x \ SOME y. (x,y)\R" lemma for_in_RI(*[intro]*): "x\Domain R \ (x,fun_of_rel R x)\R" unfolding fun_of_rel_def by (auto intro: someI) lemma Field_not_elem: "v \ Field R \ \(x,y) \ R. x \ v \ y \ v" unfolding Field_def by auto lemma Sigma_UNIV_cancel[simp]: "(A \ X - A \ UNIV) = {}" by auto lemma same_fst_trancl[simp]: "(same_fst P R)\<^sup>+ = same_fst P (\x. (R x)\<^sup>+)" proof - { fix x y assume "(x,y)\(same_fst P R)\<^sup>+" hence "(x,y)\same_fst P (\x. (R x)\<^sup>+)" by induction (auto simp: same_fst_def) } moreover { fix f f' x y assume "((f,x),(f',y))\same_fst P (\x. (R x)\<^sup>+)" hence [simp]: "f'=f" "P f" and 1: "(x,y)\(R f)\<^sup>+" by (auto simp: same_fst_def) from 1 have "((f,x),(f',y))\(same_fst P R)\<^sup>+" apply induction subgoal by (rule r_into_trancl) auto subgoal by (erule trancl_into_trancl) auto done } ultimately show ?thesis by auto qed subsection \\option\ Datatype\ lemma le_some_optE: "\Some m\x; !!m'. \x=Some m'; m\m'\ \ P\ \ P" by (cases x) auto lemma not_Some_eq2[simp]: "(\x y. v \ Some (x,y)) = (v = None)" by (cases v) auto subsection "Maps" primrec the_default where "the_default _ (Some x) = x" | "the_default x None = x" declare map_add_dom_app_simps[simp] declare map_add_upd_left[simp] lemma ran_add[simp]: "dom f \ dom g = {} \ ran (f++g) = ran f \ ran g" by (fastforce simp add: ran_def map_add_def split: option.split_asm option.split) lemma nempty_dom: "\e\Map.empty; !!m. m\dom e \ P \ \ P" by (subgoal_tac "dom e \ {}") (blast, auto) lemma le_map_dom_mono: "m\m' \ dom m \ dom m'" apply (safe) apply (drule_tac x=x in le_funD) apply simp apply (erule le_some_optE) apply simp done lemma map_add_first_le: fixes m::"'a\('b::order)" shows "\ m\m' \ \ m++n \ m'++n" apply (rule le_funI) apply (auto simp add: map_add_def split: option.split elim: le_funE) done lemma map_add_distinct_le: shows "\ m\m'; n\n'; dom m' \ dom n' = {} \ \ m++n \ m'++n'" apply (rule le_funI) apply (auto simp add: map_add_def split: option.split) apply (fastforce elim: le_funE) apply (drule le_map_dom_mono) apply (drule le_map_dom_mono) apply (case_tac "m x") apply simp apply (force) apply (fastforce dest!: le_map_dom_mono) apply (erule le_funE) apply (erule_tac x=x in le_funE) apply simp done lemma map_add_left_comm: assumes A: "dom A \ dom B = {}" shows "A ++ (B ++ C) = B ++ (A ++ C)" proof - have "A ++ (B ++ C) = (A++B)++C" by simp also have "\ = (B++A)++C" by (simp add: map_add_comm[OF A]) also have "\ = B++(A++C)" by simp finally show ?thesis . qed lemmas map_add_ac = map_add_assoc map_add_comm map_add_left_comm lemma le_map_restrict[simp]: fixes m :: "'a \ ('b::order)" shows "m |` X \ m" by (rule le_funI) (simp add: restrict_map_def) lemma map_of_distinct_upd: "x \ set (map fst xs) \ [x \ y] ++ map_of xs = (map_of xs) (x \ y)" by (induct xs) (auto simp add: fun_upd_twist) lemma map_of_distinct_upd2: assumes "x \ set(map fst xs)" "x \ set (map fst ys)" shows "map_of (xs @ (x,y) # ys) = (map_of (xs @ ys))(x \ y)" apply(insert assms) apply(induct xs) apply (auto intro: ext) done lemma map_of_distinct_upd3: assumes "x \ set(map fst xs)" "x \ set (map fst ys)" shows "map_of (xs @ (x,y) # ys) = (map_of (xs @ (x,y') # ys))(x \ y)" apply(insert assms) apply(induct xs) apply (auto intro: ext) done lemma map_of_distinct_upd4: assumes "x \ set(map fst xs)" "x \ set (map fst ys)" shows "map_of (xs @ ys) = (map_of (xs @ (x,y) # ys))(x := None)" using assms by (induct xs) (auto simp: map_of_eq_None_iff) lemma map_of_distinct_lookup: assumes "x \ set(map fst xs)" "x \ set (map fst ys)" shows "map_of (xs @ (x,y) # ys) x = Some y" proof - have "map_of (xs @ (x,y) # ys) = (map_of (xs @ ys)) (x \ y)" using assms map_of_distinct_upd2 by simp thus ?thesis by simp qed lemma ran_distinct: assumes dist: "distinct (map fst al)" shows "ran (map_of al) = snd ` set al" using assms proof (induct al) case Nil then show ?case by simp next case (Cons kv al) then have "ran (map_of al) = snd ` set al" by simp moreover from Cons.prems have "map_of al (fst kv) = None" by (simp add: map_of_eq_None_iff) ultimately show ?case by (simp only: map_of.simps ran_map_upd) simp qed lemma ran_is_image: "ran M = (the \ M) ` (dom M)" unfolding ran_def dom_def image_def by auto lemma map_card_eq_iff: assumes finite: "finite (dom M)" and card_eq: "card (dom M) = card (ran M)" and indom: "x \ dom M" shows "(M x = M y) \ (x = y)" proof - from ran_is_image finite card_eq have *: "inj_on (the \ M) (dom M)" using eq_card_imp_inj_on by metis thus ?thesis proof (cases "y \ dom M") case False with indom show ?thesis by auto next case True with indom have "the (M x) = the (M y) \ (x = y)" using inj_on_eq_iff[OF *] by auto thus ?thesis by auto qed qed lemma map_dom_ran_finite: "finite (dom M) \ finite (ran M)" by (simp add: ran_is_image) lemma map_update_eta_repair[simp]: (* An update operation may get simplified, if it happens to be eta-expanded. This lemma tries to repair some common expressions *) "dom (\x. if x=k then Some v else m x) = insert k (dom m)" "m k = None \ ran (\x. if x=k then Some v else m x) = insert v (ran m)" apply auto [] apply (force simp: ran_def) done lemma map_leI[intro?]: "\\x v. m1 x = Some v \ m2 x = Some v\ \ m1\\<^sub>mm2" unfolding map_le_def by force lemma map_leD: "m1\\<^sub>mm2 \ m1 k = Some v \ m2 k = Some v" unfolding map_le_def by force lemma map_restrict_insert_none_simp: "m x = None \ m|`(-insert x s) = m|`(-s)" by (auto intro!: ext simp:restrict_map_def) (* TODO: Move *) lemma eq_f_restr_conv: "s\dom (f A) \ A = f A |` (-s) \ A \\<^sub>m f A \ s = dom (f A) - dom A" apply auto subgoal by (metis map_leI restrict_map_eq(2)) subgoal by (metis ComplD restrict_map_eq(2)) subgoal by (metis Compl_iff restrict_in) subgoal by (force simp: map_le_def restrict_map_def) done corollary eq_f_restr_ss_eq: "\ s\dom (f A) \ \ A = f A |` (-s) \ A \\<^sub>m f A \ s = dom (f A) - dom A" using eq_f_restr_conv by blast subsubsection \Simultaneous Map Update\ definition "map_mmupd m K v k \ if k\K then Some v else m k" lemma map_mmupd_empty[simp]: "map_mmupd m {} v = m" by (auto simp: map_mmupd_def) lemma mmupd_in_upd[simp]: "k\K \ map_mmupd m K v k = Some v" by (auto simp: map_mmupd_def) lemma mmupd_notin_upd[simp]: "k\K \ map_mmupd m K v k = m k" by (auto simp: map_mmupd_def) lemma map_mmupdE: assumes "map_mmupd m K v k = Some x" obtains "k\K" "m k = Some x" | "k\K" "x=v" using assms by (auto simp: map_mmupd_def split: if_split_asm) lemma dom_mmupd[simp]: "dom (map_mmupd m K v) = dom m \ K" by (auto simp: map_mmupd_def split: if_split_asm) lemma le_map_mmupd_not_dom[simp, intro!]: "m \\<^sub>m map_mmupd m (K-dom m) v" by (auto simp: map_le_def) lemma map_mmupd_update_less: "K\K' \ map_mmupd m (K - dom m) v \\<^sub>m map_mmupd m (K'-dom m) v" by (auto simp: map_le_def map_mmupd_def) subsection\Connection between Maps and Sets of Key-Value Pairs\ definition map_to_set where "map_to_set m = {(k, v) . m k = Some v}" definition set_to_map where "set_to_map S k = Eps_Opt (\v. (k, v) \ S)" lemma set_to_map_simp : assumes inj_on_fst: "inj_on fst S" shows "(set_to_map S k = Some v) \ (k, v) \ S" proof (cases "\v. (k, v) \ S") case True note kv_ex = this then obtain v' where kv'_in: "(k, v') \ S" by blast with inj_on_fst have kv''_in: "\v''. (k, v'') \ S \ v' = v''" unfolding inj_on_def Ball_def by auto show ?thesis unfolding set_to_map_def by (simp add: kv_ex kv''_in) next case False hence kv''_nin: "\v''. (k, v'') \ S" by simp thus ?thesis by (simp add: set_to_map_def) qed lemma inj_on_fst_map_to_set : "inj_on fst (map_to_set m)" unfolding map_to_set_def inj_on_def by simp lemma map_to_set_inverse : "set_to_map (map_to_set m) = m" proof fix k show "set_to_map (map_to_set m) k = m k" proof (cases "m k") case None note mk_eq = this hence "\v. (k, v) \ map_to_set m" unfolding map_to_set_def by simp with set_to_map_simp [OF inj_on_fst_map_to_set, of m k] show ?thesis unfolding mk_eq by auto next case (Some v) note mk_eq = this hence "(k, v) \ map_to_set m" unfolding map_to_set_def by simp with set_to_map_simp [OF inj_on_fst_map_to_set, of m k v] show ?thesis unfolding mk_eq by auto qed qed lemma set_to_map_inverse : assumes inj_on_fst_S: "inj_on fst S" shows "map_to_set (set_to_map S) = S" proof (rule set_eqI) fix kv from set_to_map_simp [OF inj_on_fst_S, of "fst kv" "snd kv"] show "(kv \ map_to_set (set_to_map S)) = (kv \ S)" unfolding map_to_set_def by auto qed lemma map_to_set_empty[simp]: "map_to_set Map.empty = {}" unfolding map_to_set_def by simp lemma set_to_map_empty[simp]: "set_to_map {} = Map.empty" unfolding set_to_map_def[abs_def] by simp lemma map_to_set_empty_iff: "map_to_set m = {} \ m = Map.empty" "{} = map_to_set m \ m = Map.empty" unfolding map_to_set_def by auto lemma set_to_map_empty_iff: "set_to_map S = Map.empty \ S = {}" (is ?T1) "Map.empty = set_to_map S \ S = {}" (is ?T2) proof - show T1: ?T1 apply (simp only: set_eq_iff) apply (simp only: fun_eq_iff) apply (simp add: set_to_map_def) apply auto done from T1 show ?T2 by auto qed lemma map_to_set_upd[simp]: "map_to_set (m (k \ v)) = insert (k, v) (map_to_set m - {(k, v') |v'. True})" unfolding map_to_set_def apply (simp add: set_eq_iff) apply metis done lemma set_to_map_insert: assumes k_nin: "fst kv \ fst ` S" shows "set_to_map (insert kv S) = (set_to_map S) (fst kv \ snd kv)" proof fix k' obtain k v where kv_eq[simp]: "kv = (k, v)" by (rule prod.exhaust) from k_nin have k_nin': "\v'. (k, v') \ S" by (auto simp add: image_iff Ball_def) - show "set_to_map (insert kv S) k' = (set_to_map S(fst kv \ snd kv)) k'" + show "set_to_map (insert kv S) k' = ((set_to_map S)(fst kv \ snd kv)) k'" by (simp add: set_to_map_def k_nin') qed lemma map_to_set_dom : "dom m = fst ` (map_to_set m)" unfolding dom_def map_to_set_def by (auto simp add: image_iff) lemma map_to_set_ran : "ran m = snd ` (map_to_set m)" unfolding ran_def map_to_set_def by (auto simp add: image_iff) lemma set_to_map_dom : "dom (set_to_map S) = fst ` S" unfolding set_to_map_def[abs_def] dom_def by (auto simp add: image_iff Bex_def) lemma set_to_map_ran : "ran (set_to_map S) \ snd ` S" unfolding set_to_map_def[abs_def] ran_def subset_iff by (auto simp add: image_iff Bex_def) (metis Eps_Opt_eq_Some) lemma finite_map_to_set: "finite (map_to_set m) = finite (dom m)" unfolding map_to_set_def map_to_set_dom apply (intro iffI finite_imageI) apply assumption apply (rule finite_imageD[of fst]) apply assumption apply (simp add: inj_on_def) done lemma card_map_to_set : "card (map_to_set m) = card (dom m)" unfolding map_to_set_def map_to_set_dom apply (rule card_image[symmetric]) apply (simp add: inj_on_def) done lemma map_of_map_to_set : "distinct (map fst l) \ map_of l = m \ set l = map_to_set m" proof (induct l arbitrary: m) case Nil thus ?case by (simp add: map_to_set_empty_iff) blast next case (Cons kv l m) obtain k v where kv_eq[simp]: "kv = (k, v)" by (rule prod.exhaust) from Cons(2) have dist_l: "distinct (map fst l)" and kv'_nin: "\v'. (k, v') \ set l" by (auto simp add: image_iff) note ind_hyp = Cons(1)[OF dist_l] from kv'_nin have l_eq: "set (kv # l) = map_to_set m \ (set l = map_to_set (m (k := None))) \ m k = Some v" apply (simp add: map_to_set_def restrict_map_def set_eq_iff) apply (auto) apply (metis) apply (metis option.inject) done from kv'_nin have m_eq: "map_of (kv # l) = m \ map_of l = (m (k := None)) \ m k = Some v" apply (simp add: fun_eq_iff restrict_map_def map_of_eq_None_iff image_iff Ball_def) apply metis done show ?case unfolding m_eq l_eq using ind_hyp[of "m (k := None)"] by metis qed lemma map_to_set_map_of : "distinct (map fst l) \ map_to_set (map_of l) = set l" by (metis map_of_map_to_set) subsubsection \Mapping empty set to None\ definition "dflt_None_set S \ if S={} then None else Some S" lemma the_dflt_None_empty[simp]: "dflt_None_set {} = None" unfolding dflt_None_set_def by simp lemma the_dflt_None_nonempty[simp]: "S\{} \ dflt_None_set S = Some S" unfolding dflt_None_set_def by simp lemma the_dflt_None_set[simp]: "the_default {} (dflt_None_set x) = x" unfolding dflt_None_set_def by auto subsection \Orderings\ lemma (in order) min_arg_le[simp]: "n \ min m n \ min m n = n" "m \ min m n \ min m n = m" by (auto simp: min_def) lemma (in linorder) min_arg_not_ge[simp]: "\ min m n < m \ min m n = m" "\ min m n < n \ min m n = n" by (auto simp: min_def) lemma (in linorder) min_eq_arg[simp]: "min m n = m \ m\n" "min m n = n \ n\m" by (auto simp: min_def) lemma min_simps[simp]: "a<(b::'a::order) \ min a b = a" "b<(a::'a::order) \ min a b = b" by (auto simp add: min_def dest: less_imp_le) lemma (in -) min_less_self_conv[simp]: "min a b < a \ b < (a::_::linorder)" "min a b < b \ a < (b::_::linorder)" by (auto simp: min_def) lemma ord_eq_le_eq_trans: "\ a=b; b\c; c=d \ \ a\d" by auto lemma zero_comp_diff_simps[simp]: "(0::'a::linordered_idom) \ a - b \ b \ a" "(0::'a::linordered_idom) < a - b \ b < a" by auto subsubsection \Termination Measures\ text \Lexicographic measure, assuming upper bound for second component\ lemma mlex_fst_decrI: fixes a a' b b' N :: nat assumes "a a*N + N" using \b by linarith also have "\ \ a'*N" using \a by (metis Suc_leI ab_semigroup_add_class.add.commute ab_semigroup_mult_class.mult.commute mult_Suc_right mult_le_mono2) also have "\ \ a'*N + b'" by auto finally show ?thesis by auto qed lemma mlex_leI: fixes a a' b b' N :: nat assumes "a\a'" assumes "b\b'" shows "a*N + b \ a'*N + b'" using assms by (auto intro!: add_mono) lemma mlex_snd_decrI: fixes a a' b b' N :: nat assumes "a=a'" assumes "bCCPOs\ context ccpo begin lemma ccpo_Sup_mono: assumes C: "Complete_Partial_Order.chain (\) A" "Complete_Partial_Order.chain (\) B" assumes B: "\x\A. \y\B. x\y" shows "Sup A \ Sup B" proof (rule ccpo_Sup_least) fix x assume "x\A" with B obtain y where I: "y\B" and L: "x\y" by blast note L also from I ccpo_Sup_upper have "y\Sup B" by (blast intro: C) finally show "x\Sup B" . qed (rule C) lemma fixp_mono: assumes M: "monotone (\) (\) f" "monotone (\) (\) g" assumes LE: "\Z. f Z \ g Z" shows "ccpo_class.fixp f \ ccpo_class.fixp g" unfolding fixp_def[abs_def] apply (rule ccpo_Sup_mono) apply (rule chain_iterates M)+ proof rule fix x assume "x\ccpo_class.iterates f" thus "\y\ccpo_class.iterates g. x\y" proof (induct) case (step x) then obtain y where I: "y\ccpo_class.iterates g" and L: "x\y" by blast hence "g y \ ccpo_class.iterates g" and "f x \ g y" apply - apply (erule iterates.step) apply (rule order_trans) apply (erule monotoneD[OF M(1)]) apply (rule LE) done thus "\y\ccpo_class.iterates g. f x \ y" .. next case (Sup M) define N where "N = {SOME y. y\ccpo_class.iterates g \ x\y | x. x\M}" have N1: "\y\N. y\ccpo_class.iterates g \ (\x\M. x\y)" unfolding N_def apply auto apply (metis (lifting) Sup.hyps(2) tfl_some) by (metis (lifting) Sup.hyps(2) tfl_some) have N2: "\x\M. \y\N. x\y" unfolding N_def apply auto by (metis (lifting) Sup.hyps(2) tfl_some) have "N \ ccpo_class.iterates g" using Sup using N1 by auto hence C_chain: "Complete_Partial_Order.chain (\) N" using chain_iterates[OF M(2)] unfolding chain_def by auto have "Sup N \ ccpo_class.iterates g" and "Sup M \ Sup N" apply - apply (rule iterates.Sup[OF C_chain]) using N1 apply blast apply (rule ccpo_Sup_mono) apply (rule Sup.hyps) apply (rule C_chain) apply (rule N2) done thus ?case by blast qed qed end subsection \Code\ text \Constant for code-abort. If that gets executed, an abort-exception is raised.\ definition [simp]: "CODE_ABORT f = f ()" declare [[code abort: CODE_ABORT]] end diff --git a/thys/Concurrent_Revisions/Renaming.thy b/thys/Concurrent_Revisions/Renaming.thy --- a/thys/Concurrent_Revisions/Renaming.thy +++ b/thys/Concurrent_Revisions/Renaming.thy @@ -1,292 +1,292 @@ section Renaming text \Similar to the bound variables of lambda calculus, location and revision identifiers are meaningless names. This theory contains all of the definitions and results required for renaming data structures and proving renaming-equivalence.\ theory Renaming imports Occurrences begin subsection Definitions abbreviation rename_val :: "('r \ 'r) \ ('l \ 'l) \ ('r,'l,'v) val \ ('r,'l,'v) val" ("\\<^sub>V") where "\\<^sub>V \ \ v \ map_val \ \ id v" abbreviation rename_expr :: "('r \ 'r) \ ('l \ 'l) \ ('r,'l,'v) expr \ ('r,'l,'v) expr" ("\\<^sub>E") where "\\<^sub>E \ \ e \ map_expr \ \ id e" abbreviation rename_cntxt :: "('r \ 'r) \ ('l \ 'l) \ ('r,'l,'v) cntxt \ ('r,'l,'v) cntxt" ("\\<^sub>C") where "\\<^sub>C \ \ \ \ map_cntxt \ \ id \" definition is_store_renaming :: "('r \ 'r) \ ('l \ 'l) \ ('r,'l,'v) store \ ('r,'l,'v) store \ bool" where "is_store_renaming \ \ \ \' \ \l. case \ l of None \ \' (\ l) = None | Some v \ \' (\ l) = Some (\\<^sub>V \ \ v)" notation Option.bind (infixl "\" 80) fun \\<^sub>S :: "('r \ 'r) \ ('l \ 'l) \ ('r,'l,'v) store \ ('r,'l,'v) store" where "\\<^sub>S \ \ \ l = \ (inv \ l) \ (\v. Some (\\<^sub>V \ \ v))" lemma \\<^sub>S_implements_renaming: "bij \ \ is_store_renaming \ \ \ (\\<^sub>S \ \ \)" proof - assume "bij \" hence "inj \" using bij_def by auto thus ?thesis by (auto simp add: is_store_renaming_def option.case_eq_if) qed fun \\<^sub>L :: "('r \ 'r) \ ('l \ 'l) \ ('r,'l,'v) local_state \ ('r,'l,'v) local_state" where "\\<^sub>L \ \ (\,\,e) = (\\<^sub>S \ \ \, \\<^sub>S \ \ \, \\<^sub>E \ \ e)" definition is_global_renaming :: "('r \ 'r) \ ('l \ 'l) \ ('r,'l,'v) global_state \ ('r,'l,'v) global_state \ bool" where "is_global_renaming \ \ s s' \ \r. case s r of None \ s' (\ r) = None | Some ls \ s' (\ r) = Some (\\<^sub>L \ \ ls)" fun \\<^sub>G :: "('r \ 'r) \ ('l \ 'l) \ ('r,'l,'v) global_state \ ('r,'l,'v) global_state" where "\\<^sub>G \ \ s r = s (inv \ r) \ (\ls. Some (\\<^sub>L \ \ ls))" lemma \\<^sub>G_implements_renaming: "bij \ \ is_global_renaming \ \ s (\\<^sub>G \ \ s)" proof - assume "bij \" hence "inj \" using bij_def by auto thus ?thesis by (auto simp add: is_global_renaming_def option.case_eq_if) qed subsection \Introduction rules\ lemma \\<^sub>SI [intro]: assumes bij_\: "bij \" and none_case: "\l. \ l = None \ \' (\ l) = None" and some_case: "\l v. \ l = Some v \ \' (\ l) = Some (\\<^sub>V \ \ v)" shows "\\<^sub>S \ \ \ = \'" proof (rule ext, subst \\<^sub>S.simps) fix l show "\ (inv \ l) \ (\v. Some (\\<^sub>V \ \ v)) = \' l" (is "?lhs = \' l") proof (cases "\ (inv \ l) = None") case True have lhs_none: "?lhs = None" by (simp add: True) have "\' (\ (inv \ l)) = None" by (simp add: none_case True) hence rhs_none: "\' l = None" by (simp add: bij_\ bijection.intro bijection.inv_right) show ?thesis by (simp add: lhs_none rhs_none) next case False from this obtain v where is_some: "\ (inv \ l) = Some v" by blast hence lhs_some: "?lhs = Some (\\<^sub>V \ \ v)" by auto have "\' (\ (inv \ l)) = Some (\\<^sub>V \ \ v)" by (simp add: is_some some_case) hence rhs_some: "\' l = Some (\\<^sub>V \ \ v)" by (simp add: bij_\ bijection.intro bijection.inv_right) then show ?thesis by (simp add: lhs_some) qed qed lemma \\<^sub>GI [intro]: assumes bij_\: "bij \" and none_case: "\r. s r = None \ s' (\ r) = None" and some_case: "\r \ \ e. s r = Some (\,\,e) \ s' (\ r) = Some (\\<^sub>L \ \ (\,\,e))" shows "\\<^sub>G \ \ s = s'" proof (rule ext, subst \\<^sub>G.simps) fix r show "s (inv \ r) \ (\ls. Some (\\<^sub>L \ \ ls)) = s' r" (is "?lhs = s' r") proof (cases "s (inv \ r) = None") case True have lhs_none: "?lhs = None" by (simp add: True) have "s' (\ (inv \ r)) = None" by (simp add: none_case True) hence rhs_none: "s' r = None" by (simp add: bij_\ bijection.intro bijection.inv_right) show ?thesis by (simp add: lhs_none rhs_none) next case False from this obtain ls where "s (inv \ r) = Some ls" by blast from this obtain \ \ e where is_some: "s (inv \ r) = Some (\, \, e)" by (cases ls) blast hence lhs_some: "?lhs = Some (\\<^sub>L \ \ (\, \, e))" by auto have "s' (\ (inv \ r)) = Some (\\<^sub>L \ \ (\, \, e))" by (simp add: is_some some_case) hence rhs_some: "s' r = Some (\\<^sub>L \ \ (\, \, e))" by (simp add: bij_\ bijection.intro bijection.inv_right) then show ?thesis by (simp add: lhs_some) qed qed subsection \Renaming-equivalence\ subsubsection Identity declare val.map_id [simp] declare expr.map_id [simp] declare cntxt.map_id [simp] lemma \\<^sub>S_id [simp]: "\\<^sub>S id id \ = \" by auto lemma \\<^sub>L_id [simp]: "\\<^sub>L id id ls = ls" by (cases ls) simp lemma \\<^sub>G_id [simp]: "\\<^sub>G id id s = s" by auto subsubsection Composition declare val.map_comp [simp] declare expr.map_comp [simp] declare cntxt.map_comp [simp] lemma \\<^sub>S_comp [simp]: "\ bij \; bij \' \ \ \\<^sub>S \' \' (\\<^sub>S \ \ s) = \\<^sub>S (\' \ \) (\' \ \) s" by (auto simp add: o_inv_distrib) lemma \\<^sub>L_comp [simp]: "\ bij \; bij \' \ \ \\<^sub>L \' \' (\\<^sub>L \ \ ls) = \\<^sub>L (\' \ \) (\' \ \) ls" by (cases ls) simp lemma \\<^sub>G_comp [simp]: "\ bij \; bij \'; bij \; bij \' \ \ \\<^sub>G \' \' (\\<^sub>G \ \ s) = \\<^sub>G (\' \ \) (\' \ \) s" by (rule ext) (auto simp add: o_inv_distrib) subsubsection Inverse lemma \\<^sub>V_inv [simp]: "\ bij \; bij \ \ \ (\\<^sub>V (inv \) (inv \) v' = v) = (\\<^sub>V \ \ v = v')" by (auto simp add: bijection.intro bijection.inv_comp_right bijection.inv_comp_left) lemma \\<^sub>E_inv [simp]: "\ bij \; bij \ \ \ (\\<^sub>E (inv \) (inv \) e' = e) = (\\<^sub>E \ \ e = e')" by (auto simp add: bijection.intro bijection.inv_comp_right bijection.inv_comp_left) lemma \\<^sub>C_inv [simp]: "\ bij \; bij \ \ \ (\\<^sub>C (inv \) (inv \) \' = \) = (\\<^sub>C \ \ \ = \')" by (auto simp add: bijection.intro bijection.inv_comp_right bijection.inv_comp_left) lemma \\<^sub>S_inv [simp]: "\ bij \; bij \ \ \ (\\<^sub>S (inv \) (inv \) \' = \) = (\\<^sub>S \ \ \ = \')" by (auto simp add: bij_imp_bij_inv bijection.intro bijection.inv_comp_right bijection.inv_comp_left) lemma \\<^sub>L_inv [simp]: "\ bij \; bij \ \ \ (\\<^sub>L (inv \) (inv \) ls' = ls) = (\\<^sub>L \ \ ls = ls')" by (auto simp add: bij_imp_bij_inv bijection.intro bijection.inv_comp_right bijection.inv_comp_left) lemma \\<^sub>G_inv [simp]: "\ bij \; bij \ \ \ (\\<^sub>G (inv \) (inv \) s' = s) = (\\<^sub>G \ \ s = s')" by (auto simp add: bij_imp_bij_inv bijection.intro bijection.inv_comp_right bijection.inv_comp_left) subsubsection Equivalence definition eq_states :: "('r,'l,'v) global_state \ ('r,'l,'v) global_state \ bool" ("_ \ _" [100, 100]) where "s \ s' \ \\ \. bij \ \ bij \ \ \\<^sub>G \ \ s = s'" lemma eq_statesI [intro]: "\\<^sub>G \ \ s = s' \ bij \ \ bij \ \ s \ s'" using eq_states_def by auto lemma eq_statesE [elim]: "s \ s' \ (\\ \. \\<^sub>G \ \ s = s' \ bij \ \ bij \ \ P) \ P" using eq_states_def by blast lemma \\_refl: "s \ s" by (rule eq_statesI[of id id s]) auto lemma \\_trans: "s \ s' \ s' \ s'' \ s \ s''" proof - assume "s \ s'" from this obtain \ \ where s_s': "bij \" "bij \" "\\<^sub>G \ \ s = s'" by blast assume "s' \ s''" from this obtain \' \' where s'_s'': "bij \'" "bij \'" "\\<^sub>G \' \' s' = s''" by blast show "s \ s''" by (rule eq_statesI[of "\' \ \" "\' \ \"]) (use s_s' s'_s'' in \auto simp add: bij_comp\) qed lemma \\_sym: "s \ s' \ s' \ s" proof - assume "s \ s'" from this obtain \ \ where s_s': "bij \" "bij \" "\\<^sub>G \ \ s = s'" by blast show "s' \ s" by (rule eq_statesI[of "inv \" "inv \"]) (use s_s' in \auto simp add: bij_imp_bij_inv\) qed subsection \Distributive laws\ subsubsection Expression lemma renaming_distr_completion [simp]: "\\<^sub>E \ \ (\[e]) = ((\\<^sub>C \ \ \)[\\<^sub>E \ \ e])" by (induct \) simp+ subsubsection Store lemma renaming_distr_combination [simp]: "\\<^sub>S \ \ (\;;\) = (\\<^sub>S \ \ \;;\\<^sub>S \ \ \)" by (rule ext) auto lemma renaming_distr_store [simp]: - "bij \ \ \\<^sub>S \ \ (\(l \ v)) = \\<^sub>S \ \ \(\ l \ \\<^sub>V \ \ v)" + "bij \ \ \\<^sub>S \ \ (\(l \ v)) = (\\<^sub>S \ \ \)(\ l \ \\<^sub>V \ \ v)" by (auto simp add: bijection.intro bijection.inv_left_eq_iff) (* distribution law for local follows from the definition *) subsubsection Global lemma renaming_distr_global [simp]: - "bij \ \ \\<^sub>G \ \ (s(r \ ls)) = \\<^sub>G \ \ s(\ r \ \\<^sub>L \ \ ls)" + "bij \ \ \\<^sub>G \ \ (s(r \ ls)) = (\\<^sub>G \ \ s)(\ r \ \\<^sub>L \ \ ls)" "bij \ \ \\<^sub>G \ \ (s(r := None)) = (\\<^sub>G \ \ s)(\ r := None)" by (auto simp add: bijection.intro bijection.inv_left_eq_iff) subsection \Miscellaneous laws\ lemma rename_empty [simp]: "\\<^sub>S \ \ \ = \" "\\<^sub>G \ \ \ = \" by auto subsection Swaps lemma swap_bij: "bij (id(x := x', x' := x))" (is "bij ?f") proof (rule bijI) show "inj ?f" by (simp add: inj_on_def) show "surj ?f" proof show "UNIV \ range (id(x := x', x' := x))" proof (rule subsetI) fix y assume "y \ (UNIV :: 'a set)" show "y \ range (id(x := x', x' := x))" by (cases "y = x"; cases "y = x'") auto qed qed simp qed lemma id_trivial_update [simp]: "id(x := x) = id" by auto (* for solving trivial peaks *) lemma eliminate_renaming_val_expr [simp]: fixes v :: "('r,'l,'v) val" and e :: "('r,'l,'v) expr" shows "l \ LID\<^sub>V v \ \\<^sub>V \ (\(l := l')) v = \\<^sub>V \ \ v" "l \ LID\<^sub>E e \ \\<^sub>E \ (\(l := l')) e = \\<^sub>E \ \ e" "r \ RID\<^sub>V v \ \\<^sub>V (\(r := r')) \ v = \\<^sub>V \ \ v" "r \ RID\<^sub>E e \ \\<^sub>E (\(r := r')) \ e = \\<^sub>E \ \ e" proof - have "(\\ \ r r'. r \ RID\<^sub>V v \ \\<^sub>V (\(r := r')) \ v = \\<^sub>V \ \ v) \ (\\ \ r r'. r \ RID\<^sub>E e \ \\<^sub>E (\(r := r')) \ e = \\<^sub>E \ \ e)" by (induct rule: val_expr.induct) simp+ thus "r \ RID\<^sub>V v \ \\<^sub>V (\(r := r')) \ v = \\<^sub>V \ \ v" "r \ RID\<^sub>E e \ \\<^sub>E (\(r := r')) \ e = \\<^sub>E \ \ e" by simp+ have "(\\ \ l l'. l \ LID\<^sub>V v \ \\<^sub>V \ (\(l := l')) v = \\<^sub>V \ \ v) \ (\\ \ l l'. l \ LID\<^sub>E e \ \\<^sub>E \ (\(l := l')) e = \\<^sub>E \ \ e)" by (induct rule: val_expr.induct) simp+ thus "l \ LID\<^sub>V v \ \\<^sub>V \ (\(l := l')) v = \\<^sub>V \ \ v" and "l \ LID\<^sub>E e \ \\<^sub>E \ (\(l := l')) e = \\<^sub>E \ \ e" by simp+ qed lemma eliminate_renaming_cntxt [simp]: "r \ RID\<^sub>C \ \ \\<^sub>C (\(r := r')) \ \ = \\<^sub>C \ \ \" "l \ LID\<^sub>C \ \ \\<^sub>C \ (\(l := l')) \ = \\<^sub>C \ \ \" by (induct \ rule: cntxt.induct) auto lemma eliminate_swap_val [simp, intro]: "r \ RID\<^sub>V v \ r' \ RID\<^sub>V v \ \\<^sub>V (id(r := r', r' := r)) id v = v" "l \ LID\<^sub>V v \ l' \ LID\<^sub>V v \ \\<^sub>V id (id(l := l', l' := l)) v = v" by simp+ lemma eliminate_swap_expr [simp, intro]: "r \ RID\<^sub>E e \ r' \ RID\<^sub>E e \ \\<^sub>E (id(r := r', r' := r)) id e = e" "l \ LID\<^sub>E e \ l' \ LID\<^sub>E e \ \\<^sub>E id (id(l := l', l' := l)) e = e" by simp+ lemma eliminate_swap_cntxt [simp, intro]: "r \ RID\<^sub>C \ \ r' \ RID\<^sub>C \ \ \\<^sub>C (id(r := r', r' := r)) id \ = \" "l \ LID\<^sub>C \ \ l' \ LID\<^sub>C \ \ \\<^sub>C id (id(l := l', l' := l)) \ = \" by simp+ lemma eliminate_swap_store_rid [simp, intro]: "r \ RID\<^sub>S \ \ r' \ RID\<^sub>S \ \ \\<^sub>S (id(r := r', r' := r)) id \ = \" by (rule \\<^sub>SI) (auto simp add: swap_bij RID\<^sub>S_def domIff ranI) lemma eliminate_swap_store_lid [simp, intro]: "l \ LID\<^sub>S \ \ l' \ LID\<^sub>S \ \ \\<^sub>S id (id(l := l', l' := l)) \ = \" by (rule \\<^sub>SI) (auto simp add: swap_bij LID\<^sub>S_def domIff ranI) lemma eliminate_swap_global_rid [simp, intro]: "r \ RID\<^sub>G s \ r' \ RID\<^sub>G s \ \\<^sub>G (id(r := r', r' := r)) id s = s" by (rule \\<^sub>GI[OF swap_bij], ((rule sym, auto)[1])+) lemma eliminate_swap_global_lid [simp, intro]: "l \ LID\<^sub>G s \ l' \ LID\<^sub>G s \ \\<^sub>G id (id(l := l', l' := l)) s = s" by (rule \\<^sub>GI) (auto simp add: ID_distr_global_conditional) end diff --git a/thys/Containers/RBT_ext.thy b/thys/Containers/RBT_ext.thy --- a/thys/Containers/RBT_ext.thy +++ b/thys/Containers/RBT_ext.thy @@ -1,358 +1,358 @@ (* Title: Containers/RBT_ext.thy Author: Andreas Lochbihler, KIT *) theory RBT_ext imports "HOL-Library.RBT_Impl" Containers_Auxiliary List_Fusion begin section \More on red-black trees\ subsection \More lemmas\ context linorder begin lemma is_rbt_fold_rbt_insert_impl: "is_rbt t \ is_rbt (RBT_Impl.fold rbt_insert t' t)" by(simp add: rbt_insert_def is_rbt_fold_rbt_insertwk) lemma rbt_sorted_fold_insert: "rbt_sorted t \ rbt_sorted (RBT_Impl.fold rbt_insert t' t)" by(induct t' arbitrary: t)(simp_all add: rbt_insert_rbt_sorted) -lemma rbt_lookup_rbt_insert': "rbt_sorted t \ rbt_lookup (rbt_insert k v t) = rbt_lookup t(k \ v)" +lemma rbt_lookup_rbt_insert': "rbt_sorted t \ rbt_lookup (rbt_insert k v t) = (rbt_lookup t)(k \ v)" by(simp add: rbt_insert_def rbt_lookup_rbt_insertwk fun_eq_iff split: option.split) lemma rbt_lookup_fold_rbt_insert_impl: "rbt_sorted t2 \ rbt_lookup (RBT_Impl.fold rbt_insert t1 t2) = rbt_lookup t2 ++ map_of (rev (RBT_Impl.entries t1))" proof(induction t1 arbitrary: t2) case Empty thus ?case by simp next case (Branch c l x k r) show ?case using Branch.prems by(simp add: map_add_def Branch.IH rbt_insert_rbt_sorted rbt_sorted_fold_insert rbt_lookup_rbt_insert' fun_eq_iff split: option.split) qed end subsection \Build the cross product of two RBTs\ context fixes f :: "'a \ 'b \ 'c \ 'd \ 'e" begin definition alist_product :: "('a \ 'b) list \ ('c \ 'd) list \ (('a \ 'c) \ 'e) list" where "alist_product xs ys = concat (map (\(a, b). map (\(c, d). ((a, c), f a b c d)) ys) xs)" lemma alist_product_simps [simp]: "alist_product [] ys = []" "alist_product xs [] = []" "alist_product ((a, b) # xs) ys = map (\(c, d). ((a, c), f a b c d)) ys @ alist_product xs ys" by(simp_all add: alist_product_def) lemma append_alist_product_conv_fold: "zs @ alist_product xs ys = rev (fold (\(a, b). fold (\(c, d) rest. ((a, c), f a b c d) # rest) ys) xs (rev zs))" proof(induction xs arbitrary: zs) case Nil thus ?case by simp next case (Cons x xs) obtain a b where x: "x = (a, b)" by(cases x) have "\zs. fold (\(c, d). (#) ((a, c), f a b c d)) ys zs = rev (map (\(c, d). ((a, c), f a b c d)) ys) @ zs" by(induct ys) auto with Cons.IH[of "zs @ map (\(c, d). ((a, c), f a b c d)) ys"] x show ?case by simp qed lemma alist_product_code [code]: "alist_product xs ys = rev (fold (\(a, b). fold (\(c, d) rest. ((a, c), f a b c d) # rest) ys) xs [])" using append_alist_product_conv_fold[of "[]" xs ys] by simp lemma set_alist_product: "set (alist_product xs ys) = (\((a, b), (c, d)). ((a, c), f a b c d)) ` (set xs \ set ys)" by(auto 4 3 simp add: alist_product_def intro: rev_image_eqI rev_bexI) lemma distinct_alist_product: "\ distinct (map fst xs); distinct (map fst ys) \ \ distinct (map fst (alist_product xs ys))" proof(induct xs) case Nil thus ?case by simp next case (Cons x xs) obtain a b where x: "x = (a, b)" by(cases x) have "distinct (map (fst \ (\(c, d). ((a, c), f a b c d))) ys)" using \distinct (map fst ys)\ by(induct ys)(auto intro: rev_image_eqI) thus ?case using Cons x by(auto simp add: set_alist_product intro: rev_image_eqI) qed lemma map_of_alist_product: "map_of (alist_product xs ys) (a, c) = (case map_of xs a of None \ None | Some b \ map_option (f a b c) (map_of ys c))" proof(induction xs) case Nil thus ?case by simp next case (Cons x xs) obtain a b where x: "x = (a, b)" by (cases x) let ?map = "map (\(c, d). ((a, c), f a b c d)) ys" have "map_of ?map (a, c) = map_option (f a b c) (map_of ys c)" by(induct ys) auto moreover { fix a' assume "a' \ a" hence "map_of ?map (a', c) = None" by(induct ys) auto } ultimately show ?case using x Cons.IH by(auto simp add: map_add_def split: option.split) qed definition rbt_product :: "('a, 'b) rbt \ ('c, 'd) rbt \ ('a \ 'c, 'e) rbt" where "rbt_product rbt1 rbt2 = rbtreeify (alist_product (RBT_Impl.entries rbt1) (RBT_Impl.entries rbt2))" lemma rbt_product_code [code]: "rbt_product rbt1 rbt2 = rbtreeify (rev (RBT_Impl.fold (\a b. RBT_Impl.fold (\c d rest. ((a, c), f a b c d) # rest) rbt2) rbt1 []))" unfolding rbt_product_def alist_product_code RBT_Impl.fold_def .. end context fixes leq_a :: "'a \ 'a \ bool" (infix "\\<^sub>a" 50) and less_a :: "'a \ 'a \ bool" (infix "\\<^sub>a" 50) and leq_b :: "'b \ 'b \ bool" (infix "\\<^sub>b" 50) and less_b :: "'b \ 'b \ bool" (infix "\\<^sub>b" 50) assumes lin_a: "class.linorder leq_a less_a" and lin_b: "class.linorder leq_b less_b" begin abbreviation (input) less_eq_prod' :: "('a \ 'b) \ ('a \ 'b) \ bool" (infix "\" 50) where "less_eq_prod' \ less_eq_prod leq_a less_a leq_b" abbreviation (input) less_prod' :: "('a \ 'b) \ ('a \ 'b) \ bool" (infix "\" 50) where "less_prod' \ less_prod leq_a less_a less_b" lemmas linorder_prod = linorder_prod[OF lin_a lin_b] lemma sorted_alist_product: assumes xs: "linorder.sorted leq_a (map fst xs)" "distinct (map fst xs)" and ys: "linorder.sorted (\\<^sub>b) (map fst ys)" shows "linorder.sorted (\) (map fst (alist_product f xs ys))" proof - interpret a: linorder "(\\<^sub>a)" "(\\<^sub>a)" by(fact lin_a) note [simp] = linorder.sorted0[OF linorder_prod] linorder.sorted1[OF linorder_prod] linorder.sorted_append[OF linorder_prod] linorder.sorted1[OF lin_b] show ?thesis using xs proof(induction xs) case Nil show ?case by simp next case (Cons x xs) obtain a b where x: "x = (a, b)" by(cases x) have "linorder.sorted (\) (map fst (map (\(c, d). ((a, c), f a b c d)) ys))" using ys by(induct ys) auto thus ?case using x Cons by(fastforce simp add: set_alist_product a.not_less dest: bspec a.order_antisym intro: rev_image_eqI) qed qed lemma is_rbt_rbt_product: "\ ord.is_rbt (\\<^sub>a) rbt1; ord.is_rbt (\\<^sub>b) rbt2 \ \ ord.is_rbt (\) (rbt_product f rbt1 rbt2)" unfolding rbt_product_def by(blast intro: linorder.is_rbt_rbtreeify[OF linorder_prod] sorted_alist_product linorder.rbt_sorted_entries[OF lin_a] ord.is_rbt_rbt_sorted linorder.distinct_entries[OF lin_a] linorder.rbt_sorted_entries[OF lin_b] distinct_alist_product linorder.distinct_entries[OF lin_b]) lemma rbt_lookup_rbt_product: "\ ord.is_rbt (\\<^sub>a) rbt1; ord.is_rbt (\\<^sub>b) rbt2 \ \ ord.rbt_lookup (\) (rbt_product f rbt1 rbt2) (a, c) = (case ord.rbt_lookup (\\<^sub>a) rbt1 a of None \ None | Some b \ map_option (f a b c) (ord.rbt_lookup (\\<^sub>b) rbt2 c))" by(simp add: rbt_product_def linorder.rbt_lookup_rbtreeify[OF linorder_prod] linorder.is_rbt_rbtreeify[OF linorder_prod] sorted_alist_product linorder.rbt_sorted_entries[OF lin_a] ord.is_rbt_rbt_sorted linorder.distinct_entries[OF lin_a] linorder.rbt_sorted_entries[OF lin_b] distinct_alist_product linorder.distinct_entries[OF lin_b] map_of_alist_product linorder.map_of_entries[OF lin_a] linorder.map_of_entries[OF lin_b] cong: option.case_cong) end hide_const less_eq_prod' less_prod' subsection \Build an RBT where keys are paired with themselves\ primrec RBT_Impl_diag :: "('a, 'b) rbt \ ('a \ 'a, 'b) rbt" where "RBT_Impl_diag rbt.Empty = rbt.Empty" | "RBT_Impl_diag (rbt.Branch c l k v r) = rbt.Branch c (RBT_Impl_diag l) (k, k) v (RBT_Impl_diag r)" lemma entries_RBT_Impl_diag: "RBT_Impl.entries (RBT_Impl_diag t) = map (\(k, v). ((k, k), v)) (RBT_Impl.entries t)" by(induct t) simp_all lemma keys_RBT_Impl_diag: "RBT_Impl.keys (RBT_Impl_diag t) = map (\k. (k, k)) (RBT_Impl.keys t)" by(simp add: RBT_Impl.keys_def entries_RBT_Impl_diag split_beta) lemma rbt_sorted_RBT_Impl_diag: "ord.rbt_sorted lt t \ ord.rbt_sorted (less_prod leq lt lt) (RBT_Impl_diag t)" by(induct t)(auto simp add: ord.rbt_sorted.simps ord.rbt_less_prop ord.rbt_greater_prop keys_RBT_Impl_diag) lemma bheight_RBT_Impl_diag: "bheight (RBT_Impl_diag t) = bheight t" by(induct t) simp_all lemma inv_RBT_Impl_diag: assumes "inv1 t" "inv2 t" shows "inv1 (RBT_Impl_diag t)" "inv2 (RBT_Impl_diag t)" and "color_of t = color.B \ color_of (RBT_Impl_diag t) = color.B" using assms by(induct t)(auto simp add: bheight_RBT_Impl_diag) lemma is_rbt_RBT_Impl_diag: "ord.is_rbt lt t \ ord.is_rbt (less_prod leq lt lt) (RBT_Impl_diag t)" by(simp add: ord.is_rbt_def rbt_sorted_RBT_Impl_diag inv_RBT_Impl_diag) lemma (in linorder) rbt_lookup_RBT_Impl_diag: "ord.rbt_lookup (less_prod (\) (<) (<)) (RBT_Impl_diag t) = (\(k, k'). if k = k' then ord.rbt_lookup (<) t k else None)" by(induct t)(auto simp add: ord.rbt_lookup.simps fun_eq_iff) subsection \Folding and quantifiers over RBTs\ definition RBT_Impl_fold1 :: "('a \ 'a \ 'a) \ ('a, unit) RBT_Impl.rbt \ 'a" where "RBT_Impl_fold1 f rbt = fold f (tl (RBT_Impl.keys rbt)) (hd (RBT_Impl.keys rbt))" lemma RBT_Impl_fold1_simps [simp, code]: "RBT_Impl_fold1 f rbt.Empty = undefined" "RBT_Impl_fold1 f (Branch c rbt.Empty k v r) = RBT_Impl.fold (\k v. f k) r k" "RBT_Impl_fold1 f (Branch c (Branch c' l' k' v' r') k v r) = RBT_Impl.fold (\k v. f k) r (f k (RBT_Impl_fold1 f (Branch c' l' k' v' r')))" by(simp_all add: RBT_Impl_fold1_def RBT_Impl.keys_def fold_map RBT_Impl.fold_def split_def o_def tl_append hd_def split: list.split) definition RBT_Impl_rbt_all :: "('a \ 'b \ bool) \ ('a, 'b) rbt \ bool" where [code del]: "RBT_Impl_rbt_all P rbt = (\(k, v) \ set (RBT_Impl.entries rbt). P k v)" lemma RBT_Impl_rbt_all_simps [simp, code]: "RBT_Impl_rbt_all P rbt.Empty \ True" "RBT_Impl_rbt_all P (Branch c l k v r) \ P k v \ RBT_Impl_rbt_all P l \ RBT_Impl_rbt_all P r" by(auto simp add: RBT_Impl_rbt_all_def) definition RBT_Impl_rbt_ex :: "('a \ 'b \ bool) \ ('a, 'b) rbt \ bool" where [code del]: "RBT_Impl_rbt_ex P rbt = (\(k, v) \ set (RBT_Impl.entries rbt). P k v)" lemma RBT_Impl_rbt_ex_simps [simp, code]: "RBT_Impl_rbt_ex P rbt.Empty \ False" "RBT_Impl_rbt_ex P (Branch c l k v r) \ P k v \ RBT_Impl_rbt_ex P l \ RBT_Impl_rbt_ex P r" by(auto simp add: RBT_Impl_rbt_ex_def) subsection \List fusion for RBTs\ type_synonym ('a, 'b, 'c) rbt_generator_state = "('c \ ('a, 'b) RBT_Impl.rbt) list \ ('a, 'b) RBT_Impl.rbt" fun rbt_has_next :: "('a, 'b, 'c) rbt_generator_state \ bool" where "rbt_has_next ([], rbt.Empty) = False" | "rbt_has_next _ = True" fun rbt_keys_next :: "('a, 'b, 'a) rbt_generator_state \ 'a \ ('a, 'b, 'a) rbt_generator_state" where "rbt_keys_next ((k, t) # kts, rbt.Empty) = (k, kts, t)" | "rbt_keys_next (kts, rbt.Branch c l k v r) = rbt_keys_next ((k, r) # kts, l)" lemma rbt_generator_induct [case_names empty split shuffle]: assumes "P ([], rbt.Empty)" and "\k t kts. P (kts, t) \ P ((k, t) # kts, rbt.Empty)" and "\kts c l k v r. P ((f k v, r) # kts, l) \ P (kts, Branch c l k v r)" shows "P ktst" using assms apply(induction_schema) apply pat_completeness apply(relation "measure (\(kts, t). size_list (\(k, t). size_rbt (\_. 1) (\_. 1) t) kts + size_rbt (\_. 1) (\_. 1) t)") apply simp_all done lemma terminates_rbt_keys_generator: "terminates (rbt_has_next, rbt_keys_next)" proof fix ktst :: "('a \ ('a, 'b) rbt) list \ ('a, 'b) rbt" show "ktst \ terminates_on (rbt_has_next, rbt_keys_next)" by(induct ktst taking: "\k _. k" rule: rbt_generator_induct)(auto 4 3 intro: terminates_on.intros elim: terminates_on.cases) qed lift_definition rbt_keys_generator :: "('a, ('a, 'b, 'a) rbt_generator_state) generator" is "(rbt_has_next, rbt_keys_next)" by(rule terminates_rbt_keys_generator) definition rbt_init :: "('a, 'b) rbt \ ('a, 'b, 'c) rbt_generator_state" where "rbt_init = Pair []" lemma has_next_rbt_keys_generator [simp]: "list.has_next rbt_keys_generator = rbt_has_next" by transfer simp lemma next_rbt_keys_generator [simp]: "list.next rbt_keys_generator = rbt_keys_next" by transfer simp lemma unfoldr_rbt_keys_generator_aux: "list.unfoldr rbt_keys_generator (kts, t) = RBT_Impl.keys t @ concat (map (\(k, t). k # RBT_Impl.keys t) kts)" proof(induct "(kts, t)" arbitrary: kts t taking: "\k _. k" rule: rbt_generator_induct) case empty thus ?case by(simp add: list.unfoldr.simps) next case split thus ?case by(subst list.unfoldr.simps) simp next case shuffle thus ?case by(subst list.unfoldr.simps)(subst (asm) list.unfoldr.simps, simp) qed corollary unfoldr_rbt_keys_generator: "list.unfoldr rbt_keys_generator (rbt_init t) = RBT_Impl.keys t" by(simp add: unfoldr_rbt_keys_generator_aux rbt_init_def) fun rbt_entries_next :: "('a, 'b, 'a \ 'b) rbt_generator_state \ ('a \ 'b) \ ('a, 'b, 'a \ 'b) rbt_generator_state" where "rbt_entries_next ((kv, t) # kts, rbt.Empty) = (kv, kts, t)" | "rbt_entries_next (kts, rbt.Branch c l k v r) = rbt_entries_next (((k, v), r) # kts, l)" lemma terminates_rbt_entries_generator: "terminates (rbt_has_next, rbt_entries_next)" proof(rule terminatesI) fix ktst :: "('a, 'b, 'a \ 'b) rbt_generator_state" show "ktst \ terminates_on (rbt_has_next, rbt_entries_next)" by(induct ktst taking: Pair rule: rbt_generator_induct)(auto 4 3 intro: terminates_on.intros elim: terminates_on.cases) qed lift_definition rbt_entries_generator :: "('a \ 'b, ('a, 'b, 'a \ 'b) rbt_generator_state) generator" is "(rbt_has_next, rbt_entries_next)" by(rule terminates_rbt_entries_generator) lemma has_next_rbt_entries_generator [simp]: "list.has_next rbt_entries_generator = rbt_has_next" by transfer simp lemma next_rbt_entries_generator [simp]: "list.next rbt_entries_generator = rbt_entries_next" by transfer simp lemma unfoldr_rbt_entries_generator_aux: "list.unfoldr rbt_entries_generator (kts, t) = RBT_Impl.entries t @ concat (map (\(k, t). k # RBT_Impl.entries t) kts)" proof(induct "(kts, t)" arbitrary: kts t taking: Pair rule: rbt_generator_induct) case empty thus ?case by(simp add: list.unfoldr.simps) next case split thus ?case by(subst list.unfoldr.simps) simp next case shuffle thus ?case by(subst list.unfoldr.simps)(subst (asm) list.unfoldr.simps, simp) qed corollary unfoldr_rbt_entries_generator: "list.unfoldr rbt_entries_generator (rbt_init t) = RBT_Impl.entries t" by(simp add: unfoldr_rbt_entries_generator_aux rbt_init_def) end diff --git a/thys/CoreC++/Determinism.thy b/thys/CoreC++/Determinism.thy --- a/thys/CoreC++/Determinism.thy +++ b/thys/CoreC++/Determinism.thy @@ -1,2115 +1,2115 @@ (* Title: CoreC++ Author: Daniel Wasserrab Maintainer: Daniel Wasserrab *) section \Determinism Proof\ theory Determinism imports TypeSafe begin subsection\Some lemmas\ lemma maps_nth: "\(E(xs [\] ys)) x = Some y; length xs = length ys; distinct xs\ \ \i. x = xs!i \ i < length xs \ y = ys!i" proof (induct xs arbitrary: ys E) case Nil thus ?case by simp next case (Cons x' xs') have map:"(E(x' # xs' [\] ys)) x = Some y" and length:"length (x'#xs') = length ys" and dist:"distinct (x'#xs')" and IH:"\ys E. \(E(xs' [\] ys)) x = Some y; length xs' = length ys; distinct xs'\ \ \i. x = xs'!i \ i < length xs' \ y = ys!i" by fact+ from length obtain y' ys' where ys:"ys = y'#ys'" by(cases ys) auto { fix i assume x:"x = (x'#xs')!i" and i:"i < length(x'#xs')" have "y = ys!i" proof(cases i) case 0 with x map ys dist show ?thesis by simp next case (Suc n) with x i have x':"x = xs'!n" and n:"n < length xs'" by simp_all - from map ys have map':"(E(x' \ y')(xs' [\] ys')) x = Some y" by simp + from map ys have map':"(E(x' \ y', xs' [\] ys')) x = Some y" by simp from length ys have length':"length xs' = length ys'" by simp from dist have dist':"distinct xs'" by simp from IH[OF map' length' dist'] have "\i. x = xs'!i \ i < length xs' \ y = ys'!i" . with x' n have "y = ys'!n" by simp with ys n Suc show ?thesis by simp qed } thus ?case by simp qed lemma nth_maps:"\length pns = length Ts; distinct pns; i < length Ts\ \ (E(pns [\] Ts)) (pns!i) = Some (Ts!i)" proof (induct i arbitrary: E pns Ts) case 0 have dist:"distinct pns" and length:"length pns = length Ts" and i_length:"0 < length Ts" by fact+ from i_length obtain T' Ts' where Ts:"Ts = T'#Ts'" by(cases Ts) auto with length obtain p' pns' where "pns = p'#pns'" by(cases pns) auto with Ts dist show ?case by simp next case (Suc n) have i_length:"Suc n < length Ts" and dist:"distinct pns" and length:"length pns = length Ts" by fact+ from Suc obtain T' Ts' where Ts:"Ts = T'#Ts'" by(cases Ts) auto with length obtain p' pns' where pns:"pns = p'#pns'" by(cases pns) auto with Ts length dist have length':"length pns' = length Ts'" and dist':"distinct pns'" and notin:"p' \ set pns'" by simp_all from i_length Ts have n_length:"n < length Ts'" by simp - with length' dist' have map:"(E(p' \ T')(pns' [\] Ts')) (pns'!n) = Some(Ts'!n)" by fact - with notin have "(E(p' \ T')(pns' [\] Ts')) p' = Some T'" by simp + with length' dist' have map:"(E(p' \ T', pns' [\] Ts')) (pns'!n) = Some(Ts'!n)" by fact + with notin have "(E(p' \ T', pns' [\] Ts')) p' = Some T'" by simp with pns Ts map show ?case by simp qed lemma casts_casts_eq_result: fixes s :: state assumes casts:"P \ T casts v to v'" and casts':"P \ T casts v to w'" and type:"is_type P T" and wte:"P,E \ e :: T'" and leq:"P \ T' \ T" and eval:"P,E \ \e,s\ \ \Val v,(h,l)\" and sconf:"P,E \ s \" and wf:"wf_C_prog P" shows "v' = w'" proof(cases "\C. T \ Class C") case True with casts casts' show ?thesis by(auto elim:casts_to.cases) next case False then obtain C where T:"T = Class C" by auto with type have "is_class P C" by simp with wf T leq have "T' = NT \ (\D. T' = Class D \ P \ Path D to C unique)" by(simp add:widen_Class) thus ?thesis proof(rule disjE) assume "T' = NT" with wf eval sconf wte have "v = Null" by(fastforce dest:eval_preserves_type) with casts casts' show ?thesis by(fastforce elim:casts_to.cases) next assume "\D. T' = Class D \ P \ Path D to C unique" then obtain D where T':"T' = Class D" and path_unique:"P \ Path D to C unique" by auto with wf eval sconf wte have "P,E,h \ Val v : T' \ P,E,h \ Val v : NT" by(fastforce dest:eval_preserves_type) thus ?thesis proof(rule disjE) assume "P,E,h \ Val v : T'" with T' obtain a Cs C' S where h:"h a = Some(C',S)" and v:"v = Ref(a,Cs)" and last:"last Cs = D" by(fastforce dest:typeof_Class_Subo) from casts' v last T obtain Cs' Ds where "P \ Path D to C via Cs'" and "Ds = Cs@\<^sub>pCs'" and "w' = Ref(a,Ds)" by(auto elim:casts_to.cases) with casts T v last path_unique show ?thesis by auto(erule casts_to.cases,auto simp:path_via_def path_unique_def) next assume "P,E,h \ Val v : NT" with wf eval sconf wte have "v = Null" by(fastforce dest:eval_preserves_type) with casts casts' show ?thesis by(fastforce elim:casts_to.cases) qed qed qed lemma Casts_Casts_eq_result: assumes wf:"wf_C_prog P" shows "\P \ Ts Casts vs to vs'; P \ Ts Casts vs to ws'; \T \ set Ts. is_type P T; P,E \ es [::] Ts'; P \ Ts' [\] Ts; P,E \ \es,s\ [\] \map Val vs,(h,l)\; P,E \ s \\ \ vs' = ws'" proof (induct vs arbitrary: vs' ws' Ts Ts' es s) case Nil thus ?case by (auto elim!:Casts_to.cases) next case (Cons x xs) have CastsCons:"P \ Ts Casts x # xs to vs'" and CastsCons':"P \ Ts Casts x # xs to ws'" and type:"\T \ set Ts. is_type P T" and wtes:"P,E \ es [::] Ts'" and subs:"P \ Ts' [\] Ts" and evals:"P,E \ \es,s\ [\] \map Val (x#xs),(h,l)\" and sconf:"P,E \ s \" and IH:"\vs' ws' Ts Ts' es s. \P \ Ts Casts xs to vs'; P \ Ts Casts xs to ws'; \T \ set Ts. is_type P T; P,E \ es [::] Ts'; P \ Ts' [\] Ts; P,E \ \es,s\ [\] \map Val xs,(h,l)\; P,E \ s \\ \ vs' = ws'" by fact+ from CastsCons obtain y ys S Ss where vs':"vs' = y#ys" and Ts:"Ts = S#Ss" apply - apply(frule length_Casts_vs,cases Ts,auto) apply(frule length_Casts_vs',cases vs',auto) done with CastsCons have casts:"P \ S casts x to y" and Casts:"P \ Ss Casts xs to ys" by(auto elim:Casts_to.cases) from Ts type have type':"is_type P S" and types':"\T \ set Ss. is_type P T" by auto from Ts CastsCons' obtain z zs where ws':"ws' = z#zs" by simp(frule length_Casts_vs',cases ws',auto) with Ts CastsCons' have casts':"P \ S casts x to z" and Casts':"P \ Ss Casts xs to zs" by(auto elim:Casts_to.cases) from Ts subs obtain U Us where Ts':"Ts' = U#Us" and subs':"P \ Us [\] Ss" and sub:"P \ U \ S" by(cases Ts',auto simp:fun_of_def) from wtes Ts' obtain e' es' where es:"es = e'#es'" and wte':"P,E \ e' :: U" and wtes':"P,E \ es' [::] Us" by(cases es) auto with evals obtain h' l' where eval:"P,E \ \e',s\ \ \Val x,(h',l')\" and evals':"P,E \ \es',(h',l')\ [\] \map Val xs,(h,l)\" by (auto elim:evals.cases) from wf eval wte' sconf have "P,E \ (h',l') \" by(rule eval_preserves_sconf) from IH[OF Casts Casts' types' wtes' subs' evals' this] have eq:"ys = zs" . from casts casts' type' wte' sub eval sconf wf have "y = z" by(rule casts_casts_eq_result) with eq vs' ws' show ?case by simp qed lemma Casts_conf: assumes wf: "wf_C_prog P" shows "P \ Ts Casts vs to vs' \ (\es s Ts'. \ P,E \ es [::] Ts'; P,E \ \es,s\ [\] \map Val vs,(h,l)\; P,E \ s \; P \ Ts' [\] Ts\ \ \i < length Ts. P,h \ vs'!i :\ Ts!i)" proof(induct rule:Casts_to.induct) case Casts_Nil thus ?case by simp next case (Casts_Cons T v v' Ts vs vs') have casts:"P \ T casts v to v'" and wtes:"P,E \ es [::] Ts'" and evals:"P,E \ \es,s\ [\] \map Val (v#vs),(h,l)\" and subs:"P \ Ts' [\] (T#Ts)" and sconf:"P,E \ s \" and IH:"\es s Ts'.\P,E \ es [::] Ts'; P,E \ \es,s\ [\] \map Val vs,(h,l)\; P,E \ s \; P \ Ts' [\] Ts\ \ \i vs' ! i :\ Ts ! i" by fact+ from subs obtain U Us where Ts':"Ts' = U#Us" by(cases Ts') auto with subs have sub':"P \ U \ T" and subs':"P \ Us [\] Ts" by (simp_all add:fun_of_def) from wtes Ts' obtain e' es' where es:"es = e'#es'" by(cases es) auto with Ts' wtes have wte':"P,E \ e' :: U" and wtes':"P,E \ es' [::] Us" by auto from es evals obtain s' where eval':"P,E \ \e',s\ \ \Val v,s'\" and evals':"P,E \ \es',s'\ [\] \map Val vs,(h,l)\" by(auto elim:evals.cases) from wf eval' wte' sconf have sconf':"P,E \ s' \" by(rule eval_preserves_sconf) from evals' have hext:"hp s' \ h" by(cases s',auto intro:evals_hext) from wf eval' sconf wte' have "P,E,(hp s') \ Val v :\<^bsub>NT\<^esub> U" by(rule eval_preserves_type) with hext have wtrt:"P,E,h \ Val v :\<^bsub>NT\<^esub> U" by(cases U,auto intro:hext_typeof_mono) from casts wtrt sub' have "P,h \ v' :\ T" proof(induct rule:casts_to.induct) case (casts_prim T'' v'') have "\C. T'' \ Class C" and "P,E,h \ Val v'' :\<^bsub>NT\<^esub> U" and "P \ U \ T''" by fact+ thus ?case by(cases T'') auto next case (casts_null C) thus ?case by simp next case (casts_ref Cs C Cs' Ds a) have path:"P \ Path last Cs to C via Cs'" and Ds:"Ds = Cs @\<^sub>p Cs'" and wtref:"P,E,h \ ref (a, Cs) :\<^bsub>NT\<^esub> U" by fact+ from wtref obtain D S where subo:"Subobjs P D Cs" and h:"h a = Some(D,S)" by(cases U,auto split:if_split_asm) from path Ds have last:"C = last Ds" by(fastforce intro!:appendPath_last Subobjs_nonempty simp:path_via_def) from subo path Ds wf have "Subobjs P D Ds" by(fastforce intro:Subobjs_appendPath simp:path_via_def) with last h show ?case by simp qed with IH[OF wtes' evals' sconf' subs'] show ?case by(auto simp:nth_Cons,case_tac i,auto) qed lemma map_Val_throw_False:"map Val vs = map Val ws @ throw ex # es \ False" proof (induct vs arbitrary: ws) case Nil thus ?case by simp next case (Cons v' vs') have eq:"map Val (v'#vs') = map Val ws @ throw ex # es" and IH:"\ws'. map Val vs' = map Val ws' @ throw ex # es \ False" by fact+ from eq obtain w' ws' where ws:"ws = w'#ws'" by(cases ws) auto from eq have "tl(map Val (v'#vs')) = tl(map Val ws @ throw ex # es)" by simp hence "map Val vs' = tl(map Val ws @ throw ex # es)" by simp with ws have "map Val vs' = map Val ws' @ throw ex # es" by simp from IH[OF this] show ?case . qed lemma map_Val_throw_eq:"map Val vs @ throw ex # es = map Val ws @ throw ex' # es' \ vs = ws \ ex = ex' \ es = es'" apply(clarsimp simp:append_eq_append_conv2) apply(erule disjE) apply(case_tac us) apply(fastforce elim:map_injective simp:inj_on_def) apply(fastforce dest:map_Val_throw_False) apply(case_tac us) apply(fastforce elim:map_injective simp:inj_on_def) apply(fastforce dest:sym[THEN map_Val_throw_False]) done subsection \The proof\ lemma deterministic_big_step: assumes wf:"wf_C_prog P" shows "P,E \ \e,s\ \ \e\<^sub>1,s\<^sub>1\ \ (\e\<^sub>2 s\<^sub>2 T. \P,E \ \e,s\ \ \e\<^sub>2,s\<^sub>2\; P,E \ e :: T; P,E \ s \\ \ e\<^sub>1 = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2)" and "P,E \ \es,s\ [\] \es\<^sub>1,s\<^sub>1\ \ (\es\<^sub>2 s\<^sub>2 Ts. \P,E \ \es,s\ [\] \es\<^sub>2,s\<^sub>2\; P,E \ es [::] Ts; P,E \ s \\ \ es\<^sub>1 = es\<^sub>2 \ s\<^sub>1 = s\<^sub>2)" proof (induct rule:eval_evals.inducts) case New thus ?case by(auto elim: eval_cases) next case NewFail thus ?case by(auto elim: eval_cases) next case (StaticUpCast E e s\<^sub>0 a Cs s\<^sub>1 C Cs' Ds e\<^sub>2 s\<^sub>2) have eval:"P,E \ \\C\e,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\" and path_via:"P \ Path last Cs to C via Cs'" and Ds:"Ds = Cs @\<^sub>p Cs'" and wt:"P,E \ \C\e :: T" and sconf:"P,E \ s\<^sub>0 \" and IH:"\e\<^sub>2 s\<^sub>2 T. \P,E \ \e,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\; P,E \ e :: T; P,E \ s\<^sub>0 \\ \ ref (a,Cs) = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by fact+ from wt obtain D where "class":"is_class P C" and wte:"P,E \ e :: Class D" and disj:"P \ Path D to C unique \ (P \ C \\<^sup>* D \ (\Cs. P \ Path C to D via Cs \ Subobjs\<^sub>R P C Cs))" by auto from eval show ?case proof(rule eval_cases) fix Xs Xs' a' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref (a',Xs),s\<^sub>2\" and path_via':"P \ Path last Xs to C via Xs'" and ref:"e\<^sub>2 = ref (a',Xs@\<^sub>pXs')" from IH[OF eval_ref wte sconf] have eq:"a = a' \ Cs = Xs \ s\<^sub>1 = s\<^sub>2" by simp with wf eval_ref sconf wte have last:"last Cs = D" by(auto dest:eval_preserves_type split:if_split_asm) from disj show "ref (a,Ds) = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" proof (rule disjE) assume "P \ Path D to C unique" with path_via path_via' eq last have "Cs' = Xs'" by(fastforce simp add:path_via_def path_unique_def) with eq ref Ds show ?thesis by simp next assume "P \ C \\<^sup>* D \ (\Cs. P \ Path C to D via Cs \ Subobjs\<^sub>R P C Cs)" with "class" wf obtain Cs'' where "P \ Path C to D via Cs''" by(auto dest:leq_implies_path) with path_via path_via' wf eq last have "Cs' = Xs'" by(auto dest:path_via_reverse) with eq ref Ds show ?thesis by simp qed next fix Xs Xs' a' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs@C#Xs'),s\<^sub>2\" and ref:"e\<^sub>2 = ref (a',Xs@[C])" from IH[OF eval_ref wte sconf] have eq:"a = a' \ Cs = Xs@C#Xs' \ s\<^sub>1 = s\<^sub>2" by simp with wf eval_ref sconf wte obtain C' where last:"last Cs = D" and "Subobjs P C' (Xs@C#Xs')" by(auto dest:eval_preserves_type split:if_split_asm) hence subo:"Subobjs P C (C#Xs')" by(fastforce intro:Subobjs_Subobjs) with eq last have leq:"P \ C \\<^sup>* D" by(fastforce dest:Subobjs_subclass) from path_via last have "P \ D \\<^sup>* C" by(auto dest:Subobjs_subclass simp:path_via_def) with leq wf have CeqD:"C = D" by(rule subcls_asym2) with last path_via wf have "Cs' = [D]" by(fastforce intro:path_via_C) with Ds last have Ds':"Ds = Cs" by(simp add:appendPath_def) from subo CeqD last eq wf have "Xs' = []" by(auto dest:mdc_eq_last) with eq Ds' ref show "ref (a,Ds) = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next assume eval_null:"P,E \ \e,s\<^sub>0\ \ \null,s\<^sub>2\" from IH[OF eval_null wte sconf] show "ref (a,Ds) = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix Xs a' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs),s\<^sub>2\" and notin:"C \ set Xs" and notleq:"\ P \ last Xs \\<^sup>* C" and throw:"e\<^sub>2 = THROW ClassCast" from IH[OF eval_ref wte sconf] have eq:"a = a' \ Cs = Xs \ s\<^sub>1 = s\<^sub>2" by simp with wf eval_ref sconf wte have last:"last Cs = D" and notempty:"Cs \ []" by(auto dest!:eval_preserves_type Subobjs_nonempty split:if_split_asm) from disj have "C = D" proof(rule disjE) assume path_unique:"P \ Path D to C unique" with last have "P \ D \\<^sup>* C" by(fastforce dest:Subobjs_subclass simp:path_unique_def) with notleq last eq show ?thesis by simp next assume ass:"P \ C \\<^sup>* D \ (\Cs. P \ Path C to D via Cs \ Subobjs\<^sub>R P C Cs)" with "class" wf obtain Cs'' where path_via':"P \ Path C to D via Cs''" by(auto dest:leq_implies_path) with path_via wf eq last have "Cs'' = [D]" by(fastforce dest:path_via_reverse) with ass path_via' have "Subobjs\<^sub>R P C [D]" by simp thus ?thesis by(fastforce dest:hd_SubobjsR) qed with last notin eq notempty show "ref (a,Ds) = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by(fastforce intro:last_in_set) next fix e' assume eval_throw:"P,E \ \e,s\<^sub>0\ \ \throw e',s\<^sub>2\" from IH[OF eval_throw wte sconf] show "ref (a,Ds) = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp qed next case (StaticDownCast E e s\<^sub>0 a Cs C Cs' s\<^sub>1 e\<^sub>2 s\<^sub>2 T) have eval:"P,E \ \\C\e,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\" and eval':"P,E \ \e,s\<^sub>0\ \ \ref(a,Cs@[C]@Cs'),s\<^sub>1\" and wt:"P,E \ \C\e :: T" and sconf:"P,E \ s\<^sub>0 \" and IH:"\e\<^sub>2 s\<^sub>2 T. \P,E \ \e,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\; P,E \ e :: T; P,E \ s\<^sub>0 \\ \ ref(a,Cs@[C]@Cs') = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by fact+ from wt obtain D where wte:"P,E \ e :: Class D" and disj:"P \ Path D to C unique \ (P \ C \\<^sup>* D \ (\Cs. P \ Path C to D via Cs \ Subobjs\<^sub>R P C Cs))" by auto from eval show ?case proof(rule eval_cases) fix Xs Xs' a' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs),s\<^sub>2\" and path_via:"P \ Path last Xs to C via Xs'" and ref:"e\<^sub>2 = ref (a',Xs@\<^sub>pXs')" from IH[OF eval_ref wte sconf] have eq:"a = a' \ Cs@[C]@Cs' = Xs \ s\<^sub>1 = s\<^sub>2" by simp with wf eval_ref sconf wte obtain C' where last:"last(C#Cs') = D" and "Subobjs P C' (Cs@[C]@Cs')" by(auto dest:eval_preserves_type split:if_split_asm) hence "P \ Path C to D via C#Cs'" by(fastforce intro:Subobjs_Subobjs simp:path_via_def) with eq last path_via wf have "Xs' = [C] \ Cs' = [] \ C = D" apply clarsimp apply(split if_split_asm) by(simp,drule path_via_reverse,simp,simp)+ with ref eq show "ref(a,Cs@[C]) = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by(fastforce simp:appendPath_def) next fix Xs Xs' a' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs@C#Xs'),s\<^sub>2\" and ref:"e\<^sub>2 = ref (a',Xs@[C])" from IH[OF eval_ref wte sconf] have eq:"a = a' \ Cs@[C]@Cs' = Xs@C#Xs' \ s\<^sub>1 = s\<^sub>2" by simp with wf eval_ref sconf wte obtain C' where last:"last(C#Xs') = D" and subo:"Subobjs P C' (Cs@[C]@Cs')" by(auto dest:eval_preserves_type split:if_split_asm) from subo wf have notin:"C \ set Cs" by -(rule unique2,simp) from subo wf have "C \ set Cs'" by -(rule unique1,simp,simp) with notin eq have "Cs = Xs \ Cs' = Xs'" by -(rule only_one_append,simp+) with eq ref show "ref(a,Cs@[C]) = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next assume eval_null:"P,E \ \e,s\<^sub>0\ \ \null,s\<^sub>2\" from IH[OF eval_null wte sconf] show "ref (a,Cs@[C]) = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix Xs a' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs),s\<^sub>2\" and notin:"C \ set Xs" from IH[OF eval_ref wte sconf] have "a = a' \ Cs@[C]@Cs' = Xs \ s\<^sub>1 = s\<^sub>2" by simp with notin show "ref(a,Cs@[C]) = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by fastforce next fix e' assume eval_throw:"P,E \ \e,s\<^sub>0\ \ \throw e',s\<^sub>2\" from IH[OF eval_throw wte sconf] show "ref (a,Cs@[C]) = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp qed next case (StaticCastNull E e s\<^sub>0 s\<^sub>1 C e\<^sub>2 s\<^sub>2 T) have eval:"P,E \ \\C\e,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\" and wt:"P,E \ \C\e :: T" and sconf:"P,E \ s\<^sub>0 \" and IH:"\e\<^sub>2 s\<^sub>2 T. \P,E \ \e,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\; P,E \ e :: T; P,E \ s\<^sub>0 \\ \ null = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by fact+ from wt obtain D where wte:"P,E \ e :: Class D" by auto from eval show ?case proof(rule eval_cases) fix Xs Xs' a' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref (a',Xs),s\<^sub>2\" from IH[OF eval_ref wte sconf] show "null = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix Xs Xs' a' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs@C#Xs'),s\<^sub>2\" from IH[OF eval_ref wte sconf] show "null = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next assume eval_null:"P,E \ \e,s\<^sub>0\ \ \null,s\<^sub>2\" and "e\<^sub>2 = null" with IH[OF eval_null wte sconf] show "null = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix Xs a' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs),s\<^sub>2\" from IH[OF eval_ref wte sconf] show "null = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix e' assume eval_throw:"P,E \ \e,s\<^sub>0\ \ \throw e',s\<^sub>2\" from IH[OF eval_throw wte sconf] show "null = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp qed next case (StaticCastFail E e s\<^sub>0 a Cs s\<^sub>1 C e\<^sub>2 s\<^sub>2 T) have eval:"P,E \ \\C\e,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\" and notleq:"\ P \ last Cs \\<^sup>* C" and notin:"C \ set Cs" and wt:"P,E \ \C\e :: T" and sconf:"P,E \ s\<^sub>0 \" and IH:"\e\<^sub>2 s\<^sub>2 T. \P,E \ \e,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\; P,E \ e :: T; P,E \ s\<^sub>0 \\ \ref (a, Cs) = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by fact+ from wt obtain D where wte:"P,E \ e :: Class D" by auto from eval show ?case proof(rule eval_cases) fix Xs Xs' a' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref (a',Xs),s\<^sub>2\" and path_via:"P \ Path last Xs to C via Xs'" from IH[OF eval_ref wte sconf] have eq:"a = a' \ Cs = Xs \ s\<^sub>1 = s\<^sub>2" by simp with path_via wf have "P \ last Cs \\<^sup>* C" by(auto dest:Subobjs_subclass simp:path_via_def) with notleq show "THROW ClassCast = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix Xs Xs' a' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs@C#Xs'),s\<^sub>2\" from IH[OF eval_ref wte sconf] have "a = a' \ Cs = Xs@C#Xs' \ s\<^sub>1 = s\<^sub>2" by simp with notin show "THROW ClassCast = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next assume eval_null:"P,E \ \e,s\<^sub>0\ \ \null,s\<^sub>2\" from IH[OF eval_null wte sconf] show "THROW ClassCast = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix Xs a' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs),s\<^sub>2\" and throw:"e\<^sub>2 = THROW ClassCast" from IH[OF eval_ref wte sconf] have "a = a' \ Cs = Xs \ s\<^sub>1 = s\<^sub>2" by simp with throw show "THROW ClassCast = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix e' assume eval_throw:"P,E \ \e,s\<^sub>0\ \ \throw e',s\<^sub>2\" from IH[OF eval_throw wte sconf] show "THROW ClassCast = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp qed next case (StaticCastThrow E e s\<^sub>0 e' s\<^sub>1 C e\<^sub>2 s\<^sub>2 T) have eval:"P,E \ \\C\e,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\" and wt:"P,E \ \C\e :: T" and sconf:"P,E \ s\<^sub>0 \" and IH:"\e\<^sub>2 s\<^sub>2 T. \P,E \ \e,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\; P,E \ e :: T; P,E \ s\<^sub>0 \\ \ throw e' = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by fact+ from wt obtain D where wte:"P,E \ e :: Class D" by auto from eval show ?case proof(rule eval_cases) fix Xs Xs' a' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref (a',Xs),s\<^sub>2\" from IH[OF eval_ref wte sconf] show " throw e' = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix Xs Xs' a' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs@C#Xs'),s\<^sub>2\" from IH[OF eval_ref wte sconf] show "throw e' = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next assume eval_null:"P,E \ \e,s\<^sub>0\ \ \null,s\<^sub>2\" from IH[OF eval_null wte sconf] show "throw e' = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix Xs a' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs),s\<^sub>2\" from IH[OF eval_ref wte sconf] show "throw e' = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix e'' assume eval_throw:"P,E \ \e,s\<^sub>0\ \ \throw e'',s\<^sub>2\" and throw:"e\<^sub>2 = throw e''" from IH[OF eval_throw wte sconf] throw show "throw e' = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp qed next case (StaticUpDynCast E e s\<^sub>0 a Cs s\<^sub>1 C Cs' Ds e\<^sub>2 s\<^sub>2 T) have eval:"P,E \ \Cast C e,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\" and path_via:"P \ Path last Cs to C via Cs'" and path_unique:"P \ Path last Cs to C unique" and Ds:"Ds = Cs@\<^sub>pCs'" and wt:"P,E \ Cast C e :: T" and sconf:"P,E \ s\<^sub>0 \" and IH:"\e\<^sub>2 s\<^sub>2 T. \P,E \ \e,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\; P,E \ e :: T; P,E \ s\<^sub>0 \\ \ ref(a,Cs) = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by fact+ from wt obtain D where wte:"P,E \ e :: Class D" by auto from eval show ?case proof(rule eval_cases) fix Xs Xs' a' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref (a',Xs),s\<^sub>2\" and path_via':"P \ Path last Xs to C via Xs'" and ref:"e\<^sub>2 = ref (a',Xs@\<^sub>pXs')" from IH[OF eval_ref wte sconf] have eq:"a = a' \ Cs = Xs \ s\<^sub>1 = s\<^sub>2" by simp with wf eval_ref sconf wte have last:"last Cs = D" by(auto dest:eval_preserves_type split:if_split_asm) with path_unique path_via path_via' eq have "Xs' = Cs'" by(fastforce simp:path_via_def path_unique_def) with eq Ds ref show "ref (a, Ds) = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix Xs Xs' a' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs@C#Xs'),s\<^sub>2\" and ref:"e\<^sub>2 = ref (a',Xs@[C])" from IH[OF eval_ref wte sconf] have eq:"a = a' \ Cs = Xs@C#Xs' \ s\<^sub>1 = s\<^sub>2" by simp with wf eval_ref sconf wte obtain C' where last:"last Cs = D" and "Subobjs P C' (Xs@C#Xs')" by(auto dest:eval_preserves_type split:if_split_asm) hence "Subobjs P C (C#Xs')" by(fastforce intro:Subobjs_Subobjs) with last eq have "P \ Path C to D via C#Xs'" by(simp add:path_via_def) with path_via wf last have "Xs' = [] \ Cs' = [C] \ C = D" by(fastforce dest:path_via_reverse) with eq Ds ref show "ref (a, Ds) = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by (simp add:appendPath_def) next fix Xs Xs' D' S a' h l assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs),(h,l)\" and h:"h a' = Some(D',S)" and path_via':"P \ Path D' to C via Xs'" and path_unique':"P \ Path D' to C unique" and s2:"s\<^sub>2 = (h,l)" and ref:"e\<^sub>2 = ref(a',Xs')" from IH[OF eval_ref wte sconf] s2 have eq:"a = a' \ Cs = Xs \ s\<^sub>1 = s\<^sub>2" by simp with wf eval_ref sconf wte h have "last Cs = D" and "Subobjs P D' Cs" by(auto dest:eval_preserves_type split:if_split_asm) with path_via wf have "P \ Path D' to C via Cs@\<^sub>pCs'" by(fastforce intro:Subobjs_appendPath appendPath_last[THEN sym] dest:Subobjs_nonempty simp:path_via_def) with path_via' path_unique' Ds have "Xs' = Ds" by(fastforce simp:path_via_def path_unique_def) with eq ref show "ref (a, Ds) = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next assume eval_null:"P,E \ \e,s\<^sub>0\ \ \null,s\<^sub>2\" from IH[OF eval_null wte sconf] show "ref (a, Ds) = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix Xs D' S a' h l assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs),(h,l)\" and not_unique:"\ P \ Path last Xs to C unique" and s2:"s\<^sub>2 = (h,l)" from IH[OF eval_ref wte sconf] s2 have eq:"a = a' \ Cs = Xs \ s\<^sub>1 = s\<^sub>2" by simp with path_unique not_unique show "ref (a, Ds) = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix e' assume eval_throw:"P,E \ \e,s\<^sub>0\ \ \throw e',s\<^sub>2\" from IH[OF eval_throw wte sconf] show "ref (a, Ds) = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp qed next case (StaticDownDynCast E e s\<^sub>0 a Cs C Cs' s\<^sub>1 e\<^sub>2 s\<^sub>2 T) have eval:"P,E \ \Cast C e,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\" and wt:"P,E \ Cast C e :: T" and sconf:"P,E \ s\<^sub>0 \" and IH:"\e\<^sub>2 s\<^sub>2 T. \P,E \ \e,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\; P,E \ e :: T; P,E \ s\<^sub>0 \\ \ ref(a,Cs@[C]@Cs') = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by fact+ from wt obtain D where wte:"P,E \ e :: Class D" by auto from eval show ?case proof(rule eval_cases) fix Xs Xs' a' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs),s\<^sub>2\" and path_via:"P \ Path last Xs to C via Xs'" and ref:"e\<^sub>2 = ref (a',Xs@\<^sub>pXs')" from IH[OF eval_ref wte sconf] have eq:"a = a' \ Cs@[C]@Cs' = Xs \ s\<^sub>1 = s\<^sub>2" by simp with wf eval_ref sconf wte obtain C' where last:"last(C#Cs') = D" and "Subobjs P C' (Cs@[C]@Cs')" by(auto dest:eval_preserves_type split:if_split_asm) hence "P \ Path C to D via C#Cs'" by(fastforce intro:Subobjs_Subobjs simp:path_via_def) with eq last path_via wf have "Xs' = [C] \ Cs' = [] \ C = D" apply clarsimp apply(split if_split_asm) by(simp,drule path_via_reverse,simp,simp)+ with ref eq show "ref(a,Cs@[C]) = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by(fastforce simp:appendPath_def) next fix Xs Xs' a' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs@C#Xs'),s\<^sub>2\" and ref:"e\<^sub>2 = ref (a',Xs@[C])" from IH[OF eval_ref wte sconf] have eq:"a = a' \ Cs@[C]@Cs' = Xs@C#Xs' \ s\<^sub>1 = s\<^sub>2" by simp with wf eval_ref sconf wte obtain C' where last:"last(C#Xs') = D" and subo:"Subobjs P C' (Cs@[C]@Cs')" by(auto dest:eval_preserves_type split:if_split_asm) from subo wf have notin:"C \ set Cs" by -(rule unique2,simp) from subo wf have "C \ set Cs'" by -(rule unique1,simp,simp) with notin eq have "Cs = Xs \ Cs' = Xs'" by -(rule only_one_append,simp+) with eq ref show "ref(a,Cs@[C]) = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix Xs Xs' D' S a' h l assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs),(h,l)\" and h:"h a' = Some(D',S)" and path_via:"P \ Path D' to C via Xs'" and path_unique:"P \ Path D' to C unique" and s2:"s\<^sub>2 = (h,l)" and ref:"e\<^sub>2 = ref(a',Xs')" from IH[OF eval_ref wte sconf] s2 have eq:"a = a' \ Cs@[C]@Cs' = Xs \ s\<^sub>1 = s\<^sub>2" by simp with wf eval_ref sconf wte h have "Subobjs P D' (Cs@[C]@Cs')" by(auto dest:eval_preserves_type split:if_split_asm) hence "Subobjs P D' (Cs@[C])" by(fastforce intro:appendSubobj) with path_via path_unique have "Xs' = Cs@[C]" by(fastforce simp:path_via_def path_unique_def) with eq ref show "ref(a,Cs@[C]) = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next assume eval_null:"P,E \ \e,s\<^sub>0\ \ \null,s\<^sub>2\" from IH[OF eval_null wte sconf] show "ref (a,Cs@[C]) = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix Xs D' S a' h l assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs),(h,l)\" and notin:"C \ set Xs" and s2:"s\<^sub>2 = (h,l)" from IH[OF eval_ref wte sconf] s2 have "a = a' \ Cs@[C]@Cs' = Xs \ s\<^sub>1 = s\<^sub>2" by simp with notin show "ref (a,Cs@[C]) = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by fastforce next fix e' assume eval_throw:"P,E \ \e,s\<^sub>0\ \ \throw e',s\<^sub>2\" from IH[OF eval_throw wte sconf] show "ref (a,Cs@[C]) = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp qed next case (DynCast E e s\<^sub>0 a Cs h l D S C Cs' e\<^sub>2 s\<^sub>2 T) have eval:"P,E \ \Cast C e,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\" and path_via:"P \ Path D to C via Cs'" and path_unique:"P \ Path D to C unique" and h:"h a = Some(D,S)" and wt:"P,E \ Cast C e :: T" and sconf:"P,E \ s\<^sub>0 \" and IH:"\e\<^sub>2 s\<^sub>2 T. \P,E \ \e,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\; P,E \ e :: T; P,E \ s\<^sub>0 \\ \ ref(a,Cs) = e\<^sub>2 \ (h,l) = s\<^sub>2" by fact+ from wt obtain D' where wte:"P,E \ e :: Class D'" by auto from eval show ?case proof(rule eval_cases) fix Xs Xs' a' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs),s\<^sub>2\" and path_via':"P \ Path last Xs to C via Xs'" and ref:"e\<^sub>2 = ref (a',Xs@\<^sub>pXs')" from IH[OF eval_ref wte sconf] have eq:"a = a' \ Cs = Xs \ (h,l) = s\<^sub>2" by simp with wf eval_ref sconf wte h have "last Cs = D'" and "Subobjs P D Cs" by(auto dest:eval_preserves_type split:if_split_asm) with path_via' wf eq have "P \ Path D to C via Xs@\<^sub>pXs'" by(fastforce intro:Subobjs_appendPath appendPath_last[THEN sym] dest:Subobjs_nonempty simp:path_via_def) with path_via path_unique have "Cs' = Xs@\<^sub>pXs'" by(fastforce simp:path_via_def path_unique_def) with ref eq show "ref(a,Cs') = e\<^sub>2 \ (h, l) = s\<^sub>2" by simp next fix Xs Xs' a' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs@C#Xs'),s\<^sub>2\" and ref:"e\<^sub>2 = ref (a',Xs@[C])" from IH[OF eval_ref wte sconf] have eq:"a = a' \ Cs = Xs@C#Xs' \ (h,l) = s\<^sub>2" by simp with wf eval_ref sconf wte h have "Subobjs P D (Xs@[C]@Xs')" by(auto dest:eval_preserves_type split:if_split_asm) hence "Subobjs P D (Xs@[C])" by(fastforce intro:appendSubobj) with path_via path_unique have "Cs' = Xs@[C]" by(fastforce simp:path_via_def path_unique_def) with eq ref show "ref(a,Cs') = e\<^sub>2 \ (h, l) = s\<^sub>2" by simp next fix Xs Xs' D'' S' a' h' l' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs),(h',l')\" and h':"h' a' = Some(D'',S')" and path_via':"P \ Path D'' to C via Xs'" and s2:"s\<^sub>2 = (h',l')" and ref:"e\<^sub>2 = ref(a',Xs')" from IH[OF eval_ref wte sconf] have eq:"a = a' \ Cs = Xs \ h = h' \ l = l'" by simp with h h' path_via path_via' path_unique s2 ref show "ref(a,Cs') = e\<^sub>2 \ (h,l) = s\<^sub>2" by(fastforce simp:path_via_def path_unique_def) next assume eval_null:"P,E \ \e,s\<^sub>0\ \ \null,s\<^sub>2\" from IH[OF eval_null wte sconf] show "ref(a,Cs') = e\<^sub>2 \ (h,l) = s\<^sub>2" by simp next fix Xs D'' S' a' h' l' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs),(h',l')\" and h':"h' a' = Some(D'',S')" and not_unique:"\ P \ Path D'' to C unique" from IH[OF eval_ref wte sconf] have eq:"a = a' \ Cs = Xs \ h = h' \ l = l'" by simp with h h' path_unique not_unique show "ref(a,Cs') = e\<^sub>2 \ (h,l) = s\<^sub>2" by simp next fix e' assume eval_throw:"P,E \ \e,s\<^sub>0\ \ \throw e',s\<^sub>2\" from IH[OF eval_throw wte sconf] show "ref (a,Cs') = e\<^sub>2 \ (h,l) = s\<^sub>2" by simp qed next case (DynCastNull E e s\<^sub>0 s\<^sub>1 C e\<^sub>2 s\<^sub>2 T) have eval:"P,E \ \Cast C e,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\" and wt:"P,E \ Cast C e :: T" and sconf:"P,E \ s\<^sub>0 \" and IH:"\e\<^sub>2 s\<^sub>2 T. \P,E \ \e,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\; P,E \ e :: T; P,E \ s\<^sub>0 \\ \ null = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by fact+ from wt obtain D where wte:"P,E \ e :: Class D" by auto from eval show ?case proof(rule eval_cases) fix Xs Xs' a' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref (a',Xs),s\<^sub>2\" from IH[OF eval_ref wte sconf] show "null = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix Xs Xs' a' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs@C#Xs'),s\<^sub>2\" from IH[OF eval_ref wte sconf] show "null = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix Xs Xs' D' S a' h l assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs),(h,l)\" from IH[OF eval_ref wte sconf] show "null = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next assume eval_null:"P,E \ \e,s\<^sub>0\ \ \null,s\<^sub>2\" and "e\<^sub>2 = null" with IH[OF eval_null wte sconf] show "null = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix Xs D' S a' h l assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs),(h,l)\" and s2:"s\<^sub>2 = (h,l)" from IH[OF eval_ref wte sconf] s2 show "null = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix e' assume eval_throw:"P,E \ \e,s\<^sub>0\ \ \throw e',s\<^sub>2\" from IH[OF eval_throw wte sconf] show "null = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp qed next case (DynCastFail E e s\<^sub>0 a Cs h l D S C e\<^sub>2 s\<^sub>2 T) have eval:"P,E \ \Cast C e,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\" and h:"h a = Some(D,S)" and not_unique1:"\ P \ Path D to C unique" and not_unique2:"\ P \ Path last Cs to C unique" and notin:"C \ set Cs" and wt:"P,E \ Cast C e :: T" and sconf:"P,E \ s\<^sub>0 \" and IH:"\e\<^sub>2 s\<^sub>2 T. \P,E \ \e,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\; P,E \ e :: T; P,E \ s\<^sub>0 \\ \ ref (a, Cs) = e\<^sub>2 \ (h,l) = s\<^sub>2" by fact+ from wt obtain D' where wte:"P,E \ e :: Class D'" by auto from eval show ?case proof(rule eval_cases) fix Xs Xs' a' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs),s\<^sub>2\" and path_unique:"P \ Path last Xs to C unique" from IH[OF eval_ref wte sconf] have eq:"a = a' \ Cs = Xs \ (h,l) = s\<^sub>2" by simp with path_unique not_unique2 show "null = e\<^sub>2 \ (h,l) = s\<^sub>2" by simp next fix Xs Xs' a' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs@C#Xs'),s\<^sub>2\" from IH[OF eval_ref wte sconf] have eq:"a = a' \ Cs = Xs@C#Xs' \ (h,l) = s\<^sub>2" by simp with notin show "null = e\<^sub>2 \ (h,l) = s\<^sub>2" by fastforce next fix Xs Xs' D'' S' a' h' l' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs),(h',l')\" and h':"h' a' = Some(D'',S')" and path_unique:"P \ Path D'' to C unique" from IH[OF eval_ref wte sconf] have "a = a' \ Cs = Xs \ h = h' \ l = l'" by simp with h h' not_unique1 path_unique show "null = e\<^sub>2 \ (h,l) = s\<^sub>2" by simp next assume eval_null:"P,E \ \e,s\<^sub>0\ \ \null,s\<^sub>2\" from IH[OF eval_null wte sconf] show "null = e\<^sub>2 \ (h,l) = s\<^sub>2" by simp next fix Xs D'' S' a' h' l' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs),(h',l')\" and null:"e\<^sub>2 = null" and s2:"s\<^sub>2 = (h',l')" from IH[OF eval_ref wte sconf] null s2 show "null = e\<^sub>2 \ (h,l) = s\<^sub>2" by simp next fix e' assume eval_throw:"P,E \ \e,s\<^sub>0\ \ \throw e',s\<^sub>2\" from IH[OF eval_throw wte sconf] show "null = e\<^sub>2 \ (h,l) = s\<^sub>2" by simp qed next case (DynCastThrow E e s\<^sub>0 e' s\<^sub>1 C e\<^sub>2 s\<^sub>2 T) have eval:"P,E \ \Cast C e,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\" and wt:"P,E \ Cast C e :: T" and sconf:"P,E \ s\<^sub>0 \" and IH:"\e\<^sub>2 s\<^sub>2 T. \P,E \ \e,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\; P,E \ e :: T; P,E \ s\<^sub>0 \\ \ throw e' = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by fact+ from wt obtain D where wte:"P,E \ e :: Class D" by auto from eval show ?case proof(rule eval_cases) fix Xs Xs' a' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref (a',Xs),s\<^sub>2\" from IH[OF eval_ref wte sconf] show " throw e' = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix Xs Xs' a' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs@C#Xs'),s\<^sub>2\" from IH[OF eval_ref wte sconf] show "throw e' = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix Xs Xs' D'' S' a' h' l' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs),(h',l')\" from IH[OF eval_ref wte sconf] show "throw e' = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next assume eval_null:"P,E \ \e,s\<^sub>0\ \ \null,s\<^sub>2\" from IH[OF eval_null wte sconf] show "throw e' = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix Xs D'' S' a' h' l' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs),(h',l')\" from IH[OF eval_ref wte sconf] show "throw e' = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix e'' assume eval_throw:"P,E \ \e,s\<^sub>0\ \ \throw e'',s\<^sub>2\" and throw:"e\<^sub>2 = throw e''" from IH[OF eval_throw wte sconf] throw show "throw e' = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp qed next case Val thus ?case by(auto elim: eval_cases) next case (BinOp E e\<^sub>1 s\<^sub>0 v\<^sub>1 s\<^sub>1 e\<^sub>2 v\<^sub>2 s\<^sub>2 bop v e\<^sub>2' s\<^sub>2' T) have eval:"P,E \ \e\<^sub>1 \bop\ e\<^sub>2,s\<^sub>0\ \ \e\<^sub>2',s\<^sub>2'\" and binop:"binop (bop, v\<^sub>1, v\<^sub>2) = Some v" and wt:"P,E \ e\<^sub>1 \bop\ e\<^sub>2 :: T" and sconf:"P,E \ s\<^sub>0 \" and IH1:"\ei si T. \P,E \ \e\<^sub>1,s\<^sub>0\ \ \ei,si\; P,E \ e\<^sub>1 :: T; P,E \ s\<^sub>0 \\ \ Val v\<^sub>1 = ei \ s\<^sub>1 = si" and IH2:"\ei si T. \P,E \ \e\<^sub>2,s\<^sub>1\ \ \ei,si\; P,E \ e\<^sub>2 :: T; P,E \ s\<^sub>1 \\ \ Val v\<^sub>2 = ei \ s\<^sub>2 = si" by fact+ from wt obtain T\<^sub>1 T\<^sub>2 where wte1:"P,E \ e\<^sub>1 :: T\<^sub>1" and wte2:"P,E \ e\<^sub>2 :: T\<^sub>2" by auto from eval show ?case proof(rule eval_cases) fix s w w\<^sub>1 w\<^sub>2 assume eval_val1:"P,E \ \e\<^sub>1,s\<^sub>0\ \ \Val w\<^sub>1,s\" and eval_val2:"P,E \ \e\<^sub>2,s\ \ \Val w\<^sub>2,s\<^sub>2'\" and binop':"binop(bop,w\<^sub>1,w\<^sub>2) = Some w" and e2':"e\<^sub>2' = Val w" from IH1[OF eval_val1 wte1 sconf] have w1:"v\<^sub>1 = w\<^sub>1" and s:"s = s\<^sub>1" by simp_all with wf eval_val1 wte1 sconf have "P,E \ s\<^sub>1 \" by(fastforce intro:eval_preserves_sconf) from IH2[OF eval_val2[simplified s] wte2 this] have "v\<^sub>2 = w\<^sub>2" and s2:"s\<^sub>2 = s\<^sub>2'" by simp_all with w1 binop binop' have "w = v" by simp with e2' s2 show "Val v = e\<^sub>2' \ s\<^sub>2 = s\<^sub>2'" by simp next fix e assume eval_throw:"P,E \ \e\<^sub>1,s\<^sub>0\ \ \throw e,s\<^sub>2'\" from IH1[OF eval_throw wte1 sconf] show "Val v = e\<^sub>2' \ s\<^sub>2 = s\<^sub>2'" by simp next fix e s w assume eval_val:"P,E \ \e\<^sub>1,s\<^sub>0\ \ \Val w,s\" and eval_throw:"P,E \ \e\<^sub>2,s\ \ \throw e,s\<^sub>2'\" from IH1[OF eval_val wte1 sconf] have s:"s = s\<^sub>1" by simp_all with wf eval_val wte1 sconf have "P,E \ s\<^sub>1 \" by(fastforce intro:eval_preserves_sconf) from IH2[OF eval_throw[simplified s] wte2 this] show "Val v = e\<^sub>2' \ s\<^sub>2 = s\<^sub>2'" by simp qed next case (BinOpThrow1 E e\<^sub>1 s\<^sub>0 e s\<^sub>1 bop e\<^sub>2 e\<^sub>2' s\<^sub>2 T) have eval:"P,E \ \e\<^sub>1 \bop\ e\<^sub>2,s\<^sub>0\ \ \e\<^sub>2',s\<^sub>2\" and wt:"P,E \ e\<^sub>1 \bop\ e\<^sub>2 :: T" and sconf:"P,E \ s\<^sub>0 \" and IH:"\ei si T. \P,E \ \e\<^sub>1,s\<^sub>0\ \ \ei,si\; P,E \ e\<^sub>1 :: T; P,E \ s\<^sub>0 \\ \ throw e = ei \ s\<^sub>1 = si" by fact+ from wt obtain T\<^sub>1 T\<^sub>2 where wte1:"P,E \ e\<^sub>1 :: T\<^sub>1" by auto from eval show ?case proof(rule eval_cases) fix s w w\<^sub>1 w\<^sub>2 assume eval_val:"P,E \ \e\<^sub>1,s\<^sub>0\ \ \Val w\<^sub>1,s\" from IH[OF eval_val wte1 sconf] show "throw e = e\<^sub>2' \ s\<^sub>1 = s\<^sub>2" by simp next fix e' assume eval_throw:"P,E \ \e\<^sub>1,s\<^sub>0\ \ \throw e',s\<^sub>2\" and throw:"e\<^sub>2' = throw e'" from IH[OF eval_throw wte1 sconf] throw show "throw e = e\<^sub>2' \ s\<^sub>1 = s\<^sub>2" by simp next fix e s w assume eval_val:"P,E \ \e\<^sub>1,s\<^sub>0\ \ \Val w,s\" from IH[OF eval_val wte1 sconf] show "throw e = e\<^sub>2' \ s\<^sub>1 = s\<^sub>2" by simp qed next case (BinOpThrow2 E e\<^sub>1 s\<^sub>0 v\<^sub>1 s\<^sub>1 e\<^sub>2 e s\<^sub>2 bop e\<^sub>2' s\<^sub>2' T) have eval:"P,E \ \e\<^sub>1 \bop\ e\<^sub>2,s\<^sub>0\ \ \e\<^sub>2',s\<^sub>2'\" and wt:"P,E \ e\<^sub>1 \bop\ e\<^sub>2 :: T" and sconf:"P,E \ s\<^sub>0 \" and IH1:"\ei si T. \P,E \ \e\<^sub>1,s\<^sub>0\ \ \ei,si\; P,E \ e\<^sub>1 :: T; P,E \ s\<^sub>0 \\ \ Val v\<^sub>1 = ei \ s\<^sub>1 = si" and IH2:"\ei si T. \P,E \ \e\<^sub>2,s\<^sub>1\ \ \ei,si\; P,E \ e\<^sub>2 :: T; P,E \ s\<^sub>1 \\ \ throw e = ei \ s\<^sub>2 = si" by fact+ from wt obtain T\<^sub>1 T\<^sub>2 where wte1:"P,E \ e\<^sub>1 :: T\<^sub>1" and wte2:"P,E \ e\<^sub>2 :: T\<^sub>2" by auto from eval show ?case proof(rule eval_cases) fix s w w\<^sub>1 w\<^sub>2 assume eval_val1:"P,E \ \e\<^sub>1,s\<^sub>0\ \ \Val w\<^sub>1,s\" and eval_val2:"P,E \ \e\<^sub>2,s\ \ \Val w\<^sub>2,s\<^sub>2'\" from IH1[OF eval_val1 wte1 sconf] have s:"s = s\<^sub>1" by simp_all with wf eval_val1 wte1 sconf have "P,E \ s\<^sub>1 \" by(fastforce intro:eval_preserves_sconf) from IH2[OF eval_val2[simplified s] wte2 this] show "throw e = e\<^sub>2' \ s\<^sub>2 = s\<^sub>2'" by simp next fix e' assume eval_throw:"P,E \ \e\<^sub>1,s\<^sub>0\ \ \throw e',s\<^sub>2'\" from IH1[OF eval_throw wte1 sconf] show "throw e = e\<^sub>2' \ s\<^sub>2 = s\<^sub>2'" by simp next fix e' s w assume eval_val:"P,E \ \e\<^sub>1,s\<^sub>0\ \ \Val w,s\" and eval_throw:"P,E \ \e\<^sub>2,s\ \ \throw e',s\<^sub>2'\" and throw:"e\<^sub>2' = throw e'" from IH1[OF eval_val wte1 sconf] have s:"s = s\<^sub>1" by simp_all with wf eval_val wte1 sconf have "P,E \ s\<^sub>1 \" by(fastforce intro:eval_preserves_sconf) from IH2[OF eval_throw[simplified s] wte2 this] throw show "throw e = e\<^sub>2' \ s\<^sub>2 = s\<^sub>2'" by simp qed next case Var thus ?case by(auto elim: eval_cases) next case (LAss E e s\<^sub>0 v h l V T v' l' e\<^sub>2 s\<^sub>2 T') have eval:"P,E \ \V:=e,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\" and env:"E V = Some T" and casts:"P \ T casts v to v'" and l':"l' = l(V \ v')" and wt:"P,E \ V:=e :: T'" and sconf:"P,E \ s\<^sub>0 \" and IH:"\e\<^sub>2 s\<^sub>2 T. \P,E \ \e,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\; P,E \ e :: T; P,E \ s\<^sub>0 \\ \ Val v = e\<^sub>2 \ (h,l) = s\<^sub>2" by fact+ from wt env obtain T'' where wte:"P,E \ e :: T''" and leq:"P \ T'' \ T" by auto from eval show ?case proof(rule eval_cases) fix U h' l'' w w' assume eval_val:"P,E \ \e,s\<^sub>0\ \ \Val w,(h',l'')\" and env':"E V = Some U" and casts':"P \ U casts w to w'" and e2:"e\<^sub>2 = Val w'" and s2:"s\<^sub>2 = (h',l''(V \ w'))" from env env' have UeqT:"U = T" by simp from IH[OF eval_val wte sconf] have eq:"v = w \ h = h' \ l = l''" by simp from sconf env have "is_type P T" by(clarsimp simp:sconf_def envconf_def) with casts casts' eq UeqT wte leq eval_val sconf wf have "v' = w'" by(auto intro:casts_casts_eq_result) with e2 s2 l' eq show "Val v' = e\<^sub>2 \ (h, l') = s\<^sub>2" by simp next fix e' assume eval_throw:"P,E \ \e,s\<^sub>0\ \ \throw e',s\<^sub>2\" from IH[OF eval_throw wte sconf] show "Val v' = e\<^sub>2 \ (h, l') = s\<^sub>2" by simp qed next case (LAssThrow E e s\<^sub>0 e' s\<^sub>1 V e\<^sub>2 s\<^sub>2 T) have eval:"P,E \ \V:=e,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\" and wt:"P,E \ V:=e :: T" and sconf:"P,E \ s\<^sub>0 \" and IH:"\e\<^sub>2 s\<^sub>2 T. \P,E \ \e,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\; P,E \ e :: T; P,E \ s\<^sub>0 \\ \ throw e' = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by fact+ from wt obtain T'' where wte:"P,E \ e :: T''" by auto from eval show ?case proof(rule eval_cases) fix U h' l'' w w' assume eval_val:"P,E \ \e,s\<^sub>0\ \ \Val w,(h',l'')\" from IH[OF eval_val wte sconf] show "throw e' = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix ex assume eval_throw:"P,E \ \e,s\<^sub>0\ \ \throw ex,s\<^sub>2\" and e2:"e\<^sub>2 = throw ex" from IH[OF eval_throw wte sconf] e2 show "throw e' = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp qed next case (FAcc E e s\<^sub>0 a Cs' h l D S Ds Cs fs F v e\<^sub>2 s\<^sub>2 T) have eval:"P,E \ \e\F{Cs},s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\" and eval':"P,E \ \e,s\<^sub>0\ \ \ref (a, Cs'),(h,l)\" and h:"h a = Some(D,S)" and Ds:"Ds = Cs'@\<^sub>pCs" and S:"(Ds,fs) \ S" and fs:"fs F = Some v" and wt:"P,E \ e\F{Cs} :: T" and sconf:"P,E \ s\<^sub>0 \" and IH:"\e\<^sub>2 s\<^sub>2 T. \P,E \ \e,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\; P,E \ e :: T; P,E \ s\<^sub>0 \\ \ ref (a, Cs') = e\<^sub>2 \ (h,l) = s\<^sub>2" by fact+ from wt obtain C where wte:"P,E \ e :: Class C" by auto from eval_preserves_sconf[OF wf eval' wte sconf] h have oconf:"P,h \ (D,S) \" by(simp add:sconf_def hconf_def) from eval show ?case proof(rule eval_cases) fix Xs' D' S' a' fs' h' l' v' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs'),(h',l')\" and h':"h' a' = Some(D',S')" and S':"(Xs'@\<^sub>pCs,fs') \ S'" and fs':"fs' F = Some v'" and e2:"e\<^sub>2 = Val v'" and s2:"s\<^sub>2 = (h',l')" from IH[OF eval_ref wte sconf] h h' have eq:"a = a' \ Cs' = Xs' \ h = h' \ l = l' \ D = D' \ S = S'" by simp with oconf S S' Ds have "fs = fs'" by (auto simp:oconf_def) with fs fs' have "v = v'" by simp with e2 s2 eq show "Val v = e\<^sub>2 \ (h,l) = s\<^sub>2" by simp next assume eval_null:"P,E \ \e,s\<^sub>0\ \ \null,s\<^sub>2\" from IH[OF eval_null wte sconf] show "Val v = e\<^sub>2 \ (h,l) = s\<^sub>2" by simp next fix e' assume eval_throw:"P,E \ \e,s\<^sub>0\ \ \throw e',s\<^sub>2\" from IH[OF eval_throw wte sconf] show "Val v = e\<^sub>2 \ (h,l) = s\<^sub>2" by simp qed next case (FAccNull E e s\<^sub>0 s\<^sub>1 F Cs e\<^sub>2 s\<^sub>2 T) have eval:"P,E \ \e\F{Cs},s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\" and wt:"P,E \ e\F{Cs} :: T" and sconf:"P,E \ s\<^sub>0 \" and IH:"\e\<^sub>2 s\<^sub>2 T. \P,E \ \e,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\; P,E \ e :: T; P,E \ s\<^sub>0 \\ \ null = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by fact+ from wt obtain C where wte:"P,E \ e :: Class C" by auto from eval show ?case proof(rule eval_cases) fix Xs' D' S' a' fs' h' l' v' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs'),(h',l')\" from IH[OF eval_ref wte sconf] show "THROW NullPointer = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next assume eval_null:"P,E \ \e,s\<^sub>0\ \ \null,s\<^sub>2\" and e2:"e\<^sub>2 = THROW NullPointer" from IH[OF eval_null wte sconf] e2 show "THROW NullPointer = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix e' assume eval_throw:"P,E \ \e,s\<^sub>0\ \ \throw e',s\<^sub>2\" from IH[OF eval_throw wte sconf] show "THROW NullPointer = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp qed next case (FAccThrow E e s\<^sub>0 e' s\<^sub>1 F Cs e\<^sub>2 s\<^sub>2 T) have eval:"P,E \ \e\F{Cs},s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\" and wt:"P,E \ e\F{Cs} :: T" and sconf:"P,E \ s\<^sub>0 \" and IH:"\e\<^sub>2 s\<^sub>2 T. \P,E \ \e,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\; P,E \ e :: T; P,E \ s\<^sub>0 \\ \ throw e' = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by fact+ from wt obtain C where wte:"P,E \ e :: Class C" by auto from eval show ?case proof(rule eval_cases) fix Xs' D' S' a' fs' h' l' v' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs'),(h',l')\" from IH[OF eval_ref wte sconf] show "throw e' = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next assume eval_null:"P,E \ \e,s\<^sub>0\ \ \null,s\<^sub>2\" from IH[OF eval_null wte sconf] show "throw e' = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix ex assume eval_throw:"P,E \ \e,s\<^sub>0\ \ \throw ex,s\<^sub>2\" and e2:"e\<^sub>2 = throw ex" from IH[OF eval_throw wte sconf] e2 show "throw e' = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp qed next case (FAss E e\<^sub>1 s\<^sub>0 a Cs' s\<^sub>1 e\<^sub>2 v h\<^sub>2 l\<^sub>2 D S F T Cs v' Ds fs fs' S' h\<^sub>2' e\<^sub>2' s\<^sub>2 T') have eval:"P,E \ \e\<^sub>1\F{Cs} := e\<^sub>2,s\<^sub>0\ \ \e\<^sub>2',s\<^sub>2\" and eval':"P,E \ \e\<^sub>1,s\<^sub>0\ \ \ref(a,Cs'),s\<^sub>1\" and eval'':"P,E \ \e\<^sub>2,s\<^sub>1\ \ \Val v,(h\<^sub>2,l\<^sub>2)\" and h2:"h\<^sub>2 a = Some(D, S)" and has_least:"P \ last Cs' has least F:T via Cs" and casts:"P \ T casts v to v'" and Ds:"Ds = Cs'@\<^sub>pCs" and S:"(Ds, fs) \ S" and fs':"fs' = fs(F \ v')" and S':"S' = S - {(Ds, fs)} \ {(Ds, fs')}" and h2':"h\<^sub>2' = h\<^sub>2(a \ (D, S'))" and wt:"P,E \ e\<^sub>1\F{Cs} := e\<^sub>2 :: T'" and sconf:"P,E \ s\<^sub>0 \" and IH1:"\ei si T. \P,E \ \e\<^sub>1,s\<^sub>0\ \ \ei,si\; P,E \ e\<^sub>1 :: T; P,E \ s\<^sub>0 \\ \ ref(a,Cs') = ei \ s\<^sub>1 = si" and IH2:"\ei si T. \P,E \ \e\<^sub>2,s\<^sub>1\ \ \ei,si\; P,E \ e\<^sub>2 :: T; P,E \ s\<^sub>1 \\ \ Val v = ei \ (h\<^sub>2,l\<^sub>2) = si" by fact+ from wt obtain C T'' where wte1:"P,E \ e\<^sub>1 :: Class C" and has_least':"P \ C has least F:T' via Cs" and wte2:"P,E \ e\<^sub>2 :: T''" and leq:"P \ T'' \ T'" by auto from wf eval' wte1 sconf have "last Cs' = C" by(auto dest!:eval_preserves_type split:if_split_asm) with has_least has_least' have TeqT':"T = T'" by (fastforce intro:sees_field_fun) from eval show ?case proof(rule eval_cases) fix Xs D' S'' U a' fs'' h l s w w' assume eval_ref:"P,E \ \e\<^sub>1,s\<^sub>0\ \ \ref(a',Xs),s\" and eval_val:"P,E \ \e\<^sub>2,s\ \ \Val w,(h,l)\" and h:"h a' = Some(D',S'')" and has_least'':"P \ last Xs has least F:U via Cs" and casts':"P \ U casts w to w'" and S'':"(Xs@\<^sub>pCs,fs'') \ S''" and e2':"e\<^sub>2' = Val w'" and s2:"s\<^sub>2 = (h(a'\(D',insert (Xs@\<^sub>pCs,fs''(F \ w')) (S''-{(Xs@\<^sub>pCs,fs'')}))),l)" from IH1[OF eval_ref wte1 sconf] have eq:"a = a' \ Cs' = Xs \ s\<^sub>1 = s" by simp with wf eval_ref wte1 sconf have sconf':"P,E \ s\<^sub>1 \" by(fastforce intro:eval_preserves_sconf) from IH2[OF _ wte2 this] eval_val eq have eq':"v = w \ h = h\<^sub>2 \ l = l\<^sub>2" by auto from has_least'' eq has_least have UeqT:"U = T" by (fastforce intro:sees_field_fun) from has_least wf have "is_type P T" by(rule least_field_is_type) with casts casts' eq eq' UeqT TeqT' wte2 leq eval_val sconf' wf have v':"v' = w'" by(auto intro!:casts_casts_eq_result) from eval_preserves_sconf[OF wf eval'' wte2 sconf'] h2 have oconf:"P,h\<^sub>2 \ (D,S) \" by(simp add:sconf_def hconf_def) from eq eq' h2 h have "S = S''" by simp with oconf eq S S' S'' Ds have "fs = fs''" by (auto simp:oconf_def) with h2' h h2 eq eq' s2 S' Ds fs' v' e2' show "Val v' = e\<^sub>2' \ (h\<^sub>2',l\<^sub>2) = s\<^sub>2" by simp next fix s w assume eval_null:"P,E \ \e\<^sub>1,s\<^sub>0\ \ \null,s\" from IH1[OF eval_null wte1 sconf] show "Val v' = e\<^sub>2' \ (h\<^sub>2',l\<^sub>2) = s\<^sub>2" by simp next fix ex assume eval_throw:"P,E \ \e\<^sub>1,s\<^sub>0\ \ \throw ex,s\<^sub>2\" from IH1[OF eval_throw wte1 sconf] show "Val v' = e\<^sub>2' \ (h\<^sub>2',l\<^sub>2) = s\<^sub>2" by simp next fix ex s w assume eval_val:"P,E \ \e\<^sub>1,s\<^sub>0\ \ \Val w,s\" and eval_throw:"P,E \ \e\<^sub>2,s\ \ \throw ex,s\<^sub>2\" from IH1[OF eval_val wte1 sconf] have eq:"s = s\<^sub>1" by simp with wf eval_val wte1 sconf have sconf':"P,E \ s\<^sub>1 \" by(fastforce intro:eval_preserves_sconf) from IH2[OF eval_throw[simplified eq] wte2 this] show "Val v' = e\<^sub>2' \ (h\<^sub>2',l\<^sub>2) = s\<^sub>2" by simp qed next case (FAssNull E e\<^sub>1 s\<^sub>0 s\<^sub>1 e\<^sub>2 v s\<^sub>2 F Cs e\<^sub>2' s\<^sub>2' T) have eval:"P,E \ \e\<^sub>1\F{Cs} := e\<^sub>2,s\<^sub>0\ \ \e\<^sub>2',s\<^sub>2'\" and wt:"P,E \ e\<^sub>1\F{Cs} := e\<^sub>2 :: T" and sconf:"P,E \ s\<^sub>0 \" and IH1:"\ei si T. \P,E \ \e\<^sub>1,s\<^sub>0\ \ \ei,si\; P,E \ e\<^sub>1 :: T; P,E \ s\<^sub>0 \\ \ null = ei \ s\<^sub>1 = si" and IH2:"\ei si T. \P,E \ \e\<^sub>2,s\<^sub>1\ \ \ei,si\; P,E \ e\<^sub>2 :: T; P,E \ s\<^sub>1 \\ \ Val v = ei \ s\<^sub>2 = si" by fact+ from wt obtain C T'' where wte1:"P,E \ e\<^sub>1 :: Class C" and wte2:"P,E \ e\<^sub>2 :: T''" by auto from eval show ?case proof(rule eval_cases) fix Xs D' S'' U a' fs'' h l s w w' assume eval_ref:"P,E \ \e\<^sub>1,s\<^sub>0\ \ \ref(a',Xs),s\" from IH1[OF eval_ref wte1 sconf] show "THROW NullPointer = e\<^sub>2' \ s\<^sub>2 = s\<^sub>2'" by simp next fix s w assume eval_null:"P,E \ \e\<^sub>1,s\<^sub>0\ \ \null,s\" and eval_val:"P,E \ \e\<^sub>2,s\ \ \Val w,s\<^sub>2'\" and e2':"e\<^sub>2' = THROW NullPointer" from IH1[OF eval_null wte1 sconf] have eq:"s = s\<^sub>1" by simp with wf eval_null wte1 sconf have sconf':"P,E \ s\<^sub>1 \" by(fastforce intro:eval_preserves_sconf) from IH2[OF eval_val[simplified eq] wte2 this] e2' show "THROW NullPointer = e\<^sub>2' \ s\<^sub>2 = s\<^sub>2'" by simp next fix ex assume eval_throw:"P,E \ \e\<^sub>1,s\<^sub>0\ \ \throw ex,s\<^sub>2'\" from IH1[OF eval_throw wte1 sconf] show "THROW NullPointer = e\<^sub>2' \ s\<^sub>2 = s\<^sub>2'" by simp next fix ex s w assume eval_val:"P,E \ \e\<^sub>1,s\<^sub>0\ \ \Val w,s\" and eval_throw:"P,E \ \e\<^sub>2,s\ \ \throw ex,s\<^sub>2'\" from IH1[OF eval_val wte1 sconf] have eq:"s = s\<^sub>1" by simp with wf eval_val wte1 sconf have sconf':"P,E \ s\<^sub>1 \" by(fastforce intro:eval_preserves_sconf) from IH2[OF eval_throw[simplified eq] wte2 this] show "THROW NullPointer = e\<^sub>2' \ s\<^sub>2 = s\<^sub>2'" by simp qed next case (FAssThrow1 E e\<^sub>1 s\<^sub>0 e' s\<^sub>1 F Cs e\<^sub>2 e\<^sub>2' s\<^sub>2 T) have eval:"P,E \ \e\<^sub>1\F{Cs} := e\<^sub>2,s\<^sub>0\ \ \e\<^sub>2',s\<^sub>2\" and wt:"P,E \ e\<^sub>1\F{Cs} := e\<^sub>2 :: T" and sconf:"P,E \ s\<^sub>0 \" and IH:"\ei si T. \P,E \ \e\<^sub>1,s\<^sub>0\ \ \ei,si\; P,E \ e\<^sub>1 :: T; P,E \ s\<^sub>0 \\ \ throw e' = ei \ s\<^sub>1 = si" by fact+ from wt obtain C T'' where wte1:"P,E \ e\<^sub>1 :: Class C" by auto from eval show ?case proof(rule eval_cases) fix Xs D' S'' U a' fs'' h l s w w' assume eval_ref:"P,E \ \e\<^sub>1,s\<^sub>0\ \ \ref(a',Xs),s\" from IH[OF eval_ref wte1 sconf] show "throw e' = e\<^sub>2' \ s\<^sub>1 = s\<^sub>2" by simp next fix s w assume eval_null:"P,E \ \e\<^sub>1,s\<^sub>0\ \ \null,s\" from IH[OF eval_null wte1 sconf] show "throw e' = e\<^sub>2' \ s\<^sub>1 = s\<^sub>2" by simp next fix ex assume eval_throw:"P,E \ \e\<^sub>1,s\<^sub>0\ \ \throw ex,s\<^sub>2\" and e2':"e\<^sub>2' = throw ex" from IH[OF eval_throw wte1 sconf] e2' show "throw e' = e\<^sub>2' \ s\<^sub>1 = s\<^sub>2" by simp next fix ex s w assume eval_val:"P,E \ \e\<^sub>1,s\<^sub>0\ \ \Val w,s\" from IH[OF eval_val wte1 sconf] show "throw e' = e\<^sub>2' \ s\<^sub>1 = s\<^sub>2" by simp qed next case (FAssThrow2 E e\<^sub>1 s\<^sub>0 v s\<^sub>1 e\<^sub>2 e' s\<^sub>2 F Cs e\<^sub>2' s\<^sub>2' T) have eval:"P,E \ \e\<^sub>1\F{Cs} := e\<^sub>2,s\<^sub>0\ \ \e\<^sub>2',s\<^sub>2'\" and wt:"P,E \ e\<^sub>1\F{Cs} := e\<^sub>2 :: T" and sconf:"P,E \ s\<^sub>0 \" and IH1:"\ei si T. \P,E \ \e\<^sub>1,s\<^sub>0\ \ \ei,si\; P,E \ e\<^sub>1 :: T; P,E \ s\<^sub>0 \\ \ Val v = ei \ s\<^sub>1 = si" and IH2:"\ei si T. \P,E \ \e\<^sub>2,s\<^sub>1\ \ \ei,si\; P,E \ e\<^sub>2 :: T; P,E \ s\<^sub>1 \\ \ throw e' = ei \ s\<^sub>2 = si" by fact+ from wt obtain C T'' where wte1:"P,E \ e\<^sub>1 :: Class C" and wte2:"P,E \ e\<^sub>2 :: T''" by auto from eval show ?case proof(rule eval_cases) fix Xs D' S'' U a' fs'' h l s w w' assume eval_ref:"P,E \ \e\<^sub>1,s\<^sub>0\ \ \ref(a',Xs),s\" and eval_val:"P,E \ \e\<^sub>2,s\ \ \Val w,(h,l)\" from IH1[OF eval_ref wte1 sconf] have eq:"s = s\<^sub>1" by simp with wf eval_ref wte1 sconf have sconf':"P,E \ s\<^sub>1 \" by(fastforce intro:eval_preserves_sconf) from IH2[OF eval_val[simplified eq] wte2 this] show "throw e' = e\<^sub>2' \ s\<^sub>2 = s\<^sub>2'" by simp next fix s w assume eval_null:"P,E \ \e\<^sub>1,s\<^sub>0\ \ \null,s\" and eval_val:"P,E \ \e\<^sub>2,s\ \ \Val w,s\<^sub>2'\" from IH1[OF eval_null wte1 sconf] have eq:"s = s\<^sub>1" by simp with wf eval_null wte1 sconf have sconf':"P,E \ s\<^sub>1 \" by(fastforce intro:eval_preserves_sconf) from IH2[OF eval_val[simplified eq] wte2 this] show "throw e' = e\<^sub>2' \ s\<^sub>2 = s\<^sub>2'" by simp next fix ex assume eval_throw:"P,E \ \e\<^sub>1,s\<^sub>0\ \ \throw ex,s\<^sub>2'\" from IH1[OF eval_throw wte1 sconf] show "throw e' = e\<^sub>2' \ s\<^sub>2 = s\<^sub>2'" by simp next fix ex s w assume eval_val:"P,E \ \e\<^sub>1,s\<^sub>0\ \ \Val w,s\" and eval_throw:"P,E \ \e\<^sub>2,s\ \ \throw ex,s\<^sub>2'\" and e2':"e\<^sub>2' = throw ex" from IH1[OF eval_val wte1 sconf] have eq:"s = s\<^sub>1" by simp with wf eval_val wte1 sconf have sconf':"P,E \ s\<^sub>1 \" by(fastforce intro:eval_preserves_sconf) from IH2[OF eval_throw[simplified eq] wte2 this] e2' show "throw e' = e\<^sub>2' \ s\<^sub>2 = s\<^sub>2'" by simp qed next case (CallObjThrow E e s\<^sub>0 e' s\<^sub>1 Copt M es e\<^sub>2 s\<^sub>2 T) have eval:"P,E \ \Call e Copt M es,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\" and wt:"P,E \ Call e Copt M es :: T" and sconf:"P,E \ s\<^sub>0 \" and IH:"\e\<^sub>2 s\<^sub>2 T. \P,E \ \e,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\; P,E \ e :: T; P,E \ s\<^sub>0 \\ \ throw e' = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by fact+ from wt obtain C where wte:"P,E \ e :: Class C" by(cases Copt)auto show ?case proof(cases Copt) assume "Copt = None" with eval have "P,E \ \e\M(es),s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\" by simp thus ?thesis proof(rule eval_cases) fix ex assume eval_throw:"P,E \ \e,s\<^sub>0\ \ \throw ex,s\<^sub>2\" and e2:"e\<^sub>2 = throw ex" from IH[OF eval_throw wte sconf] e2 show "throw e' = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix es' ex' s w ws assume eval_val:"P,E \ \e,s\<^sub>0\ \ \Val w,s\" from IH[OF eval_val wte sconf] show "throw e' = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix C' Xs Xs' Ds' S' U U' Us Us' a' body'' body''' h h' l l' pns'' pns''' s ws ws' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs),s\" from IH[OF eval_ref wte sconf] show "throw e' = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix s ws assume eval_null:"P,E \ \e,s\<^sub>0\ \ \null,s\" from IH[OF eval_null wte sconf] show "throw e' = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp qed next fix C' assume "Copt = Some C'" with eval have "P,E \ \e\(C'::)M(es),s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\" by simp thus ?thesis proof(rule eval_cases) fix ex assume eval_throw:"P,E \ \e,s\<^sub>0\ \ \throw ex,s\<^sub>2\" and e2:"e\<^sub>2 = throw ex" from IH[OF eval_throw wte sconf] e2 show "throw e' = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix es' ex' s w ws assume eval_val:"P,E \ \e,s\<^sub>0\ \ \Val w,s\" from IH[OF eval_val wte sconf] show "throw e' = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix C'' Xs Xs' Ds' S' U U' Us Us' a' body'' body''' h h' l l' pns'' pns''' s ws ws' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs),s\" from IH[OF eval_ref wte sconf] show "throw e' = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix s ws assume eval_null:"P,E \ \e,s\<^sub>0\ \ \null,s\" from IH[OF eval_null wte sconf] show "throw e' = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp qed qed next case (CallParamsThrow E e s\<^sub>0 v s\<^sub>1 es vs ex es' s\<^sub>2 Copt M e\<^sub>2 s\<^sub>2' T) have eval:"P,E \ \Call e Copt M es,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2'\" and wt:"P,E \ Call e Copt M es :: T" and sconf:"P,E \ s\<^sub>0 \" and IH1:"\ei si T. \P,E \ \e,s\<^sub>0\ \ \ei,si\; P,E \ e :: T; P,E \ s\<^sub>0 \\ \ Val v = ei \ s\<^sub>1 = si" and IH2:"\esi si Ts. \P,E \ \es,s\<^sub>1\ [\] \esi,si\; P,E \ es [::] Ts; P,E \ s\<^sub>1 \\ \ map Val vs @ throw ex # es' = esi \ s\<^sub>2 = si" by fact+ from wt obtain C Ts where wte:"P,E \ e :: Class C" and wtes:"P,E \ es [::] Ts" by(cases Copt)auto show ?case proof(cases Copt) assume "Copt = None" with eval have "P,E \ \e\M(es),s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2'\" by simp thus ?thesis proof(rule eval_cases) fix ex' assume eval_throw:"P,E \ \e,s\<^sub>0\ \ \throw ex',s\<^sub>2'\" from IH1[OF eval_throw wte sconf] show "throw ex = e\<^sub>2 \ s\<^sub>2 = s\<^sub>2'" by simp next fix es'' ex' s w ws assume eval_val:"P,E \ \e,s\<^sub>0\ \ \Val w,s\" and evals_throw:"P,E \ \es,s\ [\] \map Val ws@throw ex'#es'',s\<^sub>2'\" and e2:"e\<^sub>2 = throw ex'" from IH1[OF eval_val wte sconf] have eq:"s = s\<^sub>1" by simp with wf eval_val wte sconf have sconf':"P,E \ s\<^sub>1 \" by(fastforce intro:eval_preserves_sconf) from IH2[OF evals_throw[simplified eq] wtes this] e2 have "vs = ws \ ex = ex' \ es' = es'' \ s\<^sub>2 = s\<^sub>2'" by(fastforce dest:map_Val_throw_eq) with e2 show "throw ex = e\<^sub>2 \ s\<^sub>2 = s\<^sub>2'" by simp next fix C' Xs Xs' Ds' S' U U' Us Us' a' body'' body''' h h' l l' pns'' pns''' s ws ws' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs),s\" and evals_vals:"P,E \ \es,s\ [\] \map Val ws,(h,l)\" from IH1[OF eval_ref wte sconf] have eq:"s = s\<^sub>1" by simp with wf eval_ref wte sconf have sconf':"P,E \ s\<^sub>1 \" by(fastforce intro:eval_preserves_sconf) from IH2[OF evals_vals[simplified eq] wtes this] show "throw ex = e\<^sub>2 \ s\<^sub>2 = s\<^sub>2'" by(fastforce dest:sym[THEN map_Val_throw_False]) next fix s ws assume eval_null:"P,E \ \e,s\<^sub>0\ \ \null,s\" and evals_vals:"P,E \ \es,s\ [\] \map Val ws,s\<^sub>2'\" and e2:"e\<^sub>2 = THROW NullPointer" from IH1[OF eval_null wte sconf] have eq:"s = s\<^sub>1" by simp with wf eval_null wte sconf have sconf':"P,E \ s\<^sub>1 \" by(fastforce intro:eval_preserves_sconf) from IH2[OF evals_vals[simplified eq] wtes this] show "throw ex = e\<^sub>2 \ s\<^sub>2 = s\<^sub>2'" by(fastforce dest:sym[THEN map_Val_throw_False]) qed next fix C' assume "Copt = Some C'" with eval have "P,E \ \e\(C'::)M(es),s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2'\" by simp thus ?thesis proof(rule eval_cases) fix ex' assume eval_throw:"P,E \ \e,s\<^sub>0\ \ \throw ex',s\<^sub>2'\" from IH1[OF eval_throw wte sconf] show "throw ex = e\<^sub>2 \ s\<^sub>2 = s\<^sub>2'" by simp next fix es'' ex' s w ws assume eval_val:"P,E \ \e,s\<^sub>0\ \ \Val w,s\" and evals_throw:"P,E \ \es,s\ [\] \map Val ws@throw ex'#es'',s\<^sub>2'\" and e2:"e\<^sub>2 = throw ex'" from IH1[OF eval_val wte sconf] have eq:"s = s\<^sub>1" by simp with wf eval_val wte sconf have sconf':"P,E \ s\<^sub>1 \" by(fastforce intro:eval_preserves_sconf) from IH2[OF evals_throw[simplified eq] wtes this] e2 have "vs = ws \ ex = ex' \ es' = es'' \ s\<^sub>2 = s\<^sub>2'" by(fastforce dest:map_Val_throw_eq) with e2 show "throw ex = e\<^sub>2 \ s\<^sub>2 = s\<^sub>2'" by simp next fix C' Xs Xs' Ds' S' U U' Us Us' a' body'' body''' h h' l l' pns'' pns''' s ws ws' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs),s\" and evals_vals:"P,E \ \es,s\ [\] \map Val ws,(h,l)\" from IH1[OF eval_ref wte sconf] have eq:"s = s\<^sub>1" by simp with wf eval_ref wte sconf have sconf':"P,E \ s\<^sub>1 \" by(fastforce intro:eval_preserves_sconf) from IH2[OF evals_vals[simplified eq] wtes this] show "throw ex = e\<^sub>2 \ s\<^sub>2 = s\<^sub>2'" by(fastforce dest:sym[THEN map_Val_throw_False]) next fix s ws assume eval_null:"P,E \ \e,s\<^sub>0\ \ \null,s\" and evals_vals:"P,E \ \es,s\ [\] \map Val ws,s\<^sub>2'\" and e2:"e\<^sub>2 = THROW NullPointer" from IH1[OF eval_null wte sconf] have eq:"s = s\<^sub>1" by simp with wf eval_null wte sconf have sconf':"P,E \ s\<^sub>1 \" by(fastforce intro:eval_preserves_sconf) from IH2[OF evals_vals[simplified eq] wtes this] show "throw ex = e\<^sub>2 \ s\<^sub>2 = s\<^sub>2'" by(fastforce dest:sym[THEN map_Val_throw_False]) qed qed next case (Call E e s\<^sub>0 a Cs s\<^sub>1 es vs h\<^sub>2 l\<^sub>2 C S M Ts' T' pns' body' Ds Ts T pns body Cs' vs' l\<^sub>2' new_body e' h\<^sub>3 l\<^sub>3 e\<^sub>2 s\<^sub>2 T'') have eval:"P,E \ \e\M(es),s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\" and eval':"P,E \ \e,s\<^sub>0\ \ \ref(a,Cs),s\<^sub>1\" and eval'':"P,E \ \es,s\<^sub>1\ [\] \map Val vs,(h\<^sub>2,l\<^sub>2)\" and h2:"h\<^sub>2 a = Some(C,S)" and has_least:"P \ last Cs has least M = (Ts',T',pns',body') via Ds" and selects:"P \ (C,Cs@\<^sub>pDs) selects M = (Ts,T,pns,body) via Cs'" and length:"length vs = length pns" and Casts:"P \ Ts Casts vs to vs'" and l2':"l\<^sub>2' = [this \ Ref (a, Cs'), pns [\] vs']" and new_body:"new_body = (case T' of Class D \ \D\body | _ \ body)" and eval_body:"P,E(this \ Class (last Cs'), pns [\] Ts) \ \new_body,(h\<^sub>2,l\<^sub>2')\ \ \e',(h\<^sub>3,l\<^sub>3)\" and wt:"P,E \ e\M(es) :: T''" and sconf:"P,E \ s\<^sub>0 \" and IH1:"\ei si T. \P,E \ \e,s\<^sub>0\ \ \ei,si\; P,E \ e :: T; P,E \ s\<^sub>0 \\ \ ref (a,Cs) = ei \ s\<^sub>1 = si" and IH2:"\esi si Ts. \P,E \ \es,s\<^sub>1\ [\] \esi,si\; P,E \ es [::] Ts; P,E \ s\<^sub>1 \\ \ map Val vs = esi \ (h\<^sub>2,l\<^sub>2) = si" and IH3:"\ei si T. \P,E(this \ Class (last Cs'), pns [\] Ts) \ \new_body,(h\<^sub>2,l\<^sub>2')\ \ \ei,si\; P,E(this \ Class (last Cs'), pns [\] Ts) \ new_body :: T; P,E(this \ Class (last Cs'), pns [\] Ts) \ (h\<^sub>2,l\<^sub>2') \\ \ e' = ei \ (h\<^sub>3, l\<^sub>3) = si" by fact+ from wt obtain D Ss Ss' m Cs'' where wte:"P,E \ e :: Class D" and has_least':"P \ D has least M = (Ss,T'',m) via Cs''" and wtes:"P,E \ es [::] Ss'" and subs:"P \ Ss' [\] Ss" by auto from eval_preserves_type[OF wf eval' sconf wte] have last:"last Cs = D" by (auto split:if_split_asm) with has_least has_least' wf have eq:"Ts' = Ss \ T' = T'' \ (pns',body') = m \ Ds = Cs''" by(fastforce dest:wf_sees_method_fun) from wf selects have param_type:"\T \ set Ts. is_type P T" and return_type:"is_type P T" and TnotNT:"T \ NT" by(auto dest:select_method_wf_mdecl simp:wf_mdecl_def) from selects wf have subo:"Subobjs P C Cs'" by(induct rule:SelectMethodDef.induct, auto simp:FinalOverriderMethodDef_def OverriderMethodDefs_def MinimalMethodDefs_def LeastMethodDef_def MethodDefs_def) with wf have "class":"is_class P (last Cs')" by(auto intro!:Subobj_last_isClass) from eval'' have hext:"hp s\<^sub>1 \ h\<^sub>2" by (cases s\<^sub>1,auto intro: evals_hext) from wf eval' sconf wte last have "P,E,(hp s\<^sub>1) \ ref(a,Cs) :\<^bsub>NT\<^esub> Class(last Cs)" by -(rule eval_preserves_type,simp_all) with hext have "P,E,h\<^sub>2 \ ref(a,Cs) :\<^bsub>NT\<^esub> Class(last Cs)" by(auto intro:WTrt_hext_mono dest:hext_objD split:if_split_asm) with h2 have "Subobjs P C Cs" by (auto split:if_split_asm) hence "P \ Path C to (last Cs) via Cs" by (auto simp:path_via_def split:if_split_asm) with selects has_least wf have param_types:"Ts' = Ts \ P \ T \ T'" by -(rule select_least_methods_subtypes,simp_all) from wf selects have wt_body:"P,[this\Class(last Cs'),pns[\]Ts] \ body :: T" and this_not_pns:"this \ set pns" and length:"length pns = length Ts" and dist:"distinct pns" by(auto dest!:select_method_wf_mdecl simp:wf_mdecl_def) have "P,[this\Class(last Cs'),pns[\]Ts] \ new_body :: T'" proof(cases "\C. T' = Class C") case False with wt_body new_body param_types show ?thesis by(cases T') auto next case True then obtain D' where T':"T' = Class D'" by auto with wf has_least have "class":"is_class P D'" by(fastforce dest:has_least_wf_mdecl simp:wf_mdecl_def) with wf T' TnotNT param_types obtain D'' where T:"T = Class D''" by(fastforce dest:widen_Class) with wf return_type T' param_types have "P \ Path D'' to D' unique" by(simp add:Class_widen_Class) with wt_body "class" T T' new_body show ?thesis by auto qed hence wt_new_body:"P,E(this\Class(last Cs'),pns[\]Ts) \ new_body :: T'" by(fastforce intro:wt_env_mono) from eval show ?case proof(rule eval_cases) fix ex' assume eval_throw:"P,E \ \e,s\<^sub>0\ \ \throw ex',s\<^sub>2\" from IH1[OF eval_throw wte sconf] show "e' = e\<^sub>2 \ (h\<^sub>3, l\<^sub>2) = s\<^sub>2" by simp next fix es'' ex' s w ws assume eval_val:"P,E \ \e,s\<^sub>0\ \ \Val w,s\" and evals_throw:"P,E \ \es,s\ [\] \map Val ws@throw ex'#es'',s\<^sub>2\" from IH1[OF eval_val wte sconf] have eq:"s = s\<^sub>1" by simp with wf eval_val wte sconf have sconf':"P,E \ s\<^sub>1 \" by(fastforce intro:eval_preserves_sconf) from IH2[OF evals_throw[simplified eq] wtes this] show "e' = e\<^sub>2 \ (h\<^sub>3, l\<^sub>2) = s\<^sub>2" by(fastforce dest:map_Val_throw_False) next fix C' Xs Xs' Ds' S' U U' Us Us' a' body'' body''' h h' l l' pns'' pns''' s ws ws' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs),s\" and evals_vals:"P,E \ \es,s\ [\] \map Val ws,(h,l)\" and h:"h a' = Some(C',S')" and has_least'':"P \ last Xs has least M = (Us',U',pns''',body''') via Ds'" and selects':"P \ (C',Xs@\<^sub>pDs') selects M = (Us,U,pns'',body'') via Xs'" and length':"length ws = length pns''" and Casts':"P \ Us Casts ws to ws'" and eval_body':"P,E(this \ Class (last Xs'), pns'' [\] Us) \ \case U' of Class D \ \D\body'' | _ \ body'', (h,[this \ Ref(a',Xs'), pns'' [\] ws'])\ \ \e\<^sub>2,(h',l')\" and s2:"s\<^sub>2 = (h',l)" from IH1[OF eval_ref wte sconf] have eq1:"a = a' \ Cs = Xs" and s:"s = s\<^sub>1" by simp_all with has_least has_least'' wf have eq2:"T' = U' \ Ts' = Us' \ Ds = Ds'" by(fastforce dest:wf_sees_method_fun) from s wf eval_ref wte sconf have sconf':"P,E \ s\<^sub>1 \" by(fastforce intro:eval_preserves_sconf) from IH2[OF evals_vals[simplified s] wtes this] have eq3:"vs = ws \ h\<^sub>2 = h \ l\<^sub>2 = l" by(fastforce elim:map_injective simp:inj_on_def) with eq1 h2 h have eq4:"C = C' \ S = S'" by simp with eq1 eq2 selects selects' wf have eq5:"Ts = Us \ T = U \ pns'' = pns \ body'' = body \ Cs' = Xs'" by simp(drule_tac mthd'="(Us,U,pns'',body'')" in wf_select_method_fun,auto) with subs eq param_types have "P \ Ss' [\] Us" by simp with wf Casts Casts' param_type wtes evals_vals sconf' s eq eq2 eq3 eq5 have eq6:"vs' = ws'" by(fastforce intro:Casts_Casts_eq_result) with eval_body' l2' eq1 eq2 eq3 eq5 new_body have eval_body'':"P,E(this \ Class(last Cs'), pns [\] Ts) \ \new_body,(h\<^sub>2,l\<^sub>2')\ \ \e\<^sub>2,(h',l')\" by fastforce from wf evals_vals wtes sconf' s eq3 have sconf'':"P,E \ (h\<^sub>2,l\<^sub>2) \" by(fastforce intro:evals_preserves_sconf) have "P,E(this \ Class (last Cs'), pns [\] Ts) \ (h\<^sub>2,l\<^sub>2') \" proof(auto simp:sconf_def) from sconf'' show "P \ h\<^sub>2 \" by(simp add:sconf_def) next { fix V v assume map:"[this \ Ref (a,Cs'), pns [\] vs'] V = Some v" have "\T. (E(this \ Class (last Cs'), pns [\] Ts)) V = Some T \ P,h\<^sub>2 \ v :\ T" proof(cases "V \ set (this#pns)") case False with map show ?thesis by simp next case True hence "V = this \ V \ set pns" by simp thus ?thesis proof(rule disjE) assume V:"V = this" with map this_not_pns have "v = Ref(a,Cs')" by simp with V h2 subo this_not_pns have "(E(this \ Class (last Cs'),pns [\] Ts)) V = Some(Class (last Cs'))" and "P,h\<^sub>2 \ v :\ Class (last Cs')" by simp_all thus ?thesis by simp next assume "V \ set pns" then obtain i where V:"V = pns!i" and length_i:"i < length pns" by(auto simp:in_set_conv_nth) from Casts have "length Ts = length vs'" by(induct rule:Casts_to.induct,auto) with length have "length pns = length vs'" by simp with map dist V length_i have v:"v = vs'!i" by(fastforce dest:maps_nth) from length dist length_i - have env:"(E(this \ Class (last Cs'))(pns [\] Ts)) (pns!i) = Some(Ts!i)" + have env:"(E(this \ Class (last Cs'), pns [\] Ts)) (pns!i) = Some(Ts!i)" by(rule_tac E="E(this \ Class (last Cs'))" in nth_maps,simp_all) from wf Casts wtes subs eq param_types eval'' sconf' have "\i < length Ts. P,h\<^sub>2 \ vs'!i :\ Ts!i" by simp(rule Casts_conf,auto) with length_i length env V v show ?thesis by simp qed qed } thus "P,h\<^sub>2 \ l\<^sub>2' (:\)\<^sub>w E(this \ Class (last Cs'), pns [\] Ts)" using l2' by(simp add:lconf_def) next { fix V Tx assume env:"(E(this \ Class (last Cs'), pns [\] Ts)) V = Some Tx" have "is_type P Tx" proof(cases "V \ set (this#pns)") case False with env sconf'' show ?thesis by(clarsimp simp:sconf_def envconf_def) next case True hence "V = this \ V \ set pns" by simp thus ?thesis proof(rule disjE) assume "V = this" with env this_not_pns have "Tx = Class(last Cs')" by simp with "class" show ?thesis by simp next assume "V \ set pns" then obtain i where V:"V = pns!i" and length_i:"i < length pns" by(auto simp:in_set_conv_nth) with dist length env have "Tx = Ts!i" by(fastforce dest:maps_nth) with length_i length have "Tx \ set Ts" by(fastforce simp:in_set_conv_nth) with param_type show ?thesis by simp qed qed } thus "P \ E(this \ Class (last Cs'), pns [\] Ts) \" by (simp add:envconf_def) qed from IH3[OF eval_body'' wt_new_body this] have "e' = e\<^sub>2 \ (h\<^sub>3, l\<^sub>3) = (h',l')" . with eq3 s2 show "e' = e\<^sub>2 \ (h\<^sub>3,l\<^sub>2) = s\<^sub>2" by simp next fix s ws assume eval_null:"P,E \ \e,s\<^sub>0\ \ \null,s\" from IH1[OF eval_null wte sconf] show "e' = e\<^sub>2 \ (h\<^sub>3,l\<^sub>2) = s\<^sub>2" by simp qed next case (StaticCall E e s\<^sub>0 a Cs s\<^sub>1 es vs h\<^sub>2 l\<^sub>2 C Cs'' M Ts T pns body Cs' Ds vs' l\<^sub>2' e' h\<^sub>3 l\<^sub>3 e\<^sub>2 s\<^sub>2 T') have eval:"P,E \ \e\(C::)M(es),s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\" and eval':"P,E \ \e,s\<^sub>0\ \ \ref(a,Cs),s\<^sub>1\" and eval'':"P,E \ \es,s\<^sub>1\ [\] \map Val vs,(h\<^sub>2, l\<^sub>2)\" and path_unique:"P \ Path last Cs to C unique" and path_via:"P \ Path last Cs to C via Cs''" and has_least:"P \ C has least M = (Ts,T,pns,body) via Cs'" and Ds:"Ds = (Cs@\<^sub>pCs'')@\<^sub>pCs'" and length:"length vs = length pns" and Casts:"P \ Ts Casts vs to vs'" and l2':"l\<^sub>2' = [this \ Ref (a, Ds), pns [\] vs']" and eval_body:"P,E(this \ Class (last Ds), pns [\] Ts) \ \body,(h\<^sub>2,l\<^sub>2')\ \ \e',(h\<^sub>3,l\<^sub>3)\" and wt:"P,E \ e\(C::)M(es) :: T'" and sconf:"P,E \ s\<^sub>0 \" and IH1:"\ei si T. \P,E \ \e,s\<^sub>0\ \ \ei,si\; P,E \ e :: T; P,E \ s\<^sub>0 \\ \ ref (a,Cs) = ei \ s\<^sub>1 = si" and IH2:"\esi si Ts. \P,E \ \es,s\<^sub>1\ [\] \esi,si\; P,E \ es [::] Ts; P,E \ s\<^sub>1 \\ \ map Val vs = esi \ (h\<^sub>2,l\<^sub>2) = si" and IH3:"\ei si T. \P,E(this \ Class (last Ds), pns [\] Ts) \ \body,(h\<^sub>2,l\<^sub>2')\ \ \ei,si\; P,E(this \ Class (last Ds), pns [\] Ts) \ body :: T; P,E(this \ Class (last Ds), pns [\] Ts) \ (h\<^sub>2,l\<^sub>2') \\ \ e' = ei \ (h\<^sub>3, l\<^sub>3) = si" by fact+ from wt has_least wf obtain C' Ts' where wte:"P,E \ e :: Class C'" and wtes:"P,E \ es [::] Ts'" and subs:"P \ Ts' [\] Ts" by(auto dest:wf_sees_method_fun) from eval_preserves_type[OF wf eval' sconf wte] have last:"last Cs = C'" by (auto split:if_split_asm) from wf has_least have param_type:"\T \ set Ts. is_type P T" and return_type:"is_type P T" and TnotNT:"T \ NT" by(auto dest:has_least_wf_mdecl simp:wf_mdecl_def) from path_via have last':"last Cs'' = last(Cs@\<^sub>pCs'')" by(fastforce intro!:appendPath_last Subobjs_nonempty simp:path_via_def) from eval'' have hext:"hp s\<^sub>1 \ h\<^sub>2" by (cases s\<^sub>1,auto intro: evals_hext) from wf eval' sconf wte last have "P,E,(hp s\<^sub>1) \ ref(a,Cs) :\<^bsub>NT\<^esub> Class(last Cs)" by -(rule eval_preserves_type,simp_all) with hext have "P,E,h\<^sub>2 \ ref(a,Cs) :\<^bsub>NT\<^esub> Class(last Cs)" by(auto intro:WTrt_hext_mono dest:hext_objD split:if_split_asm) then obtain D S where h2:"h\<^sub>2 a = Some(D,S)" and "Subobjs P D Cs" by (auto split:if_split_asm) with path_via wf have "Subobjs P D (Cs@\<^sub>pCs'')" and "last Cs'' = C" by(auto intro:Subobjs_appendPath simp:path_via_def) with has_least wf last' Ds have subo:"Subobjs P D Ds" by(fastforce intro:Subobjs_appendPath simp:LeastMethodDef_def MethodDefs_def) with wf have "class":"is_class P (last Ds)" by(auto intro!:Subobj_last_isClass) from has_least wf obtain D' where "Subobjs P D' Cs'" by(auto simp:LeastMethodDef_def MethodDefs_def) with Ds have last_Ds:"last Cs' = last Ds" by(fastforce intro!:appendPath_last Subobjs_nonempty) with wf has_least have "P,[this\Class(last Ds),pns[\]Ts] \ body :: T" and this_not_pns:"this \ set pns" and length:"length pns = length Ts" and dist:"distinct pns" by(auto dest!:has_least_wf_mdecl simp:wf_mdecl_def) hence wt_body:"P,E(this\Class(last Ds),pns[\]Ts) \ body :: T" by(fastforce intro:wt_env_mono) from eval show ?case proof(rule eval_cases) fix ex' assume eval_throw:"P,E \ \e,s\<^sub>0\ \ \throw ex',s\<^sub>2\" from IH1[OF eval_throw wte sconf] show "e' = e\<^sub>2 \ (h\<^sub>3, l\<^sub>2) = s\<^sub>2" by simp next fix es'' ex' s w ws assume eval_val:"P,E \ \e,s\<^sub>0\ \ \Val w,s\" and evals_throw:"P,E \ \es,s\ [\] \map Val ws@throw ex'#es'',s\<^sub>2\" from IH1[OF eval_val wte sconf] have eq:"s = s\<^sub>1" by simp with wf eval_val wte sconf have sconf':"P,E \ s\<^sub>1 \" by(fastforce intro:eval_preserves_sconf) from IH2[OF evals_throw[simplified eq] wtes this] show "e' = e\<^sub>2 \ (h\<^sub>3, l\<^sub>2) = s\<^sub>2" by(fastforce dest:map_Val_throw_False) next fix Xs Xs' Xs'' U Us a' body' h h' l l' pns' s ws ws' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs),s\" and evals_vals:"P,E \ \es,s\ [\] \map Val ws,(h,l)\" and path_unique':"P \ Path last Xs to C unique" and path_via':"P \ Path last Xs to C via Xs''" and has_least':"P \ C has least M = (Us,U,pns',body') via Xs'" and length':"length ws = length pns'" and Casts':"P \ Us Casts ws to ws'" and eval_body':"P,E(this \ Class(last((Xs@\<^sub>pXs'')@\<^sub>pXs')),pns' [\] Us) \ \body',(h,[this \ Ref(a',(Xs@\<^sub>pXs'')@\<^sub>pXs'),pns' [\] ws'])\ \ \e\<^sub>2,(h',l')\" and s2:"s\<^sub>2 = (h',l)" from IH1[OF eval_ref wte sconf] have eq1:"a = a' \ Cs = Xs" and s:"s = s\<^sub>1" by simp_all from has_least has_least' wf have eq2:"T = U \ Ts = Us \ Cs' = Xs' \ pns = pns' \ body = body'" by(fastforce dest:wf_sees_method_fun) from s wf eval_ref wte sconf have sconf':"P,E \ s\<^sub>1 \" by(fastforce intro:eval_preserves_sconf) from IH2[OF evals_vals[simplified s] wtes this] have eq3:"vs = ws \ h\<^sub>2 = h \ l\<^sub>2 = l" by(fastforce elim:map_injective simp:inj_on_def) from path_unique path_via path_via' eq1 have "Cs'' = Xs''" by(fastforce simp:path_unique_def path_via_def) with Ds eq1 eq2 have Ds':"Ds = (Xs@\<^sub>pXs'')@\<^sub>pXs'" by simp from wf Casts Casts' param_type wtes subs evals_vals sconf' s eq2 eq3 have eq4:"vs' = ws'" by(fastforce intro:Casts_Casts_eq_result) with eval_body' Ds' l2' eq1 eq2 eq3 have eval_body'':"P,E(this \ Class(last Ds),pns [\] Ts) \ \body,(h\<^sub>2,l\<^sub>2')\ \ \e\<^sub>2,(h',l')\" by simp from wf evals_vals wtes sconf' s eq3 have sconf'':"P,E \ (h\<^sub>2,l\<^sub>2) \" by(fastforce intro:evals_preserves_sconf) have "P,E(this \ Class (last Ds), pns [\] Ts) \ (h\<^sub>2,l\<^sub>2') \" proof(auto simp:sconf_def) from sconf'' show "P \ h\<^sub>2 \" by(simp add:sconf_def) next { fix V v assume map:"[this \ Ref (a,Ds), pns [\] vs'] V = Some v" have "\T. (E(this \ Class (last Ds), pns [\] Ts)) V = Some T \ P,h\<^sub>2 \ v :\ T" proof(cases "V \ set (this#pns)") case False with map show ?thesis by simp next case True hence "V = this \ V \ set pns" by simp thus ?thesis proof(rule disjE) assume V:"V = this" with map this_not_pns have "v = Ref(a,Ds)" by simp with V h2 subo this_not_pns have "(E(this \ Class (last Ds),pns [\] Ts)) V = Some(Class (last Ds))" and "P,h\<^sub>2 \ v :\ Class (last Ds)" by simp_all thus ?thesis by simp next assume "V \ set pns" then obtain i where V:"V = pns!i" and length_i:"i < length pns" by(auto simp:in_set_conv_nth) from Casts have "length Ts = length vs'" by(induct rule:Casts_to.induct,auto) with length have "length pns = length vs'" by simp with map dist V length_i have v:"v = vs'!i" by(fastforce dest:maps_nth) from length dist length_i - have env:"(E(this \ Class (last Ds))(pns [\] Ts)) (pns!i) = Some(Ts!i)" + have env:"(E(this \ Class (last Ds), pns [\] Ts)) (pns!i) = Some(Ts!i)" by(rule_tac E="E(this \ Class (last Ds))" in nth_maps,simp_all) from wf Casts wtes subs eval'' sconf' have "\i < length Ts. P,h\<^sub>2 \ vs'!i :\ Ts!i" by -(rule Casts_conf,auto) with length_i length env V v show ?thesis by simp qed qed } thus "P,h\<^sub>2 \ l\<^sub>2' (:\)\<^sub>w E(this \ Class (last Ds), pns [\] Ts)" using l2' by(simp add:lconf_def) next { fix V Tx assume env:"(E(this \ Class (last Ds), pns [\] Ts)) V = Some Tx" have "is_type P Tx" proof(cases "V \ set (this#pns)") case False with env sconf'' show ?thesis by(clarsimp simp:sconf_def envconf_def) next case True hence "V = this \ V \ set pns" by simp thus ?thesis proof(rule disjE) assume "V = this" with env this_not_pns have "Tx = Class(last Ds)" by simp with "class" show ?thesis by simp next assume "V \ set pns" then obtain i where V:"V = pns!i" and length_i:"i < length pns" by(auto simp:in_set_conv_nth) with dist length env have "Tx = Ts!i" by(fastforce dest:maps_nth) with length_i length have "Tx \ set Ts" by(fastforce simp:in_set_conv_nth) with param_type show ?thesis by simp qed qed } thus "P \ E(this \ Class (last Ds), pns [\] Ts) \" by (simp add:envconf_def) qed from IH3[OF eval_body'' wt_body this] have "e' = e\<^sub>2 \ (h\<^sub>3, l\<^sub>3) = (h',l')" . with eq3 s2 show "e' = e\<^sub>2 \ (h\<^sub>3, l\<^sub>2) = s\<^sub>2" by simp next fix s ws assume eval_null:"P,E \ \e,s\<^sub>0\ \ \null,s\" from IH1[OF eval_null wte sconf] show "e' = e\<^sub>2 \ (h\<^sub>3,l\<^sub>2) = s\<^sub>2" by simp qed next case (CallNull E e s\<^sub>0 s\<^sub>1 es vs s\<^sub>2 Copt M e\<^sub>2 s\<^sub>2' T) have eval:"P,E \ \Call e Copt M es,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2'\" and wt:"P,E \ Call e Copt M es :: T" and sconf:"P,E \ s\<^sub>0 \" and IH1:"\ei si T. \P,E \ \e,s\<^sub>0\ \ \ei,si\; P,E \ e :: T; P,E \ s\<^sub>0 \\ \ null = ei \ s\<^sub>1 = si" and IH2:"\esi si Ts. \P,E \ \es,s\<^sub>1\ [\] \esi,si\; P,E \ es [::] Ts; P,E \ s\<^sub>1 \\ \ map Val vs = esi \ s\<^sub>2 = si" by fact+ from wt obtain C Ts where wte:"P,E \ e :: Class C" and wtes:"P,E \ es [::] Ts" by(cases Copt)auto show ?case proof(cases Copt) assume "Copt = None" with eval have "P,E \ \e\M(es),s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2'\" by simp thus ?thesis proof(rule eval_cases) fix ex' assume eval_throw:"P,E \ \e,s\<^sub>0\ \ \throw ex',s\<^sub>2'\" from IH1[OF eval_throw wte sconf] show "THROW NullPointer = e\<^sub>2 \ s\<^sub>2 = s\<^sub>2'" by simp next fix es' ex' s w ws assume eval_val:"P,E \ \e,s\<^sub>0\ \ \Val w,s\" and evals_throw:"P,E \ \es,s\ [\] \map Val ws@throw ex'#es',s\<^sub>2'\" from IH1[OF eval_val wte sconf] have eq:"s = s\<^sub>1" by simp with wf eval_val wte sconf have sconf':"P,E \ s\<^sub>1 \" by(fastforce intro:eval_preserves_sconf) from IH2[OF evals_throw[simplified eq] wtes this] show "THROW NullPointer = e\<^sub>2 \ s\<^sub>2 = s\<^sub>2'" by(fastforce dest:map_Val_throw_False) next fix C' Xs Xs' Ds' S' U U' Us Us' a' body'' body''' h h' l l' pns'' pns''' s ws ws' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs),s\" from IH1[OF eval_ref wte sconf] show "THROW NullPointer = e\<^sub>2 \ s\<^sub>2 = s\<^sub>2'" by simp next fix s ws assume eval_null:"P,E \ \e,s\<^sub>0\ \ \null,s\" and evals_vals:"P,E \ \es,s\ [\] \map Val ws,s\<^sub>2'\" and e2:"e\<^sub>2 = THROW NullPointer" from IH1[OF eval_null wte sconf] have eq:"s = s\<^sub>1" by simp with wf eval_null wte sconf have sconf':"P,E \ s\<^sub>1 \" by(fastforce intro:eval_preserves_sconf) from IH2[OF evals_vals[simplified eq] wtes this] e2 show "THROW NullPointer = e\<^sub>2 \ s\<^sub>2 = s\<^sub>2'" by simp qed next fix C' assume "Copt = Some C'" with eval have "P,E \ \e\(C'::)M(es),s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2'\" by simp thus ?thesis proof(rule eval_cases) fix ex' assume eval_throw:"P,E \ \e,s\<^sub>0\ \ \throw ex',s\<^sub>2'\" from IH1[OF eval_throw wte sconf] show "THROW NullPointer = e\<^sub>2 \ s\<^sub>2 = s\<^sub>2'" by simp next fix es' ex' s w ws assume eval_val:"P,E \ \e,s\<^sub>0\ \ \Val w,s\" and evals_throw:"P,E \ \es,s\ [\] \map Val ws@throw ex'#es',s\<^sub>2'\" from IH1[OF eval_val wte sconf] have eq:"s = s\<^sub>1" by simp with wf eval_val wte sconf have sconf':"P,E \ s\<^sub>1 \" by(fastforce intro:eval_preserves_sconf) from IH2[OF evals_throw[simplified eq] wtes this] show "THROW NullPointer = e\<^sub>2 \ s\<^sub>2 = s\<^sub>2'" by(fastforce dest:map_Val_throw_False) next fix C' Xs Xs' Ds' S' U U' Us Us' a' body'' body''' h h' l l' pns'' pns''' s ws ws' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref(a',Xs),s\" from IH1[OF eval_ref wte sconf] show "THROW NullPointer = e\<^sub>2 \ s\<^sub>2 = s\<^sub>2'" by simp next fix s ws assume eval_null:"P,E \ \e,s\<^sub>0\ \ \null,s\" and evals_vals:"P,E \ \es,s\ [\] \map Val ws,s\<^sub>2'\" and e2:"e\<^sub>2 = THROW NullPointer" from IH1[OF eval_null wte sconf] have eq:"s = s\<^sub>1" by simp with wf eval_null wte sconf have sconf':"P,E \ s\<^sub>1 \" by(fastforce intro:eval_preserves_sconf) from IH2[OF evals_vals[simplified eq] wtes this] e2 show "THROW NullPointer = e\<^sub>2 \ s\<^sub>2 = s\<^sub>2'" by simp qed qed next case (Block E V T e\<^sub>0 h\<^sub>0 l\<^sub>0 e\<^sub>1 h\<^sub>1 l\<^sub>1 e\<^sub>2 s\<^sub>2 T') have eval:"P,E \ \{V:T; e\<^sub>0},(h\<^sub>0, l\<^sub>0)\ \ \e\<^sub>2,s\<^sub>2\" and wt:"P,E \ {V:T; e\<^sub>0} :: T'" and sconf:"P,E \ (h\<^sub>0, l\<^sub>0) \" and IH:"\e\<^sub>2 s\<^sub>2 T'. \P,E(V \ T) \ \e\<^sub>0,(h\<^sub>0, l\<^sub>0(V := None))\ \ \e\<^sub>2,s\<^sub>2\; P,E(V \ T) \ e\<^sub>0 :: T'; P,E(V \ T) \ (h\<^sub>0, l\<^sub>0(V := None)) \\ \ e\<^sub>1 = e\<^sub>2 \ (h\<^sub>1, l\<^sub>1) = s\<^sub>2" by fact+ from wt have type:"is_type P T" and wte:"P,E(V \ T) \ e\<^sub>0 :: T'" by auto from sconf type have sconf':"P,E(V \ T) \ (h\<^sub>0, l\<^sub>0(V := None)) \" by(auto simp:sconf_def lconf_def envconf_def) from eval obtain h l where eval':"P,E(V \ T) \ \e\<^sub>0,(h\<^sub>0,l\<^sub>0(V:=None))\ \ \e\<^sub>2,(h,l)\" and s2:"s\<^sub>2 = (h,l(V:=l\<^sub>0 V))" by (auto elim:eval_cases) from IH[OF eval' wte sconf'] s2 show ?case by simp next case (Seq E e\<^sub>0 s\<^sub>0 v s\<^sub>1 e\<^sub>1 e\<^sub>2 s\<^sub>2 e\<^sub>2' s\<^sub>2' T) have eval:"P,E \ \e\<^sub>0;; e\<^sub>1,s\<^sub>0\ \ \e\<^sub>2',s\<^sub>2'\" and wt:"P,E \ e\<^sub>0;; e\<^sub>1 :: T" and sconf:"P,E \ s\<^sub>0 \" and IH1:"\ei si T. \P,E \ \e\<^sub>0,s\<^sub>0\ \ \ei,si\; P,E \ e\<^sub>0 :: T; P,E \ s\<^sub>0 \\ \ Val v = ei \ s\<^sub>1 = si" and IH2:"\ei si T. \P,E \ \e\<^sub>1,s\<^sub>1\ \ \ei,si\; P,E \ e\<^sub>1 :: T; P,E \ s\<^sub>1 \\ \ e\<^sub>2 = ei \ s\<^sub>2 = si" by fact+ from wt obtain T' where wte0:"P,E \ e\<^sub>0 :: T'" and wte1:"P,E \ e\<^sub>1 :: T" by auto from eval show ?case proof(rule eval_cases) fix s w assume eval_val:"P,E \ \e\<^sub>0,s\<^sub>0\ \ \Val w,s\" and eval':"P,E \ \e\<^sub>1,s\ \ \e\<^sub>2',s\<^sub>2'\" from IH1[OF eval_val wte0 sconf] have eq:"s = s\<^sub>1" by simp with wf eval_val wte0 sconf have "P,E \ s\<^sub>1 \" by(fastforce intro:eval_preserves_sconf) from IH2[OF eval'[simplified eq] wte1 this] show "e\<^sub>2 = e\<^sub>2' \ s\<^sub>2 = s\<^sub>2'" . next fix ex assume eval_throw:"P,E \ \e\<^sub>0,s\<^sub>0\ \ \throw ex,s\<^sub>2'\" from IH1[OF eval_throw wte0 sconf] show "e\<^sub>2 = e\<^sub>2' \ s\<^sub>2 = s\<^sub>2'" by simp qed next case (SeqThrow E e\<^sub>0 s\<^sub>0 e s\<^sub>1 e\<^sub>1 e\<^sub>2 s\<^sub>2 T) have eval:"P,E \ \e\<^sub>0;; e\<^sub>1,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\" and wt:"P,E \ e\<^sub>0;; e\<^sub>1 :: T" and sconf:"P,E \ s\<^sub>0 \" and IH:"\ei si T. \P,E \ \e\<^sub>0,s\<^sub>0\ \ \ei,si\; P,E \ e\<^sub>0 :: T; P,E \ s\<^sub>0 \\ \ throw e = ei \ s\<^sub>1 = si" by fact+ from wt obtain T' where wte0:"P,E \ e\<^sub>0 :: T'" by auto from eval show ?case proof(rule eval_cases) fix s w assume eval_val:"P,E \ \e\<^sub>0,s\<^sub>0\ \ \Val w,s\" from IH[OF eval_val wte0 sconf] show "throw e = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix ex assume eval_throw:"P,E \ \e\<^sub>0,s\<^sub>0\ \ \throw ex,s\<^sub>2\" and e2:"e\<^sub>2 = throw ex" from IH[OF eval_throw wte0 sconf] e2 show "throw e = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp qed next case (CondT E e s\<^sub>0 s\<^sub>1 e\<^sub>1 e' s\<^sub>2 e\<^sub>2 e\<^sub>2' s\<^sub>2' T) have eval:"P,E \ \if (e) e\<^sub>1 else e\<^sub>2,s\<^sub>0\ \ \e\<^sub>2',s\<^sub>2'\" and wt:"P,E \ if (e) e\<^sub>1 else e\<^sub>2 :: T" and sconf:"P,E \ s\<^sub>0 \" and IH1:"\ei si T. \P,E \ \e,s\<^sub>0\ \ \ei,si\; P,E \ e :: T; P,E \ s\<^sub>0 \\ \ true = ei \ s\<^sub>1 = si" and IH2:"\ei si T. \P,E \ \e\<^sub>1,s\<^sub>1\ \ \ei,si\; P,E \ e\<^sub>1 :: T; P,E \ s\<^sub>1 \\ \ e' = ei \ s\<^sub>2 = si" by fact+ from wt have wte:"P,E \ e :: Boolean" and wte1:"P,E \ e\<^sub>1 :: T" by auto from eval show ?case proof(rule eval_cases) fix s assume eval_true:"P,E \ \e,s\<^sub>0\ \ \true,s\" and eval':"P,E \ \e\<^sub>1,s\ \ \e\<^sub>2',s\<^sub>2'\" from IH1[OF eval_true wte sconf] have eq:"s = s\<^sub>1" by simp with wf eval_true wte sconf have "P,E \ s\<^sub>1 \" by(fastforce intro:eval_preserves_sconf) from IH2[OF eval'[simplified eq] wte1 this] show "e' = e\<^sub>2' \ s\<^sub>2 = s\<^sub>2'" . next fix s assume eval_false:"P,E \ \e,s\<^sub>0\ \ \false,s\" from IH1[OF eval_false wte sconf] show "e' = e\<^sub>2' \ s\<^sub>2 = s\<^sub>2'" by simp next fix ex assume eval_throw:"P,E \ \e,s\<^sub>0\ \ \throw ex,s\<^sub>2'\" from IH1[OF eval_throw wte sconf] show "e' = e\<^sub>2' \ s\<^sub>2 = s\<^sub>2'" by simp qed next case (CondF E e s\<^sub>0 s\<^sub>1 e\<^sub>2 e' s\<^sub>2 e\<^sub>1 e\<^sub>2' s\<^sub>2' T) have eval:"P,E \ \if (e) e\<^sub>1 else e\<^sub>2,s\<^sub>0\ \ \e\<^sub>2',s\<^sub>2'\" and wt:"P,E \ if (e) e\<^sub>1 else e\<^sub>2 :: T" and sconf:"P,E \ s\<^sub>0 \" and IH1:"\ei si T. \P,E \ \e,s\<^sub>0\ \ \ei,si\; P,E \ e :: T; P,E \ s\<^sub>0 \\ \ false = ei \ s\<^sub>1 = si" and IH2:"\ei si T. \P,E \ \e\<^sub>2,s\<^sub>1\ \ \ei,si\; P,E \ e\<^sub>2 :: T; P,E \ s\<^sub>1 \\ \ e' = ei \ s\<^sub>2 = si" by fact+ from wt have wte:"P,E \ e :: Boolean" and wte2:"P,E \ e\<^sub>2 :: T" by auto from eval show ?case proof(rule eval_cases) fix s assume eval_true:"P,E \ \e,s\<^sub>0\ \ \true,s\" from IH1[OF eval_true wte sconf] show "e' = e\<^sub>2' \ s\<^sub>2 = s\<^sub>2'" by simp next fix s assume eval_false:"P,E \ \e,s\<^sub>0\ \ \false,s\" and eval':"P,E \ \e\<^sub>2,s\ \ \e\<^sub>2',s\<^sub>2'\" from IH1[OF eval_false wte sconf] have eq:"s = s\<^sub>1" by simp with wf eval_false wte sconf have "P,E \ s\<^sub>1 \" by(fastforce intro:eval_preserves_sconf) from IH2[OF eval'[simplified eq] wte2 this] show "e' = e\<^sub>2' \ s\<^sub>2 = s\<^sub>2'" . next fix ex assume eval_throw:"P,E \ \e,s\<^sub>0\ \ \throw ex,s\<^sub>2'\" from IH1[OF eval_throw wte sconf] show "e' = e\<^sub>2' \ s\<^sub>2 = s\<^sub>2'" by simp qed next case (CondThrow E e s\<^sub>0 e' s\<^sub>1 e\<^sub>1 e\<^sub>2 e\<^sub>2' s\<^sub>2 T) have eval:"P,E \ \if (e) e\<^sub>1 else e\<^sub>2,s\<^sub>0\ \ \e\<^sub>2',s\<^sub>2\" and wt:"P,E \ if (e) e\<^sub>1 else e\<^sub>2 :: T" and sconf:"P,E \ s\<^sub>0 \" and IH:"\ei si T. \P,E \ \e,s\<^sub>0\ \ \ei,si\; P,E \ e :: T; P,E \ s\<^sub>0 \\ \ throw e' = ei \ s\<^sub>1 = si" by fact+ from wt have wte:"P,E \ e :: Boolean" by auto from eval show ?case proof(rule eval_cases) fix s assume eval_true:"P,E \ \e,s\<^sub>0\ \ \true,s\" from IH[OF eval_true wte sconf] show "throw e' = e\<^sub>2' \ s\<^sub>1 = s\<^sub>2" by simp next fix s assume eval_false:"P,E \ \e,s\<^sub>0\ \ \false,s\" from IH[OF eval_false wte sconf] show "throw e' = e\<^sub>2' \ s\<^sub>1 = s\<^sub>2" by simp next fix ex assume eval_throw:"P,E \ \e,s\<^sub>0\ \ \throw ex,s\<^sub>2\" and e2':"e\<^sub>2' = throw ex" from IH[OF eval_throw wte sconf] e2' show "throw e' = e\<^sub>2' \ s\<^sub>1 = s\<^sub>2" by simp qed next case (WhileF E e s\<^sub>0 s\<^sub>1 c e\<^sub>2 s\<^sub>2 T) have eval:"P,E \ \while (e) c,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\" and wt:"P,E \ while (e) c :: T" and sconf:"P,E \ s\<^sub>0 \" and IH:"\e\<^sub>2 s\<^sub>2 T. \P,E \ \e,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\; P,E \ e :: T; P,E \ s\<^sub>0 \\ \ false = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by fact+ from wt have wte:"P,E \ e :: Boolean" by auto from eval show ?case proof(rule eval_cases) assume eval_false:"P,E \ \e,s\<^sub>0\ \ \false,s\<^sub>2\" and e2:"e\<^sub>2 = unit" from IH[OF eval_false wte sconf] e2 show "unit = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix s s' w assume eval_true:"P,E \ \e,s\<^sub>0\ \ \true,s\" from IH[OF eval_true wte sconf] show "unit = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix ex assume eval_throw:"P,E \ \e,s\<^sub>0\ \ \throw ex,s\<^sub>2\" from IH[OF eval_throw wte sconf] show "unit = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix ex s assume eval_true:"P,E \ \e,s\<^sub>0\ \ \true,s\" from IH[OF eval_true wte sconf] show "unit = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp qed next case (WhileT E e s\<^sub>0 s\<^sub>1 c v\<^sub>1 s\<^sub>2 e\<^sub>3 s\<^sub>3 e\<^sub>2 s\<^sub>2' T) have eval:"P,E \ \while (e) c,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2'\" and wt:"P,E \ while (e) c :: T" and sconf:"P,E \ s\<^sub>0 \" and IH1:"\ei si T. \P,E \ \e,s\<^sub>0\ \ \ei,si\; P,E \ e :: T; P,E \ s\<^sub>0 \\ \ true = ei \ s\<^sub>1 = si" and IH2:"\ei si T. \P,E \ \c,s\<^sub>1\ \ \ei,si\; P,E \ c :: T; P,E \ s\<^sub>1 \\ \ Val v\<^sub>1 = ei \ s\<^sub>2 = si" and IH3:"\ei si T. \P,E \ \while (e) c,s\<^sub>2\ \ \ei,si\; P,E \ while (e) c :: T; P,E \ s\<^sub>2 \\ \ e\<^sub>3 = ei \ s\<^sub>3 = si" by fact+ from wt obtain T' where wte:"P,E \ e :: Boolean" and wtc:"P,E \ c :: T'" by auto from eval show ?case proof(rule eval_cases) assume eval_false:"P,E \ \e,s\<^sub>0\ \ \false,s\<^sub>2'\" from IH1[OF eval_false wte sconf] show "e\<^sub>3 = e\<^sub>2 \ s\<^sub>3 = s\<^sub>2'" by simp next fix s s' w assume eval_true:"P,E \ \e,s\<^sub>0\ \ \true,s\" and eval_val:"P,E \ \c,s\ \ \Val w,s'\" and eval_while:"P,E \ \while (e) c,s'\ \ \e\<^sub>2,s\<^sub>2'\" from IH1[OF eval_true wte sconf] have eq:"s = s\<^sub>1" by simp with wf eval_true wte sconf have sconf':"P,E \ s\<^sub>1 \" by(fastforce intro:eval_preserves_sconf) from IH2[OF eval_val[simplified eq] wtc this] have eq':"s' = s\<^sub>2" by simp with wf eval_val wtc sconf' eq have "P,E \ s\<^sub>2 \" by(fastforce intro:eval_preserves_sconf) from IH3[OF eval_while[simplified eq'] wt this] show "e\<^sub>3 = e\<^sub>2 \ s\<^sub>3 = s\<^sub>2'" . next fix ex assume eval_throw:"P,E \ \e,s\<^sub>0\ \ \throw ex,s\<^sub>2'\" from IH1[OF eval_throw wte sconf] show "e\<^sub>3 = e\<^sub>2 \ s\<^sub>3 = s\<^sub>2'" by simp next fix ex s assume eval_true:"P,E \ \e,s\<^sub>0\ \ \true,s\" and eval_throw:"P,E \ \c,s\ \ \throw ex,s\<^sub>2'\" from IH1[OF eval_true wte sconf] have eq:"s = s\<^sub>1" by simp with wf eval_true wte sconf have sconf':"P,E \ s\<^sub>1 \" by(fastforce intro:eval_preserves_sconf) from IH2[OF eval_throw[simplified eq] wtc this] show "e\<^sub>3 = e\<^sub>2 \ s\<^sub>3 = s\<^sub>2'" by simp qed next case (WhileCondThrow E e s\<^sub>0 e' s\<^sub>1 c e\<^sub>2 s\<^sub>2 T) have eval:"P,E \ \while (e) c,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\" and wt:"P,E \ while (e) c :: T" and sconf:"P,E \ s\<^sub>0 \" and IH:"\ei si T. \P,E \ \e,s\<^sub>0\ \ \ei,si\; P,E \ e :: T; P,E \ s\<^sub>0 \\ \ throw e' = ei \ s\<^sub>1 = si" by fact+ from wt have wte:"P,E \ e :: Boolean" by auto from eval show ?case proof(rule eval_cases) assume eval_false:"P,E \ \e,s\<^sub>0\ \ \false,s\<^sub>2\" from IH[OF eval_false wte sconf] show "throw e' = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix s s' w assume eval_true:"P,E \ \e,s\<^sub>0\ \ \true,s\" from IH[OF eval_true wte sconf] show "throw e' = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix ex assume eval_throw:"P,E \ \e,s\<^sub>0\ \ \throw ex,s\<^sub>2\" and e2:"e\<^sub>2 = throw ex" from IH[OF eval_throw wte sconf] e2 show "throw e' = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix ex s assume eval_true:"P,E \ \e,s\<^sub>0\ \ \true,s\" from IH[OF eval_true wte sconf] show "throw e' = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp qed next case (WhileBodyThrow E e s\<^sub>0 s\<^sub>1 c e' s\<^sub>2 e\<^sub>2 s\<^sub>2' T) have eval:"P,E \ \while (e) c,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2'\" and wt:"P,E \ while (e) c :: T" and sconf:"P,E \ s\<^sub>0 \" and IH1:"\ei si T. \P,E \ \e,s\<^sub>0\ \ \ei,si\; P,E \ e :: T; P,E \ s\<^sub>0 \\ \ true = ei \ s\<^sub>1 = si" and IH2:"\ei si T. \P,E \ \c,s\<^sub>1\ \ \ei,si\; P,E \ c :: T; P,E \ s\<^sub>1 \\ \ throw e' = ei \ s\<^sub>2 = si" by fact+ from wt obtain T' where wte:"P,E \ e :: Boolean" and wtc:"P,E \ c :: T'" by auto from eval show ?case proof(rule eval_cases) assume eval_false:"P,E \ \e,s\<^sub>0\ \ \false,s\<^sub>2'\" from IH1[OF eval_false wte sconf] show "throw e' = e\<^sub>2 \ s\<^sub>2 = s\<^sub>2'" by simp next fix s s' w assume eval_true:"P,E \ \e,s\<^sub>0\ \ \true,s\" and eval_val:"P,E \ \c,s\ \ \Val w,s'\" from IH1[OF eval_true wte sconf] have eq:"s = s\<^sub>1" by simp with wf eval_true wte sconf have sconf':"P,E \ s\<^sub>1 \" by(fastforce intro:eval_preserves_sconf) from IH2[OF eval_val[simplified eq] wtc this] show "throw e' = e\<^sub>2 \ s\<^sub>2 = s\<^sub>2'" by simp next fix ex assume eval_throw:"P,E \ \e,s\<^sub>0\ \ \throw ex,s\<^sub>2'\" from IH1[OF eval_throw wte sconf] show "throw e' = e\<^sub>2 \ s\<^sub>2 = s\<^sub>2'" by simp next fix ex s assume eval_true:"P,E \ \e,s\<^sub>0\ \ \true,s\" and eval_throw:"P,E \ \c,s\ \ \throw ex,s\<^sub>2'\" and e2:"e\<^sub>2 = throw ex" from IH1[OF eval_true wte sconf] have eq:"s = s\<^sub>1" by simp with wf eval_true wte sconf have sconf':"P,E \ s\<^sub>1 \" by(fastforce intro:eval_preserves_sconf) from IH2[OF eval_throw[simplified eq] wtc this] e2 show "throw e' = e\<^sub>2 \ s\<^sub>2 = s\<^sub>2'" by simp qed next case (Throw E e s\<^sub>0 r s\<^sub>1 e\<^sub>2 s\<^sub>2 T) have eval:"P,E \ \throw e,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\" and wt:"P,E \ throw e :: T" and sconf:"P,E \ s\<^sub>0 \" and IH:"\ei si T. \P,E \ \e,s\<^sub>0\ \ \ei,si\; P,E \ e :: T; P,E \ s\<^sub>0 \\ \ ref r = ei \ s\<^sub>1 = si" by fact+ from wt obtain C where wte:"P,E \ e :: Class C" by auto from eval show ?case proof(rule eval_cases) fix r' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref r',s\<^sub>2\" and e2:"e\<^sub>2 = Throw r'" from IH[OF eval_ref wte sconf] e2 show "Throw r = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next assume eval_null:"P,E \ \e,s\<^sub>0\ \ \null,s\<^sub>2\" from IH[OF eval_null wte sconf] show "Throw r = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix ex assume eval_throw:"P,E \ \e,s\<^sub>0\ \ \throw ex,s\<^sub>2\" from IH[OF eval_throw wte sconf] show "Throw r = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp qed next case (ThrowNull E e s\<^sub>0 s\<^sub>1 e\<^sub>2 s\<^sub>2 T) have eval:"P,E \ \throw e,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\" and wt:"P,E \ throw e :: T" and sconf:"P,E \ s\<^sub>0 \" and IH:"\ei si T. \P,E \ \e,s\<^sub>0\ \ \ei,si\; P,E \ e :: T; P,E \ s\<^sub>0 \\ \ null = ei \ s\<^sub>1 = si" by fact+ from wt obtain C where wte:"P,E \ e :: Class C" by auto from eval show ?case proof(rule eval_cases) fix r' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref r',s\<^sub>2\" from IH[OF eval_ref wte sconf] show "THROW NullPointer = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next assume eval_null:"P,E \ \e,s\<^sub>0\ \ \null,s\<^sub>2\" and e2:"e\<^sub>2 = THROW NullPointer" from IH[OF eval_null wte sconf] e2 show "THROW NullPointer = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix ex assume eval_throw:"P,E \ \e,s\<^sub>0\ \ \throw ex,s\<^sub>2\" from IH[OF eval_throw wte sconf] show "THROW NullPointer = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp qed next case (ThrowThrow E e s\<^sub>0 e' s\<^sub>1 e\<^sub>2 s\<^sub>2 T) have eval:"P,E \ \throw e,s\<^sub>0\ \ \e\<^sub>2,s\<^sub>2\" and wt:"P,E \ throw e :: T" and sconf:"P,E \ s\<^sub>0 \" and IH:"\ei si T. \P,E \ \e,s\<^sub>0\ \ \ei,si\; P,E \ e :: T; P,E \ s\<^sub>0 \\ \ throw e' = ei \ s\<^sub>1 = si" by fact+ from wt obtain C where wte:"P,E \ e :: Class C" by auto from eval show ?case proof(rule eval_cases) fix r' assume eval_ref:"P,E \ \e,s\<^sub>0\ \ \ref r',s\<^sub>2\" from IH[OF eval_ref wte sconf] show "throw e' = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next assume eval_null:"P,E \ \e,s\<^sub>0\ \ \null,s\<^sub>2\" from IH[OF eval_null wte sconf] show "throw e' = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix ex assume eval_throw:"P,E \ \e,s\<^sub>0\ \ \throw ex,s\<^sub>2\" and e2:"e\<^sub>2 = throw ex" from IH[OF eval_throw wte sconf] e2 show "throw e' = e\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp qed next case Nil thus ?case by (auto elim:evals_cases) next case (Cons E e s\<^sub>0 v s\<^sub>1 es es' s\<^sub>2 es\<^sub>2 s\<^sub>2' Ts) have evals:"P,E \ \e#es,s\<^sub>0\ [\] \es\<^sub>2,s\<^sub>2'\" and wt:"P,E \ e#es [::] Ts" and sconf:"P,E \ s\<^sub>0 \" and IH1:"\ei si T. \P,E \ \e,s\<^sub>0\ \ \ei,si\; P,E \ e :: T; P,E \ s\<^sub>0 \\ \ Val v = ei \ s\<^sub>1 = si" and IH2:"\esi si Ts. \P,E \ \es,s\<^sub>1\ [\] \esi,si\; P,E \ es [::] Ts; P,E \ s\<^sub>1 \\ \ es' = esi \ s\<^sub>2 = si" by fact+ from wt obtain T' Ts' where Ts:"Ts = T'#Ts'" by(cases Ts) auto with wt have wte:"P,E \ e :: T'" and wtes:"P,E \ es [::] Ts'" by auto from evals show ?case proof(rule evals_cases) fix es'' s w assume eval_val:"P,E \ \e,s\<^sub>0\ \ \Val w,s\" and evals_vals:"P,E \ \es,s\ [\] \es'',s\<^sub>2'\" and es2:"es\<^sub>2 = Val w#es''" from IH1[OF eval_val wte sconf] have s:"s = s\<^sub>1" and v:"v = w" by simp_all with wf eval_val wte sconf have "P,E \ s\<^sub>1 \" by(fastforce intro:eval_preserves_sconf) from IH2[OF evals_vals[simplified s] wtes this] have "es' = es'' \ s\<^sub>2 = s\<^sub>2'" . with es2 v show "Val v # es' = es\<^sub>2 \ s\<^sub>2 = s\<^sub>2'" by simp next fix ex assume eval_throw:"P,E \ \e,s\<^sub>0\ \ \throw ex,s\<^sub>2'\" from IH1[OF eval_throw wte sconf] show "Val v # es' = es\<^sub>2 \ s\<^sub>2 = s\<^sub>2'" by simp qed next case (ConsThrow E e s\<^sub>0 e' s\<^sub>1 es es\<^sub>2 s\<^sub>2 Ts) have evals:"P,E \ \e#es,s\<^sub>0\ [\] \es\<^sub>2,s\<^sub>2\" and wt:"P,E \ e#es [::] Ts" and sconf:"P,E \ s\<^sub>0 \" and IH:"\ei si T. \P,E \ \e,s\<^sub>0\ \ \ei,si\; P,E \ e :: T; P,E \ s\<^sub>0 \\ \ throw e' = ei \ s\<^sub>1 = si" by fact+ from wt obtain T' Ts' where Ts:"Ts = T'#Ts'" by(cases Ts) auto with wt have wte:"P,E \ e :: T'" by auto from evals show ?case proof(rule evals_cases) fix es'' s w assume eval_val:"P,E \ \e,s\<^sub>0\ \ \Val w,s\" from IH[OF eval_val wte sconf] show "throw e'#es = es\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp next fix ex assume eval_throw:"P,E \ \e,s\<^sub>0\ \ \throw ex,s\<^sub>2\" and es2:"es\<^sub>2 = throw ex#es" from IH[OF eval_throw wte sconf] es2 show "throw e'#es = es\<^sub>2 \ s\<^sub>1 = s\<^sub>2" by simp qed qed end diff --git a/thys/CoreC++/Equivalence.thy b/thys/CoreC++/Equivalence.thy --- a/thys/CoreC++/Equivalence.thy +++ b/thys/CoreC++/Equivalence.thy @@ -1,2835 +1,2835 @@ (* Title: CoreC++ Author: Daniel Wasserrab Maintainer: Daniel Wasserrab Based on the Jinja theory J/Equivalence.thy by Tobias Nipkow *) section \Equivalence of Big Step and Small Step Semantics\ theory Equivalence imports BigStep SmallStep WWellForm begin subsection\Some casts-lemmas\ lemma assumes wf:"wf_prog wf_md P" shows casts_casts: "P \ T casts v to v' \ P \ T casts v' to v'" proof(induct rule:casts_to.induct) case casts_prim thus ?case by(rule casts_to.casts_prim) next case (casts_null C) thus ?case by(rule casts_to.casts_null) next case (casts_ref Cs C Cs' Ds a) have path_via:"P \ Path last Cs to C via Cs'" and Ds:"Ds = Cs @\<^sub>p Cs'" by fact+ with wf have "last Cs' = C" and "Cs' \ []" and "class": "is_class P C" by(auto intro!:Subobjs_nonempty Subobj_last_isClass simp:path_via_def) with Ds have last:"last Ds = C" by -(drule_tac Cs' = "Cs" in appendPath_last,simp) hence Ds':"Ds = Ds @\<^sub>p [C]" by(simp add:appendPath_def) from last "class" have "P \ Path last Ds to C via [C]" by(fastforce intro:Subobjs_Base simp:path_via_def) with Ds' show ?case by(fastforce intro:casts_to.casts_ref) qed lemma casts_casts_eq: "\ P \ T casts v to v; P \ T casts v to v'; wf_prog wf_md P \ \ v = v'" apply - apply(erule casts_to.cases) apply clarsimp apply(erule casts_to.cases) apply simp apply simp apply (simp (asm_lr)) apply(erule casts_to.cases) apply simp apply simp apply simp apply simp apply(erule casts_to.cases) apply simp apply simp apply clarsimp apply(erule appendPath_path_via) by auto lemma assumes wf:"wf_prog wf_md P" shows None_lcl_casts_values: "P,E \ \e,(h,l)\ \ \e',(h',l')\ \ (\V. \l V = None; E V = Some T; l' V = Some v'\ \ P \ T casts v' to v')" and "P,E \ \es,(h,l)\ [\] \es',(h',l')\ \ (\V. \l V = None; E V = Some T; l' V = Some v'\ \ P \ T casts v' to v')" proof(induct rule:red_reds_inducts) case (RedLAss E V T' w w' h l V') have env:"E V = Some T'" and env':"E V' = Some T" and l:"l V' = None" and lupd:"(l(V \ w')) V' = Some v'" and casts:"P \ T' casts w to w'" by fact+ show ?case proof(cases "V = V'") case True with lupd have v':"v' = w'" by simp from True env env' have "T = T'" by simp with v' casts wf show ?thesis by(fastforce intro:casts_casts) next case False with lupd have "l V' = Some v'" by(fastforce split:if_split_asm) with l show ?thesis by simp qed next case (BlockRedNone E V T' e h l e' h' l' V') have l:"l V' = None" and l'upd:"(l'(V := l V)) V' = Some v'" and env:"E V' = Some T" and IH:"\V'. \(l(V := None)) V' = None; (E(V \ T')) V' = Some T; l' V' = Some v'\ \ P \ T casts v' to v'" by fact+ show ?case proof(cases "V = V'") case True with l'upd l show ?thesis by fastforce next case False with l l'upd have lnew:"(l(V := None)) V' = None" and l'new:"l' V' = Some v'" by (auto split:if_split_asm) from env False have env':"(E(V \ T')) V' = Some T" by fastforce from IH[OF lnew env' l'new] show ?thesis . qed next case (BlockRedSome E V T' e h l e' h' l' v V') have l:"l V' = None" and l'upd:"(l'(V := l V)) V' = Some v'" and env:"E V' = Some T" and IH:"\V'. \(l(V := None)) V' = None; (E(V \ T')) V' = Some T; l' V' = Some v'\ \ P \ T casts v' to v'" by fact+ show ?case proof(cases "V = V'") case True with l l'upd show ?thesis by fastforce next case False with l l'upd have lnew:"(l(V := None)) V' = None" and l'new:"l' V' = Some v'" by (auto split:if_split_asm) from env False have env':"(E(V \ T')) V' = Some T" by fastforce from IH[OF lnew env' l'new] show ?thesis . qed next case (InitBlockRed E V T' e h l w' e' h' l' w'' w V') have l:"l V' = None" and l'upd:"(l'(V := l V)) V' = Some v'" and env:"E V' = Some T" and IH:"\V'. \(l(V \ w')) V' = None; (E(V \ T')) V' = Some T; l' V' = Some v'\ \ P \ T casts v' to v'" by fact+ show ?case proof(cases "V = V'") case True with l l'upd show ?thesis by fastforce next case False with l l'upd have lnew:"(l(V \ w')) V' = None" and l'new:"l' V' = Some v'" by (auto split:if_split_asm) from env False have env':"(E(V \ T')) V' = Some T" by fastforce from IH[OF lnew env' l'new] show ?thesis . qed qed (auto intro:casts_casts wf) lemma assumes wf:"wf_prog wf_md P" shows Some_lcl_casts_values: "P,E \ \e,(h,l)\ \ \e',(h',l')\ \ (\V. \l V = Some v; E V = Some T; P \ T casts v'' to v; l' V = Some v'\ \ P \ T casts v' to v')" and "P,E \ \es,(h,l)\ [\] \es',(h',l')\ \ (\V. \l V = Some v; E V = Some T; P \ T casts v'' to v; l' V = Some v'\ \ P \ T casts v' to v')" proof(induct rule:red_reds_inducts) case (RedNew h a h' C' E l V) have l1:"l V = Some v" and l2:"l V = Some v'" and casts:"P \ T casts v'' to v " by fact+ from l1 l2 have eq:"v = v'" by simp with casts wf show ?case by(fastforce intro:casts_casts) next case (RedLAss E V T' w w' h l V') have l:"l V' = Some v" and lupd:"(l(V \ w')) V' = Some v'" and T'casts:"P \ T' casts w to w'" and env:"E V = Some T'" and env':"E V' = Some T" and casts:"P \ T casts v'' to v" by fact+ show ?case proof (cases "V = V'") case True with lupd have v':"v' = w'" by simp from True env env' have "T = T'" by simp with T'casts v' wf show ?thesis by(fastforce intro:casts_casts) next case False with l lupd have "v = v'" by (auto split:if_split_asm) with casts wf show ?thesis by(fastforce intro:casts_casts) qed next case (RedFAss h a D S Cs' F T' Cs w w' Ds fs E l V) have l1:"l V = Some v" and l2:"l V = Some v'" and hp:"h a = Some(D, S)" and T'casts:"P \ T' casts w to w'" and casts:"P \ T casts v'' to v" by fact+ from l1 l2 have eq:"v = v'" by simp with casts wf show ?case by(fastforce intro:casts_casts) next case (BlockRedNone E V T' e h l e' h' l' V') have l':"l' V = None" and l:"l V' = Some v" and l'upd:"(l'(V := l V)) V' = Some v'" and env:"E V' = Some T" and casts:"P \ T casts v'' to v" and IH:"\V'. \(l(V := None)) V' = Some v; (E(V \ T')) V' = Some T; P \ T casts v'' to v ; l' V' = Some v'\ \ P \ T casts v' to v'" by fact+ show ?case proof(cases "V = V'") case True with l' l'upd have "l V = Some v'" by auto with True l have eq:"v = v'" by simp with casts wf show ?thesis by(fastforce intro:casts_casts) next case False with l l'upd have lnew:"(l(V := None)) V' = Some v" and l'new:"l' V' = Some v'" by (auto split:if_split_asm) from env False have env':"(E(V \ T')) V' = Some T" by fastforce from IH[OF lnew env' casts l'new] show ?thesis . qed next case (BlockRedSome E V T' e h l e' h' l' w V') have l':"l' V = Some w" and l:"l V' = Some v" and l'upd:"(l'(V := l V)) V' = Some v'" and env:"E V' = Some T" and casts:"P \ T casts v'' to v" and IH:"\V'. \(l(V := None)) V' = Some v; (E(V \ T')) V' = Some T; P \ T casts v'' to v ; l' V' = Some v'\ \ P \ T casts v' to v'" by fact+ show ?case proof(cases "V = V'") case True with l' l'upd have "l V = Some v'" by auto with True l have eq:"v = v'" by simp with casts wf show ?thesis by(fastforce intro:casts_casts) next case False with l l'upd have lnew:"(l(V := None)) V' = Some v" and l'new:"l' V' = Some v'" by (auto split:if_split_asm) from env False have env':"(E(V \ T')) V' = Some T" by fastforce from IH[OF lnew env' casts l'new] show ?thesis . qed next case (InitBlockRed E V T' e h l w' e' h' l' w'' w V') have l:"l V' = Some v" and l':"l' V = Some w''" and l'upd:"(l'(V := l V)) V' = Some v'" and env:"E V' = Some T" and casts:"P \ T casts v'' to v" and IH:"\V'. \(l(V \ w')) V' = Some v; (E(V \ T')) V' = Some T; P \ T casts v'' to v ; l' V' = Some v'\ \ P \ T casts v' to v'" by fact+ show ?case proof(cases "V = V'") case True with l' l'upd have "l V = Some v'" by auto with True l have eq:"v = v'" by simp with casts wf show ?thesis by(fastforce intro:casts_casts) next case False with l l'upd have lnew:"(l(V \ w')) V' = Some v" and l'new:"l' V' = Some v'" by (auto split:if_split_asm) from env False have env':"(E(V \ T')) V' = Some T" by fastforce from IH[OF lnew env' casts l'new] show ?thesis . qed qed (auto intro:casts_casts wf) subsection\Small steps simulate big step\ subsection \Cast\ lemma StaticCastReds: "P,E \ \e,s\ \* \e',s'\ \ P,E \ \\C\e,s\ \* \\C\e',s'\" apply(erule rtrancl_induct2) apply blast apply(erule rtrancl_into_rtrancl) apply (simp add:StaticCastRed) done lemma StaticCastRedsNull: "P,E \ \e,s\ \* \null,s'\ \ P,E \ \\C\e,s\ \* \null,s'\" apply(rule rtrancl_into_rtrancl) apply(erule StaticCastReds) apply(simp add:RedStaticCastNull) done lemma StaticUpCastReds: "\ P,E \ \e,s\ \* \ref(a,Cs),s'\; P \ Path last Cs to C via Cs'; Ds = Cs@\<^sub>pCs' \ \ P,E \ \\C\e,s\ \* \ref(a,Ds),s'\" apply(rule rtrancl_into_rtrancl) apply(erule StaticCastReds) apply(fastforce intro:RedStaticUpCast) done lemma StaticDownCastReds: "P,E \ \e,s\ \* \ref(a,Cs@[C]@Cs'),s'\ \ P,E \ \\C\e,s\ \* \ref(a,Cs@[C]),s'\" apply(rule rtrancl_into_rtrancl) apply(erule StaticCastReds) apply simp apply(subgoal_tac "P,E \ \\C\ref(a,Cs@[C]@Cs'),s'\ \ \ref(a,Cs@[C]),s'\") apply simp apply(rule RedStaticDownCast) done lemma StaticCastRedsFail: "\ P,E \ \e,s\ \* \ref(a,Cs),s'\; C \ set Cs; \ P \ (last Cs) \\<^sup>* C \ \ P,E \ \\C\e,s\ \* \THROW ClassCast,s'\" apply(rule rtrancl_into_rtrancl) apply(erule StaticCastReds) apply(fastforce intro:RedStaticCastFail) done lemma StaticCastRedsThrow: "\ P,E \ \e,s\ \* \Throw r,s'\ \ \ P,E \ \\C\e,s\ \* \Throw r,s'\" apply(rule rtrancl_into_rtrancl) apply(erule StaticCastReds) apply(simp add:red_reds.StaticCastThrow) done lemma DynCastReds: "P,E \ \e,s\ \* \e',s'\ \ P,E \ \Cast C e,s\ \* \Cast C e',s'\" apply(erule rtrancl_induct2) apply blast apply(erule rtrancl_into_rtrancl) apply (simp add:DynCastRed) done lemma DynCastRedsNull: "P,E \ \e,s\ \* \null,s'\ \ P,E \ \Cast C e,s\ \* \null,s'\" apply(rule rtrancl_into_rtrancl) apply(erule DynCastReds) apply(simp add:RedDynCastNull) done lemma DynCastRedsRef: "\ P,E \ \e,s\ \* \ref(a,Cs),s'\; hp s' a = Some (D,S); P \ Path D to C via Cs'; P \ Path D to C unique \ \ P,E \ \Cast C e,s\ \* \ref(a,Cs'),s'\" apply(rule rtrancl_into_rtrancl) apply(erule DynCastReds) apply(fastforce intro:RedDynCast) done lemma StaticUpDynCastReds: "\ P,E \ \e,s\ \* \ref(a,Cs),s'\; P \ Path last Cs to C unique; P \ Path last Cs to C via Cs'; Ds = Cs@\<^sub>pCs' \ \ P,E \ \Cast C e,s\ \* \ref(a,Ds),s'\" apply(rule rtrancl_into_rtrancl) apply(erule DynCastReds) apply(fastforce intro:RedStaticUpDynCast) done lemma StaticDownDynCastReds: "P,E \ \e,s\ \* \ref(a,Cs@[C]@Cs'),s'\ \ P,E \ \Cast C e,s\ \* \ref(a,Cs@[C]),s'\" apply(rule rtrancl_into_rtrancl) apply(erule DynCastReds) apply simp apply(subgoal_tac "P,E \ \Cast C (ref(a,Cs@[C]@Cs')),s'\ \ \ref(a,Cs@[C]),s'\") apply simp apply(rule RedStaticDownDynCast) done lemma DynCastRedsFail: "\ P,E \ \e,s\ \* \ref(a,Cs),s'\; hp s' a = Some (D,S); \ P \ Path D to C unique; \ P \ Path last Cs to C unique; C \ set Cs \ \ P,E \ \Cast C e,s\ \* \null,s'\" apply(rule rtrancl_into_rtrancl) apply(erule DynCastReds) apply(fastforce intro:RedDynCastFail) done lemma DynCastRedsThrow: "\ P,E \ \e,s\ \* \Throw r,s'\ \ \ P,E \ \Cast C e,s\ \* \Throw r,s'\" apply(rule rtrancl_into_rtrancl) apply(erule DynCastReds) apply(simp add:red_reds.DynCastThrow) done subsection \LAss\ lemma LAssReds: "P,E \ \e,s\ \* \e',s'\ \ P,E \ \V:=e,s\ \* \V:=e',s'\" apply(erule rtrancl_induct2) apply blast apply(erule rtrancl_into_rtrancl) apply(simp add:LAssRed) done lemma LAssRedsVal: "\ P,E \ \e,s\ \* \Val v,(h',l')\; E V = Some T; P \ T casts v to v'\ \ P,E \ \ V:=e,s\ \* \Val v',(h',l'(V\v'))\" apply(rule rtrancl_into_rtrancl) apply(erule LAssReds) apply(simp add:RedLAss) done lemma LAssRedsThrow: "\ P,E \ \e,s\ \* \Throw r,s'\ \ \ P,E \ \ V:=e,s\ \* \Throw r,s'\" apply(rule rtrancl_into_rtrancl) apply(erule LAssReds) apply(simp add:red_reds.LAssThrow) done subsection \BinOp\ lemma BinOp1Reds: "P,E \ \e,s\ \* \e',s'\ \ P,E \ \ e \bop\ e\<^sub>2, s\ \* \e' \bop\ e\<^sub>2, s'\" apply(erule rtrancl_induct2) apply blast apply(erule rtrancl_into_rtrancl) apply(simp add:BinOpRed1) done lemma BinOp2Reds: "P,E \ \e,s\ \* \e',s'\ \ P,E \ \(Val v) \bop\ e, s\ \* \(Val v) \bop\ e', s'\" apply(erule rtrancl_induct2) apply blast apply(erule rtrancl_into_rtrancl) apply(simp add:BinOpRed2) done lemma BinOpRedsVal: "\ P,E \ \e\<^sub>1,s\<^sub>0\ \* \Val v\<^sub>1,s\<^sub>1\; P,E \ \e\<^sub>2,s\<^sub>1\ \* \Val v\<^sub>2,s\<^sub>2\; binop(bop,v\<^sub>1,v\<^sub>2) = Some v \ \ P,E \ \e\<^sub>1 \bop\ e\<^sub>2, s\<^sub>0\ \* \Val v,s\<^sub>2\" apply(rule rtrancl_trans) apply(erule BinOp1Reds) apply(rule rtrancl_into_rtrancl) apply(erule BinOp2Reds) apply(simp add:RedBinOp) done lemma BinOpRedsThrow1: "P,E \ \e,s\ \* \Throw r,s'\ \ P,E \ \e \bop\ e\<^sub>2, s\ \* \Throw r, s'\" apply(rule rtrancl_into_rtrancl) apply(erule BinOp1Reds) apply(simp add:red_reds.BinOpThrow1) done lemma BinOpRedsThrow2: "\ P,E \ \e\<^sub>1,s\<^sub>0\ \* \Val v\<^sub>1,s\<^sub>1\; P,E \ \e\<^sub>2,s\<^sub>1\ \* \Throw r,s\<^sub>2\\ \ P,E \ \e\<^sub>1 \bop\ e\<^sub>2, s\<^sub>0\ \* \Throw r,s\<^sub>2\" apply(rule rtrancl_trans) apply(erule BinOp1Reds) apply(rule rtrancl_into_rtrancl) apply(erule BinOp2Reds) apply(simp add:red_reds.BinOpThrow2) done subsection \FAcc\ lemma FAccReds: "P,E \ \e,s\ \* \e',s'\ \ P,E \ \e\F{Cs}, s\ \* \e'\F{Cs}, s'\" apply(erule rtrancl_induct2) apply blast apply(erule rtrancl_into_rtrancl) apply(simp add:FAccRed) done lemma FAccRedsVal: "\P,E \ \e,s\ \* \ref(a,Cs'),s'\; hp s' a = Some(D,S); Ds = Cs'@\<^sub>pCs; (Ds,fs) \ S; fs F = Some v \ \ P,E \ \e\F{Cs},s\ \* \Val v,s'\" apply(rule rtrancl_into_rtrancl) apply(erule FAccReds) apply (fastforce intro:RedFAcc) done lemma FAccRedsNull: "P,E \ \e,s\ \* \null,s'\ \ P,E \ \e\F{Cs},s\ \* \THROW NullPointer,s'\" apply(rule rtrancl_into_rtrancl) apply(erule FAccReds) apply(simp add:RedFAccNull) done lemma FAccRedsThrow: "P,E \ \e,s\ \* \Throw r,s'\ \ P,E \ \e\F{Cs},s\ \* \Throw r,s'\" apply(rule rtrancl_into_rtrancl) apply(erule FAccReds) apply(simp add:red_reds.FAccThrow) done subsection \FAss\ lemma FAssReds1: "P,E \ \e,s\ \* \e',s'\ \ P,E \ \e\F{Cs}:=e\<^sub>2, s\ \* \e'\F{Cs}:=e\<^sub>2, s'\" apply(erule rtrancl_induct2) apply blast apply(erule rtrancl_into_rtrancl) apply(simp add:FAssRed1) done lemma FAssReds2: "P,E \ \e,s\ \* \e',s'\ \ P,E \ \Val v\F{Cs}:=e, s\ \* \Val v\F{Cs}:=e', s'\" apply(erule rtrancl_induct2) apply blast apply(erule rtrancl_into_rtrancl) apply(simp add:FAssRed2) done lemma FAssRedsVal: "\ P,E \ \e\<^sub>1,s\<^sub>0\ \* \ref(a,Cs'),s\<^sub>1\; P,E \ \e\<^sub>2,s\<^sub>1\ \* \Val v,(h\<^sub>2,l\<^sub>2)\; h\<^sub>2 a = Some(D,S); P \ (last Cs') has least F:T via Cs; P \ T casts v to v'; Ds = Cs'@\<^sub>pCs; (Ds,fs) \ S \ \ P,E \ \e\<^sub>1\F{Cs}:=e\<^sub>2, s\<^sub>0\ \* \Val v',(h\<^sub>2(a\(D,insert (Ds,fs(F\v')) (S - {(Ds,fs)}))),l\<^sub>2)\" apply(rule rtrancl_trans) apply(erule FAssReds1) apply(rule rtrancl_into_rtrancl) apply(erule FAssReds2) apply(fastforce intro:RedFAss) done lemma FAssRedsNull: "\ P,E \ \e\<^sub>1,s\<^sub>0\ \* \null,s\<^sub>1\; P,E \ \e\<^sub>2,s\<^sub>1\ \* \Val v,s\<^sub>2\ \ \ P,E \ \e\<^sub>1\F{Cs}:=e\<^sub>2, s\<^sub>0\ \* \THROW NullPointer, s\<^sub>2\" apply(rule rtrancl_trans) apply(erule FAssReds1) apply(rule rtrancl_into_rtrancl) apply(erule FAssReds2) apply(simp add:RedFAssNull) done lemma FAssRedsThrow1: "P,E \ \e,s\ \* \Throw r,s'\ \ P,E \ \e\F{Cs}:=e\<^sub>2, s\ \* \Throw r, s'\" apply(rule rtrancl_into_rtrancl) apply(erule FAssReds1) apply(simp add:red_reds.FAssThrow1) done lemma FAssRedsThrow2: "\ P,E \ \e\<^sub>1,s\<^sub>0\ \* \Val v,s\<^sub>1\; P,E \ \e\<^sub>2,s\<^sub>1\ \* \Throw r,s\<^sub>2\ \ \ P,E \ \e\<^sub>1\F{Cs}:=e\<^sub>2,s\<^sub>0\ \* \Throw r,s\<^sub>2\" apply(rule rtrancl_trans) apply(erule FAssReds1) apply(rule rtrancl_into_rtrancl) apply(erule FAssReds2) apply(simp add:red_reds.FAssThrow2) done subsection \;;\ lemma SeqReds: "P,E \ \e,s\ \* \e',s'\ \ P,E \ \e;;e\<^sub>2, s\ \* \e';;e\<^sub>2, s'\" apply(erule rtrancl_induct2) apply blast apply(erule rtrancl_into_rtrancl) apply(simp add:SeqRed) done lemma SeqRedsThrow: "P,E \ \e,s\ \* \Throw r,s'\ \ P,E \ \e;;e\<^sub>2, s\ \* \Throw r, s'\" apply(rule rtrancl_into_rtrancl) apply(erule SeqReds) apply(simp add:red_reds.SeqThrow) done lemma SeqReds2: "\ P,E \ \e\<^sub>1,s\<^sub>0\ \* \Val v\<^sub>1,s\<^sub>1\; P,E \ \e\<^sub>2,s\<^sub>1\ \* \e\<^sub>2',s\<^sub>2\ \ \ P,E \ \e\<^sub>1;;e\<^sub>2, s\<^sub>0\ \* \e\<^sub>2',s\<^sub>2\" apply(rule rtrancl_trans) apply(erule SeqReds) apply(rule_tac b="(e\<^sub>2,s\<^sub>1)" in converse_rtrancl_into_rtrancl) apply(simp add:RedSeq) apply assumption done subsection \If\ lemma CondReds: "P,E \ \e,s\ \* \e',s'\ \ P,E \ \if (e) e\<^sub>1 else e\<^sub>2,s\ \* \if (e') e\<^sub>1 else e\<^sub>2,s'\" apply(erule rtrancl_induct2) apply blast apply(erule rtrancl_into_rtrancl) apply(simp add:CondRed) done lemma CondRedsThrow: "P,E \ \e,s\ \* \Throw r,s'\ \ P,E \ \if (e) e\<^sub>1 else e\<^sub>2, s\ \* \Throw r,s'\" apply(rule rtrancl_into_rtrancl) apply(erule CondReds) apply(simp add:red_reds.CondThrow) done lemma CondReds2T: "\P,E \ \e,s\<^sub>0\ \* \true,s\<^sub>1\; P,E \ \e\<^sub>1, s\<^sub>1\ \* \e',s\<^sub>2\ \ \ P,E \ \if (e) e\<^sub>1 else e\<^sub>2, s\<^sub>0\ \* \e',s\<^sub>2\" apply(rule rtrancl_trans) apply(erule CondReds) apply(rule_tac b="(e\<^sub>1, s\<^sub>1)" in converse_rtrancl_into_rtrancl) apply(simp add:RedCondT) apply assumption done lemma CondReds2F: "\P,E \ \e,s\<^sub>0\ \* \false,s\<^sub>1\; P,E \ \e\<^sub>2, s\<^sub>1\ \* \e',s\<^sub>2\ \ \ P,E \ \if (e) e\<^sub>1 else e\<^sub>2, s\<^sub>0\ \* \e',s\<^sub>2\" apply(rule rtrancl_trans) apply(erule CondReds) apply(rule_tac b="(e\<^sub>2, s\<^sub>1)" in converse_rtrancl_into_rtrancl) apply(simp add:RedCondF) apply assumption done subsection \While\ lemma WhileFReds: "P,E \ \b,s\ \* \false,s'\ \ P,E \ \while (b) c,s\ \* \unit,s'\" apply(rule_tac b="(if(b) (c;;while(b) c) else unit, s)" in converse_rtrancl_into_rtrancl) apply(simp add:RedWhile) apply(rule rtrancl_into_rtrancl) apply(erule CondReds) apply(simp add:RedCondF) done lemma WhileRedsThrow: "P,E \ \b,s\ \* \Throw r,s'\ \ P,E \ \while (b) c,s\ \* \Throw r,s'\" apply(rule_tac b="(if(b) (c;;while(b) c) else unit, s)" in converse_rtrancl_into_rtrancl) apply(simp add:RedWhile) apply(rule rtrancl_into_rtrancl) apply(erule CondReds) apply(simp add:red_reds.CondThrow) done lemma WhileTReds: "\ P,E \ \b,s\<^sub>0\ \* \true,s\<^sub>1\; P,E \ \c,s\<^sub>1\ \* \Val v\<^sub>1,s\<^sub>2\; P,E \ \while (b) c,s\<^sub>2\ \* \e,s\<^sub>3\ \ \ P,E \ \while (b) c,s\<^sub>0\ \* \e,s\<^sub>3\" apply(rule_tac b="(if(b) (c;;while(b) c) else unit, s\<^sub>0)" in converse_rtrancl_into_rtrancl) apply(simp add:RedWhile) apply(rule rtrancl_trans) apply(erule CondReds) apply(rule_tac b="(c;;while(b) c,s\<^sub>1)" in converse_rtrancl_into_rtrancl) apply(simp add:RedCondT) apply(rule rtrancl_trans) apply(erule SeqReds) apply(rule_tac b="(while(b) c,s\<^sub>2)" in converse_rtrancl_into_rtrancl) apply(simp add:RedSeq) apply assumption done lemma WhileTRedsThrow: "\ P,E \ \b,s\<^sub>0\ \* \true,s\<^sub>1\; P,E \ \c,s\<^sub>1\ \* \Throw r,s\<^sub>2\ \ \ P,E \ \while (b) c,s\<^sub>0\ \* \Throw r,s\<^sub>2\" apply(rule_tac b="(if(b) (c;;while(b) c) else unit, s\<^sub>0)" in converse_rtrancl_into_rtrancl) apply(simp add:RedWhile) apply(rule rtrancl_trans) apply(erule CondReds) apply(rule_tac b="(c;;while(b) c,s\<^sub>1)" in converse_rtrancl_into_rtrancl) apply(simp add:RedCondT) apply(rule rtrancl_trans) apply(erule SeqReds) apply(rule_tac b="(Throw r,s\<^sub>2)" in converse_rtrancl_into_rtrancl) apply(simp add:red_reds.SeqThrow) apply simp done subsection \Throw\ lemma ThrowReds: "P,E \ \e,s\ \* \e',s'\ \ P,E \ \throw e,s\ \* \throw e',s'\" apply(erule rtrancl_induct2) apply blast apply(erule rtrancl_into_rtrancl) apply(simp add:ThrowRed) done lemma ThrowRedsNull: "P,E \ \e,s\ \* \null,s'\ \ P,E \ \throw e,s\ \* \THROW NullPointer,s'\" apply(rule rtrancl_into_rtrancl) apply(erule ThrowReds) apply(simp add:RedThrowNull) done lemma ThrowRedsThrow: "P,E \ \e,s\ \* \Throw r,s'\ \ P,E \ \throw e,s\ \* \Throw r,s'\" apply(rule rtrancl_into_rtrancl) apply(erule ThrowReds) apply(simp add:red_reds.ThrowThrow) done subsection \InitBlock\ lemma assumes wf:"wf_prog wf_md P" shows InitBlockReds_aux: "P,E(V \ T) \ \e,s\ \* \e',s'\ \ \h l h' l' v v'. s = (h,l(V\v')) \ P \ T casts v to v' \ s' = (h',l') \ (\v'' w. P,E \ \{V:T := Val v; e},(h,l)\ \* \{V:T := Val v''; e'},(h',l'(V:=(l V)))\ \ P \ T casts v'' to w)" proof (erule converse_rtrancl_induct2) { fix h l h' l' v v' assume "s' = (h, l(V \ v'))" and "s' = (h', l')" hence h:"h = h'" and l':"l' = l(V \ v')" by simp_all hence "P,E \ \{V:T; V:=Val v;; e'},(h, l)\ \* \{V:T; V:=Val v;; e'},(h', l'(V := l V))\" by(fastforce simp: fun_upd_same simp del:fun_upd_apply) } hence "\h l h' l' v v'. s' = (h, l(V \ v')) \ P \ T casts v to v' \ s' = (h', l') \ P,E \ \{V:T; V:=Val v;; e'},(h, l)\ \* \{V:T; V:=Val v;; e'},(h', l'(V := l V))\ \ P \ T casts v to v'" by auto thus "\h l h' l' v v'. s' = (h, l(V \ v')) \ P \ T casts v to v' \ s' = (h', l') \ (\v'' w. P,E \ \{V:T; V:=Val v;; e'},(h, l)\ \* \{V:T; V:=Val v'';; e'},(h', l'(V := l V))\ \ P \ T casts v'' to w)" by auto next fix e s e'' s'' assume Red:"((e,s),e'',s'') \ Red P (E(V \ T))" and reds:"P,E(V \ T) \ \e'',s''\ \* \e',s'\" and IH:"\h l h' l' v v'. s'' = (h, l(V \ v')) \ P \ T casts v to v' \ s' = (h', l') \ (\v'' w. P,E \ \{V:T; V:=Val v;; e''},(h, l)\ \* \{V:T; V:=Val v'';; e'},(h', l'(V := l V))\ \ P \ T casts v'' to w)" { fix h l h' l' v v' assume s:"s = (h, l(V \ v'))" and s':"s' = (h', l')" and casts:"P \ T casts v to v'" obtain h'' l'' where s'':"s'' = (h'',l'')" by (cases s'') auto with Red s have "V \ dom l''" by (fastforce dest:red_lcl_incr) then obtain v'' where l'':"l'' V = Some v''" by auto with Red s s'' casts have step:"P,E \ \{V:T := Val v; e},(h, l)\ \ \{V:T := Val v''; e''}, (h'',l''(V := l V))\" by(fastforce intro:InitBlockRed) from Red s s'' l'' casts wf have casts':"P \ T casts v'' to v''" by(fastforce intro:Some_lcl_casts_values) with IH s'' s' l'' obtain v''' w where "P,E \ \{V:T := Val v''; e''}, (h'',l''(V := l V))\ \* \{V:T := Val v'''; e'},(h', l'(V := l V))\ \ P \ T casts v''' to w" apply simp apply (erule_tac x = "l''(V := l V)" in allE) apply (erule_tac x = "v''" in allE) apply (erule_tac x = "v''" in allE) by(auto intro:ext) with step have "\v'' w. P,E \ \{V:T; V:=Val v;; e},(h, l)\ \* \{V:T; V:=Val v'';; e'},(h', l'(V := l V))\ \ P \ T casts v'' to w" apply(rule_tac x="v'''" in exI) apply auto apply (rule converse_rtrancl_into_rtrancl) by simp_all } thus "\h l h' l' v v'. s = (h, l(V \ v')) \ P \ T casts v to v' \ s' = (h', l') \ (\v'' w. P,E \ \{V:T; V:=Val v;; e},(h, l)\ \* \{V:T; V:=Val v'';; e'},(h', l'(V := l V))\ \ P \ T casts v'' to w)" by auto qed lemma InitBlockReds: "\P,E(V \ T) \ \e, (h,l(V\v'))\ \* \e', (h',l')\; P \ T casts v to v'; wf_prog wf_md P \ \ \v'' w. P,E \ \{V:T := Val v; e}, (h,l)\ \* \{V:T := Val v''; e'}, (h',l'(V:=(l V)))\ \ P \ T casts v'' to w" by(blast dest:InitBlockReds_aux) lemma InitBlockRedsFinal: assumes reds:"P,E(V \ T) \ \e,(h,l(V\v'))\ \* \e',(h',l')\" and final:"final e'" and casts:"P \ T casts v to v'" and wf:"wf_prog wf_md P" shows "P,E \ \{V:T := Val v; e},(h,l)\ \* \e',(h', l'(V := l V))\" proof - from reds casts wf obtain v'' and w where steps:"P,E \ \{V:T := Val v; e},(h,l)\ \* \{V:T := Val v''; e'}, (h',l'(V:=(l V)))\" and casts':"P \ T casts v'' to w" by (auto dest:InitBlockReds) from final casts casts' have step:"P,E \ \{V:T := Val v''; e'}, (h',l'(V:=(l V)))\ \ \e',(h',l'(V := l V))\" by(auto elim!:finalE intro:RedInitBlock InitBlockThrow) from step steps show ?thesis by(fastforce intro:rtrancl_into_rtrancl) qed subsection \Block\ lemma BlockRedsFinal: assumes reds: "P,E(V \ T) \ \e\<^sub>0,s\<^sub>0\ \* \e\<^sub>2,(h\<^sub>2,l\<^sub>2)\" and fin: "final e\<^sub>2" and wf:"wf_prog wf_md P" shows "\h\<^sub>0 l\<^sub>0. s\<^sub>0 = (h\<^sub>0,l\<^sub>0(V:=None)) \ P,E \ \{V:T; e\<^sub>0},(h\<^sub>0,l\<^sub>0)\ \* \e\<^sub>2,(h\<^sub>2,l\<^sub>2(V:=l\<^sub>0 V))\" using reds proof (induct rule:converse_rtrancl_induct2) case refl thus ?case by(fastforce intro:finalE[OF fin] RedBlock BlockThrow simp del:fun_upd_apply) next case (step e\<^sub>0 s\<^sub>0 e\<^sub>1 s\<^sub>1) have Red: "((e\<^sub>0,s\<^sub>0),e\<^sub>1,s\<^sub>1) \ Red P (E(V \ T))" and reds: "P,E(V \ T) \ \e\<^sub>1,s\<^sub>1\ \* \e\<^sub>2,(h\<^sub>2,l\<^sub>2)\" and IH: "\h l. s\<^sub>1 = (h,l(V := None)) \ P,E \ \{V:T; e\<^sub>1},(h,l)\ \* \e\<^sub>2,(h\<^sub>2, l\<^sub>2(V := l V))\" and s\<^sub>0: "s\<^sub>0 = (h\<^sub>0, l\<^sub>0(V := None))" by fact+ obtain h\<^sub>1 l\<^sub>1 where s\<^sub>1: "s\<^sub>1 = (h\<^sub>1,l\<^sub>1)" by fastforce show ?case proof cases assume "assigned V e\<^sub>0" then obtain v e where e\<^sub>0: "e\<^sub>0 = V:= Val v;; e" by (unfold assigned_def)blast from Red e\<^sub>0 s\<^sub>0 obtain v' where e\<^sub>1: "e\<^sub>1 = Val v';;e" and s\<^sub>1: "s\<^sub>1 = (h\<^sub>0, l\<^sub>0(V \ v'))" and casts:"P \ T casts v to v'" by auto from e\<^sub>1 fin have "e\<^sub>1 \ e\<^sub>2" by (auto simp:final_def) then obtain e' s' where red1: "P,E(V \ T) \ \e\<^sub>1,s\<^sub>1\ \ \e',s'\" and reds': "P,E(V \ T) \ \e',s'\ \* \e\<^sub>2,(h\<^sub>2,l\<^sub>2)\" using converse_rtranclE2[OF reds] by simp blast from red1 e\<^sub>1 have es': "e' = e" "s' = s\<^sub>1" by auto show ?thesis using e\<^sub>0 s\<^sub>1 es' reds' by(fastforce intro!: InitBlockRedsFinal[OF _ fin casts wf] simp del:fun_upd_apply) next assume unass: "\ assigned V e\<^sub>0" show ?thesis proof (cases "l\<^sub>1 V") assume None: "l\<^sub>1 V = None" hence "P,E \ \{V:T; e\<^sub>0},(h\<^sub>0,l\<^sub>0)\ \ \{V:T; e\<^sub>1},(h\<^sub>1, l\<^sub>1(V := l\<^sub>0 V))\" using s\<^sub>0 s\<^sub>1 Red by(simp add: BlockRedNone[OF _ _ unass]) moreover have "P,E \ \{V:T; e\<^sub>1},(h\<^sub>1, l\<^sub>1(V := l\<^sub>0 V))\ \* \e\<^sub>2,(h\<^sub>2, l\<^sub>2(V := l\<^sub>0 V))\" using IH[of _ "l\<^sub>1(V := l\<^sub>0 V)"] s\<^sub>1 None by(simp add:fun_upd_idem) ultimately show ?case by(rule_tac b="({V:T; e\<^sub>1},(h\<^sub>1, l\<^sub>1(V := l\<^sub>0 V)))" in converse_rtrancl_into_rtrancl,simp) next fix v assume Some: "l\<^sub>1 V = Some v" with Red Some s\<^sub>0 s\<^sub>1 wf have casts:"P \ T casts v to v" by(fastforce intro:None_lcl_casts_values) from Some have "P,E \ \{V:T;e\<^sub>0},(h\<^sub>0,l\<^sub>0)\ \ \{V:T := Val v; e\<^sub>1},(h\<^sub>1,l\<^sub>1(V := l\<^sub>0 V))\" using s\<^sub>0 s\<^sub>1 Red by(simp add: BlockRedSome[OF _ _ unass]) moreover have "P,E \ \{V:T := Val v; e\<^sub>1},(h\<^sub>1,l\<^sub>1(V:= l\<^sub>0 V))\ \* \e\<^sub>2,(h\<^sub>2,l\<^sub>2(V:=l\<^sub>0 V))\" using InitBlockRedsFinal[OF _ fin casts wf,of _ _ "l\<^sub>1(V:=l\<^sub>0 V)" V] Some reds s\<^sub>1 by (simp add:fun_upd_idem) ultimately show ?case by(rule_tac b="({V:T; V:=Val v;; e\<^sub>1},(h\<^sub>1, l\<^sub>1(V := l\<^sub>0 V)))" in converse_rtrancl_into_rtrancl,simp) qed qed qed subsection \List\ lemma ListReds1: "P,E \ \e,s\ \* \e',s'\ \ P,E \ \e#es,s\ [\]* \e' # es,s'\" apply(erule rtrancl_induct2) apply blast apply(erule rtrancl_into_rtrancl) apply(simp add:ListRed1) done lemma ListReds2: "P,E \ \es,s\ [\]* \es',s'\ \ P,E \ \Val v # es,s\ [\]* \Val v # es',s'\" apply(erule rtrancl_induct2) apply blast apply(erule rtrancl_into_rtrancl) apply(simp add:ListRed2) done lemma ListRedsVal: "\ P,E \ \e,s\<^sub>0\ \* \Val v,s\<^sub>1\; P,E \ \es,s\<^sub>1\ [\]* \es',s\<^sub>2\ \ \ P,E \ \e#es,s\<^sub>0\ [\]* \Val v # es',s\<^sub>2\" apply(rule rtrancl_trans) apply(erule ListReds1) apply(erule ListReds2) done subsection \Call\ text\First a few lemmas on what happens to free variables during redction.\ lemma assumes wf: "wwf_prog P" shows Red_fv: "P,E \ \e,(h,l)\ \ \e',(h',l')\ \ fv e' \ fv e" and "P,E \ \es,(h,l)\ [\] \es',(h',l')\ \ fvs es' \ fvs es" proof (induct rule:red_reds_inducts) case (RedCall h l a C S Cs M Ts' T' pns' body' Ds Ts T pns body Cs' vs bs new_body E) hence "fv body \ {this} \ set pns" using assms by(fastforce dest!:select_method_wf_mdecl simp:wf_mdecl_def) with RedCall.hyps show ?case by(cases T') auto next case (RedStaticCall Cs C Cs'' M Ts T pns body Cs' Ds vs E a a' b) hence "fv body \ {this} \ set pns" using assms by(fastforce dest!:has_least_wf_mdecl simp:wf_mdecl_def) with RedStaticCall.hyps show ?case by auto qed auto lemma Red_dom_lcl: "P,E \ \e,(h,l)\ \ \e',(h',l')\ \ dom l' \ dom l \ fv e" and "P,E \ \es,(h,l)\ [\] \es',(h',l')\ \ dom l' \ dom l \ fvs es" proof (induct rule:red_reds_inducts) case RedLAss thus ?case by(force split:if_splits) next case CallParams thus ?case by(force split:if_splits) next case BlockRedNone thus ?case by clarsimp (fastforce split:if_splits) next case BlockRedSome thus ?case by clarsimp (fastforce split:if_splits) next case InitBlockRed thus ?case by clarsimp (fastforce split:if_splits) qed auto lemma Reds_dom_lcl: "\ wwf_prog P; P,E \ \e,(h,l)\ \* \e',(h',l')\ \ \ dom l' \ dom l \ fv e" apply(erule converse_rtrancl_induct_red) apply blast apply(blast dest: Red_fv Red_dom_lcl) done text\Now a few lemmas on the behaviour of blocks during reduction.\ lemma override_on_upd_lemma: "(override_on f (g(a\b)) A)(a := g a) = override_on f g (insert a A)" apply(rule ext) apply(simp add:override_on_def) done declare fun_upd_apply[simp del] map_upds_twist[simp del] lemma assumes wf:"wf_prog wf_md P" shows blocksReds: "\l\<^sub>0 E vs'. \ length Vs = length Ts; length vs = length Ts; distinct Vs; \<^cancel>\\T\set Ts. is_type P T;\ P \ Ts Casts vs to vs'; P,E(Vs [\] Ts) \ \e, (h\<^sub>0,l\<^sub>0(Vs [\] vs'))\ \* \e', (h\<^sub>1,l\<^sub>1)\ \ \ \vs''. P,E \ \blocks(Vs,Ts,vs,e), (h\<^sub>0,l\<^sub>0)\ \* \blocks(Vs,Ts,vs'',e'), (h\<^sub>1,override_on l\<^sub>1 l\<^sub>0 (set Vs))\ \ (\ws. P \ Ts Casts vs'' to ws) \ length vs = length vs''" proof(induct Vs Ts vs e rule:blocks_old_induct) case (5 V Vs T Ts v vs e) have length1:"length (V#Vs) = length (T#Ts)" and length2:"length (v#vs) = length (T#Ts)" and dist:"distinct (V#Vs)" and casts:"P \ (T#Ts) Casts (v#vs) to vs'" and reds:"P,E(V#Vs [\] T#Ts) \ \e,(h\<^sub>0,l\<^sub>0(V#Vs [\] vs'))\ \* \e',(h\<^sub>1,l\<^sub>1)\" and IH:"\l\<^sub>0 E vs''. \length Vs = length Ts; length vs = length Ts; distinct Vs; P \ Ts Casts vs to vs''; P,E(Vs [\] Ts) \ \e,(h\<^sub>0,l\<^sub>0(Vs [\] vs''))\ \* \e',(h\<^sub>1,l\<^sub>1)\\ \ \vs''. P,E \ \blocks (Vs,Ts,vs,e),(h\<^sub>0,l\<^sub>0)\ \* \blocks (Vs,Ts,vs'',e'),(h\<^sub>1, override_on l\<^sub>1 l\<^sub>0 (set Vs))\ \ (\ws. P \ Ts Casts vs'' to ws) \ length vs = length vs''" by fact+ from length1 have length1':"length Vs = length Ts" by simp from length2 have length2':"length vs = length Ts" by simp from dist have dist':"distinct Vs" by simp from casts obtain x xs where vs':"vs' = x#xs" by(cases vs',auto dest:length_Casts_vs') with reds - have reds':"P,E(V \ T)(Vs [\] Ts) \ \e,(h\<^sub>0,l\<^sub>0(V \ x)(Vs [\] xs))\ + have reds':"P,E(V \ T, Vs [\] Ts) \ \e,(h\<^sub>0,l\<^sub>0(V \ x, Vs [\] xs))\ \* \e',(h\<^sub>1,l\<^sub>1)\" by simp from casts vs' have casts':"P \ Ts Casts vs to xs" and cast':"P \ T casts v to x" by(auto elim:Casts_to.cases) from IH[OF length1' length2' dist' casts' reds'] obtain vs'' ws where blocks:"P,E(V \ T) \ \blocks (Vs, Ts, vs, e),(h\<^sub>0, l\<^sub>0(V \ x))\ \* \blocks (Vs, Ts, vs'', e'),(h\<^sub>1, override_on l\<^sub>1 (l\<^sub>0(V \ x)) (set Vs))\" and castsws:"P \ Ts Casts vs'' to ws" and lengthvs'':"length vs = length vs''" by auto from InitBlockReds[OF blocks cast' wf] obtain v'' w where blocks':"P,E \ \{V:T; V:=Val v;; blocks (Vs, Ts, vs, e)},(h\<^sub>0, l\<^sub>0)\ \* \{V:T; V:=Val v'';; blocks (Vs, Ts, vs'', e')}, (h\<^sub>1, (override_on l\<^sub>1 (l\<^sub>0(V \ x)) (set Vs))(V := l\<^sub>0 V))\" and "P \ T casts v'' to w" by auto with castsws have "P \ T#Ts Casts v''#vs'' to w#ws" by -(rule Casts_Cons) with blocks' lengthvs'' show ?case by(rule_tac x="v''#vs''" in exI,auto simp:override_on_upd_lemma) next case (6 e) have casts:"P \ [] Casts [] to vs'" and step:"P,E([] [\] []) \ \e,(h\<^sub>0, l\<^sub>0([] [\] vs'))\ \* \e',(h\<^sub>1, l\<^sub>1)\" by fact+ from casts have "vs' = []" by(fastforce dest:length_Casts_vs') with step have "P,E \ \e,(h\<^sub>0, l\<^sub>0)\ \* \e',(h\<^sub>1, l\<^sub>1)\" by simp with casts show ?case by auto qed simp_all lemma assumes wf:"wf_prog wf_md P" shows blocksFinal: "\E l vs'. \ length Vs = length Ts; length vs = length Ts; \<^cancel>\\T\set Ts. is_type P T;\ final e; P \ Ts Casts vs to vs' \ \ P,E \ \blocks(Vs,Ts,vs,e), (h,l)\ \* \e, (h,l)\" proof(induct Vs Ts vs e rule:blocks_old_induct) case (5 V Vs T Ts v vs e) have length1:"length (V # Vs) = length (T # Ts)" and length2:"length (v # vs) = length (T # Ts)" and final:"final e" and casts:"P \ T # Ts Casts v # vs to vs'" and IH:"\E l vs'. \length Vs = length Ts; length vs = length Ts; final e; P \ Ts Casts vs to vs' \ \ P,E \ \blocks (Vs, Ts, vs, e),(h, l)\ \* \e,(h, l)\" by fact+ from length1 length2 have length1':"length Vs = length Ts" and length2':"length vs = length Ts" by simp_all from casts obtain x xs where vs':"vs' = x#xs" by(cases vs',auto dest:length_Casts_vs') with casts have casts':"P \ Ts Casts vs to xs" and cast':"P \ T casts v to x" by(auto elim:Casts_to.cases) from InitBlockReds[OF IH[OF length1' length2' final casts'] cast' wf, of V l] obtain v'' w where blocks:"P,E \ \{V:T; V:=Val v;; blocks (Vs, Ts, vs, e)},(h, l)\ \* \{V:T; V:=Val v'';; e},(h,l)\" and "P \ T casts v'' to w" by auto blast with final have "P,E \ \{V:T; V:=Val v'';; e},(h,l)\ \ \e,(h,l)\" by(auto elim!:finalE intro:RedInitBlock InitBlockThrow) with blocks show ?case by -(rule_tac b="({V:T; V:=Val v'';; e},(h, l))" in rtrancl_into_rtrancl,simp_all) qed auto lemma assumes wfmd:"wf_prog wf_md P" and wf: "length Vs = length Ts" "length vs = length Ts" "distinct Vs" and casts:"P \ Ts Casts vs to vs'" and reds: "P,E(Vs [\] Ts) \ \e, (h\<^sub>0, l\<^sub>0(Vs [\] vs'))\ \* \e', (h\<^sub>1, l\<^sub>1)\" and fin: "final e'" and l2: "l\<^sub>2 = override_on l\<^sub>1 l\<^sub>0 (set Vs)" shows blocksRedsFinal: "P,E \ \blocks(Vs,Ts,vs,e), (h\<^sub>0, l\<^sub>0)\ \* \e', (h\<^sub>1,l\<^sub>2)\" proof - obtain vs'' ws where blocks:"P,E \ \blocks(Vs,Ts,vs,e), (h\<^sub>0, l\<^sub>0)\ \* \blocks(Vs,Ts,vs'',e'), (h\<^sub>1,l\<^sub>2)\" and length:"length vs = length vs''" and casts':"P \ Ts Casts vs'' to ws" using l2 blocksReds[OF wfmd wf casts reds] by auto have "P,E \ \blocks(Vs,Ts,vs'',e'), (h\<^sub>1,l\<^sub>2)\ \* \e', (h\<^sub>1,l\<^sub>2)\" using blocksFinal[OF wfmd _ _ fin casts'] wf length by simp with blocks show ?thesis by simp qed text\An now the actual method call reduction lemmas.\ lemma CallRedsObj: "P,E \ \e,s\ \* \e',s'\ \ P,E \ \Call e Copt M es,s\ \* \Call e' Copt M es,s'\" apply(erule rtrancl_induct2) apply blast apply(erule rtrancl_into_rtrancl) apply(simp add:CallObj) done lemma CallRedsParams: "P,E \ \es,s\ [\]* \es',s'\ \ P,E \ \Call (Val v) Copt M es,s\ \* \Call (Val v) Copt M es',s'\" apply(erule rtrancl_induct2) apply blast apply(erule rtrancl_into_rtrancl) apply(simp add:CallParams) done lemma cast_lcl: "P,E \ \\C\(Val v),(h,l)\ \ \Val v',(h,l)\ \ P,E \ \\C\(Val v),(h,l')\ \ \Val v',(h,l')\" apply(erule red.cases) apply(auto intro:red_reds.intros) apply(subgoal_tac "P,E \ \\C\ref (a,Cs@[C]@Cs'),(h,l')\ \ \ref (a,Cs@[C]),(h,l')\") apply simp apply(rule RedStaticDownCast) done lemma cast_env: "P,E \ \\C\(Val v),(h,l)\ \ \Val v',(h,l)\ \ P,E' \ \\C\(Val v),(h,l)\ \ \Val v',(h,l)\" apply(erule red.cases) apply(auto intro:red_reds.intros) apply(subgoal_tac "P,E' \ \\C\ref (a,Cs@[C]@Cs'),(h,l)\ \ \ref (a,Cs@[C]),(h,l)\") apply simp apply(rule RedStaticDownCast) done lemma Cast_step_Cast_or_fin: "P,E \ \\C\e,s\ \* \e',s'\ \ final e' \ (\e''. e' = \C\e'')" by(induct rule:rtrancl_induct2,auto elim:red.cases simp:final_def) lemma Cast_red:"P,E \ \e,s\ \* \e',s'\ \ (\e\<^sub>1. \e = \C\e\<^sub>0; e' = \C\e\<^sub>1\ \ P,E \ \e\<^sub>0,s\ \* \e\<^sub>1,s'\)" proof(induct rule:rtrancl_induct2) case refl thus ?case by simp next case (step e'' s'' e' s') have step:"P,E \ \e,s\ \* \e'',s''\" and Red:"((e'', s''), e', s') \ Red P E" and cast:"e = \C\e\<^sub>0" and cast':"e' = \C\e\<^sub>1" and IH:"\e\<^sub>1. \e = \C\e\<^sub>0; e'' = \C\e\<^sub>1\ \ P,E \ \e\<^sub>0,s\ \* \e\<^sub>1,s''\" by fact+ from Red have red:"P,E \ \e'',s''\ \ \e',s'\" by simp from step cast have "final e'' \ (\ex. e'' = \C\ex)" by simp(rule Cast_step_Cast_or_fin) thus ?case proof(rule disjE) assume "final e''" with red show ?thesis by(auto simp:final_def) next assume "\ex. e'' = \C\ex" then obtain ex where e'':"e'' = \C\ex" by blast with cast' red have "P,E \ \ex,s''\ \ \e\<^sub>1,s'\" by(auto elim:red.cases) with IH[OF cast e''] show ?thesis by(rule_tac b="(ex,s'')" in rtrancl_into_rtrancl,simp_all) qed qed lemma Cast_final:"\P,E \ \\C\e,s\ \* \e',s'\; final e'\ \ \e'' s''. P,E \ \e,s\ \* \e'',s''\ \ P,E \ \\C\e'',s''\ \ \e',s'\ \ final e''" proof(induct rule:rtrancl_induct2) case refl thus ?case by (simp add:final_def) next case (step e'' s'' e' s') have step:"P,E \ \\C\e,s\ \* \e'',s''\" and Red:"((e'', s''), e', s') \ Red P E" and final:"final e'" and IH:"final e'' \ \ex sx. P,E \ \e,s\ \* \ex,sx\ \ P,E \ \\C\ex,sx\ \ \e'',s''\ \ final ex" by fact+ from Red have red:"P,E \ \e'',s''\ \ \e',s'\" by simp from step have "final e'' \ (\ex. e'' = \C\ex)" by(rule Cast_step_Cast_or_fin) thus ?case proof(rule disjE) assume "final e''" with red show ?thesis by(auto simp:final_def) next assume "\ex. e'' = \C\ex" then obtain ex where e'':"e'' = \C\ex" by blast with red final have final':"final ex" by(auto elim:red.cases simp:final_def) from step e'' have "P,E \ \e,s\ \* \ex,s''\" by(fastforce intro:Cast_red) with e'' red final' show ?thesis by blast qed qed lemma Cast_final_eq: assumes red:"P,E \ \\C\e,(h,l)\ \ \e',(h,l)\" and final:"final e" and final':"final e'" shows "P,E' \ \\C\e,(h,l')\ \ \e',(h,l')\" proof - from red final show ?thesis proof(auto simp:final_def) fix v assume "P,E \ \\C\(Val v),(h,l)\ \ \e',(h,l)\" with final' show "P,E' \ \\C\(Val v),(h,l')\ \ \e',(h,l')\" proof(auto simp:final_def) fix v' assume "P,E \ \\C\(Val v),(h,l)\ \ \Val v',(h,l)\" thus "P,E' \ \\C\(Val v),(h,l')\ \ \Val v',(h,l')\" by(auto intro:cast_lcl cast_env) next fix a Cs assume "P,E \ \\C\(Val v),(h,l)\ \ \Throw (a,Cs),(h,l)\" thus "P,E' \ \\C\(Val v),(h,l')\ \ \Throw (a,Cs),(h,l')\" by(auto elim:red.cases intro!:RedStaticCastFail) qed next fix a Cs assume "P,E \ \\C\(Throw (a,Cs)),(h,l)\ \ \e',(h,l)\" with final' show "P,E' \ \\C\(Throw (a,Cs)),(h,l')\ \ \e',(h,l')\" proof(auto simp:final_def) fix v assume "P,E \ \\C\(Throw (a,Cs)),(h,l)\ \ \Val v,(h,l)\" thus "P,E' \ \\C\(Throw (a,Cs)),(h,l')\ \ \Val v,(h,l')\" by(auto elim:red.cases) next fix a' Cs' assume "P,E \ \\C\(Throw (a,Cs)),(h,l)\ \ \Throw (a',Cs'),(h,l)\" thus "P,E' \ \\C\(Throw (a,Cs)),(h,l')\ \ \Throw (a',Cs'),(h,l')\" by(auto elim:red.cases intro:red_reds.StaticCastThrow) qed qed qed lemma CallRedsFinal: assumes wwf: "wwf_prog P" and "P,E \ \e,s\<^sub>0\ \* \ref(a,Cs),s\<^sub>1\" "P,E \ \es,s\<^sub>1\ [\]* \map Val vs,(h\<^sub>2,l\<^sub>2)\" and hp: "h\<^sub>2 a = Some(C,S)" and "method": "P \ last Cs has least M = (Ts',T',pns',body') via Ds" and select: "P \ (C,Cs@\<^sub>pDs) selects M = (Ts,T,pns,body) via Cs'" and size: "size vs = size pns" and casts: "P \ Ts Casts vs to vs'" and l\<^sub>2': "l\<^sub>2' = [this \ Ref(a,Cs'), pns[\]vs']" and body_case:"new_body = (case T' of Class D \ \D\body | _ \ body)" and body: "P,E(this \ Class (last Cs'), pns [\] Ts) \ \new_body,(h\<^sub>2,l\<^sub>2')\ \* \ef,(h\<^sub>3,l\<^sub>3)\" and final:"final ef" shows "P,E \ \e\M(es), s\<^sub>0\ \* \ef,(h\<^sub>3,l\<^sub>2)\" proof - have wf: "size Ts = size pns \ distinct pns \ this \ set pns" and wt: "fv body \ {this} \ set pns" using assms by(fastforce dest!:select_method_wf_mdecl simp:wf_mdecl_def)+ have "dom l\<^sub>3 \ {this} \ set pns" using Reds_dom_lcl[OF wwf body] wt l\<^sub>2' set_take_subset body_case by (cases T') force+ hence eql\<^sub>2: "override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 ({this} \ set pns) = l\<^sub>2" by(fastforce simp add:map_add_def override_on_def fun_eq_iff) from wwf select have "is_class P (last Cs')" by (auto elim!:SelectMethodDef.cases intro:Subobj_last_isClass simp:LeastMethodDef_def FinalOverriderMethodDef_def OverriderMethodDefs_def MinimalMethodDefs_def MethodDefs_def) hence "P \ Class (last Cs') casts Ref(a,Cs') to Ref(a,Cs')" by(auto intro!:casts_ref Subobjs_Base simp:path_via_def appendPath_def) with casts have casts':"P \ Class (last Cs')#Ts Casts Ref(a,Cs')#vs to Ref(a,Cs')#vs'" by -(rule Casts_Cons) have 1:"P,E \ \e\M(es),s\<^sub>0\ \* \(ref(a,Cs))\M(es),s\<^sub>1\" by(rule CallRedsObj)(rule assms(2)) have 2:"P,E \ \(ref(a,Cs))\M(es),s\<^sub>1\ \* \(ref(a,Cs))\M(map Val vs),(h\<^sub>2,l\<^sub>2)\" by(rule CallRedsParams)(rule assms(3)) from body[THEN Red_lcl_add, of l\<^sub>2] have body': "P,E(this \ Class (last Cs'), pns [\] Ts) \ \new_body,(h\<^sub>2,l\<^sub>2(this\ Ref(a,Cs'), pns[\]vs'))\ \* \ef,(h\<^sub>3,l\<^sub>2++l\<^sub>3)\" by (simp add:l\<^sub>2') show ?thesis proof(cases "\C. T'\ Class C") case True hence "P,E \ \(ref(a,Cs))\M(map Val vs), (h\<^sub>2,l\<^sub>2)\ \ \blocks(this#pns,Class(last Cs')#Ts,Ref(a,Cs')#vs,body), (h\<^sub>2,l\<^sub>2)\" using hp "method" select size wf by -(rule RedCall,auto,cases T',auto) hence 3:"P,E \ \(ref(a,Cs))\M(map Val vs), (h\<^sub>2,l\<^sub>2)\ \* \blocks(this#pns,Class(last Cs')#Ts,Ref(a,Cs')#vs,body), (h\<^sub>2,l\<^sub>2)\" by(simp add:r_into_rtrancl) have "P,E \ \blocks(this#pns,Class(last Cs')#Ts,Ref(a,Cs')#vs,body),(h\<^sub>2,l\<^sub>2)\ \* \ef,(h\<^sub>3,override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 ({this} \ set pns))\" using True wf body' wwf size final casts' body_case by -(rule_tac vs'="Ref(a,Cs')#vs'" in blocksRedsFinal,simp_all,cases T',auto) with 1 2 3 show ?thesis using eql\<^sub>2 by simp next case False then obtain D where T':"T' = Class D" by auto with final body' body_case obtain s' e' where body'':"P,E(this \ Class (last Cs'),pns [\] Ts) \ \body,(h\<^sub>2,l\<^sub>2(this\ Ref(a,Cs'), pns[\]vs'))\ \* \e',s'\" and final':"final e'" and cast:"P,E(this \ Class (last Cs'), pns [\] Ts) \ \\D\e',s'\ \ \ef,(h\<^sub>3,l\<^sub>2++l\<^sub>3)\" by(cases T')(auto dest:Cast_final) from T' have "P,E \ \(ref(a,Cs))\M(map Val vs), (h\<^sub>2,l\<^sub>2)\ \ \\D\blocks(this#pns,Class(last Cs')#Ts,Ref(a,Cs')#vs,body), (h\<^sub>2,l\<^sub>2)\" using hp "method" select size wf by -(rule RedCall,auto) hence 3:"P,E \ \(ref(a,Cs))\M(map Val vs), (h\<^sub>2,l\<^sub>2)\ \* \\D\blocks(this#pns,Class(last Cs')#Ts,Ref(a,Cs')#vs,body),(h\<^sub>2,l\<^sub>2)\" by(simp add:r_into_rtrancl) from cast final have eq:"s' = (h\<^sub>3,l\<^sub>2++l\<^sub>3)" by(auto elim:red.cases simp:final_def) hence "P,E \ \blocks(this#pns, Class (last Cs')#Ts, Ref(a,Cs')#vs,body), (h\<^sub>2,l\<^sub>2)\ \* \e',(h\<^sub>3,override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 ({this} \ set pns))\" using wf body'' wwf size final' casts' by -(rule_tac vs'="Ref(a,Cs')#vs'" in blocksRedsFinal,simp_all) hence "P,E \ \\D\(blocks(this#pns,Class(last Cs')#Ts,Ref(a,Cs')#vs,body)),(h\<^sub>2,l\<^sub>2)\ \* \\D\e',(h\<^sub>3,override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 ({this} \ set pns))\" by(rule StaticCastReds) moreover have "P,E \ \\D\e',(h\<^sub>3,override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 ({this} \ set pns))\ \ \ef,(h\<^sub>3,override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 ({this} \ set pns))\" using eq cast final final' by(fastforce intro:Cast_final_eq) ultimately have "P,E \ \\D\(blocks(this#pns, Class (last Cs')#Ts, Ref(a,Cs')#vs,body)), (h\<^sub>2,l\<^sub>2)\ \* \ef,(h\<^sub>3,override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 ({this} \ set pns))\" by(rule_tac b="(\D\e',(h\<^sub>3,override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 ({this} \ set pns)))" in rtrancl_into_rtrancl,simp_all) with 1 2 3 show ?thesis using eql\<^sub>2 by simp qed qed lemma StaticCallRedsFinal: assumes wwf: "wwf_prog P" and "P,E \ \e,s\<^sub>0\ \* \ref(a,Cs),s\<^sub>1\" "P,E \ \es,s\<^sub>1\ [\]* \map Val vs,(h\<^sub>2,l\<^sub>2)\" and path_unique: "P \ Path (last Cs) to C unique" and path_via: "P \ Path (last Cs) to C via Cs''" and Ds: "Ds = (Cs@\<^sub>pCs'')@\<^sub>pCs'" and least: "P \ C has least M = (Ts,T,pns,body) via Cs'" and size: "size vs = size pns" and casts: "P \ Ts Casts vs to vs'" and l\<^sub>2': "l\<^sub>2' = [this \ Ref(a,Ds), pns[\]vs']" and body: "P,E(this\Class(last Ds), pns[\]Ts) \ \body,(h\<^sub>2,l\<^sub>2')\ \* \ef,(h\<^sub>3,l\<^sub>3)\" and final:"final ef" shows "P,E \ \e\(C::)M(es), s\<^sub>0\ \* \ef,(h\<^sub>3,l\<^sub>2)\" proof - have wf: "size Ts = size pns \ distinct pns \ this \ set pns \ (\T\set Ts. is_type P T)" and wt: "fv body \ {this} \ set pns" using assms by(fastforce dest!:has_least_wf_mdecl simp:wf_mdecl_def)+ have "dom l\<^sub>3 \ {this} \ set pns" using Reds_dom_lcl[OF wwf body] wt l\<^sub>2' set_take_subset by force hence eql\<^sub>2: "override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 ({this} \ set pns) = l\<^sub>2" by(fastforce simp add:map_add_def override_on_def fun_eq_iff) from wwf least have "Cs' \ []" by (auto elim!:Subobjs_nonempty simp:LeastMethodDef_def MethodDefs_def) with Ds have "last Cs' = last Ds" by(fastforce intro:appendPath_last) with wwf least have "is_class P (last Ds)" by(auto dest:Subobj_last_isClass simp:LeastMethodDef_def MethodDefs_def) hence "P \ Class (last Ds) casts Ref(a,Ds) to Ref(a,Ds)" by(auto intro!:casts_ref Subobjs_Base simp:path_via_def appendPath_def) with casts have casts':"P \ Class (last Ds)#Ts Casts Ref(a,Ds)#vs to Ref(a,Ds)#vs'" by -(rule Casts_Cons) have 1:"P,E \ \e\(C::)M(es),s\<^sub>0\ \* \(ref(a,Cs))\(C::)M(es),s\<^sub>1\" by(rule CallRedsObj)(rule assms(2)) have 2:"P,E \ \(ref(a,Cs))\(C::)M(es),s\<^sub>1\ \* \(ref(a,Cs))\(C::)M(map Val vs),(h\<^sub>2,l\<^sub>2)\" by(rule CallRedsParams)(rule assms(3)) from body[THEN Red_lcl_add, of l\<^sub>2] have body': "P,E(this\Class(last Ds), pns[\]Ts) \ \body,(h\<^sub>2,l\<^sub>2(this\ Ref(a,Ds), pns[\]vs'))\ \* \ef,(h\<^sub>3,l\<^sub>2++l\<^sub>3)\" by (simp add:l\<^sub>2') have "P,E \ \(ref(a,Cs))\(C::)M(map Val vs), (h\<^sub>2,l\<^sub>2)\ \ \blocks(this#pns, Class (last Ds)#Ts, Ref(a,Ds)#vs, body), (h\<^sub>2,l\<^sub>2)\" using path_unique path_via least size wf Ds by -(rule RedStaticCall,auto) hence 3:"P,E \ \(ref(a,Cs))\(C::)M(map Val vs), (h\<^sub>2,l\<^sub>2)\ \* \blocks(this#pns,Class(last Ds)#Ts,Ref(a,Ds)#vs,body), (h\<^sub>2,l\<^sub>2)\" by(simp add:r_into_rtrancl) have "P,E \ \blocks(this#pns,Class(last Ds)#Ts,Ref(a,Ds)#vs,body),(h\<^sub>2,l\<^sub>2)\ \* \ef,(h\<^sub>3,override_on (l\<^sub>2++l\<^sub>3) l\<^sub>2 ({this} \ set pns))\" using wf body' wwf size final casts' by -(rule_tac vs'="Ref(a,Ds)#vs'" in blocksRedsFinal,simp_all) with 1 2 3 show ?thesis using eql\<^sub>2 by simp qed lemma CallRedsThrowParams: "\ P,E \ \e,s\<^sub>0\ \* \Val v,s\<^sub>1\; P,E \ \es,s\<^sub>1\ [\]* \map Val vs\<^sub>1 @ Throw ex # es\<^sub>2,s\<^sub>2\ \ \ P,E \ \Call e Copt M es,s\<^sub>0\ \* \Throw ex,s\<^sub>2\" apply(rule rtrancl_trans) apply(erule CallRedsObj) apply(rule rtrancl_into_rtrancl) apply(erule CallRedsParams) apply(simp add:CallThrowParams) done lemma CallRedsThrowObj: "P,E \ \e,s\<^sub>0\ \* \Throw ex,s\<^sub>1\ \ P,E \ \Call e Copt M es,s\<^sub>0\ \* \Throw ex,s\<^sub>1\" apply(rule rtrancl_into_rtrancl) apply(erule CallRedsObj) apply(simp add:CallThrowObj) done lemma CallRedsNull: "\ P,E \ \e,s\<^sub>0\ \* \null,s\<^sub>1\; P,E \ \es,s\<^sub>1\ [\]* \map Val vs,s\<^sub>2\ \ \ P,E \ \Call e Copt M es,s\<^sub>0\ \* \THROW NullPointer,s\<^sub>2\" apply(rule rtrancl_trans) apply(erule CallRedsObj) apply(rule rtrancl_into_rtrancl) apply(erule CallRedsParams) apply(simp add:RedCallNull) done subsection \The main Theorem\ lemma assumes wwf: "wwf_prog P" shows big_by_small: "P,E \ \e,s\ \ \e',s'\ \ P,E \ \e,s\ \* \e',s'\" and bigs_by_smalls: "P,E \ \es,s\ [\] \es',s'\ \ P,E \ \es,s\ [\]* \es',s'\" proof (induct rule: eval_evals.inducts) case New thus ?case by (auto simp:RedNew) next case NewFail thus ?case by (auto simp:RedNewFail) next case StaticUpCast thus ?case by(simp add:StaticUpCastReds) next case StaticDownCast thus ?case by(simp add:StaticDownCastReds) next case StaticCastNull thus ?case by(simp add:StaticCastRedsNull) next case StaticCastFail thus ?case by(simp add:StaticCastRedsFail) next case StaticCastThrow thus ?case by(auto dest!:eval_final simp:StaticCastRedsThrow) next case StaticUpDynCast thus ?case by(simp add:StaticUpDynCastReds) next case StaticDownDynCast thus ?case by(simp add:StaticDownDynCastReds) next case DynCast thus ?case by(fastforce intro:DynCastRedsRef) next case DynCastNull thus ?case by(simp add:DynCastRedsNull) next case DynCastFail thus ?case by(fastforce intro!:DynCastRedsFail) next case DynCastThrow thus ?case by(auto dest!:eval_final simp:DynCastRedsThrow) next case Val thus ?case by simp next case BinOp thus ?case by(fastforce simp:BinOpRedsVal) next case BinOpThrow1 thus ?case by(fastforce dest!:eval_final simp: BinOpRedsThrow1) next case BinOpThrow2 thus ?case by(fastforce dest!:eval_final simp: BinOpRedsThrow2) next case Var thus ?case by (fastforce simp:RedVar) next case LAss thus ?case by(fastforce simp: LAssRedsVal) next case LAssThrow thus ?case by(fastforce dest!:eval_final simp: LAssRedsThrow) next case FAcc thus ?case by(fastforce intro:FAccRedsVal) next case FAccNull thus ?case by(simp add:FAccRedsNull) next case FAccThrow thus ?case by(fastforce dest!:eval_final simp:FAccRedsThrow) next case FAss thus ?case by(fastforce simp:FAssRedsVal) next case FAssNull thus ?case by(fastforce simp:FAssRedsNull) next case FAssThrow1 thus ?case by(fastforce dest!:eval_final simp:FAssRedsThrow1) next case FAssThrow2 thus ?case by(fastforce dest!:eval_final simp:FAssRedsThrow2) next case CallObjThrow thus ?case by(fastforce dest!:eval_final simp:CallRedsThrowObj) next case CallNull thus ?case thm CallRedsNull by(simp add:CallRedsNull) next case CallParamsThrow thus ?case by(fastforce dest!:evals_final simp:CallRedsThrowParams) next case (Call E e s\<^sub>0 a Cs s\<^sub>1 ps vs h\<^sub>2 l\<^sub>2 C S M Ts' T' pns' body' Ds Ts T pns body Cs' vs' l\<^sub>2' new_body e' h\<^sub>3 l\<^sub>3) have IHe: "P,E \ \e,s\<^sub>0\ \* \ref(a,Cs),s\<^sub>1\" and IHes: "P,E \ \ps,s\<^sub>1\ [\]* \map Val vs,(h\<^sub>2,l\<^sub>2)\" and h\<^sub>2a: "h\<^sub>2 a = Some(C,S)" and "method": "P \ last Cs has least M = (Ts',T',pns',body') via Ds" and select: "P \ (C,Cs@\<^sub>pDs) selects M = (Ts,T,pns,body) via Cs'" and same_length: "length vs = length pns" and casts: "P \ Ts Casts vs to vs'" and l\<^sub>2': "l\<^sub>2' = [this \ Ref (a,Cs'), pns[\]vs']" and body_case: "new_body = (case T' of Class D \ \D\body | _ \ body)" and eval_body: "P,E(this \ Class (last Cs'), pns [\] Ts) \ \new_body,(h\<^sub>2, l\<^sub>2')\ \ \e',(h\<^sub>3, l\<^sub>3)\" and IHbody: "P,E(this \ Class (last Cs'), pns [\] Ts) \ \new_body,(h\<^sub>2, l\<^sub>2')\ \* \e',(h\<^sub>3, l\<^sub>3)\" by fact+ from wwf select same_length have lengthTs:"length Ts = length vs" by (fastforce dest!:select_method_wf_mdecl simp:wf_mdecl_def) show "P,E \ \e\M(ps),s\<^sub>0\ \* \e',(h\<^sub>3, l\<^sub>2)\" using "method" select same_length l\<^sub>2' h\<^sub>2a casts body_case IHbody eval_final[OF eval_body] by(fastforce intro!:CallRedsFinal[OF wwf IHe IHes]) next case (StaticCall E e s\<^sub>0 a Cs s\<^sub>1 ps vs h\<^sub>2 l\<^sub>2 C Cs'' M Ts T pns body Cs' Ds vs' l\<^sub>2' e' h\<^sub>3 l\<^sub>3) have IHe: "P,E \ \e,s\<^sub>0\ \* \ref(a,Cs),s\<^sub>1\" and IHes: "P,E \ \ps,s\<^sub>1\ [\]* \map Val vs,(h\<^sub>2,l\<^sub>2)\" and path_unique: "P \ Path last Cs to C unique" and path_via: "P \ Path last Cs to C via Cs''" and least: "P \ C has least M = (Ts, T, pns, body) via Cs'" and Ds: "Ds = (Cs @\<^sub>p Cs'') @\<^sub>p Cs'" and same_length: "length vs = length pns" and casts: "P \ Ts Casts vs to vs'" and l\<^sub>2': "l\<^sub>2' = [this \ Ref (a,Ds), pns[\]vs']" and eval_body: "P,E(this \ Class (last Ds), pns [\] Ts) \ \body,(h\<^sub>2, l\<^sub>2')\ \ \e',(h\<^sub>3, l\<^sub>3)\" and IHbody: "P,E(this \ Class (last Ds), pns [\] Ts) \ \body,(h\<^sub>2, l\<^sub>2')\ \* \e',(h\<^sub>3, l\<^sub>3)\" by fact+ from wwf least same_length have lengthTs:"length Ts = length vs" by (fastforce dest!:has_least_wf_mdecl simp:wf_mdecl_def) show "P,E \ \e\(C::)M(ps),s\<^sub>0\ \* \e',(h\<^sub>3, l\<^sub>2)\" using path_unique path_via least Ds same_length l\<^sub>2' casts IHbody eval_final[OF eval_body] by(fastforce intro!:StaticCallRedsFinal[OF wwf IHe IHes]) next case Block with wwf show ?case by(fastforce simp: BlockRedsFinal dest:eval_final) next case Seq thus ?case by(fastforce simp:SeqReds2) next case SeqThrow thus ?case by(fastforce dest!:eval_final simp: SeqRedsThrow) next case CondT thus ?case by(fastforce simp:CondReds2T) next case CondF thus ?case by(fastforce simp:CondReds2F) next case CondThrow thus ?case by(fastforce dest!:eval_final simp:CondRedsThrow) next case WhileF thus ?case by(fastforce simp:WhileFReds) next case WhileT thus ?case by(fastforce simp: WhileTReds) next case WhileCondThrow thus ?case by(fastforce dest!:eval_final simp: WhileRedsThrow) next case WhileBodyThrow thus ?case by(fastforce dest!:eval_final simp: WhileTRedsThrow) next case Throw thus ?case by(fastforce simp:ThrowReds) next case ThrowNull thus ?case by(fastforce simp:ThrowRedsNull) next case ThrowThrow thus ?case by(fastforce dest!:eval_final simp:ThrowRedsThrow) next case Nil thus ?case by simp next case Cons thus ?case by(fastforce intro!:Cons_eq_appendI[OF refl refl] ListRedsVal) next case ConsThrow thus ?case by(fastforce elim: ListReds1) qed subsection\Big steps simulates small step\ text \The big step equivalent of \RedWhile\:\ lemma unfold_while: "P,E \ \while(b) c,s\ \ \e',s'\ = P,E \ \if(b) (c;;while(b) c) else (unit),s\ \ \e',s'\" proof assume "P,E \ \while (b) c,s\ \ \e',s'\" thus "P,E \ \if (b) (c;; while (b) c) else unit,s\ \ \e',s'\" by cases (fastforce intro: eval_evals.intros)+ next assume "P,E \ \if (b) (c;; while (b) c) else unit,s\ \ \e',s'\" thus "P,E \ \while (b) c,s\ \ \e',s'\" proof (cases) fix ex assume e': "e' = throw ex" assume "P,E \ \b,s\ \ \throw ex,s'\" hence "P,E \ \while(b) c,s\ \ \throw ex,s'\" by (rule WhileCondThrow) with e' show ?thesis by simp next fix s\<^sub>1 assume eval_false: "P,E \ \b,s\ \ \false,s\<^sub>1\" and eval_unit: "P,E \ \unit,s\<^sub>1\ \ \e',s'\" with eval_unit have "s' = s\<^sub>1" "e' = unit" by (auto elim: eval_cases) moreover from eval_false have "P,E \ \while (b) c,s\ \ \unit,s\<^sub>1\" by - (rule WhileF, simp) ultimately show ?thesis by simp next fix s\<^sub>1 assume eval_true: "P,E \ \b,s\ \ \true,s\<^sub>1\" and eval_rest: "P,E \ \c;; while (b) c,s\<^sub>1\\\e',s'\" from eval_rest show ?thesis proof (cases) fix s\<^sub>2 v\<^sub>1 assume "P,E \ \c,s\<^sub>1\ \ \Val v\<^sub>1,s\<^sub>2\" "P,E \ \while (b) c,s\<^sub>2\ \ \e',s'\" with eval_true show "P,E \ \while(b) c,s\ \ \e',s'\" by (rule WhileT) next fix ex assume "P,E \ \c,s\<^sub>1\ \ \throw ex,s'\" "e' = throw ex" with eval_true show "P,E \ \while(b) c,s\ \ \e',s'\" by (iprover intro: WhileBodyThrow) qed qed qed lemma blocksEval: "\Ts vs l l' E. \size ps = size Ts; size ps = size vs; P,E \ \blocks(ps,Ts,vs,e),(h,l)\ \ \e',(h',l')\ \ \ \ l'' vs'. P,E(ps [\] Ts) \ \e,(h,l(ps[\]vs'))\ \ \e',(h',l'')\ \ P \ Ts Casts vs to vs' \ length vs' = length vs" proof (induct ps) case Nil then show ?case by(fastforce intro:Casts_Nil) next case (Cons p ps') have length_eqs: "length (p # ps') = length Ts" "length (p # ps') = length vs" and IH:"\Ts vs l l' E. \length ps' = length Ts; length ps' = length vs; P,E \ \blocks (ps',Ts,vs,e),(h,l)\ \ \e',(h',l')\\ \ \l'' vs'. P,E(ps' [\] Ts) \ \e,(h,l(ps' [\] vs'))\ \ \e',(h', l'')\ \ P \ Ts Casts vs to vs' \ length vs' = length vs" by fact+ then obtain T Ts' where Ts: "Ts = T#Ts'" by (cases "Ts") simp obtain v vs' where vs: "vs = v#vs'" using length_eqs by (cases "vs") simp with length_eqs Ts have length1:"length ps' = length Ts'" and length2:"length ps' = length vs'" by simp_all have "P,E \ \blocks (p # ps', Ts, vs, e),(h,l)\ \ \e',(h', l')\" by fact with Ts vs have blocks:"P,E \ \{p:T := Val v; blocks (ps',Ts',vs',e)},(h,l)\ \ \e',(h',l')\" by simp then obtain l''' v' where eval_ps': "P,E(p \ T) \ \blocks (ps',Ts',vs',e),(h,l(p\v'))\ \ \e',(h',l''')\" and l''': "l'=l'''(p:=l p)" and casts:"P \ T casts v to v'" by(auto elim!: eval_cases simp:fun_upd_same) from IH[OF length1 length2 eval_ps'] obtain l'' vs'' where - "P,E(p \ T)(ps' [\] Ts') \ \e,(h, l(p\v')(ps'[\]vs''))\ \ + "P,E(p \ T, ps' [\] Ts') \ \e,(h, l(p\v', ps'[\]vs''))\ \ \e',(h',l'')\" and "P \ Ts' Casts vs' to vs''" and "length vs'' = length vs'" by auto with Ts vs casts show ?case by -(rule_tac x="l''" in exI,rule_tac x="v'#vs''" in exI,simp, rule Casts_Cons) qed lemma CastblocksEval: "\Ts vs l l' E. \size ps = size Ts; size ps = size vs; P,E \ \\C'\(blocks(ps,Ts,vs,e)),(h,l)\ \ \e',(h',l')\ \ \ \ l'' vs'. P,E(ps [\] Ts) \ \\C'\e,(h,l(ps[\]vs'))\ \ \e',(h',l'')\ \ P \ Ts Casts vs to vs' \ length vs' = length vs" proof (induct ps) case Nil then show ?case by(fastforce intro:Casts_Nil) next case (Cons p ps') have length_eqs: "length (p # ps') = length Ts" "length (p # ps') = length vs" by fact+ then obtain T Ts' where Ts: "Ts = T#Ts'" by (cases "Ts") simp obtain v vs' where vs: "vs = v#vs'" using length_eqs by (cases "vs") simp with length_eqs Ts have length1:"length ps' = length Ts'" and length2:"length ps' = length vs'" by simp_all have "P,E \ \\C'\(blocks (p # ps', Ts, vs, e)),(h,l)\ \ \e',(h', l')\" by fact moreover { fix a Cs Cs' assume blocks:"P,E \ \blocks(p#ps',Ts,vs,e),(h,l)\ \ \ref (a,Cs),(h',l')\" and path_via:"P \ Path last Cs to C' via Cs'" and e':"e' = ref(a,Cs@\<^sub>pCs')" from blocks length_eqs obtain l'' vs'' where eval:"P,E(p#ps' [\] Ts) \ \e,(h,l(p#ps'[\]vs''))\ \ \ref (a,Cs),(h',l'')\" and casts:"P \ Ts Casts vs to vs''" and length:"length vs'' = length vs" by -(drule blocksEval,auto) from eval path_via have "P,E(p#ps'[\]Ts) \ \\C'\e,(h,l(p#ps'[\]vs''))\ \ \ref(a,Cs@\<^sub>pCs'),(h',l'')\" by(auto intro:StaticUpCast) with e' casts length have ?case by simp blast } moreover { fix a Cs Cs' assume blocks:"P,E \ \blocks(p#ps',Ts,vs,e),(h,l)\ \ \ref (a,Cs@C'# Cs'),(h',l')\" and e':"e' = ref (a,Cs@[C'])" from blocks length_eqs obtain l'' vs'' where eval:"P,E(p#ps' [\] Ts) \ \e,(h,l(p#ps'[\]vs''))\ \ \ref (a,Cs@C'# Cs'),(h',l'')\" and casts:"P \ Ts Casts vs to vs''" and length:"length vs'' = length vs" by -(drule blocksEval,auto) from eval have "P,E(p#ps'[\]Ts) \ \\C'\e,(h,l(p#ps'[\]vs''))\ \ \ref(a,Cs@[C']),(h',l'')\" by(auto intro:StaticDownCast) with e' casts length have ?case by simp blast } moreover { assume "P,E \ \blocks(p#ps',Ts,vs,e),(h,l)\ \ \null,(h',l')\" and e':"e' = null" with length_eqs obtain l'' vs'' where eval:"P,E(p#ps' [\] Ts) \ \e,(h,l(p#ps'[\]vs''))\ \ \null,(h',l'')\" and casts:"P \ Ts Casts vs to vs''" and length:"length vs'' = length vs" by -(drule blocksEval,auto) from eval have "P,E(p#ps' [\] Ts) \ \\C'\e,(h,l(p#ps'[\]vs''))\ \ \null,(h',l'')\" by(auto intro:StaticCastNull) with e' casts length have ?case by simp blast } moreover { fix a Cs assume blocks:"P,E \ \blocks(p#ps',Ts,vs,e),(h,l)\ \ \ref (a,Cs),(h',l')\" and notin:"C' \ set Cs" and leq:"\ P \ (last Cs) \\<^sup>* C'" and e':"e' = THROW ClassCast" from blocks length_eqs obtain l'' vs'' where eval:"P,E(p#ps' [\] Ts) \ \e,(h,l(p#ps'[\]vs''))\ \ \ref (a,Cs),(h',l'')\" and casts:"P \ Ts Casts vs to vs''" and length:"length vs'' = length vs" by -(drule blocksEval,auto) from eval notin leq have "P,E(p#ps'[\]Ts) \ \\C'\e,(h,l(p#ps'[\]vs''))\ \ \THROW ClassCast,(h',l'')\" by(auto intro:StaticCastFail) with e' casts length have ?case by simp blast } moreover { fix r assume "P,E \ \blocks(p#ps',Ts,vs,e),(h,l)\ \ \throw r,(h',l')\" and e':"e' = throw r" with length_eqs obtain l'' vs'' where eval:"P,E(p#ps' [\] Ts) \ \e,(h,l(p#ps'[\]vs''))\ \ \throw r,(h',l'')\" and casts:"P \ Ts Casts vs to vs''" and length:"length vs'' = length vs" by -(drule blocksEval,auto) from eval have "P,E(p#ps'[\]Ts) \ \\C'\e,(h,l(p#ps'[\]vs''))\ \ \throw r,(h',l'')\" by(auto intro:eval_evals.StaticCastThrow) with e' casts length have ?case by simp blast } ultimately show ?case by -(erule eval_cases,fastforce+) qed lemma assumes wf: "wwf_prog P" shows eval_restrict_lcl: "P,E \ \e,(h,l)\ \ \e',(h',l')\ \ (\W. fv e \ W \ P,E \ \e,(h,l|`W)\ \ \e',(h',l'|`W)\)" and "P,E \ \es,(h,l)\ [\] \es',(h',l')\ \ (\W. fvs es \ W \ P,E \ \es,(h,l|`W)\ [\] \es',(h',l'|`W)\)" proof(induct rule:eval_evals_inducts) case (Block E V T e\<^sub>0 h\<^sub>0 l\<^sub>0 e\<^sub>1 h\<^sub>1 l\<^sub>1) have IH: "\W. fv e\<^sub>0 \ W \ P,E(V \ T) \ \e\<^sub>0,(h\<^sub>0,l\<^sub>0(V:=None)|`W)\ \ \e\<^sub>1,(h\<^sub>1,l\<^sub>1|`W)\" by fact (*have type:"is_type P T" .*) have "fv({V:T; e\<^sub>0}) \ W" by fact hence "fv e\<^sub>0 - {V} \ W" by simp_all hence "fv e\<^sub>0 \ insert V W" by fast with IH[OF this] have "P,E(V \ T) \ \e\<^sub>0,(h\<^sub>0, (l\<^sub>0|`W)(V := None))\ \ \e\<^sub>1,(h\<^sub>1, l\<^sub>1|`insert V W)\" by fastforce from eval_evals.Block[OF this] show ?case by fastforce next case Seq thus ?case by simp (blast intro:eval_evals.Seq) next case New thus ?case by(simp add:eval_evals.intros) next case NewFail thus ?case by(simp add:eval_evals.intros) next case StaticUpCast thus ?case by simp (blast intro:eval_evals.StaticUpCast) next case (StaticDownCast E e h l a Cs C Cs' h' l') have IH:"\W. fv e \ W \ P,E \ \e,(h,l |` W)\ \ \ref(a,Cs@[C]@Cs'),(h',l' |` W)\" by fact have "fv (\C\e) \ W" by fact hence "fv e \ W" by simp from IH[OF this] show ?case by(rule eval_evals.StaticDownCast) next case StaticCastNull thus ?case by simp (blast intro:eval_evals.StaticCastNull) next case StaticCastFail thus ?case by simp (blast intro:eval_evals.StaticCastFail) next case StaticCastThrow thus ?case by(simp add:eval_evals.intros) next case DynCast thus ?case by simp (blast intro:eval_evals.DynCast) next case StaticUpDynCast thus ?case by simp (blast intro:eval_evals.StaticUpDynCast) next case (StaticDownDynCast E e h l a Cs C Cs' h' l') have IH:"\W. fv e \ W \ P,E \ \e,(h,l |` W)\ \ \ref(a,Cs@[C]@Cs'),(h',l' |` W)\" by fact have "fv (Cast C e) \ W" by fact hence "fv e \ W" by simp from IH[OF this] show ?case by(rule eval_evals.StaticDownDynCast) next case DynCastNull thus ?case by simp (blast intro:eval_evals.DynCastNull) next case DynCastFail thus ?case by simp (blast intro:eval_evals.DynCastFail) next case DynCastThrow thus ?case by(simp add:eval_evals.intros) next case Val thus ?case by(simp add:eval_evals.intros) next case BinOp thus ?case by simp (blast intro:eval_evals.BinOp) next case BinOpThrow1 thus ?case by simp (blast intro:eval_evals.BinOpThrow1) next case BinOpThrow2 thus ?case by simp (blast intro:eval_evals.BinOpThrow2) next case Var thus ?case by(simp add:eval_evals.intros) next case (LAss E e h\<^sub>0 l\<^sub>0 v h l V T v' l') have IH: "\W. fv e \ W \ P,E \ \e,(h\<^sub>0,l\<^sub>0|`W)\ \ \Val v,(h,l|`W)\" and env:"E V = \T\" and casts:"P \ T casts v to v'" and [simp]: "l' = l(V \ v')" by fact+ have "fv (V:=e) \ W" by fact hence fv: "fv e \ W" and VinW: "V \ W" by auto from eval_evals.LAss[OF IH[OF fv] _ casts] env VinW show ?case by fastforce next case LAssThrow thus ?case by(fastforce intro: eval_evals.LAssThrow) next case FAcc thus ?case by simp (blast intro: eval_evals.FAcc) next case FAccNull thus ?case by(fastforce intro: eval_evals.FAccNull) next case FAccThrow thus ?case by(fastforce intro: eval_evals.FAccThrow) next case (FAss E e\<^sub>1 h l a Cs' h' l' e\<^sub>2 v h\<^sub>2 l\<^sub>2 D S F T Cs v' Ds fs fs' S' h\<^sub>2' W) have IH1: "\W. fv e\<^sub>1 \ W \ P,E \ \e\<^sub>1,(h, l|`W)\ \ \ref (a, Cs'),(h', l'|`W)\" and IH2: "\W. fv e\<^sub>2 \ W \ P,E \ \e\<^sub>2,(h', l'|`W)\ \ \Val v,(h\<^sub>2, l\<^sub>2|`W)\" and fv:"fv (e\<^sub>1\F{Cs} := e\<^sub>2) \ W" and h:"h\<^sub>2 a = Some(D,S)" and Ds:"Ds = Cs' @\<^sub>p Cs" and S:"(Ds,fs) \ S" and fs':"fs' = fs(F \ v')" and S':"S' = S - {(Ds, fs)} \ {(Ds, fs')}" and h':"h\<^sub>2' = h\<^sub>2(a \ (D, S'))" and field:"P \ last Cs' has least F:T via Cs" and casts:"P \ T casts v to v'" by fact+ from fv have fv1:"fv e\<^sub>1 \ W" and fv2:"fv e\<^sub>2 \ W" by auto from eval_evals.FAss[OF IH1[OF fv1] IH2[OF fv2] _ field casts] h Ds S fs' S' h' show ?case by simp next case FAssNull thus ?case by simp (blast intro: eval_evals.FAssNull) next case FAssThrow1 thus ?case by simp (blast intro: eval_evals.FAssThrow1) next case FAssThrow2 thus ?case by simp (blast intro: eval_evals.FAssThrow2) next case CallObjThrow thus ?case by simp (blast intro: eval_evals.intros) next case CallNull thus ?case by simp (blast intro: eval_evals.CallNull) next case CallParamsThrow thus ?case by simp (blast intro: eval_evals.CallParamsThrow) next case (Call E e h\<^sub>0 l\<^sub>0 a Cs h\<^sub>1 l\<^sub>1 ps vs h\<^sub>2 l\<^sub>2 C S M Ts' T' pns' body' Ds Ts T pns body Cs' vs' l\<^sub>2' new_body e' h\<^sub>3 l\<^sub>3 W) have IHe: "\W. fv e \ W \ P,E \ \e,(h\<^sub>0,l\<^sub>0|`W)\ \ \ref(a,Cs),(h\<^sub>1,l\<^sub>1|`W)\" and IHps: "\W. fvs ps \ W \ P,E \ \ps,(h\<^sub>1,l\<^sub>1|`W)\ [\] \map Val vs,(h\<^sub>2,l\<^sub>2|`W)\" and IHbd: "\W. fv new_body \ W \ P,E(this \ Class (last Cs'), pns [\] Ts) \ \new_body,(h\<^sub>2,l\<^sub>2'|`W)\ \ \e',(h\<^sub>3,l\<^sub>3|`W)\" and h\<^sub>2a: "h\<^sub>2 a = Some (C,S)" and "method": "P \ last Cs has least M = (Ts',T',pns',body') via Ds" and select:"P \ (C,Cs@\<^sub>pDs) selects M = (Ts,T,pns,body) via Cs'" and same_len: "size vs = size pns" and casts:"P \ Ts Casts vs to vs'" and l\<^sub>2': "l\<^sub>2' = [this \ Ref(a,Cs'), pns [\] vs']" and body_case: "new_body = (case T' of Class D \ \D\body | _ \ body)" by fact+ have "fv (e\M(ps)) \ W" by fact hence fve: "fv e \ W" and fvps: "fvs(ps) \ W" by auto have wfmethod: "size Ts = size pns \ this \ set pns" and fvbd: "fv body \ {this} \ set pns" using select wf by(fastforce dest!:select_method_wf_mdecl simp:wf_mdecl_def)+ from fvbd body_case have fvbd':"fv new_body \ {this} \ set pns" by(cases T') auto from l\<^sub>2' have "l\<^sub>2' |` ({this} \ set pns) = [this \ Ref (a, Cs'), pns [\] vs']" by (auto intro!:ext simp:restrict_map_def fun_upd_def) with eval_evals.Call[OF IHe[OF fve] IHps[OF fvps] _ "method" select same_len casts _ body_case IHbd[OF fvbd']] h\<^sub>2a show ?case by simp next case (StaticCall E e h\<^sub>0 l\<^sub>0 a Cs h\<^sub>1 l\<^sub>1 ps vs h\<^sub>2 l\<^sub>2 C Cs'' M Ts T pns body Cs' Ds vs' l\<^sub>2' e' h\<^sub>3 l\<^sub>3 W) have IHe: "\W. fv e \ W \ P,E \ \e,(h\<^sub>0,l\<^sub>0|`W)\ \ \ref(a,Cs),(h\<^sub>1,l\<^sub>1|`W)\" and IHps: "\W. fvs ps \ W \ P,E \ \ps,(h\<^sub>1,l\<^sub>1|`W)\ [\] \map Val vs,(h\<^sub>2,l\<^sub>2|`W)\" and IHbd: "\W. fv body \ W \ P,E(this \ Class (last Ds), pns [\] Ts) \ \body,(h\<^sub>2,l\<^sub>2'|`W)\ \ \e',(h\<^sub>3,l\<^sub>3|`W)\" and path_unique: "P \ Path last Cs to C unique" and path_via: "P \ Path last Cs to C via Cs''" and least: "P \ C has least M = (Ts, T, pns, body) via Cs'" and Ds: "Ds = (Cs @\<^sub>p Cs'') @\<^sub>p Cs'" and same_len: "size vs = size pns" and casts:"P \ Ts Casts vs to vs'" and l\<^sub>2': "l\<^sub>2' = [this \ Ref(a,Ds), pns [\] vs']" by fact+ have "fv (e\(C::)M(ps)) \ W" by fact hence fve: "fv e \ W" and fvps: "fvs(ps) \ W" by auto have wfmethod: "size Ts = size pns \ this \ set pns" and fvbd: "fv body \ {this} \ set pns" using least wf by(fastforce dest!:has_least_wf_mdecl simp:wf_mdecl_def)+ from fvbd have fvbd':"fv body \ {this} \ set pns" by auto from l\<^sub>2' have "l\<^sub>2' |` ({this} \ set pns) = [this \ Ref(a,Ds), pns [\] vs']" by (auto intro!:ext simp:restrict_map_def fun_upd_def) with eval_evals.StaticCall[OF IHe[OF fve] IHps[OF fvps] path_unique path_via least Ds same_len casts _ IHbd[OF fvbd']] show ?case by simp next case SeqThrow thus ?case by simp (blast intro: eval_evals.SeqThrow) next case CondT thus ?case by simp (blast intro: eval_evals.CondT) next case CondF thus ?case by simp (blast intro: eval_evals.CondF) next case CondThrow thus ?case by simp (blast intro: eval_evals.CondThrow) next case WhileF thus ?case by simp (blast intro: eval_evals.WhileF) next case WhileT thus ?case by simp (blast intro: eval_evals.WhileT) next case WhileCondThrow thus ?case by simp (blast intro: eval_evals.WhileCondThrow) next case WhileBodyThrow thus ?case by simp (blast intro: eval_evals.WhileBodyThrow) next case Throw thus ?case by simp (blast intro: eval_evals.Throw) next case ThrowNull thus ?case by simp (blast intro: eval_evals.ThrowNull) next case ThrowThrow thus ?case by simp (blast intro: eval_evals.ThrowThrow) next case Nil thus ?case by (simp add: eval_evals.Nil) next case Cons thus ?case by simp (blast intro: eval_evals.Cons) next case ConsThrow thus ?case by simp (blast intro: eval_evals.ConsThrow) qed lemma eval_notfree_unchanged: assumes wf:"wwf_prog P" shows "P,E \ \e,(h,l)\ \ \e',(h',l')\ \ (\V. V \ fv e \ l' V = l V)" and "P,E \ \es,(h,l)\ [\] \es',(h',l')\ \ (\V. V \ fvs es \ l' V = l V)" proof(induct rule:eval_evals_inducts) case LAss thus ?case by(simp add:fun_upd_apply) next case Block thus ?case by (simp only:fun_upd_apply split:if_splits) fastforce qed simp_all lemma eval_closed_lcl_unchanged: assumes wf:"wwf_prog P" and eval:"P,E \ \e,(h,l)\ \ \e',(h',l')\" and fv:"fv e = {}" shows "l' = l" proof - from wf eval have "\V. V \ fv e \ l' V = l V" by (rule eval_notfree_unchanged) with fv have "\V. l' V = l V" by simp thus ?thesis by(simp add:fun_eq_iff) qed (* Hiermit kann man die ganze pair-Splitterei in den automatischen Taktiken abschalten. Wieder anschalten siehe nach dem Beweis. *) declare split_paired_All [simp del] split_paired_Ex [simp del] declaration \K (Simplifier.map_ss (fn ss => ss delloop "split_all_tac"))\ setup \map_theory_claset (fn ctxt => ctxt delSWrapper "split_all_tac")\ lemma list_eval_Throw: assumes eval_e: "P,E \ \throw x,s\ \ \e',s'\" shows "P,E \ \map Val vs @ throw x # es',s\ [\] \map Val vs @ e' # es',s'\" proof - from eval_e obtain a where e': "e' = Throw a" by (cases) (auto dest!: eval_final) { fix es have "\vs. es = map Val vs @ throw x # es' \ P,E \ \es,s\[\]\map Val vs @ e' # es',s'\" proof (induct es type: list) case Nil thus ?case by simp next case (Cons e es vs) have e_es: "e # es = map Val vs @ throw x # es'" by fact show "P,E \ \e # es,s\ [\] \map Val vs @ e' # es',s'\" proof (cases vs) case Nil with e_es obtain "e=throw x" "es=es'" by simp moreover from eval_e e' have "P,E \ \throw x # es,s\ [\] \Throw a # es,s'\" by (iprover intro: ConsThrow) ultimately show ?thesis using Nil e' by simp next case (Cons v vs') have vs: "vs = v # vs'" by fact with e_es obtain e: "e=Val v" and es:"es= map Val vs' @ throw x # es'" by simp from e have "P,E \ \e,s\ \ \Val v,s\" by (iprover intro: eval_evals.Val) moreover from es have "P,E \ \es,s\ [\] \map Val vs' @ e' # es',s'\" by (rule Cons.hyps) ultimately show "P,E \ \e#es,s\ [\] \map Val vs @ e' # es',s'\" using vs by (auto intro: eval_evals.Cons) qed qed } thus ?thesis by simp qed text \The key lemma:\ lemma assumes wf: "wwf_prog P" shows extend_1_eval: "P,E \ \e,s\ \ \e'',s''\ \ (\s' e'. P,E \ \e'',s''\ \ \e',s'\ \ P,E \ \e,s\ \ \e',s'\)" and extend_1_evals: "P,E \ \es,t\ [\] \es'',t''\ \ (\t' es'. P,E \ \es'',t''\ [\] \es',t'\ \ P,E \ \es,t\ [\] \es',t'\)" proof (induct rule: red_reds.inducts) case RedNew thus ?case by (iprover elim: eval_cases intro: eval_evals.intros) next case RedNewFail thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case (StaticCastRed E e s e'' s'' C s' e') thus ?case by -(erule eval_cases,auto intro:eval_evals.intros, subgoal_tac "P,E \ \e'',s''\ \ \ref(a,Cs@[C]@Cs'),s'\", rule_tac Cs'="Cs'" in StaticDownCast,auto) next case RedStaticCastNull thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case RedStaticUpCast thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case RedStaticDownCast thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case RedStaticCastFail thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case RedStaticUpDynCast thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case RedStaticDownDynCast thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case (DynCastRed E e s e'' s'' C s' e') have eval:"P,E \ \Cast C e'',s''\ \ \e',s'\" and IH:"\ex sx. P,E \ \e'',s''\ \ \ex,sx\ \ P,E \ \e,s\ \ \ex,sx\" by fact+ moreover { fix Cs Cs' a assume "P,E \ \e'',s''\ \ \ref (a, Cs @ C # Cs'),s'\" from IH[OF this] have "P,E \ \e,s\ \ \ref (a, Cs@[C]@Cs'),s'\" by simp hence "P,E \ \Cast C e,s\ \ \ref (a, Cs@[C]),s'\" by(rule StaticDownDynCast) } ultimately show ?case by -(erule eval_cases,auto intro: eval_evals.intros) next case RedDynCastNull thus ?case by (iprover elim:eval_cases intro:eval_evals.intros) next case (RedDynCast s a D S C Cs' E Cs s' e') thus ?case by (cases s)(auto elim!:eval_cases intro:eval_evals.intros) next case (RedDynCastFail s a D S C Cs E s'' e'') thus ?case by (cases s)(auto elim!: eval_cases intro: eval_evals.intros) next case BinOpRed1 thus ?case by -(erule eval_cases,auto intro: eval_evals.intros) next case BinOpRed2 thus ?case by (fastforce elim!:eval_cases intro:eval_evals.intros eval_finalId) next case RedBinOp thus ?case by (iprover elim:eval_cases intro:eval_evals.intros) next case (RedVar s V v E s' e') thus ?case by (cases s)(fastforce elim:eval_cases intro:eval_evals.intros) next case LAssRed thus ?case by -(erule eval_cases,auto intro:eval_evals.intros) next case RedLAss thus ?case by (fastforce elim:eval_cases intro:eval_evals.intros) next case FAccRed thus ?case by -(erule eval_cases,auto intro:eval_evals.intros) next case (RedFAcc s a D S Ds Cs' Cs fs F v E s' e') thus ?case by (cases s)(fastforce elim:eval_cases intro:eval_evals.intros) next case RedFAccNull thus ?case by (fastforce elim!:eval_cases intro:eval_evals.intros) next case (FAssRed1 E e\<^sub>1 s e\<^sub>1' s'' F Cs e\<^sub>2 s' e') have eval:"P,E \ \e\<^sub>1'\F{Cs} := e\<^sub>2,s''\ \ \e',s'\" and IH:"\ex sx. P,E \ \e\<^sub>1',s''\ \ \ex,sx\ \ P,E \ \e\<^sub>1,s\ \ \ex,sx\" by fact+ { fix Cs' D S T a fs h\<^sub>2 l\<^sub>2 s\<^sub>1 v v' assume ref:"P,E \ \e\<^sub>1',s''\ \ \ref (a, Cs'),s\<^sub>1\" and rest:"P,E \ \e\<^sub>2,s\<^sub>1\ \ \Val v,(h\<^sub>2, l\<^sub>2)\" "h\<^sub>2 a = \(D, S)\" "P \ last Cs' has least F:T via Cs" "P \ T casts v to v'" "(Cs' @\<^sub>p Cs, fs) \ S" from IH[OF ref] have "P,E \ \e\<^sub>1,s\ \ \ref (a, Cs'),s\<^sub>1\" . with rest have "P,E \ \e\<^sub>1\F{Cs} := e\<^sub>2,s\ \ \Val v',(h\<^sub>2(a \ (D,insert (Cs'@\<^sub>pCs,fs(F \ v'))(S - {(Cs'@\<^sub>pCs,fs)}))),l\<^sub>2)\" by-(rule FAss,simp_all) } moreover { fix s\<^sub>1 v assume null:"P,E \ \e\<^sub>1',s''\ \ \null,s\<^sub>1\" and rest:"P,E \ \e\<^sub>2,s\<^sub>1\ \ \Val v,s'\" from IH[OF null] have "P,E \ \e\<^sub>1,s\ \ \null,s\<^sub>1\" . with rest have "P,E \ \e\<^sub>1\F{Cs} := e\<^sub>2,s\ \ \THROW NullPointer,s'\" by-(rule FAssNull,simp_all) } moreover { fix e' assume throw:"P,E \ \e\<^sub>1',s''\ \ \throw e',s'\" from IH[OF throw] have "P,E \ \e\<^sub>1,s\ \ \throw e',s'\" . hence "P,E \ \e\<^sub>1\F{Cs} := e\<^sub>2,s\ \ \throw e',s'\" by-(rule eval_evals.FAssThrow1,simp_all) } moreover { fix e' s\<^sub>1 v assume val:"P,E \ \e\<^sub>1',s''\ \ \Val v,s\<^sub>1\" and rest:"P,E \ \e\<^sub>2,s\<^sub>1\ \ \throw e',s'\" from IH[OF val] have "P,E \ \e\<^sub>1,s\ \ \Val v,s\<^sub>1\" . with rest have "P,E \ \e\<^sub>1\F{Cs} := e\<^sub>2,s\ \ \throw e',s'\" by-(rule eval_evals.FAssThrow2,simp_all) } ultimately show ?case using eval by -(erule eval_cases,auto) next case (FAssRed2 E e\<^sub>2 s e\<^sub>2' s'' v F Cs s' e') have eval:"P,E \ \Val v\F{Cs} := e\<^sub>2',s''\ \ \e',s'\" and IH:"\ex sx. P,E \ \e\<^sub>2',s''\ \ \ex,sx\ \ P,E \ \e\<^sub>2,s\ \ \ex,sx\" by fact+ { fix Cs' D S T a fs h\<^sub>2 l\<^sub>2 s\<^sub>1 v' v'' assume val1:"P,E \ \Val v,s''\ \ \ref (a,Cs'),s\<^sub>1\" and val2:"P,E \ \e\<^sub>2',s\<^sub>1\ \ \Val v',(h\<^sub>2, l\<^sub>2)\" and rest:"h\<^sub>2 a = \(D, S)\" "P \ last Cs' has least F:T via Cs" "P \ T casts v' to v''" "(Cs'@\<^sub>pCs,fs) \ S" from val1 have s'':"s\<^sub>1 = s''" by -(erule eval_cases) with val1 have "P,E \ \Val v,s\ \ \ref (a,Cs'),s\" by(fastforce elim:eval_cases intro:eval_finalId) also from IH[OF val2[simplified s'']] have "P,E \ \e\<^sub>2,s\ \ \Val v',(h\<^sub>2, l\<^sub>2)\" . ultimately have "P,E \ \Val v\F{Cs} := e\<^sub>2,s\ \ \Val v'',(h\<^sub>2(a\(D,insert(Cs'@\<^sub>pCs,fs(F \ v''))(S - {(Cs'@\<^sub>pCs,fs)}))),l\<^sub>2)\" using rest by -(rule FAss,simp_all) } moreover { fix s\<^sub>1 v' assume val1:"P,E \ \Val v,s''\ \ \null,s\<^sub>1\" and val2:"P,E \ \e\<^sub>2',s\<^sub>1\ \ \Val v',s'\" from val1 have s'':"s\<^sub>1 = s''" by -(erule eval_cases) with val1 have "P,E \ \Val v,s\ \ \null,s\" by(fastforce elim:eval_cases intro:eval_finalId) also from IH[OF val2[simplified s'']] have "P,E \ \e\<^sub>2,s\ \ \Val v',s'\" . ultimately have "P,E \ \Val v\F{Cs} := e\<^sub>2,s\ \ \THROW NullPointer,s'\" by -(rule FAssNull,simp_all) } moreover { fix r assume val:"P,E \ \Val v,s''\ \ \throw r,s'\" hence s'':"s'' = s'" by -(erule eval_cases,simp) with val have "P,E \ \Val v\F{Cs} := e\<^sub>2,s\ \ \throw r,s'\" by -(rule eval_evals.FAssThrow1,erule eval_cases,simp) } moreover { fix r s\<^sub>1 v' assume val1:"P,E \ \Val v,s''\ \ \Val v',s\<^sub>1\" and val2:"P,E \ \e\<^sub>2',s\<^sub>1\ \ \throw r,s'\" from val1 have s'':"s\<^sub>1 = s''" by -(erule eval_cases) with val1 have "P,E \ \Val v,s\ \ \Val v',s\" by(fastforce elim:eval_cases intro:eval_finalId) also from IH[OF val2[simplified s'']] have "P,E \ \e\<^sub>2,s\ \ \throw r,s'\" . ultimately have "P,E \ \Val v\F{Cs} := e\<^sub>2,s\ \ \throw r,s'\" by -(rule eval_evals.FAssThrow2,simp_all) } ultimately show ?case using eval by -(erule eval_cases,auto) next case (RedFAss h a D S Cs' F T Cs v v' Ds fs E l s' e') have val:"P,E \ \Val v',(h(a \ (D,insert (Ds,fs(F \ v'))(S - {(Ds,fs)}))),l)\ \ \e',s'\" and rest:"h a = \(D, S)\" "P \ last Cs' has least F:T via Cs" "P \ T casts v to v'" "Ds = Cs' @\<^sub>p Cs" "(Ds, fs) \ S" by fact+ from val have "s' = (h(a \ (D,insert (Ds,fs(F \ v')) (S - {(Ds,fs)}))),l)" and "e' = Val v'" by -(erule eval_cases,simp_all)+ with rest show ?case apply simp by(rule FAss,simp_all)(rule eval_finalId,simp)+ next case RedFAssNull thus ?case by (fastforce elim!: eval_cases intro: eval_evals.intros) next case (CallObj E e s e' s' Copt M es s'' e'') thus ?case apply - apply(cases Copt,simp) by(erule eval_cases,auto intro:eval_evals.intros)+ next case (CallParams E es s es' s'' v Copt M s' e') have call:"P,E \ \Call (Val v) Copt M es',s''\ \ \e',s'\" and IH:"\esx sx. P,E \ \es',s''\ [\] \esx,sx\ \ P,E \ \es,s\ [\] \esx,sx\" by fact+ show ?case proof(cases Copt) case None with call have eval:"P,E \ \Val v\M(es'),s''\ \ \e',s'\" by simp from eval show ?thesis proof(rule eval_cases) fix r assume "P,E \ \Val v,s''\ \ \throw r,s'\" "e' = throw r" with None show "P,E \ \Call (Val v) Copt M es,s\ \ \e',s'\" by(fastforce elim:eval_cases) next fix es'' r sx v' vs assume val:"P,E \ \Val v,s''\ \ \Val v',sx\" and evals:"P,E \ \es',sx\ [\] \map Val vs @ throw r # es'',s'\" and e':"e' = throw r" have val':"P,E \ \Val v,s\ \ \Val v,s\" by(rule Val) from val have eq:"v' = v \ s'' = sx" by -(erule eval_cases,simp) with IH evals have "P,E \ \es,s\ [\] \map Val vs @ throw r # es'',s'\" by simp with eq CallParamsThrow[OF val'] e' None show "P,E \ \Call (Val v) Copt M es,s\ \ \e',s'\" by fastforce next fix C Cs Cs' Ds S T T' Ts Ts' a body body' h\<^sub>2 h\<^sub>3 l\<^sub>2 l\<^sub>3 pns pns' s\<^sub>1 vs vs' assume val:"P,E \ \Val v,s''\ \ \ref(a,Cs),s\<^sub>1\" and evals:"P,E \ \es',s\<^sub>1\ [\] \map Val vs,(h\<^sub>2,l\<^sub>2)\" and hp:"h\<^sub>2 a = Some(C, S)" and "method":"P \ last Cs has least M = (Ts',T',pns',body') via Ds" and select:"P \ (C,Cs@\<^sub>pDs) selects M = (Ts,T,pns,body) via Cs'" and length:"length vs = length pns" and casts:"P \ Ts Casts vs to vs'" and body:"P,E(this \ Class (last Cs'), pns [\] Ts) \ \case T' of Class D \ \D\body | _ \ body,(h\<^sub>2,[this \ Ref(a,Cs'),pns [\] vs'])\ \ \e',(h\<^sub>3, l\<^sub>3)\" and s':"s' = (h\<^sub>3, l\<^sub>2)" from val have val':"P,E \ \Val v,s\ \ \ref(a,Cs),s\" and eq:"s'' = s\<^sub>1 \ v = Ref(a,Cs)" by(auto elim:eval_cases intro:Val) from body obtain new_body where body_case:"new_body = (case T' of Class D \ \D\body | _ \ body)" and body':"P,E(this \ Class (last Cs'), pns [\] Ts) \ \new_body,(h\<^sub>2,[this \ Ref(a,Cs'),pns [\] vs'])\ \ \e',(h\<^sub>3, l\<^sub>3)\" by simp from eq IH evals have "P,E \ \es,s\ [\] \map Val vs,(h\<^sub>2,l\<^sub>2)\" by simp with eq Call[OF val' _ _ "method" select length casts _ body_case] hp body' s' None show "P,E \ \Call (Val v) Copt M es,s\ \ \e',s'\" by fastforce next fix s\<^sub>1 vs assume val:"P,E \ \Val v,s''\ \ \null,s\<^sub>1\" and evals:"P,E \ \es',s\<^sub>1\ [\] \map Val vs,s'\" and e':"e' = THROW NullPointer" from val have val':"P,E \ \Val v,s\ \ \null,s\" and eq:"s'' = s\<^sub>1 \ v = Null" by(auto elim:eval_cases intro:Val) from eq IH evals have "P,E \ \es,s\ [\] \map Val vs,s'\" by simp with eq CallNull[OF val'] e' None show "P,E \ \Call (Val v) Copt M es,s\ \ \e',s'\" by fastforce qed next case (Some C) with call have eval:"P,E \ \Val v\(C::)M(es'),s''\ \ \e',s'\" by simp from eval show ?thesis proof(rule eval_cases) fix r assume "P,E \ \Val v,s''\ \ \throw r,s'\" "e' = throw r" with Some show "P,E \ \Call (Val v) Copt M es,s\ \ \e',s'\" by(fastforce elim:eval_cases) next fix es'' r sx v' vs assume val:"P,E \ \Val v,s''\ \ \Val v',sx\" and evals:"P,E \ \es',sx\ [\] \map Val vs @ throw r # es'',s'\" and e':"e' = throw r" have val':"P,E \ \Val v,s\ \ \Val v,s\" by(rule Val) from val have eq:"v' = v \ s'' = sx" by -(erule eval_cases,simp) with IH evals have "P,E \ \es,s\ [\] \map Val vs @ throw r # es'',s'\" by simp with eq CallParamsThrow[OF val'] e' Some show "P,E \ \Call (Val v) Copt M es,s\ \ \e',s'\" by fastforce next fix Cs Cs' Cs'' T Ts a body h\<^sub>2 h\<^sub>3 l\<^sub>2 l\<^sub>3 pns s\<^sub>1 vs vs' assume val:"P,E \ \Val v,s''\ \ \ref (a,Cs),s\<^sub>1\" and evals:"P,E \ \es',s\<^sub>1\ [\] \map Val vs,(h\<^sub>2,l\<^sub>2)\" and path_unique:"P \ Path last Cs to C unique" and path_via:"P \ Path last Cs to C via Cs''" and least:"P \ C has least M = (Ts, T, pns, body) via Cs'" and length:"length vs = length pns" and casts:"P \ Ts Casts vs to vs'" and body:"P,E(this \ Class (last ((Cs @\<^sub>p Cs'') @\<^sub>p Cs')), pns [\] Ts) \ \body,(h\<^sub>2,[this \ Ref(a,(Cs@\<^sub>pCs'')@\<^sub>pCs'),pns [\] vs'])\ \ \e',(h\<^sub>3,l\<^sub>3)\" and s':"s' = (h\<^sub>3,l\<^sub>2)" from val have val':"P,E \ \Val v,s\ \ \ref(a,Cs),s\" and eq:"s'' = s\<^sub>1 \ v = Ref(a,Cs)" by(auto elim:eval_cases intro:Val) from eq IH evals have "P,E \ \es,s\ [\] \map Val vs,(h\<^sub>2,l\<^sub>2)\" by simp with eq StaticCall[OF val' _ path_unique path_via least _ _ casts _ body] length s' Some show "P,E \ \Call (Val v) Copt M es,s\ \ \e',s'\" by fastforce next fix s\<^sub>1 vs assume val:"P,E \ \Val v,s''\ \ \null,s\<^sub>1\" and evals:"P,E \ \es',s\<^sub>1\ [\] \map Val vs,s'\" and e':"e' = THROW NullPointer" from val have val':"P,E \ \Val v,s\ \ \null,s\" and eq:"s'' = s\<^sub>1 \ v = Null" by(auto elim:eval_cases intro:Val) from eq IH evals have "P,E \ \es,s\ [\] \map Val vs,s'\" by simp with eq CallNull[OF val'] e' Some show "P,E \ \Call (Val v) Copt M es,s\ \ \e',s'\" by fastforce qed qed next case (RedCall s a C S Cs M Ts' T' pns' body' Ds Ts T pns body Cs' vs bs new_body E s' e') obtain h l where "s' = (h,l)" by(cases s') auto have "P,E \ \ref(a,Cs),s\ \ \ref(a,Cs),s\" by (rule eval_evals.intros) moreover have finals: "finals(map Val vs)" by simp obtain h\<^sub>2 l\<^sub>2 where s: "s = (h\<^sub>2,l\<^sub>2)" by (cases s) with finals have "P,E \ \map Val vs,s\ [\] \map Val vs,(h\<^sub>2,l\<^sub>2)\" by (iprover intro: eval_finalsId) moreover from s have h\<^sub>2a:"h\<^sub>2 a = Some (C,S)" using RedCall by simp moreover have "method": "P \ last Cs has least M = (Ts',T',pns',body') via Ds" by fact moreover have select:"P \ (C,Cs@\<^sub>pDs) selects M = (Ts,T,pns,body) via Cs'" by fact moreover have blocks:"bs = blocks(this#pns,Class(last Cs')#Ts,Ref(a,Cs')#vs,body)" by fact moreover have body_case:"new_body = (case T' of Class D \ \D\bs | _ \ bs)" by fact moreover have same_len\<^sub>1: "length Ts = length pns" and this_distinct: "this \ set pns" and fv: "fv body \ {this} \ set pns" using select wf by (fastforce dest!:select_method_wf_mdecl simp:wf_mdecl_def)+ have same_len: "length vs = length pns" by fact moreover obtain h\<^sub>3 l\<^sub>3 where s': "s' = (h\<^sub>3,l\<^sub>3)" by (cases s') have eval_blocks:"P,E \ \new_body,s\ \ \e',s'\" by fact hence id: "l\<^sub>3 = l\<^sub>2" using fv s s' same_len\<^sub>1 same_len wf blocks body_case by(cases T')(auto elim!: eval_closed_lcl_unchanged) from same_len\<^sub>1 have same_len':"length(this#pns) = length(Class (last Cs')#Ts)" by simp from same_len\<^sub>1 same_len have same_len\<^sub>2:"length(this#pns) = length(Ref(a,Cs')#vs)" by simp from eval_blocks have eval_blocks':"P,E \ \new_body,(h\<^sub>2,l\<^sub>2)\ \ \e',(h\<^sub>3,l\<^sub>3)\" using s s' by simp have casts_unique:"\vs'. P \ Class (last Cs')#Ts Casts Ref(a,Cs')#vs to vs' \ vs' = Ref(a,Cs')#tl vs'" using wf by -(erule Casts_to.cases,auto elim!:casts_to.cases dest!:mdc_eq_last simp:path_via_def appendPath_def) have "\l'' vs' new_body'. P,E(this\Class(last Cs'), pns[\]Ts) \ \new_body',(h\<^sub>2,l\<^sub>2(this # pns[\]Ref(a,Cs')#vs'))\ \ \e',(h\<^sub>3, l'')\ \ P \ Class(last Cs')#Ts Casts Ref(a,Cs')#vs to Ref(a,Cs')#vs' \ length vs' = length vs \ fv new_body' \ {this} \ set pns \ new_body' = (case T' of Class D \ \D\body | _ \ body)" proof(cases "\C. T' \ Class C") case True with same_len' same_len\<^sub>2 eval_blocks' casts_unique body_case blocks obtain l'' vs' where body:"P,E(this\Class(last Cs'), pns[\]Ts) \ \body,(h\<^sub>2,l\<^sub>2(this # pns[\]Ref(a,Cs')#vs'))\ \ \e',(h\<^sub>3, l'')\" and casts:"P \ Class(last Cs')#Ts Casts Ref(a,Cs')#vs to Ref(a,Cs')#vs'" and lengthvs':"length vs' = length vs" by -(drule_tac vs="Ref(a,Cs')#vs" in blocksEval,assumption,cases T', auto simp:length_Suc_conv,blast) with fv True show ?thesis by(cases T') auto next case False then obtain D where T':"T' = Class D" by auto with same_len' same_len\<^sub>2 eval_blocks' casts_unique body_case blocks obtain l'' vs' where body:"P,E(this\Class(last Cs'), pns[\]Ts) \ \\D\body,(h\<^sub>2,l\<^sub>2(this # pns[\]Ref(a,Cs')#vs'))\ \ \e',(h\<^sub>3, l'')\" and casts:"P \ Class(last Cs')#Ts Casts Ref(a,Cs')#vs to Ref(a,Cs')#vs'" and lengthvs':"length vs' = length vs" by - (drule_tac vs="Ref(a,Cs')#vs" in CastblocksEval, assumption,simp,clarsimp simp:length_Suc_conv,auto) from fv have "fv (\D\body) \ {this} \ set pns" by simp with body casts lengthvs' T' show ?thesis by auto qed then obtain l'' vs' new_body' where body:"P,E(this\Class(last Cs'), pns[\]Ts) \ \new_body',(h\<^sub>2,l\<^sub>2(this # pns[\]Ref(a,Cs')#vs'))\ \ \e',(h\<^sub>3, l'')\" and casts:"P \ Class(last Cs')#Ts Casts Ref(a,Cs')#vs to Ref(a,Cs')#vs'" and lengthvs':"length vs' = length vs" and body_case':"new_body' = (case T' of Class D \ \D\body | _ \ body)" and fv':"fv new_body' \ {this} \ set pns" by auto from same_len\<^sub>2 lengthvs' have same_len\<^sub>3:"length (this # pns) = length (Ref (a, Cs') # vs')" by simp from restrict_map_upds[OF same_len\<^sub>3,of "set(this#pns)" "l\<^sub>2"] have "l\<^sub>2(this # pns[\]Ref(a,Cs')#vs')|`(set(this#pns)) = [this # pns[\]Ref(a,Cs')#vs']" by simp with eval_restrict_lcl[OF wf body fv'] this_distinct same_len\<^sub>1 same_len have "P,E(this\Class(last Cs'), pns[\]Ts) \ \new_body',(h\<^sub>2,[this # pns[\]Ref(a,Cs')#vs'])\ \ \e',(h\<^sub>3, l''|`(set(this#pns)))\" by simp with casts obtain l\<^sub>2' l\<^sub>3' vs' where "P \ Ts Casts vs to vs'" and "P,E(this\Class(last Cs'), pns[\]Ts) \ \new_body',(h\<^sub>2,l\<^sub>2')\ \ \e',(h\<^sub>3,l\<^sub>3')\" and "l\<^sub>2' = [this\Ref(a,Cs'),pns[\]vs']" by(auto elim:Casts_to.cases) ultimately have "P,E \ \(ref(a,Cs))\M(map Val vs),s\ \ \e',(h\<^sub>3,l\<^sub>2)\" using body_case' by -(rule Call,simp_all) with s' id show ?case by simp next case (RedStaticCall Cs C Cs'' M Ts T pns body Cs' Ds vs E a s s' e') have "P,E \ \ref(a,Cs),s\ \ \ref(a,Cs),s\" by (rule eval_evals.intros) moreover have finals: "finals(map Val vs)" by simp obtain h\<^sub>2 l\<^sub>2 where s: "s = (h\<^sub>2,l\<^sub>2)" by (cases s) with finals have "P,E \ \map Val vs,s\ [\] \map Val vs,(h\<^sub>2,l\<^sub>2)\" by (iprover intro: eval_finalsId) moreover have path_unique:"P \ Path last Cs to C unique" by fact moreover have path_via:"P \ Path last Cs to C via Cs''" by fact moreover have least:"P \ C has least M = (Ts, T, pns, body) via Cs'" by fact moreover have same_len\<^sub>1: "length Ts = length pns" and this_distinct: "this \ set pns" and fv: "fv body \ {this} \ set pns" using least wf by (fastforce dest!:has_least_wf_mdecl simp:wf_mdecl_def)+ moreover have same_len:"length vs = length pns" by fact moreover have Ds:"Ds = (Cs @\<^sub>p Cs'') @\<^sub>p Cs'" by fact moreover obtain h\<^sub>3 l\<^sub>3 where s': "s' = (h\<^sub>3,l\<^sub>3)" by (cases s') have eval_blocks:"P,E \ \blocks(this#pns,Class(last Ds)#Ts,Ref(a,Ds)#vs,body),s\ \ \e',s'\" by fact hence id: "l\<^sub>3 = l\<^sub>2" using fv s s' same_len\<^sub>1 same_len wf by(auto elim!: eval_closed_lcl_unchanged) from same_len\<^sub>1 have same_len':"length(this#pns) = length(Class (last Ds)#Ts)" by simp from same_len\<^sub>1 same_len have same_len\<^sub>2:"length(this#pns) = length(Ref(a,Ds)#vs)" by simp from eval_blocks have eval_blocks':"P,E \ \blocks(this#pns,Class(last Ds)#Ts,Ref(a,Ds)#vs,body), (h\<^sub>2,l\<^sub>2)\ \ \e',(h\<^sub>3,l\<^sub>3)\" using s s' by simp have casts_unique:"\vs'. P \ Class (last Ds)#Ts Casts Ref(a,Ds)#vs to vs' \ vs' = Ref(a,Ds)#tl vs'" using wf by -(erule Casts_to.cases,auto elim!:casts_to.cases dest!:mdc_eq_last simp:path_via_def appendPath_def) from same_len' same_len\<^sub>2 eval_blocks' casts_unique obtain l'' vs' where body:"P,E(this\Class(last Ds), pns[\]Ts) \ \body,(h\<^sub>2,l\<^sub>2(this # pns[\]Ref(a,Ds)#vs'))\ \ \e',(h\<^sub>3, l'')\" and casts:"P \ Class(last Ds)#Ts Casts Ref(a,Ds)#vs to Ref(a,Ds)#vs'" and lengthvs':"length vs' = length vs" by -(drule_tac vs="Ref(a,Ds)#vs" in blocksEval,auto simp:length_Suc_conv,blast) from same_len\<^sub>2 lengthvs' have same_len\<^sub>3:"length (this # pns) = length (Ref(a,Ds) # vs')" by simp from restrict_map_upds[OF same_len\<^sub>3,of "set(this#pns)" "l\<^sub>2"] have "l\<^sub>2(this # pns[\]Ref(a,Ds)#vs')|`(set(this#pns)) = [this # pns[\]Ref(a,Ds)#vs']" by simp with eval_restrict_lcl[OF wf body fv] this_distinct same_len\<^sub>1 same_len have "P,E(this\Class(last Ds), pns[\]Ts) \ \body,(h\<^sub>2,[this#pns [\] Ref(a,Ds)#vs'])\ \ \e',(h\<^sub>3, l''|`(set(this#pns)))\" by simp with casts obtain l\<^sub>2' l\<^sub>3' vs' where "P \ Ts Casts vs to vs'" and "P,E(this \ Class(last Ds),pns [\] Ts) \ \body,(h\<^sub>2,l\<^sub>2')\ \ \e',(h\<^sub>3,l\<^sub>3')\" and "l\<^sub>2' = [this \ Ref(a,Ds),pns [\] vs']" by(auto elim:Casts_to.cases) ultimately have "P,E \ \(ref(a,Cs))\(C::)M(map Val vs),s\ \ \e',(h\<^sub>3,l\<^sub>2)\" by -(rule StaticCall,simp_all) with s' id show ?case by simp next case RedCallNull thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros eval_finalsId) next case BlockRedNone thus ?case by (fastforce elim!: eval_cases intro: eval_evals.intros simp add: fun_upd_same fun_upd_idem) next case (BlockRedSome E V T e h l e'' h' l' v s' e') have eval:"P,E \ \{V:T:=Val v; e''},(h', l'(V := l V))\ \ \e',s'\" and red:"P,E(V \ T) \ \e,(h, l(V := None))\ \ \e'',(h', l')\" and notassigned:"\ assigned V e" and l':"l' V = Some v" and IH:"\ex sx. P,E(V \ T) \ \e'',(h', l')\ \ \ex,sx\ \ P,E(V \ T) \ \e,(h, l(V := None))\ \ \ex,sx\" by fact+ from l' have l'upd:"l'(V \ v) = l'" by (rule map_upd_triv) from wf red l' have casts:"P \ T casts v to v" apply - apply(erule_tac V="V" in None_lcl_casts_values) by(simp add:fun_upd_same)+ from eval obtain h'' l'' where "P,E(V \ T) \ \V:=Val v;; e'',(h',l'(V:=None))\ \ \e',(h'',l'')\ \ s' = (h'',l''(V:=l V))" by (fastforce elim:eval_cases simp:fun_upd_same fun_upd_idem) moreover { fix T' h\<^sub>0 l\<^sub>0 v' v'' assume eval':"P,E(V \ T) \ \e'',(h\<^sub>0,l\<^sub>0(V \ v''))\ \ \e',(h'', l'')\" and val:"P,E(V \ T) \ \Val v,(h', l'(V := None))\ \ \Val v',(h\<^sub>0,l\<^sub>0)\" and env:"(E(V \ T)) V = Some T'" and casts':"P \ T' casts v' to v''" from env have TeqT':"T = T'" by (simp add:fun_upd_same) from val have eq:"v = v' \ h' = h\<^sub>0 \ l'(V := None) = l\<^sub>0" by -(erule eval_cases,simp) with casts casts' wf TeqT' have "v = v''" by clarsimp(rule casts_casts_eq) with eq eval' have "P,E(V \ T) \ \e'',(h', l'(V \ v))\ \ \e',(h'', l'')\" by clarsimp } ultimately have "P,E(V \ T) \ \e'',(h',l'(V \ v))\ \ \e',(h'',l'')\" and s':"s' = (h'',l''(V:=l V))" apply auto apply(erule eval_cases) apply(erule eval_cases) apply auto apply(erule eval_cases) apply auto apply(erule eval_cases) apply auto done with l'upd have eval'':"P,E(V \ T) \ \e'',(h',l')\ \ \e',(h'',l'')\" by simp from IH[OF eval''] have "P,E(V \ T) \ \e,(h, l(V := None))\ \ \e',(h'', l'')\" . with s' show ?case by(fastforce intro:Block) next case (InitBlockRed E V T e h l v' e'' h' l' v'' v s' e') have eval:" P,E \ \{V:T:=Val v''; e''},(h', l'(V := l V))\ \ \e',s'\" and red:"P,E(V \ T) \ \e,(h, l(V \ v'))\ \ \e'',(h', l')\" and casts:"P \ T casts v to v'" and l':"l' V = Some v''" and IH:"\ex sx. P,E(V \ T) \ \e'',(h', l')\ \ \ex,sx\ \ P,E(V \ T) \ \e,(h, l(V \ v'))\ \ \ex,sx\" by fact+ from l' have l'upd:"l'(V \ v'') = l'" by (rule map_upd_triv) from wf casts have "P \ T casts v' to v'" by(rule casts_casts) with wf red l' have casts':"P \ T casts v'' to v''" apply - apply(erule_tac V="V" in Some_lcl_casts_values) by(simp add:fun_upd_same)+ from eval obtain h'' l'' where "P,E(V \ T) \ \V:=Val v'';; e'',(h',l'(V:=None))\ \ \e',(h'',l'')\ \ s' = (h'',l''(V:=l V))" by (fastforce elim:eval_cases simp:fun_upd_same fun_upd_idem) moreover { fix T' v''' assume eval':"P,E(V \ T) \ \e'',(h',l'(V \ v'''))\ \ \e',(h'', l'')\" and env:"(E(V \ T)) V = Some T'" and casts'':"P \ T' casts v'' to v'''" from env have "T = T'" by (simp add:fun_upd_same) with casts' casts'' wf have "v'' = v'''" by simp(rule casts_casts_eq) with eval' have "P,E(V \ T) \ \e'',(h', l'(V \ v''))\ \ \e',(h'', l'')\" by simp } ultimately have "P,E(V \ T) \ \e'',(h',l'(V \ v''))\ \ \e',(h'',l'')\" and s':"s' = (h'',l''(V:=l V))" by(auto elim!:eval_cases) with l'upd have eval'':"P,E(V \ T) \ \e'',(h',l')\ \ \e',(h'',l'')\" by simp from IH[OF eval''] have evale:"P,E(V \ T) \ \e,(h, l(V \ v'))\ \ \e',(h'', l'')\" . from casts have "P,E(V \ T) \ \V:=Val v,(h,l(V:=None))\ \ \Val v',(h,l(V \ v'))\" by -(rule_tac l="l(V:=None)" in LAss, auto intro:eval_evals.intros simp:fun_upd_same) with evale s' show ?case by(fastforce intro:Block Seq) next case (RedBlock E V T v s s' e') have "P,E \ \Val v,s\ \ \e',s'\" by fact then obtain s': "s'=s" and e': "e'=Val v" by cases simp obtain h l where s: "s=(h,l)" by (cases s) have "P,E(V \ T) \ \Val v,(h,l(V:=None))\ \ \Val v,(h,l(V:=None))\" by (rule eval_evals.intros) hence "P,E \ \{V:T;Val v},(h,l)\ \ \Val v,(h,(l(V:=None))(V:=l V))\" by (rule eval_evals.Block) thus "P,E \ \{V:T; Val v},s\ \ \e',s'\" using s s' e' by simp next case (RedInitBlock T v v' E V u s s' e') have "P,E \ \Val u,s\ \ \e',s'\" and casts:"P \ T casts v to v'" by fact+ then obtain s': "s' = s" and e': "e'=Val u" by cases simp obtain h l where s: "s=(h,l)" by (cases s) have val:"P,E(V \ T) \ \Val v,(h,l(V:=None))\ \ \Val v,(h,l(V:=None))\" by (rule eval_evals.intros) with casts have "P,E(V \ T) \ \V:=Val v,(h,l(V:=None))\ \ \Val v',(h,l(V \ v'))\" by -(rule_tac l="l(V:=None)" in LAss,auto simp:fun_upd_same) hence "P,E \ \{V:T :=Val v; Val u},(h,l)\ \ \Val u,(h, (l(V\v'))(V:=l V))\" by (fastforce intro!: eval_evals.intros) thus ?case using s s' e' by simp next case SeqRed thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case RedSeq thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case CondRed thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case RedCondT thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case RedCondF thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case RedWhile thus ?case by (auto simp add: unfold_while intro:eval_evals.intros elim:eval_cases) next case ThrowRed thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case RedThrowNull thus ?case by -(auto elim!:eval_cases intro!:eval_evals.ThrowNull eval_finalId) next case ListRed1 thus ?case by (fastforce elim: evals_cases intro: eval_evals.intros) next case ListRed2 thus ?case by (fastforce elim!: evals_cases eval_cases intro: eval_evals.intros eval_finalId) next case StaticCastThrow thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case DynCastThrow thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case BinOpThrow1 thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case BinOpThrow2 thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case LAssThrow thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case FAccThrow thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case FAssThrow1 thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case FAssThrow2 thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case CallThrowObj thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case (CallThrowParams es vs r es' E v Copt M s s' e') have "P,E \ \Val v,s\ \ \Val v,s\" by (rule eval_evals.intros) moreover have es: "es = map Val vs @ Throw r # es'" by fact have eval_e: "P,E \ \Throw r,s\ \ \e',s'\" by fact then obtain s': "s' = s" and e': "e' = Throw r" by cases (auto elim!:eval_cases) with list_eval_Throw [OF eval_e] es have "P,E \ \es,s\ [\] \map Val vs @ Throw r # es',s'\" by simp ultimately have "P,E \ \Call (Val v) Copt M es,s\ \ \Throw r,s'\" by (rule eval_evals.CallParamsThrow) thus ?case using e' by simp next case (BlockThrow E V T r s s' e') have "P,E \ \Throw r, s\ \ \e',s'\" by fact then obtain s': "s' = s" and e': "e' = Throw r" by cases (auto elim!:eval_cases) obtain h l where s: "s=(h,l)" by (cases s) have "P,E(V \ T) \ \Throw r, (h,l(V:=None))\ \ \Throw r, (h,l(V:=None))\" by (simp add:eval_evals.intros eval_finalId) hence "P,E \ \{V:T;Throw r},(h,l)\ \ \Throw r, (h,(l(V:=None))(V:=l V))\" by (rule eval_evals.Block) thus "P,E \ \{V:T; Throw r},s\ \ \e',s'\" using s s' e' by simp next case (InitBlockThrow T v v' E V r s s' e') have "P,E \ \Throw r,s\ \ \e',s'\" and casts:"P \ T casts v to v'" by fact+ then obtain s': "s' = s" and e': "e' = Throw r" by cases (auto elim!:eval_cases) obtain h l where s: "s = (h,l)" by (cases s) have "P,E(V \ T) \ \Val v,(h,l(V:=None))\ \ \Val v,(h,l(V:=None))\" by (rule eval_evals.intros) with casts have "P,E(V \ T) \ \V:=Val v,(h,l(V := None))\ \ \Val v',(h,l(V \ v'))\" by -(rule_tac l="l(V:=None)" in LAss,auto simp:fun_upd_same) hence "P,E \ \{V:T := Val v; Throw r},(h,l)\ \ \Throw r, (h, (l(V\v'))(V:=l V))\" by(fastforce intro:eval_evals.intros) thus "P,E \ \{V:T := Val v; Throw r},s\ \ \e',s'\" using s s' e' by simp next case SeqThrow thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case CondThrow thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) next case ThrowThrow thus ?case by (fastforce elim: eval_cases intro: eval_evals.intros) qed (* ... und wieder anschalten: *) declare split_paired_All [simp] split_paired_Ex [simp] setup \map_theory_claset (fn ctxt => ctxt addSbefore ("split_all_tac", split_all_tac))\ setup \map_theory_simpset (fn ctxt => ctxt addloop ("split_all_tac", split_all_tac))\ text \Its extension to \\*\:\ lemma extend_eval: assumes wf: "wwf_prog P" and reds: "P,E \ \e,s\ \* \e'',s''\" and eval_rest: "P,E \ \e'',s''\ \ \e',s'\" shows "P,E \ \e,s\ \ \e',s'\" using reds eval_rest apply (induct rule: converse_rtrancl_induct2) apply simp apply simp apply (rule extend_1_eval) apply (rule wf) apply assumption+ done lemma extend_evals: assumes wf: "wwf_prog P" and reds: "P,E \ \es,s\ [\]* \es'',s''\" and eval_rest: "P,E \ \es'',s''\ [\] \es',s'\" shows "P,E \ \es,s\ [\] \es',s'\" using reds eval_rest apply (induct rule: converse_rtrancl_induct2) apply simp apply simp apply (rule extend_1_evals) apply (rule wf) apply assumption+ done text \Finally, small step semantics can be simulated by big step semantics: \ theorem assumes wf: "wwf_prog P" shows small_by_big: "\P,E \ \e,s\ \* \e',s'\; final e'\ \ P,E \ \e,s\ \ \e',s'\" and "\P,E \ \es,s\ [\]* \es',s'\; finals es'\ \ P,E \ \es,s\ [\] \es',s'\" proof - note wf moreover assume "P,E \ \e,s\ \* \e',s'\" moreover assume "final e'" then have "P,E \ \e',s'\ \ \e',s'\" by (rule eval_finalId) ultimately show "P,E \ \e,s\\\e',s'\" by (rule extend_eval) next note wf moreover assume "P,E \ \es,s\ [\]* \es',s'\" moreover assume "finals es'" then have "P,E \ \es',s'\ [\] \es',s'\" by (rule eval_finalsId) ultimately show "P,E \ \es,s\ [\] \es',s'\" by (rule extend_evals) qed subsection \Equivalence\ text\And now, the crowning achievement:\ corollary big_iff_small: "wwf_prog P \ P,E \ \e,s\ \ \e',s'\ = (P,E \ \e,s\ \* \e',s'\ \ final e')" by(blast dest: big_by_small eval_final small_by_big) end diff --git a/thys/CoreC++/Exceptions.thy b/thys/CoreC++/Exceptions.thy --- a/thys/CoreC++/Exceptions.thy +++ b/thys/CoreC++/Exceptions.thy @@ -1,101 +1,101 @@ (* Title: CoreC++ Author: Gerwin Klein, Martin Strecker, Daniel Wasserrab Maintainer: Daniel Wasserrab *) section \Exceptions\ theory Exceptions imports Objects begin subsection \Exceptions\ definition NullPointer :: cname where "NullPointer \ ''NullPointer''" definition ClassCast :: cname where "ClassCast \ ''ClassCast''" definition OutOfMemory :: cname where "OutOfMemory \ ''OutOfMemory''" definition sys_xcpts :: "cname set" where "sys_xcpts \ {NullPointer, ClassCast, OutOfMemory}" definition addr_of_sys_xcpt :: "cname \ addr" where "addr_of_sys_xcpt s \ if s = NullPointer then 0 else if s = ClassCast then 1 else if s = OutOfMemory then 2 else undefined" definition start_heap :: "prog \ heap" where - "start_heap P \ Map.empty (addr_of_sys_xcpt NullPointer \ blank P NullPointer) - (addr_of_sys_xcpt ClassCast \ blank P ClassCast) - (addr_of_sys_xcpt OutOfMemory \ blank P OutOfMemory)" + "start_heap P \ Map.empty (addr_of_sys_xcpt NullPointer \ blank P NullPointer, + addr_of_sys_xcpt ClassCast \ blank P ClassCast, + addr_of_sys_xcpt OutOfMemory \ blank P OutOfMemory)" definition preallocated :: "heap \ bool" where "preallocated h \ \C \ sys_xcpts. \S. h (addr_of_sys_xcpt C) = Some (C,S)" subsection "System exceptions" lemma [simp]: "NullPointer \ sys_xcpts \ OutOfMemory \ sys_xcpts \ ClassCast \ sys_xcpts" by(simp add: sys_xcpts_def) lemma sys_xcpts_cases [consumes 1, cases set]: "\ C \ sys_xcpts; P NullPointer; P OutOfMemory; P ClassCast\ \ P C" by (auto simp add: sys_xcpts_def) subsection "@{term preallocated}" lemma preallocated_dom [simp]: "\ preallocated h; C \ sys_xcpts \ \ addr_of_sys_xcpt C \ dom h" by (fastforce simp:preallocated_def dom_def) lemma preallocatedD: "\ preallocated h; C \ sys_xcpts \ \ \S. h (addr_of_sys_xcpt C) = Some (C,S)" by(auto simp add: preallocated_def sys_xcpts_def) lemma preallocatedE [elim?]: "\ preallocated h; C \ sys_xcpts; \S. h (addr_of_sys_xcpt C) = Some(C,S) \ P h C\ \ P h C" by (fast dest: preallocatedD) lemma cname_of_xcp [simp]: "\ preallocated h; C \ sys_xcpts \ \ cname_of h (addr_of_sys_xcpt C) = C" by (auto elim: preallocatedE) lemma preallocated_start: "preallocated (start_heap P)" by (auto simp add: start_heap_def blank_def sys_xcpts_def fun_upd_apply addr_of_sys_xcpt_def preallocated_def) subsection "@{term start_heap}" lemma start_Subobj: "\start_heap P a = Some(C, S); (Cs,fs) \ S\ \ Subobjs P C Cs" by (fastforce elim:init_obj.cases simp:start_heap_def blank_def fun_upd_apply split:if_split_asm) lemma start_SuboSet: "\start_heap P a = Some(C, S); Subobjs P C Cs\ \ \fs. (Cs,fs) \ S" by (fastforce intro:init_obj.intros simp:start_heap_def blank_def split:if_split_asm) lemma start_init_obj: "start_heap P a = Some(C,S) \ S = Collect (init_obj P C)" by (auto simp:start_heap_def blank_def split:if_split_asm) lemma start_subobj: "\start_heap P a = Some(C, S); \fs. (Cs, fs) \ S\ \ Subobjs P C Cs" by (fastforce elim:init_obj.cases simp:start_heap_def blank_def split:if_split_asm) end diff --git a/thys/CoreC++/Execute.thy b/thys/CoreC++/Execute.thy --- a/thys/CoreC++/Execute.thy +++ b/thys/CoreC++/Execute.thy @@ -1,1377 +1,1377 @@ (* Title: CoreC++ Author: Daniel Wasserrab, Stefan Berghofer Maintainer: Daniel Wasserrab *) section \Code generation for Semantics and Type System\ theory Execute imports BigStep WellType "HOL-Library.AList_Mapping" "HOL-Library.Code_Target_Numeral" begin subsection\General redefinitions\ inductive app :: "'a list \ 'a list \ 'a list \ bool" where "app [] ys ys" | "app xs ys zs \ app (x # xs) ys (x # zs)" theorem app_eq1: "\ys zs. zs = xs @ ys \ app xs ys zs" apply (induct xs) apply simp apply (rule app.intros) apply simp apply (iprover intro: app.intros) done theorem app_eq2: "app xs ys zs \ zs = xs @ ys" by (erule app.induct) simp_all theorem app_eq: "app xs ys zs = (zs = xs @ ys)" apply (rule iffI) apply (erule app_eq2) apply (erule app_eq1) done code_pred (modes: i \ i \ i \ bool, i \ i \ o \ bool, i \ o \ i \ bool, o \ i \ i \ bool, o \ o \ i \ bool as reverse_app) app . declare rtranclp_rtrancl_eq[code del] lemmas [code_pred_intro] = rtranclp.rtrancl_refl converse_rtranclp_into_rtranclp code_pred (modes: (i => o => bool) => i => i => bool, (i => o => bool) => i => o => bool) rtranclp by(erule converse_rtranclpE) blast+ definition Set_project :: "('a \ 'b) set => 'a => 'b set" where "Set_project A a = {b. (a, b) \ A}" lemma Set_project_set [code]: "Set_project (set xs) a = set (List.map_filter (\(a', b). if a = a' then Some b else None) xs)" by(auto simp add: Set_project_def map_filter_def intro: rev_image_eqI split: if_split_asm) text\Redefine map Val vs\ inductive map_val :: "expr list \ val list \ bool" where Nil: "map_val [] []" | Cons: "map_val xs ys \ map_val (Val y # xs) (y # ys)" code_pred (modes: i \ i \ bool, i \ o \ bool) map_val . inductive map_val2 :: "expr list \ val list \ expr list \ bool" where Nil: "map_val2 [] [] []" | Cons: "map_val2 xs ys zs \ map_val2 (Val y # xs) (y # ys) zs" | Throw: "map_val2 (throw e # xs) [] (throw e # xs)" code_pred (modes: i \ i \ i \ bool, i \ o \ o \ bool) map_val2 . theorem map_val_conv: "(xs = map Val ys) = map_val xs ys" (*<*) proof - have "\ys. xs = map Val ys \ map_val xs ys" apply (induct xs type:list) apply (case_tac ys) apply simp apply (rule map_val.Nil) apply simp apply (case_tac ys) apply simp apply simp apply (rule map_val.Cons) apply simp done moreover have "map_val xs ys \ xs = map Val ys" by (erule map_val.induct) simp+ ultimately show ?thesis .. qed (*>*) theorem map_val2_conv: "(xs = map Val ys @ throw e # zs) = map_val2 xs ys (throw e # zs)" (*<*) proof - have "\ys. xs = map Val ys @ throw e # zs \ map_val2 xs ys (throw e # zs)" apply (induct xs type:list) apply (case_tac ys) apply simp apply simp apply simp apply (case_tac ys) apply simp apply (rule map_val2.Throw) apply simp apply (rule map_val2.Cons) apply simp done moreover have "map_val2 xs ys (throw e # zs) \ xs = map Val ys @ throw e # zs" by (erule map_val2.induct) simp+ ultimately show ?thesis .. qed (*>*) subsection\Code generation\ lemma subclsRp_code [code_pred_intro]: "\ class P C = \(Bs, rest)\; Predicate_Compile.contains (set Bs) (Repeats D) \ \ subclsRp P C D" by(auto intro: subclsRp.intros simp add: contains_def) code_pred (modes: i \ i \ i \ bool, i \ i \ o \ bool) subclsRp by(erule subclsRp.cases)(fastforce simp add: Predicate_Compile.contains_def) lemma subclsR_code [code_pred_inline]: "P \ C \\<^sub>R D \ subclsRp P C D" by(simp add: subclsR_def) lemma subclsSp_code [code_pred_intro]: "\ class P C = \(Bs, rest)\; Predicate_Compile.contains (set Bs) (Shares D) \ \ subclsSp P C D" by(auto intro: subclsSp.intros simp add: Predicate_Compile.contains_def) code_pred (modes: i \ i \ i \ bool, i \ i \ o \ bool) subclsSp by(erule subclsSp.cases)(fastforce simp add: Predicate_Compile.contains_def) declare SubobjsR_Base [code_pred_intro] lemma SubobjsR_Rep_code [code_pred_intro]: "\subclsRp P C D; Subobjs\<^sub>R P D Cs\ \ Subobjs\<^sub>R P C (C # Cs)" by(simp add: SubobjsR_Rep subclsR_def) code_pred (modes: i \ i \ i \ bool, i \ i \ o \ bool) Subobjs\<^sub>R by(erule Subobjs\<^sub>R.cases)(auto simp add: subclsR_code) lemma subcls1p_code [code_pred_intro]: "\class P C = Some (Bs,rest); Predicate_Compile.contains (baseClasses Bs) D \ \ subcls1p P C D" by(auto intro: subcls1p.intros simp add: Predicate_Compile.contains_def) code_pred (modes: i \ i \ i \ bool, i \ i \ o \ bool) subcls1p by(fastforce elim!: subcls1p.cases simp add: Predicate_Compile.contains_def) declare Subobjs_Rep [code_pred_intro] lemma Subobjs_Sh_code [code_pred_intro]: "\ (subcls1p P)^** C C'; subclsSp P C' D; Subobjs\<^sub>R P D Cs\ \ Subobjs P C Cs" by(rule Subobjs_Sh)(simp_all add: rtrancl_def subcls1_def subclsS_def) code_pred (modes: i \ i \ i \ bool, i \ i \ o \ bool) Subobjs by(erule Subobjs.cases)(auto simp add: rtrancl_def subcls1_def subclsS_def) definition widen_unique :: "prog \ cname \ cname \ path \ bool" where "widen_unique P C D Cs \ (\Cs'. Subobjs P C Cs' \ last Cs' = D \ Cs = Cs')" code_pred [inductify, skip_proof] widen_unique . lemma widen_subcls': "\Subobjs P C Cs'; last Cs' = D; widen_unique P C D Cs' \ \ P \ Class C \ Class D" by(rule widen_subcls,auto simp:path_unique_def widen_unique_def) declare widen_refl [code_pred_intro] widen_subcls' [code_pred_intro widen_subcls] widen_null [code_pred_intro] code_pred (modes: i \ i \ i \ bool) widen by(erule widen.cases)(auto simp add: path_unique_def widen_unique_def) code_pred (modes: i \ i \ i \ i \ bool, i \ i \ i \ o \ bool, i \ i \ o \ i \ bool, i \ i \ o \ o \ bool) leq_path1p . lemma leq_path_unfold: "P,C \ Cs \ Ds \ (leq_path1p P C)^** Cs Ds" by(simp add: leq_path1_def rtrancl_def) code_pred (modes: i => i => i => o => bool, i => i => i => i => bool) [inductify,skip_proof] path_via . lemma path_unique_eq [code_pred_def]: "P \ Path C to D unique \ (\Cs. Subobjs P C Cs \ last Cs = D \ (\Cs'. Subobjs P C Cs' \ last Cs' = D \ Cs = Cs'))" by(auto simp add: path_unique_def) code_pred (modes: i => i => o => bool, i => i => i => bool) [inductify, skip_proof] path_unique . text \Redefine MethodDefs and FieldDecls\ (* FIXME: These predicates should be defined inductively in the first place! *) definition MethodDefs' :: "prog \ cname \ mname \ path \ method \ bool" where "MethodDefs' P C M Cs mthd \ (Cs, mthd) \ MethodDefs P C M" lemma [code_pred_intro]: "Subobjs P C Cs \ class P (last Cs) = \(Bs,fs,ms)\ \ map_of ms M = \mthd\ \ MethodDefs' P C M Cs mthd" by (simp add: MethodDefs_def MethodDefs'_def) code_pred (modes: i \ i \ i \ o \ o \ bool, i \ i \ i \ i \ i \ bool) MethodDefs' by(fastforce simp add: MethodDefs_def MethodDefs'_def) definition FieldDecls' :: "prog \ cname \ vname \ path \ ty \ bool" where "FieldDecls' P C F Cs T \ (Cs, T) \ FieldDecls P C F" lemma [code_pred_intro]: "Subobjs P C Cs \ class P (last Cs) = \(Bs,fs,ms)\ \ map_of fs F = \T\ \ FieldDecls' P C F Cs T" by (simp add: FieldDecls_def FieldDecls'_def) code_pred (modes: i \ i \ i \ o \ o \ bool, i \ i \ i \ i \ i \ bool) FieldDecls' by(fastforce simp add: FieldDecls_def FieldDecls'_def) definition MinimalMethodDefs' :: "prog \ cname \ mname \ path \ method \ bool" where "MinimalMethodDefs' P C M Cs mthd \ (Cs, mthd) \ MinimalMethodDefs P C M" definition MinimalMethodDefs_unique :: "prog \ cname \ mname \ path \ bool" where "MinimalMethodDefs_unique P C M Cs \ (\Cs' mthd. MethodDefs' P C M Cs' mthd \ (leq_path1p P C)^** Cs' Cs \ Cs' = Cs)" code_pred [inductify, skip_proof] MinimalMethodDefs_unique . lemma [code_pred_intro]: "MethodDefs' P C M Cs mthd \ MinimalMethodDefs_unique P C M Cs \ MinimalMethodDefs' P C M Cs mthd" by (fastforce simp add:MinimalMethodDefs_def MinimalMethodDefs'_def MethodDefs'_def MinimalMethodDefs_unique_def leq_path_unfold) code_pred (modes: i \ i \ i \ o \ o \ bool) MinimalMethodDefs' by(fastforce simp add:MinimalMethodDefs_def MinimalMethodDefs'_def MethodDefs'_def MinimalMethodDefs_unique_def leq_path_unfold) definition LeastMethodDef_unique :: "prog \ cname \ mname \ path \ bool" where "LeastMethodDef_unique P C M Cs \ (\Cs' mthd'. MethodDefs' P C M Cs' mthd' \ (leq_path1p P C)^** Cs Cs')" code_pred [inductify, skip_proof] LeastMethodDef_unique . lemma LeastMethodDef_unfold: "P \ C has least M = mthd via Cs \ MethodDefs' P C M Cs mthd \ LeastMethodDef_unique P C M Cs" by(fastforce simp add: LeastMethodDef_def MethodDefs'_def leq_path_unfold LeastMethodDef_unique_def) lemma LeastMethodDef_intro [code_pred_intro]: "\ MethodDefs' P C M Cs mthd; LeastMethodDef_unique P C M Cs \ \ P \ C has least M = mthd via Cs" by(simp add: LeastMethodDef_unfold LeastMethodDef_unique_def) code_pred (modes: i => i => i => o => o => bool) LeastMethodDef by(simp add: LeastMethodDef_unfold LeastMethodDef_unique_def) definition OverriderMethodDefs' :: "prog \ subobj \ mname \ path \ method \ bool" where "OverriderMethodDefs' P R M Cs mthd \ (Cs, mthd) \ OverriderMethodDefs P R M" lemma Overrider1 [code_pred_intro]: "P \ (ldc R) has least M = mthd' via Cs' \ MinimalMethodDefs' P (mdc R) M Cs mthd \ last (snd R) = hd Cs' \ (leq_path1p P (mdc R))^** Cs (snd R @ tl Cs') \ OverriderMethodDefs' P R M Cs mthd" apply(simp add:OverriderMethodDefs_def OverriderMethodDefs'_def MinimalMethodDefs'_def appendPath_def leq_path_unfold) apply(rule_tac x="Cs'" in exI) apply clarsimp apply(cases mthd') apply blast done lemma Overrider2 [code_pred_intro]: "P \ (ldc R) has least M = mthd' via Cs' \ MinimalMethodDefs' P (mdc R) M Cs mthd \ last (snd R) \ hd Cs' \ (leq_path1p P (mdc R))^** Cs Cs' \ OverriderMethodDefs' P R M Cs mthd" by(auto simp add:OverriderMethodDefs_def OverriderMethodDefs'_def MinimalMethodDefs'_def appendPath_def leq_path_unfold simp del: split_paired_Ex) code_pred (modes: i \ i \ i \ o \ o \ bool, i \ i \ i \ i \ o \ bool, i \ i \ i \ o \ i \ bool, i \ i \ i \ i \ i \ bool) OverriderMethodDefs' apply(clarsimp simp add: OverriderMethodDefs'_def MinimalMethodDefs'_def MethodDefs'_def OverriderMethodDefs_def appendPath_def leq_path_unfold) apply(case_tac "last xb = hd Cs'") apply(simp) apply(thin_tac "PROP _") apply(simp add: leq_path1_def) done definition WTDynCast_ex :: "prog \ cname \ cname \ bool" where "WTDynCast_ex P D C \ (\Cs. P \ Path D to C via Cs)" code_pred [inductify, skip_proof] WTDynCast_ex . lemma WTDynCast_new: "\P,E \ e :: Class D; is_class P C; P \ Path D to C unique \ \ WTDynCast_ex P D C\ \ P,E \ Cast C e :: Class C" by(rule WTDynCast)(auto simp add: WTDynCast_ex_def) definition WTStaticCast_sub :: "prog \ cname \ cname \ bool" where "WTStaticCast_sub P C D \ P \ Path D to C unique \ ((subcls1p P)^** C D \ (\Cs. P \ Path C to D via Cs \ Subobjs\<^sub>R P C Cs))" code_pred [inductify, skip_proof] WTStaticCast_sub . lemma WTStaticCast_new: "\P,E \ e :: Class D; is_class P C; WTStaticCast_sub P C D \ \ P,E \ \C\e :: Class C" by (rule WTStaticCast)(auto simp add: WTStaticCast_sub_def subcls1_def rtrancl_def) lemma WTBinOp1: "\ P,E \ e\<^sub>1 :: T; P,E \ e\<^sub>2 :: T\ \ P,E \ e\<^sub>1 \Eq\ e\<^sub>2 :: Boolean" apply (rule WTBinOp) apply assumption+ apply simp done lemma WTBinOp2: "\ P,E \ e\<^sub>1 :: Integer; P,E \ e\<^sub>2 :: Integer \ \ P,E \ e\<^sub>1 \Add\ e\<^sub>2 :: Integer" apply (rule WTBinOp) apply assumption+ apply simp done lemma LeastFieldDecl_unfold [code_pred_def]: "P \ C has least F:T via Cs \ FieldDecls' P C F Cs T \ (\Cs' T'. FieldDecls' P C F Cs' T' \ (leq_path1p P C)^** Cs Cs')" by(auto simp add: LeastFieldDecl_def FieldDecls'_def leq_path_unfold) code_pred [inductify, skip_proof] LeastFieldDecl . lemmas [code_pred_intro] = WT_WTs.WTNew declare WTDynCast_new[code_pred_intro WTDynCast_new] WTStaticCast_new[code_pred_intro WTStaticCast_new] lemmas [code_pred_intro] = WT_WTs.WTVal WT_WTs.WTVar declare WTBinOp1[code_pred_intro WTBinOp1] WTBinOp2 [code_pred_intro WTBinOp2] lemmas [code_pred_intro] = WT_WTs.WTLAss WT_WTs.WTFAcc WT_WTs.WTFAss WT_WTs.WTCall WTStaticCall WT_WTs.WTBlock WT_WTs.WTSeq WT_WTs.WTCond WT_WTs.WTWhile WT_WTs.WTThrow lemmas [code_pred_intro] = WT_WTs.WTNil WT_WTs.WTCons code_pred (modes: WT: i \ i \ i \ i \ bool, i \ i \ i \ o \ bool and WTs: i \ i \ i \ i \ bool, i \ i \ i \ o \ bool) WT proof - case WT from WT.prems show thesis proof(cases (no_simp) rule: WT.cases) case WTDynCast thus thesis by(rule WT.WTDynCast_new[OF refl, unfolded WTDynCast_ex_def, simplified]) next case WTStaticCast thus ?thesis unfolding subcls1_def rtrancl_def mem_Collect_eq prod.case by(rule WT.WTStaticCast_new[OF refl, unfolded WTStaticCast_sub_def]) next case WTBinOp thus ?thesis by(split bop.split_asm)(simp_all, (erule (4) WT.WTBinOp1[OF refl] WT.WTBinOp2[OF refl])+) qed(assumption|erule (2) WT.that[OF refl])+ next case WTs from WTs.prems show thesis by(cases (no_simp) rule: WTs.cases)(assumption|erule (2) WTs.that[OF refl])+ qed lemma casts_to_code [code_pred_intro]: "(case T of Class C \ False | _ \ True) \ P \ T casts v to v" "P \ Class C casts Null to Null" "\Subobjs P (last Cs) Cs'; last Cs' = C; last Cs = hd Cs'; Cs @ tl Cs' = Ds\ \ P \ Class C casts Ref(a,Cs) to Ref(a,Ds)" "\Subobjs P (last Cs) Cs'; last Cs' = C; last Cs \ hd Cs'\ \ P \ Class C casts Ref(a,Cs) to Ref(a,Cs')" by(auto intro: casts_to.intros simp add: path_via_def appendPath_def) code_pred (modes: i \ i \ i \ o \ bool, i \ i \ i \ i \ bool) casts_to apply(erule casts_to.cases) apply(fastforce split: ty.splits) apply simp apply(fastforce simp add: appendPath_def path_via_def split: if_split_asm) done code_pred (modes: i \ i \ i \ o \ bool, i \ i \ i \ i \ bool) Casts_to . lemma card_eq_1_iff_ex1: "x \ A \ card A = 1 \ A = {x}" apply(rule iffI) apply(rule equalityI) apply(rule subsetI) apply(subgoal_tac "card {x, xa} \ card A") apply(auto intro: ccontr)[1] apply(rule card_mono) apply simp_all apply(metis Suc_n_not_n card.infinite) done lemma FinalOverriderMethodDef_unfold [code_pred_def]: "P \ R has overrider M = mthd via Cs \ OverriderMethodDefs' P R M Cs mthd \ (\Cs' mthd'. OverriderMethodDefs' P R M Cs' mthd' \ Cs = Cs' \ mthd = mthd')" by(auto simp add: FinalOverriderMethodDef_def OverriderMethodDefs'_def card_eq_1_iff_ex1 simp del: One_nat_def) code_pred (modes: i => i => i => o => o => bool) [inductify, skip_proof] FinalOverriderMethodDef . code_pred (modes: i => i => i => i => o => o => bool, i \ i \ i \ i \ i \ i \ bool) [inductify] SelectMethodDef . text \Isomorphic subo with mapping instead of a map\ type_synonym subo' = "(path \ (vname, val) mapping)" type_synonym obj' = "cname \ subo' set" lift_definition init_class_fieldmap' :: "prog \ cname \ (vname, val) mapping" is "init_class_fieldmap" . lemma init_class_fieldmap'_code [code]: "init_class_fieldmap' P C = Mapping (map (\(F,T).(F,default_val T)) (fst(snd(the(class P C)))) )" by transfer(simp add: init_class_fieldmap_def) lift_definition init_obj' :: "prog \ cname \ subo' \ bool" is init_obj . lemma init_obj'_intros [code_pred_intro]: "Subobjs P C Cs \ init_obj' P C (Cs, init_class_fieldmap' P (last Cs))" by(transfer)(rule init_obj.intros) code_pred (modes: i \ i \ o \ bool as init_obj_pred) init_obj' by transfer(erule init_obj.cases, blast) lemma init_obj_pred_conv: "set_of_pred (init_obj_pred P C) = Collect (init_obj' P C)" by(auto elim: init_obj_predE intro: init_obj_predI) lift_definition blank' :: "prog \ cname \ obj'" is "blank" . lemma blank'_code [code]: "blank' P C = (C, set_of_pred (init_obj_pred P C))" unfolding init_obj_pred_conv by transfer(simp add: blank_def) type_synonym heap' = "addr \ obj'" abbreviation cname_of' :: "heap' \ addr \ cname" where "\hp. cname_of' hp a == fst (the (hp a))" lift_definition new_Addr' :: "heap' \ addr option" is "new_Addr" . lift_definition start_heap' :: "prog \ heap'" is "start_heap" . lemma start_heap'_code [code]: - "start_heap' P = Map.empty (addr_of_sys_xcpt NullPointer \ blank' P NullPointer) - (addr_of_sys_xcpt ClassCast \ blank' P ClassCast) - (addr_of_sys_xcpt OutOfMemory \ blank' P OutOfMemory)" + "start_heap' P = Map.empty (addr_of_sys_xcpt NullPointer \ blank' P NullPointer, + addr_of_sys_xcpt ClassCast \ blank' P ClassCast, + addr_of_sys_xcpt OutOfMemory \ blank' P OutOfMemory)" by transfer(simp add: start_heap_def) type_synonym state' = "heap' \ locals" lift_definition hp' :: "state' \ heap'" is hp . lemma hp'_code [code]: "hp' = fst" by transfer simp lift_definition lcl' :: "state' \ locals" is lcl . lemma lcl_code [code]: "lcl' = snd" by transfer simp lift_definition eval' :: "prog \ env \ expr \ state' \ expr \ state' \ bool" ("_,_ \ ((1\_,/_\) \''/ (1\_,/_\))" [51,0,0,0,0] 81) is eval . lift_definition evals' :: "prog \ env \ expr list \ state' \ expr list \ state' \ bool" ("_,_ \ ((1\_,/_\) [\'']/ (1\_,/_\))" [51,0,0,0,0] 81) is evals . lemma New': "\ new_Addr' h = Some a; h' = h(a\(blank' P C)) \ \ P,E \ \new C,(h,l)\ \' \ref (a,[C]),(h',l)\" by transfer(unfold blank_def, rule New) lemma NewFail': "new_Addr' h = None \ P,E \ \new C, (h,l)\ \' \THROW OutOfMemory,(h,l)\" by transfer(rule NewFail) lemma StaticUpCast': "\ P,E \ \e,s\<^sub>0\ \' \ref (a,Cs),s\<^sub>1\; P \ Path last Cs to C via Cs'; Ds = Cs@\<^sub>pCs' \ \ P,E \ \\C\e,s\<^sub>0\ \' \ref (a,Ds),s\<^sub>1\" by transfer(rule StaticUpCast) lemma StaticDownCast'_new: (* requires reverse append *) "\P,E \ \e,s\<^sub>0\ \' \ref (a,Ds),s\<^sub>1\; app Cs [C] Ds'; app Ds' Cs' Ds\ \ P,E \ \\C\e,s\<^sub>0\ \' \ref(a,Cs@[C]),s\<^sub>1\" apply transfer apply (rule StaticDownCast) apply (simp add: app_eq) done lemma StaticCastNull': "P,E \ \e,s\<^sub>0\ \' \null,s\<^sub>1\ \ P,E \ \\C\e,s\<^sub>0\ \' \null,s\<^sub>1\" by transfer(rule StaticCastNull) lemma StaticCastFail'_new: (* manual unfolding of subcls *) "\ P,E \ \e,s\<^sub>0\\' \ref (a,Cs),s\<^sub>1\; \ (subcls1p P)^** (last Cs) C; C \ set Cs\ \ P,E \ \\C\e,s\<^sub>0\ \' \THROW ClassCast,s\<^sub>1\" apply transfer by (fastforce intro:StaticCastFail simp add: rtrancl_def subcls1_def) lemma StaticCastThrow': "P,E \ \e,s\<^sub>0\ \' \throw e',s\<^sub>1\ \ P,E \ \\C\e,s\<^sub>0\ \' \throw e',s\<^sub>1\" by transfer(rule StaticCastThrow) lemma StaticUpDynCast': "\P,E \ \e,s\<^sub>0\ \' \ref(a,Cs),s\<^sub>1\; P \ Path last Cs to C unique; P \ Path last Cs to C via Cs'; Ds = Cs@\<^sub>pCs' \ \ P,E \ \Cast C e,s\<^sub>0\ \' \ref(a,Ds),s\<^sub>1\" by transfer(rule StaticUpDynCast) lemma StaticDownDynCast'_new: (* requires reverse append *) "\P,E \ \e,s\<^sub>0\ \' \ref (a,Ds),s\<^sub>1\; app Cs [C] Ds'; app Ds' Cs' Ds\ \ P,E \ \Cast C e,s\<^sub>0\ \' \ref(a,Cs@[C]),s\<^sub>1\" apply transfer apply (rule StaticDownDynCast) apply (simp add: app_eq) done lemma DynCast': "\ P,E \ \e,s\<^sub>0\ \' \ref (a,Cs),(h,l)\; h a = Some(D,S); P \ Path D to C via Cs'; P \ Path D to C unique \ \ P,E \ \Cast C e,s\<^sub>0\ \' \ref (a,Cs'),(h,l)\" by transfer(rule DynCast) lemma DynCastNull': "P,E \ \e,s\<^sub>0\ \' \null,s\<^sub>1\ \ P,E \ \Cast C e,s\<^sub>0\ \' \null,s\<^sub>1\" by transfer(rule DynCastNull) lemma DynCastFail': "\ P,E \ \e,s\<^sub>0\\' \ref (a,Cs),(h,l)\; h a = Some(D,S); \ P \ Path D to C unique; \ P \ Path last Cs to C unique; C \ set Cs \ \ P,E \ \Cast C e,s\<^sub>0\ \' \null,(h,l)\" by transfer(rule DynCastFail) lemma DynCastThrow': "P,E \ \e,s\<^sub>0\ \' \throw e',s\<^sub>1\ \ P,E \ \Cast C e,s\<^sub>0\ \' \throw e',s\<^sub>1\" by transfer(rule DynCastThrow) lemma Val': "P,E \ \Val v,s\ \' \Val v,s\" by transfer(rule Val) lemma BinOp': "\ P,E \ \e\<^sub>1,s\<^sub>0\ \' \Val v\<^sub>1,s\<^sub>1\; P,E \ \e\<^sub>2,s\<^sub>1\ \' \Val v\<^sub>2,s\<^sub>2\; binop(bop,v\<^sub>1,v\<^sub>2) = Some v \ \ P,E \ \e\<^sub>1 \bop\ e\<^sub>2,s\<^sub>0\\'\Val v,s\<^sub>2\" by transfer(rule BinOp) lemma BinOpThrow1': "P,E \ \e\<^sub>1,s\<^sub>0\ \' \throw e,s\<^sub>1\ \ P,E \ \e\<^sub>1 \bop\ e\<^sub>2, s\<^sub>0\ \' \throw e,s\<^sub>1\" by transfer(rule BinOpThrow1) lemma BinOpThrow2': "\ P,E \ \e\<^sub>1,s\<^sub>0\ \' \Val v\<^sub>1,s\<^sub>1\; P,E \ \e\<^sub>2,s\<^sub>1\ \' \throw e,s\<^sub>2\ \ \ P,E \ \e\<^sub>1 \bop\ e\<^sub>2,s\<^sub>0\ \' \throw e,s\<^sub>2\" by transfer(rule BinOpThrow2) lemma Var': "l V = Some v \ P,E \ \Var V,(h,l)\ \' \Val v,(h,l)\" by transfer(rule Var) lemma LAss': "\ P,E \ \e,s\<^sub>0\ \' \Val v,(h,l)\; E V = Some T; P \ T casts v to v'; l' = l(V\v') \ \ P,E \ \V:=e,s\<^sub>0\ \' \Val v',(h,l')\" by (transfer) (erule (3) LAss) lemma LAssThrow': "P,E \ \e,s\<^sub>0\ \' \throw e',s\<^sub>1\ \ P,E \ \V:=e,s\<^sub>0\ \' \throw e',s\<^sub>1\" by transfer(rule LAssThrow) lemma FAcc'_new: (* iteration over set *) "\ P,E \ \e,s\<^sub>0\ \' \ref (a,Cs'),(h,l)\; h a = Some(D,S); Ds = Cs'@\<^sub>pCs; Predicate_Compile.contains (Set_project S Ds) fs; Mapping.lookup fs F = Some v \ \ P,E \ \e\F{Cs},s\<^sub>0\ \' \Val v,(h,l)\" unfolding Set_project_def mem_Collect_eq Predicate_Compile.contains_def by transfer(rule FAcc) lemma FAccNull': "P,E \ \e,s\<^sub>0\ \' \null,s\<^sub>1\ \ P,E \ \e\F{Cs},s\<^sub>0\ \' \THROW NullPointer,s\<^sub>1\" by transfer(rule FAccNull) lemma FAccThrow': "P,E \ \e,s\<^sub>0\ \' \throw e',s\<^sub>1\ \ P,E \ \e\F{Cs},s\<^sub>0\ \' \throw e',s\<^sub>1\" by transfer(rule FAccThrow) lemma FAss'_new: (* iteration over set *) "\ P,E \ \e\<^sub>1,s\<^sub>0\ \' \ref (a,Cs'),s\<^sub>1\; P,E \ \e\<^sub>2,s\<^sub>1\ \' \Val v,(h\<^sub>2,l\<^sub>2)\; h\<^sub>2 a = Some(D,S); P \ (last Cs') has least F:T via Cs; P \ T casts v to v'; Ds = Cs'@\<^sub>pCs; Predicate_Compile.contains (Set_project S Ds) fs; fs' = Mapping.update F v' fs; S' = S - {(Ds,fs)} \ {(Ds,fs')}; h\<^sub>2' = h\<^sub>2(a\(D,S'))\ \ P,E \ \e\<^sub>1\F{Cs}:=e\<^sub>2,s\<^sub>0\ \' \Val v',(h\<^sub>2',l\<^sub>2)\" unfolding Predicate_Compile.contains_def Set_project_def mem_Collect_eq by transfer(rule FAss) lemma FAssNull': "\ P,E \ \e\<^sub>1,s\<^sub>0\ \' \null,s\<^sub>1\; P,E \ \e\<^sub>2,s\<^sub>1\ \' \Val v,s\<^sub>2\ \ \ P,E \ \e\<^sub>1\F{Cs}:=e\<^sub>2,s\<^sub>0\ \' \THROW NullPointer,s\<^sub>2\" by transfer(rule FAssNull) lemma FAssThrow1': "P,E \ \e\<^sub>1,s\<^sub>0\ \' \throw e',s\<^sub>1\ \ P,E \ \e\<^sub>1\F{Cs}:=e\<^sub>2,s\<^sub>0\ \' \throw e',s\<^sub>1\" by transfer(rule FAssThrow1) lemma FAssThrow2': "\ P,E \ \e\<^sub>1,s\<^sub>0\ \' \Val v,s\<^sub>1\; P,E \ \e\<^sub>2,s\<^sub>1\ \' \throw e',s\<^sub>2\ \ \ P,E \ \e\<^sub>1\F{Cs}:=e\<^sub>2,s\<^sub>0\ \' \throw e',s\<^sub>2\" by transfer(rule FAssThrow2) lemma CallObjThrow': "P,E \ \e,s\<^sub>0\ \' \throw e',s\<^sub>1\ \ P,E \ \Call e Copt M es,s\<^sub>0\ \' \throw e',s\<^sub>1\" by transfer(rule CallObjThrow) lemma CallParamsThrow'_new: (* requires inverse map Val and append *) "\ P,E \ \e,s0\ \' \Val v,s1\; P,E \ \es,s1\ [\'] \evs,s2\; map_val2 evs vs (throw ex # es') \ \ P,E \ \Call e Copt M es,s0\ \' \throw ex,s2\" apply transfer apply(rule eval_evals.CallParamsThrow, assumption+) apply(simp add: map_val2_conv[symmetric]) done lemma Call'_new: (* requires inverse map Val *) "\ P,E \ \e,s\<^sub>0\ \' \ref (a,Cs),s\<^sub>1\; P,E \ \ps,s\<^sub>1\ [\'] \evs,(h\<^sub>2,l\<^sub>2)\; map_val evs vs; h\<^sub>2 a = Some(C,S); P \ last Cs has least M = (Ts',T',pns',body') via Ds; P \ (C,Cs@\<^sub>pDs) selects M = (Ts,T,pns,body) via Cs'; length vs = length pns; P \ Ts Casts vs to vs'; l\<^sub>2' = [this\Ref (a,Cs'), pns[\]vs']; new_body = (case T' of Class D \ \D\body | _ \ body); P,E(this\Class(last Cs'), pns[\]Ts) \ \new_body,(h\<^sub>2,l\<^sub>2')\ \' \e',(h\<^sub>3,l\<^sub>3)\ \ \ P,E \ \e\M(ps),s\<^sub>0\ \' \e',(h\<^sub>3,l\<^sub>2)\" apply transfer apply(rule Call) apply assumption+ apply(simp add: map_val_conv[symmetric]) apply assumption+ done lemma StaticCall'_new: (* requires inverse map Val *) "\ P,E \ \e,s\<^sub>0\ \' \ref (a,Cs),s\<^sub>1\; P,E \ \ps,s\<^sub>1\ [\'] \evs,(h\<^sub>2,l\<^sub>2)\; map_val evs vs; P \ Path (last Cs) to C unique; P \ Path (last Cs) to C via Cs''; P \ C has least M = (Ts,T,pns,body) via Cs'; Ds = (Cs@\<^sub>pCs'')@\<^sub>pCs'; length vs = length pns; P \ Ts Casts vs to vs'; l\<^sub>2' = [this\Ref (a,Ds), pns[\]vs']; P,E(this\Class(last Ds), pns[\]Ts) \ \body,(h\<^sub>2,l\<^sub>2')\ \' \e',(h\<^sub>3,l\<^sub>3)\ \ \ P,E \ \e\(C::)M(ps),s\<^sub>0\ \' \e',(h\<^sub>3,l\<^sub>2)\" apply transfer apply(rule StaticCall) apply(assumption)+ apply(simp add: map_val_conv[symmetric]) apply assumption+ done lemma CallNull'_new: (* requires inverse map Val *) "\ P,E \ \e,s\<^sub>0\ \' \null,s\<^sub>1\; P,E \ \es,s\<^sub>1\ [\'] \evs,s\<^sub>2\; map_val evs vs \ \ P,E \ \Call e Copt M es,s\<^sub>0\ \' \THROW NullPointer,s\<^sub>2\" apply transfer apply(rule CallNull, assumption+) apply(simp add: map_val_conv[symmetric]) done lemma Block': "\P,E(V \ T) \ \e\<^sub>0,(h\<^sub>0,l\<^sub>0(V:=None))\ \' \e\<^sub>1,(h\<^sub>1,l\<^sub>1)\ \ \ P,E \ \{V:T; e\<^sub>0},(h\<^sub>0,l\<^sub>0)\ \' \e\<^sub>1,(h\<^sub>1,l\<^sub>1(V:=l\<^sub>0 V))\" by transfer(rule Block) lemma Seq': "\ P,E \ \e\<^sub>0,s\<^sub>0\ \' \Val v,s\<^sub>1\; P,E \ \e\<^sub>1,s\<^sub>1\ \' \e\<^sub>2,s\<^sub>2\ \ \ P,E \ \e\<^sub>0;;e\<^sub>1,s\<^sub>0\ \' \e\<^sub>2,s\<^sub>2\" by transfer(rule Seq) lemma SeqThrow': "P,E \ \e\<^sub>0,s\<^sub>0\ \' \throw e,s\<^sub>1\ \ P,E \ \e\<^sub>0;;e\<^sub>1,s\<^sub>0\\'\throw e,s\<^sub>1\" by transfer(rule SeqThrow) lemma CondT': "\ P,E \ \e,s\<^sub>0\ \' \true,s\<^sub>1\; P,E \ \e\<^sub>1,s\<^sub>1\ \' \e',s\<^sub>2\ \ \ P,E \ \if (e) e\<^sub>1 else e\<^sub>2,s\<^sub>0\ \' \e',s\<^sub>2\" by transfer(rule CondT) lemma CondF': "\ P,E \ \e,s\<^sub>0\ \' \false,s\<^sub>1\; P,E \ \e\<^sub>2,s\<^sub>1\ \' \e',s\<^sub>2\ \ \ P,E \ \if (e) e\<^sub>1 else e\<^sub>2,s\<^sub>0\ \' \e',s\<^sub>2\" by transfer(rule CondF) lemma CondThrow': "P,E \ \e,s\<^sub>0\ \' \throw e',s\<^sub>1\ \ P,E \ \if (e) e\<^sub>1 else e\<^sub>2, s\<^sub>0\ \' \throw e',s\<^sub>1\" by transfer(rule CondThrow) lemma WhileF': "P,E \ \e,s\<^sub>0\ \' \false,s\<^sub>1\ \ P,E \ \while (e) c,s\<^sub>0\ \' \unit,s\<^sub>1\" by transfer(rule WhileF) lemma WhileT': "\ P,E \ \e,s\<^sub>0\ \' \true,s\<^sub>1\; P,E \ \c,s\<^sub>1\ \' \Val v\<^sub>1,s\<^sub>2\; P,E \ \while (e) c,s\<^sub>2\ \' \e\<^sub>3,s\<^sub>3\ \ \ P,E \ \while (e) c,s\<^sub>0\ \' \e\<^sub>3,s\<^sub>3\" by transfer(rule WhileT) lemma WhileCondThrow': "P,E \ \e,s\<^sub>0\ \' \ throw e',s\<^sub>1\ \ P,E \ \while (e) c,s\<^sub>0\ \' \throw e',s\<^sub>1\" by transfer(rule WhileCondThrow) lemma WhileBodyThrow': "\ P,E \ \e,s\<^sub>0\ \' \true,s\<^sub>1\; P,E \ \c,s\<^sub>1\ \' \throw e',s\<^sub>2\\ \ P,E \ \while (e) c,s\<^sub>0\ \' \throw e',s\<^sub>2\" by transfer(rule WhileBodyThrow) lemma Throw': "P,E \ \e,s\<^sub>0\ \' \ref r,s\<^sub>1\ \ P,E \ \throw e,s\<^sub>0\ \' \Throw r,s\<^sub>1\" by transfer(rule eval_evals.Throw) lemma ThrowNull': "P,E \ \e,s\<^sub>0\ \' \null,s\<^sub>1\ \ P,E \ \throw e,s\<^sub>0\ \' \THROW NullPointer,s\<^sub>1\" by transfer(rule ThrowNull) lemma ThrowThrow': "P,E \ \e,s\<^sub>0\ \' \throw e',s\<^sub>1\ \ P,E \ \throw e,s\<^sub>0\ \' \throw e',s\<^sub>1\" by transfer(rule ThrowThrow) lemma Nil': "P,E \ \[],s\ [\'] \[],s\" by transfer(rule eval_evals.Nil) lemma Cons': "\ P,E \ \e,s\<^sub>0\ \' \Val v,s\<^sub>1\; P,E \ \es,s\<^sub>1\ [\'] \es',s\<^sub>2\ \ \ P,E \ \e#es,s\<^sub>0\ [\'] \Val v # es',s\<^sub>2\" by transfer(rule eval_evals.Cons) lemma ConsThrow': "P,E \ \e, s\<^sub>0\ \' \throw e', s\<^sub>1\ \ P,E \ \e#es, s\<^sub>0\ [\'] \throw e' # es, s\<^sub>1\" by transfer(rule ConsThrow) text \Axiomatic heap address model refinement\ partial_function (option) lowest :: "(nat \ bool) \ nat \ nat option" where [code]: "lowest P n = (if P n then Some n else lowest P (Suc n))" axiomatization where new_Addr'_code [code]: "new_Addr' h = lowest (Option.is_none \ h) 0" \ \admissible: a tightening of the specification of @{const new_Addr'}\ lemma eval'_cases [consumes 1, case_names New NewFail StaticUpCast StaticDownCast StaticCastNull StaticCastFail StaticCastThrow StaticUpDynCast StaticDownDynCast DynCast DynCastNull DynCastFail DynCastThrow Val BinOp BinOpThrow1 BinOpThrow2 Var LAss LAssThrow FAcc FAccNull FAccThrow FAss FAssNull FAssThrow1 FAssThrow2 CallObjThrow CallParamsThrow Call StaticCall CallNull Block Seq SeqThrow CondT CondF CondThrow WhileF WhileT WhileCondThrow WhileBodyThrow Throw ThrowNull ThrowThrow]: assumes "P,x \ \y,z\ \' \u,v\" and "\h a h' C E l. x = E \ y = new C \ z = (h, l) \ u = ref (a, [C]) \ v = (h', l) \ new_Addr' h = \a\ \ h' = h(a \ blank' P C) \ thesis" and "\h E C l. x = E \ y = new C \ z = (h, l) \ u = Throw (addr_of_sys_xcpt OutOfMemory, [OutOfMemory]) \ v = (h, l) \ new_Addr' h = None \ thesis" and "\E e s\<^sub>0 a Cs s\<^sub>1 C Cs' Ds. x = E \ y = \C\e \ z = s\<^sub>0 \ u = ref (a, Ds) \ v = s\<^sub>1 \ P,E \ \e,s\<^sub>0\ \' \ref (a, Cs),s\<^sub>1\ \ P \ Path last Cs to C via Cs' \ Ds = Cs @\<^sub>p Cs' \ thesis" and "\E e s\<^sub>0 a Cs C Cs' s\<^sub>1. x = E \ y = \C\e \ z = s\<^sub>0 \ u = ref (a, Cs @ [C]) \ v = s\<^sub>1 \ P,E \ \e,s\<^sub>0\ \' \ref (a, Cs @ [C] @ Cs'),s\<^sub>1\ \ thesis" and "\E e s\<^sub>0 s\<^sub>1 C. x = E \ y = \C\e \ z = s\<^sub>0 \ u = null \ v = s\<^sub>1 \ P,E \ \e,s\<^sub>0\ \' \null,s\<^sub>1\ \ thesis" and "\E e s\<^sub>0 a Cs s\<^sub>1 C. x = E \ y = \C\e \ z = s\<^sub>0 \ u = Throw (addr_of_sys_xcpt ClassCast, [ClassCast]) \ v = s\<^sub>1 \ P,E \ \e,s\<^sub>0\ \' \ref (a, Cs),s\<^sub>1\ \ (last Cs, C) \ (subcls1 P)\<^sup>* \ C \ set Cs \ thesis" and "\E e s\<^sub>0 e' s\<^sub>1 C. x = E \ y = \C\e \ z = s\<^sub>0 \ u = throw e' \ v = s\<^sub>1 \ P,E \ \e,s\<^sub>0\ \' \throw e',s\<^sub>1\ \ thesis" and "\E e s\<^sub>0 a Cs s\<^sub>1 C Cs' Ds. x = E \ y = Cast C e \ z = s\<^sub>0 \ u = ref (a, Ds) \ v = s\<^sub>1 \ P,E \ \e,s\<^sub>0\ \' \ref (a, Cs),s\<^sub>1\ \ P \ Path last Cs to C unique \ P \ Path last Cs to C via Cs' \ Ds = Cs @\<^sub>p Cs' \ thesis" and "\E e s\<^sub>0 a Cs C Cs' s\<^sub>1. x = E \ y = Cast C e \ z = s\<^sub>0 \ u = ref (a, Cs @ [C]) \ v = s\<^sub>1 \ P,E \ \e,s\<^sub>0\ \' \ref (a, Cs @ [C] @ Cs'),s\<^sub>1\ \ thesis" and "\E e s\<^sub>0 a Cs h l D S C Cs'. x = E \ y = Cast C e \ z = s\<^sub>0 \ u = ref (a, Cs') \ v = (h, l) \ P,E \ \e,s\<^sub>0\ \' \ref (a, Cs),(h, l)\ \ h a = \(D, S)\ \ P \ Path D to C via Cs' \ P \ Path D to C unique \ thesis" and "\E e s\<^sub>0 s\<^sub>1 C. x = E \ y = Cast C e \ z = s\<^sub>0 \ u = null \ v = s\<^sub>1 \ P,E \ \e,s\<^sub>0\ \' \null,s\<^sub>1\ \ thesis" and "\E e s\<^sub>0 a Cs h l D S C. x = E \ y = Cast C e \ z = s\<^sub>0 \ u = null \ v = (h, l) \ P,E \ \e,s\<^sub>0\ \' \ref (a, Cs),(h, l)\ \ h a = \(D, S)\ \ \ P \ Path D to C unique \ \ P \ Path last Cs to C unique \ C \ set Cs \ thesis" and "\E e s\<^sub>0 e' s\<^sub>1 C. x = E \ y = Cast C e \ z = s\<^sub>0 \ u = throw e' \ v = s\<^sub>1 \ P,E \ \e,s\<^sub>0\ \' \throw e',s\<^sub>1\ \ thesis" and "\E va s. x = E \ y = Val va \ z = s \ u = Val va \ v = s \ thesis" and "\E e\<^sub>1 s\<^sub>0 v\<^sub>1 s\<^sub>1 e\<^sub>2 v\<^sub>2 s\<^sub>2 bop va. x = E \ y = e\<^sub>1 \bop\ e\<^sub>2 \ z = s\<^sub>0 \ u = Val va \ v = s\<^sub>2 \ P,E \ \e\<^sub>1,s\<^sub>0\ \' \Val v\<^sub>1,s\<^sub>1\ \ P,E \ \e\<^sub>2,s\<^sub>1\ \' \Val v\<^sub>2,s\<^sub>2\ \ binop (bop, v\<^sub>1, v\<^sub>2) = \va\ \ thesis" and "\E e\<^sub>1 s\<^sub>0 e s\<^sub>1 bop e\<^sub>2. x = E \ y = e\<^sub>1 \bop\ e\<^sub>2 \ z = s\<^sub>0 \ u = throw e \ v = s\<^sub>1 \ P,E \ \e\<^sub>1,s\<^sub>0\ \' \throw e,s\<^sub>1\ \ thesis" and "\E e\<^sub>1 s\<^sub>0 v\<^sub>1 s\<^sub>1 e\<^sub>2 e s\<^sub>2 bop. x = E \ y = e\<^sub>1 \bop\ e\<^sub>2 \ z = s\<^sub>0 \ u = throw e \ v = s\<^sub>2 \ P,E \ \e\<^sub>1,s\<^sub>0\ \' \Val v\<^sub>1,s\<^sub>1\ \ P,E \ \e\<^sub>2,s\<^sub>1\ \' \throw e,s\<^sub>2\ \ thesis" and "\l V va E h. x = E \ y = Var V \ z = (h, l) \ u = Val va \ v = (h, l) \ l V = \va\ \ thesis" and "\E e s\<^sub>0 va h l V T v' l'. x = E \ y = V:=e \ z = s\<^sub>0 \ u = Val v' \ v = (h, l') \ P,E \ \e,s\<^sub>0\ \' \Val va,(h, l)\ \ E V = \T\ \ P \ T casts va to v' \ l' = l(V \ v') \ thesis" and "\E e s\<^sub>0 e' s\<^sub>1 V. x = E \ y = V:=e \ z = s\<^sub>0 \ u = throw e' \ v = s\<^sub>1 \ P,E \ \e,s\<^sub>0\ \' \throw e',s\<^sub>1\ \ thesis" and "\E e s\<^sub>0 a Cs' h l D S Ds Cs fs F va. x = E \ y = e\F{Cs} \ z = s\<^sub>0 \ u = Val va \ v = (h, l) \ P,E \ \e,s\<^sub>0\ \' \ref (a, Cs'),(h, l)\ \ h a = \(D, S)\ \ Ds = Cs' @\<^sub>p Cs \ (Ds, fs) \ S \ Mapping.lookup fs F = \va\ \ thesis" and "\E e s\<^sub>0 s\<^sub>1 F Cs. x = E \ y = e\F{Cs} \ z = s\<^sub>0 \ u = Throw (addr_of_sys_xcpt NullPointer, [NullPointer]) \ v = s\<^sub>1 \ P,E \ \e,s\<^sub>0\ \' \null,s\<^sub>1\ \ thesis" and "\E e s\<^sub>0 e' s\<^sub>1 F Cs. x = E \ y = e\F{Cs} \ z = s\<^sub>0 \ u = throw e' \ v = s\<^sub>1 \ P,E \ \e,s\<^sub>0\ \' \throw e',s\<^sub>1\ \ thesis" and "\E e\<^sub>1 s\<^sub>0 a Cs' s\<^sub>1 e\<^sub>2 va h\<^sub>2 l\<^sub>2 D S F T Cs v' Ds fs fs' S' h\<^sub>2'. x = E \ y = e\<^sub>1\F{Cs} := e\<^sub>2 \ z = s\<^sub>0 \ u = Val v' \ v = (h\<^sub>2', l\<^sub>2) \ P,E \ \e\<^sub>1,s\<^sub>0\ \' \ref (a, Cs'),s\<^sub>1\ \ P,E \ \e\<^sub>2,s\<^sub>1\ \' \Val va,(h\<^sub>2, l\<^sub>2)\ \ h\<^sub>2 a = \(D, S)\ \ P \ last Cs' has least F:T via Cs \ P \ T casts va to v' \ Ds = Cs' @\<^sub>p Cs \ (Ds, fs) \ S \ fs' = Mapping.update F v' fs \ S' = S - {(Ds, fs)} \ {(Ds, fs')} \ h\<^sub>2' = h\<^sub>2(a \ (D, S')) \ thesis" and "\E e\<^sub>1 s\<^sub>0 s\<^sub>1 e\<^sub>2 va s\<^sub>2 F Cs. x = E \ y = e\<^sub>1\F{Cs} := e\<^sub>2 \ z = s\<^sub>0 \ u = Throw (addr_of_sys_xcpt NullPointer, [NullPointer]) \ v = s\<^sub>2 \ P,E \ \e\<^sub>1,s\<^sub>0\ \' \null,s\<^sub>1\ \ P,E \ \e\<^sub>2,s\<^sub>1\ \' \Val va,s\<^sub>2\ \ thesis" and "\E e\<^sub>1 s\<^sub>0 e' s\<^sub>1 F Cs e\<^sub>2. x = E \ y = e\<^sub>1\F{Cs} := e\<^sub>2 \ z = s\<^sub>0 \ u = throw e' \ v = s\<^sub>1 \ P,E \ \e\<^sub>1,s\<^sub>0\ \' \throw e',s\<^sub>1\ \ thesis" and "\E e\<^sub>1 s\<^sub>0 va s\<^sub>1 e\<^sub>2 e' s\<^sub>2 F Cs. x = E \ y = e\<^sub>1\F{Cs} := e\<^sub>2 \ z = s\<^sub>0 \ u = throw e' \ v = s\<^sub>2 \ P,E \ \e\<^sub>1,s\<^sub>0\ \' \Val va,s\<^sub>1\ \ P,E \ \e\<^sub>2,s\<^sub>1\ \' \throw e',s\<^sub>2\ \ thesis" and "\E e s\<^sub>0 e' s\<^sub>1 Copt M es. x = E \ y = Call e Copt M es \ z = s\<^sub>0 \ u = throw e' \ v = s\<^sub>1 \ P,E \ \e,s\<^sub>0\ \' \throw e',s\<^sub>1\ \ thesis" and "\E e s\<^sub>0 va s\<^sub>1 es vs ex es' s\<^sub>2 Copt M. x = E \ y = Call e Copt M es \ z = s\<^sub>0 \ u = throw ex \ v = s\<^sub>2 \ P,E \ \e,s\<^sub>0\ \' \Val va,s\<^sub>1\ \ P,E \ \es,s\<^sub>1\ [\'] \map Val vs @ throw ex # es',s\<^sub>2\ \ thesis" and "\E e s\<^sub>0 a Cs s\<^sub>1 ps vs h\<^sub>2 l\<^sub>2 C S M Ts' T' pns' body' Ds Ts T pns body Cs' vs' l\<^sub>2' new_body e' h\<^sub>3 l\<^sub>3. x = E \ y = Call e None M ps \ z = s\<^sub>0 \ u = e' \ v = (h\<^sub>3, l\<^sub>2) \ P,E \ \e,s\<^sub>0\ \' \ref (a, Cs),s\<^sub>1\ \ P,E \ \ps,s\<^sub>1\ [\'] \map Val vs,(h\<^sub>2, l\<^sub>2)\ \ h\<^sub>2 a = \(C, S)\ \ P \ last Cs has least M = (Ts', T', pns', body') via Ds \ P \ (C,Cs @\<^sub>p Ds) selects M = (Ts, T, pns, body) via Cs' \ length vs = length pns \ P \ Ts Casts vs to vs' \ l\<^sub>2' = [this \ Ref (a, Cs'), pns [\] vs'] \ new_body = (case T' of Class D \ \D\body | _ \ body) \ P,E(this \ Class (last Cs'), pns [\] Ts) \ \new_body,(h\<^sub>2, l\<^sub>2')\ \' \e',(h\<^sub>3, l\<^sub>3)\ \ thesis" and "\E e s\<^sub>0 a Cs s\<^sub>1 ps vs h\<^sub>2 l\<^sub>2 C Cs'' M Ts T pns body Cs' Ds vs' l\<^sub>2' e' h\<^sub>3 l\<^sub>3. x = E \ y = Call e \C\ M ps \ z = s\<^sub>0 \ u = e' \ v = (h\<^sub>3, l\<^sub>2) \ P,E \ \e,s\<^sub>0\ \' \ref (a, Cs),s\<^sub>1\ \ P,E \ \ps,s\<^sub>1\ [\'] \map Val vs,(h\<^sub>2, l\<^sub>2)\ \ P \ Path last Cs to C unique \ P \ Path last Cs to C via Cs'' \ P \ C has least M = (Ts, T, pns, body) via Cs' \ Ds = (Cs @\<^sub>p Cs'') @\<^sub>p Cs' \ length vs = length pns \ P \ Ts Casts vs to vs' \ l\<^sub>2' = [this \ Ref (a, Ds), pns [\] vs'] \ P,E(this \ Class (last Ds), pns [\] Ts) \ \body,(h\<^sub>2, l\<^sub>2')\ \' \e',(h\<^sub>3, l\<^sub>3)\ \ thesis" and "\E e s\<^sub>0 s\<^sub>1 es vs s\<^sub>2 Copt M. x = E \ y = Call e Copt M es \ z = s\<^sub>0 \ u = Throw (addr_of_sys_xcpt NullPointer, [NullPointer]) \ v = s\<^sub>2 \ P,E \ \e,s\<^sub>0\ \' \null,s\<^sub>1\ \ P,E \ \es,s\<^sub>1\ [\'] \map Val vs,s\<^sub>2\ \ thesis" and "\E V T e\<^sub>0 h\<^sub>0 l\<^sub>0 e\<^sub>1 h\<^sub>1 l\<^sub>1. x = E \ y = {V:T; e\<^sub>0} \ z = (h\<^sub>0, l\<^sub>0) \ u = e\<^sub>1 \ v = (h\<^sub>1, l\<^sub>1(V := l\<^sub>0 V)) \ P,E(V \ T) \ \e\<^sub>0,(h\<^sub>0, l\<^sub>0(V := None))\ \' \e\<^sub>1,(h\<^sub>1, l\<^sub>1)\ \ thesis" and "\E e\<^sub>0 s\<^sub>0 va s\<^sub>1 e\<^sub>1 e\<^sub>2 s\<^sub>2. x = E \ y = e\<^sub>0;; e\<^sub>1 \ z = s\<^sub>0 \ u = e\<^sub>2 \ v = s\<^sub>2 \ P,E \ \e\<^sub>0,s\<^sub>0\ \' \Val va,s\<^sub>1\ \ P,E \ \e\<^sub>1,s\<^sub>1\ \' \e\<^sub>2,s\<^sub>2\ \ thesis" and "\E e\<^sub>0 s\<^sub>0 e s\<^sub>1 e\<^sub>1. x = E \ y = e\<^sub>0;; e\<^sub>1 \ z = s\<^sub>0 \ u = throw e \ v = s\<^sub>1 \ P,E \ \e\<^sub>0,s\<^sub>0\ \' \throw e,s\<^sub>1\ \ thesis" and "\E e s\<^sub>0 s\<^sub>1 e\<^sub>1 e' s\<^sub>2 e\<^sub>2. x = E \ y = if (e) e\<^sub>1 else e\<^sub>2 \ z = s\<^sub>0 \ u = e' \ v = s\<^sub>2 \ P,E \ \e,s\<^sub>0\ \' \true,s\<^sub>1\ \ P,E \ \e\<^sub>1,s\<^sub>1\ \' \e',s\<^sub>2\ \ thesis" and "\E e s\<^sub>0 s\<^sub>1 e\<^sub>2 e' s\<^sub>2 e\<^sub>1. x = E \ y = if (e) e\<^sub>1 else e\<^sub>2 \ z = s\<^sub>0 \ u = e' \ v = s\<^sub>2 \ P,E \ \e,s\<^sub>0\ \' \false,s\<^sub>1\ \ P,E \ \e\<^sub>2,s\<^sub>1\ \' \e',s\<^sub>2\ \ thesis" and "\E e s\<^sub>0 e' s\<^sub>1 e\<^sub>1 e\<^sub>2. x = E \ y = if (e) e\<^sub>1 else e\<^sub>2 \ z = s\<^sub>0 \ u = throw e' \ v = s\<^sub>1 \ P,E \ \e,s\<^sub>0\ \' \throw e',s\<^sub>1\ \ thesis" and "\E e s\<^sub>0 s\<^sub>1 c. x = E \ y = while (e) c \ z = s\<^sub>0 \ u = unit \ v = s\<^sub>1 \ P,E \ \e,s\<^sub>0\ \' \false,s\<^sub>1\ \ thesis" and "\E e s\<^sub>0 s\<^sub>1 c v\<^sub>1 s\<^sub>2 e\<^sub>3 s\<^sub>3. x = E \ y = while (e) c \ z = s\<^sub>0 \ u = e\<^sub>3 \ v = s\<^sub>3 \ P,E \ \e,s\<^sub>0\ \' \true,s\<^sub>1\ \ P,E \ \c,s\<^sub>1\ \' \Val v\<^sub>1,s\<^sub>2\ \ P,E \ \while (e) c,s\<^sub>2\ \' \e\<^sub>3,s\<^sub>3\ \ thesis" and "\E e s\<^sub>0 e' s\<^sub>1 c. x = E \ y = while (e) c \ z = s\<^sub>0 \ u = throw e' \ v = s\<^sub>1 \ P,E \ \e,s\<^sub>0\ \' \throw e',s\<^sub>1\ \ thesis" and "\E e s\<^sub>0 s\<^sub>1 c e' s\<^sub>2. x = E \ y = while (e) c \ z = s\<^sub>0 \ u = throw e' \ v = s\<^sub>2 \ P,E \ \e,s\<^sub>0\ \' \true,s\<^sub>1\ \ P,E \ \c,s\<^sub>1\ \' \throw e',s\<^sub>2\ \ thesis" and "\E e s\<^sub>0 r s\<^sub>1. x = E \ y = throw e \ z = s\<^sub>0 \ u = Throw r \ v = s\<^sub>1 \ P,E \ \e,s\<^sub>0\ \' \ref r,s\<^sub>1\ \ thesis" and "\E e s\<^sub>0 s\<^sub>1. x = E \ y = throw e \ z = s\<^sub>0 \ u = Throw (addr_of_sys_xcpt NullPointer, [NullPointer]) \ v = s\<^sub>1 \ P,E \ \e,s\<^sub>0\ \' \null,s\<^sub>1\ \ thesis" and "\E e s\<^sub>0 e' s\<^sub>1. x = E \ y = throw e \ z = s\<^sub>0 \ u = throw e' \ v = s\<^sub>1 \ P,E \ \e,s\<^sub>0\ \' \throw e',s\<^sub>1\ \ thesis" shows thesis using assms by(transfer)(erule eval.cases, unfold blank_def, assumption+) lemmas [code_pred_intro] = New' NewFail' StaticUpCast' declare StaticDownCast'_new[code_pred_intro StaticDownCast'] lemmas [code_pred_intro] = StaticCastNull' declare StaticCastFail'_new[code_pred_intro StaticCastFail'] lemmas [code_pred_intro] = StaticCastThrow' StaticUpDynCast' declare StaticDownDynCast'_new[code_pred_intro StaticDownDynCast'] DynCast'[code_pred_intro DynCast'] lemmas [code_pred_intro] = DynCastNull' declare DynCastFail'[code_pred_intro DynCastFail'] lemmas [code_pred_intro] = DynCastThrow' Val' BinOp' BinOpThrow1' declare BinOpThrow2'[code_pred_intro BinOpThrow2'] lemmas [code_pred_intro] = Var' LAss' LAssThrow' declare FAcc'_new[code_pred_intro FAcc'] lemmas [code_pred_intro] = FAccNull' FAccThrow' declare FAss'_new[code_pred_intro FAss'] lemmas [code_pred_intro] = FAssNull' FAssThrow1' declare FAssThrow2'[code_pred_intro FAssThrow2'] lemmas [code_pred_intro] = CallObjThrow' declare CallParamsThrow'_new[code_pred_intro CallParamsThrow'] Call'_new[code_pred_intro Call'] StaticCall'_new[code_pred_intro StaticCall'] CallNull'_new[code_pred_intro CallNull'] lemmas [code_pred_intro] = Block' Seq' declare SeqThrow'[code_pred_intro SeqThrow'] lemmas [code_pred_intro] = CondT' declare CondF'[code_pred_intro CondF'] CondThrow'[code_pred_intro CondThrow'] lemmas [code_pred_intro] = WhileF' WhileT' declare WhileCondThrow'[code_pred_intro WhileCondThrow'] WhileBodyThrow'[code_pred_intro WhileBodyThrow'] lemmas [code_pred_intro] = Throw' declare ThrowNull'[code_pred_intro ThrowNull'] lemmas [code_pred_intro] = ThrowThrow' lemmas [code_pred_intro] = Nil' Cons' ConsThrow' code_pred (modes: eval': i \ i \ i \ i \ o \ o \ bool as big_step and evals': i \ i \ i \ i \ o \ o \ bool as big_steps) eval' proof - case eval' from eval'.prems show thesis proof(cases (no_simp) rule: eval'_cases) case (StaticDownCast E C e s\<^sub>0 a Cs Cs' s\<^sub>1) moreover have "app a [Cs] (a @ [Cs])" "app (a @ [Cs]) Cs' (a @ [Cs] @ Cs')" by(simp_all add: app_eq) ultimately show ?thesis by(rule eval'.StaticDownCast'[OF refl]) next case StaticCastFail thus ?thesis unfolding rtrancl_def subcls1_def mem_Collect_eq prod.case by(rule eval'.StaticCastFail'[OF refl]) next case (StaticDownDynCast E e s\<^sub>0 a Cs C Cs' s\<^sub>1) moreover have "app Cs [C] (Cs @ [C])" "app (Cs @ [C]) Cs' (Cs @ [C] @ Cs')" by(simp_all add: app_eq) ultimately show thesis by(rule eval'.StaticDownDynCast'[OF refl]) next case DynCast thus ?thesis by(rule eval'.DynCast'[OF refl]) next case DynCastFail thus ?thesis by(rule eval'.DynCastFail'[OF refl]) next case BinOpThrow2 thus ?thesis by(rule eval'.BinOpThrow2'[OF refl]) next case FAcc thus ?thesis by(rule eval'.FAcc'[OF refl, unfolded Predicate_Compile.contains_def Set_project_def mem_Collect_eq]) next case FAss thus ?thesis by(rule eval'.FAss'[OF refl, unfolded Predicate_Compile.contains_def Set_project_def mem_Collect_eq]) next case FAssThrow2 thus ?thesis by(rule eval'.FAssThrow2'[OF refl]) next case (CallParamsThrow E e s\<^sub>0 v s\<^sub>1 es vs ex es' s\<^sub>2 Copt M) moreover have "map_val2 (map Val vs @ throw ex # es') vs (throw ex # es')" by(simp add: map_val2_conv[symmetric]) ultimately show ?thesis by(rule eval'.CallParamsThrow'[OF refl]) next case (Call E e s\<^sub>0 a Cs s\<^sub>1 ps vs) moreover have "map_val (map Val vs) vs" by(simp add: map_val_conv[symmetric]) ultimately show ?thesis by-(rule eval'.Call'[OF refl]) next case (StaticCall E e s\<^sub>0 a Cs s\<^sub>1 ps vs) moreover have "map_val (map Val vs) vs" by(simp add: map_val_conv[symmetric]) ultimately show ?thesis by-(rule eval'.StaticCall'[OF refl]) next case (CallNull E e s\<^sub>0 s\<^sub>1 es vs) moreover have "map_val (map Val vs) vs" by(simp add: map_val_conv[symmetric]) ultimately show ?thesis by-(rule eval'.CallNull'[OF refl]) next case SeqThrow thus ?thesis by(rule eval'.SeqThrow'[OF refl]) next case CondF thus ?thesis by(rule eval'.CondF'[OF refl]) next case CondThrow thus ?thesis by(rule eval'.CondThrow'[OF refl]) next case WhileCondThrow thus ?thesis by(rule eval'.WhileCondThrow'[OF refl]) next case WhileBodyThrow thus ?thesis by(rule eval'.WhileBodyThrow'[OF refl]) next case ThrowNull thus ?thesis by(rule eval'.ThrowNull'[OF refl]) qed(assumption|erule (4) eval'.that[OF refl])+ next case evals' from evals'.prems evals'.that[OF refl] show thesis by transfer(erule evals.cases) qed subsection \Examples\ declare [[values_timeout = 180]] values [expected "{Val (Intg 5)}"] "{fst (e', s') | e' s'. [],Map.empty \ \{''V'':Integer; ''V'' := Val(Intg 5);; Var ''V''},(Map.empty,Map.empty)\ \' \e', s'\}" values [expected "{Val (Intg 11)}"] "{fst (e', s') | e' s'. [],Map.empty \ \(Val(Intg 5)) \Add\ (Val(Intg 6)),(Map.empty,Map.empty)\ \' \e', s'\}" values [expected "{Val (Intg 83)}"] "{fst (e', s') | e' s'. [],[''V''\Integer] \ \(Var ''V'') \Add\ (Val(Intg 6)), (Map.empty,[''V''\Intg 77])\ \' \e', s'\}" values [expected "{Some (Intg 6)}"] "{lcl' (snd (e', s')) ''V'' | e' s'. [],[''V''\Integer] \ \''V'' := Val(Intg 6),(Map.empty,Map.empty)\ \' \e', s'\}" values [expected "{Some (Intg 12)}"] "{lcl' (snd (e', s')) ''mult'' | e' s'. [],[''V''\Integer,''a''\Integer,''b''\Integer,''mult''\Integer] \ \(''a'' := Val(Intg 3));;(''b'' := Val(Intg 4));;(''mult'' := Val(Intg 0));; (''V'' := Val(Intg 1));; while (Var ''V'' \Eq\ Val(Intg 1))((''mult'' := Var ''mult'' \Add\ Var ''b'');; (''a'' := Var ''a'' \Add\ Val(Intg (- 1)));; (''V'' := (if(Var ''a'' \Eq\ Val(Intg 0)) Val(Intg 0) else Val(Intg 1)))), (Map.empty,Map.empty)\ \' \e', s'\}" values [expected "{Val (Intg 30)}"] "{fst (e', s') | e' s'. [],[''a''\Integer, ''b''\Integer, ''c''\ Integer, ''cond''\Boolean] \ \''a'' := Val(Intg 17);; ''b'' := Val(Intg 13);; ''c'' := Val(Intg 42);; ''cond'' := true;; if (Var ''cond'') (Var ''a'' \Add\ Var ''b'') else (Var ''a'' \Add\ Var ''c''), (Map.empty,Map.empty)\ \' \e',s'\}" text \progOverrider examples\ definition classBottom :: "cdecl" where "classBottom = (''Bottom'', [Repeats ''Left'', Repeats ''Right''], [(''x'',Integer)],[])" definition classLeft :: "cdecl" where "classLeft = (''Left'', [Repeats ''Top''],[],[(''f'', [Class ''Top'', Integer],Integer, [''V'',''W''],Var this \ ''x'' {[''Left'',''Top'']} \Add\ Val (Intg 5))])" definition classRight :: "cdecl" where "classRight = (''Right'', [Shares ''Right2''],[], [(''f'', [Class ''Top'', Integer], Integer,[''V'',''W''],Var this \ ''x'' {[''Right2'',''Top'']} \Add\ Val (Intg 7)),(''g'',[],Class ''Left'',[],new ''Left'')])" definition classRight2 :: "cdecl" where "classRight2 = (''Right2'', [Repeats ''Top''],[], [(''f'', [Class ''Top'', Integer], Integer,[''V'',''W''],Var this \ ''x'' {[''Right2'',''Top'']} \Add\ Val (Intg 9)),(''g'',[],Class ''Top'',[],new ''Top'')])" definition classTop :: "cdecl" where "classTop = (''Top'', [], [(''x'',Integer)],[])" definition progOverrider :: "cdecl list" where "progOverrider = [classBottom, classLeft, classRight, classRight2, classTop]" values [expected "{Val(Ref(0,[''Bottom'',''Left'']))}"] \ \dynCastSide\ "{fst (e', s') | e' s'. progOverrider,[''V''\Class ''Right''] \ \''V'' := new ''Bottom'' ;; Cast ''Left'' (Var ''V''),(Map.empty,Map.empty)\ \' \e', s'\}" values [expected "{Val(Ref(0,[''Right'']))}"] \ \dynCastViaSh\ "{fst (e', s') | e' s'. progOverrider,[''V''\Class ''Right2''] \ \''V'' := new ''Right'' ;; Cast ''Right'' (Var ''V''),(Map.empty,Map.empty)\ \' \e', s'\}" values [expected "{Val (Intg 42)}"] \ \block\ "{fst (e', s') | e' s'. progOverrider,[''V''\Integer] \ \''V'' := Val(Intg 42) ;; {''V'':Class ''Left''; ''V'' := new ''Bottom''} ;; Var ''V'', (Map.empty,Map.empty)\ \' \e', s'\}" values [expected "{Val (Intg 8)}"] \ \staticCall\ "{fst (e', s') | e' s'. progOverrider,[''V''\Class ''Right'',''W''\Class ''Bottom''] \ \''V'' := new ''Bottom'' ;; ''W'' := new ''Bottom'' ;; ((Cast ''Left'' (Var ''W''))\''x''{[''Left'',''Top'']} := Val(Intg 3));; (Var ''W''\(''Left''::)''f''([Var ''V'',Val(Intg 2)])),(Map.empty,Map.empty)\ \' \e', s'\}" values [expected "{Val (Intg 12)}"] \ \call\ "{fst (e', s') | e' s'. progOverrider,[''V''\Class ''Right2'',''W''\Class ''Left''] \ \''V'' := new ''Right'' ;; ''W'' := new ''Left'' ;; (Var ''V''\''f''([Var ''W'',Val(Intg 42)])) \Add\ (Var ''W''\''f''([Var ''V'',Val(Intg 13)])), (Map.empty,Map.empty)\ \' \e', s'\}" values [expected "{Val(Intg 13)}"] \ \callOverrider\ "{fst (e', s') | e' s'. progOverrider,[''V''\Class ''Right2'',''W''\Class ''Left''] \ \''V'' := new ''Bottom'';; (Var ''V'' \ ''x'' {[''Right2'',''Top'']} := Val(Intg 6));; ''W'' := new ''Left'' ;; Var ''V''\''f''([Var ''W'',Val(Intg 42)]), (Map.empty,Map.empty)\ \' \e', s'\}" values [expected "{Val(Ref(1,[''Left'',''Top'']))}"] \ \callClass\ "{fst (e', s') | e' s'. progOverrider,[''V''\Class ''Right2''] \ \''V'' := new ''Right'' ;; Var ''V''\''g''([]),(Map.empty,Map.empty)\ \' \e', s'\}" values [expected "{Val(Intg 42)}"] \ \fieldAss\ "{fst (e', s') | e' s'. progOverrider,[''V''\Class ''Right2''] \ \''V'' := new ''Right'' ;; (Var ''V''\''x''{[''Right2'',''Top'']} := (Val(Intg 42))) ;; (Var ''V''\''x''{[''Right2'',''Top'']}),(Map.empty,Map.empty)\ \' \e', s'\}" text \typing rules\ values [expected "{Class ''Bottom''}"] \ \typeNew\ "{T. progOverrider,Map.empty \ new ''Bottom'' :: T}" values [expected "{Class ''Left''}"] \ \typeDynCast\ "{T. progOverrider,Map.empty \ Cast ''Left'' (new ''Bottom'') :: T}" values [expected "{Class ''Left''}"] \ \typeStaticCast\ "{T. progOverrider,Map.empty \ \''Left''\ (new ''Bottom'') :: T}" values [expected "{Integer}"] \ \typeVal\ "{T. [],Map.empty \ Val(Intg 17) :: T}" values [expected "{Integer}"] \ \typeVar\ "{T. [],[''V'' \ Integer] \ Var ''V'' :: T}" values [expected "{Boolean}"] \ \typeBinOp\ "{T. [],Map.empty \ (Val(Intg 5)) \Eq\ (Val(Intg 6)) :: T}" values [expected "{Class ''Top''}"] \ \typeLAss\ "{T. progOverrider,[''V'' \ Class ''Top''] \ ''V'' := (new ''Left'') :: T}" values [expected "{Integer}"] \ \typeFAcc\ "{T. progOverrider,Map.empty \ (new ''Right'')\''x''{[''Right2'',''Top'']} :: T}" values [expected "{Integer}"] \ \typeFAss\ "{T. progOverrider,Map.empty \ (new ''Right'')\''x''{[''Right2'',''Top'']} :: T}" values [expected "{Integer}"] \ \typeStaticCall\ "{T. progOverrider,[''V''\Class ''Left''] \ ''V'' := new ''Left'' ;; Var ''V''\(''Left''::)''f''([new ''Top'', Val(Intg 13)]) :: T}" values [expected "{Class ''Top''}"] \ \typeCall\ "{T. progOverrider,[''V''\Class ''Right2''] \ ''V'' := new ''Right'' ;; Var ''V''\''g''([]) :: T}" values [expected "{Class ''Top''}"] \ \typeBlock\ "{T. progOverrider,Map.empty \ {''V'':Class ''Top''; ''V'' := new ''Left''} :: T}" values [expected "{Integer}"] \ \typeCond\ "{T. [],Map.empty \ if (true) Val(Intg 6) else Val(Intg 9) :: T}" values [expected "{Void}"] \ \typeWhile\ "{T. [],Map.empty \ while (false) Val(Intg 17) :: T}" values [expected "{Void}"] \ \typeThrow\ "{T. progOverrider,Map.empty \ throw (new ''Bottom'') :: T}" values [expected "{Integer}"] \ \typeBig\ "{T. progOverrider,[''V''\Class ''Right2'',''W''\Class ''Left''] \ ''V'' := new ''Right'' ;; ''W'' := new ''Left'' ;; (Var ''V''\''f''([Var ''W'', Val(Intg 7)])) \Add\ (Var ''W''\''f''([Var ''V'', Val(Intg 13)])) :: T}" text \progDiamond examples\ definition classDiamondBottom :: "cdecl" where "classDiamondBottom = (''Bottom'', [Repeats ''Left'', Repeats ''Right''],[(''x'',Integer)], [(''g'', [],Integer, [],Var this \ ''x'' {[''Bottom'']} \Add\ Val (Intg 5))])" definition classDiamondLeft :: "cdecl" where "classDiamondLeft = (''Left'', [Repeats ''TopRep'',Shares ''TopSh''],[],[])" definition classDiamondRight :: "cdecl" where "classDiamondRight = (''Right'', [Repeats ''TopRep'',Shares ''TopSh''],[], [(''f'', [Integer], Boolean,[''i''], Var ''i'' \Eq\ Val (Intg 7))])" definition classDiamondTopRep :: "cdecl" where "classDiamondTopRep = (''TopRep'', [], [(''x'',Integer)], [(''g'', [],Integer, [], Var this \ ''x'' {[''TopRep'']} \Add\ Val (Intg 10))])" definition classDiamondTopSh :: "cdecl" where "classDiamondTopSh = (''TopSh'', [], [], [(''f'', [Integer], Boolean,[''i''], Var ''i'' \Eq\ Val (Intg 3))])" definition progDiamond :: "cdecl list" where "progDiamond = [classDiamondBottom, classDiamondLeft, classDiamondRight, classDiamondTopRep, classDiamondTopSh]" values [expected "{Val(Ref(0,[''Bottom'',''Left'']))}"] \ \cast1\ "{fst (e', s') | e' s'. progDiamond,[''V''\Class ''Left''] \ \''V'' := new ''Bottom'', (Map.empty,Map.empty)\ \' \e', s'\}" values [expected "{Val(Ref(0,[''TopSh'']))}"] \ \cast2\ "{fst (e', s') | e' s'. progDiamond,[''V''\Class ''TopSh''] \ \''V'' := new ''Bottom'', (Map.empty,Map.empty)\ \' \e', s'\}" values [expected "{}"] \ \typeCast3 not typeable\ "{T. progDiamond,[''V''\Class ''TopRep''] \ ''V'' := new ''Bottom'' :: T}" values [expected "{ Val(Ref(0,[''Bottom'', ''Left'', ''TopRep''])), Val(Ref(0,[''Bottom'', ''Right'', ''TopRep''])) }"] \ \cast3\ "{fst (e', s') | e' s'. progDiamond,[''V''\Class ''TopRep''] \ \''V'' := new ''Bottom'', (Map.empty,Map.empty)\ \' \e', s'\}" values [expected "{Val(Intg 17)}"] \ \fieldAss\ "{fst (e', s') | e' s'. progDiamond,[''V''\Class ''Bottom''] \ \''V'' := new ''Bottom'' ;; ((Var ''V'')\''x''{[''Bottom'']} := (Val(Intg 17))) ;; ((Var ''V'')\''x''{[''Bottom'']}),(Map.empty,Map.empty)\ \' \e',s'\}" values [expected "{Val Null}"] \ \dynCastNull\ "{fst (e', s') | e' s'. progDiamond,Map.empty \ \Cast ''Right'' null,(Map.empty,Map.empty)\ \' \e',s'\}" values [expected "{Val (Ref(0, [''Right'']))}"] \ \dynCastViaSh\ "{fst (e', s') | e' s'. progDiamond,[''V''\Class ''TopSh''] \ \''V'' := new ''Right'' ;; Cast ''Right'' (Var ''V''),(Map.empty,Map.empty)\ \' \e',s'\}" values [expected "{Val Null}"] \ \dynCastFail\ "{fst (e', s') | e' s'. progDiamond,[''V''\Class ''TopRep''] \ \''V'' := new ''Right'' ;; Cast ''Bottom'' (Var ''V''),(Map.empty,Map.empty)\ \' \e',s'\}" values [expected "{Val (Ref(0, [''Bottom'', ''Left'']))}"] \ \dynCastSide\ "{fst (e', s') | e' s'. progDiamond,[''V''\Class ''Right''] \ \''V'' := new ''Bottom'' ;; Cast ''Left'' (Var ''V''),(Map.empty,Map.empty)\ \' \e',s'\}" text \failing g++ example\ definition classD :: "cdecl" where "classD = (''D'', [Shares ''A'', Shares ''B'', Repeats ''C''],[],[])" definition classC :: "cdecl" where "classC = (''C'', [Shares ''A'', Shares ''B''],[], [(''f'',[],Integer,[],Val(Intg 42))])" definition classB :: "cdecl" where "classB = (''B'', [],[], [(''f'',[],Integer,[],Val(Intg 17))])" definition classA :: "cdecl" where "classA = (''A'', [],[], [(''f'',[],Integer,[],Val(Intg 13))])" definition ProgFailing :: "cdecl list" where "ProgFailing = [classA,classB,classC,classD]" values [expected "{Val (Intg 42)}"] \ \callFailGplusplus\ "{fst (e', s') | e' s'. ProgFailing,Map.empty \ \{''V'':Class ''D''; ''V'' := new ''D'';; Var ''V''\''f''([])}, (Map.empty,Map.empty)\ \' \e', s'\}" end diff --git a/thys/CoreC++/TypeSafe.thy b/thys/CoreC++/TypeSafe.thy --- a/thys/CoreC++/TypeSafe.thy +++ b/thys/CoreC++/TypeSafe.thy @@ -1,1691 +1,1689 @@ (* Title: CoreC++ Author: Daniel Wasserrab Maintainer: Daniel Wasserrab Based on the Jinja theory J/TypeSafe.thy by Tobias Nipkow *) section \Type Safety Proof\ theory TypeSafe imports HeapExtension CWellForm begin subsection\Basic preservation lemmas\ lemma assumes wf:"wwf_prog P" and casts:"P \ T casts v to v'" and typeof:"P \ typeof\<^bsub>h\<^esub> v = Some T'" and leq:"P \ T' \ T" shows casts_conf:"P,h \ v' :\ T" proof - { fix a' C Cs S' assume leq:"P \ Class (last Cs) \ T" and subo:"Subobjs P C Cs" and casts':"P \ T casts Ref (a',Cs) to v'" and h:"h a' = Some(C,S')" from subo wf have "is_class P (last Cs)" by(fastforce intro:Subobj_last_isClass) with leq wf obtain C' where T:"T = Class C'" and path_unique:"P \ Path (last Cs) to C' unique" by(auto dest:Class_widen) from path_unique obtain Cs' where path_via:"P \ Path (last Cs) to C' via Cs'" by(auto simp:path_via_def path_unique_def) with T path_unique casts' have v':"v' = Ref (a',Cs@\<^sub>pCs')" by -(erule casts_to.cases,auto simp:path_unique_def path_via_def) from subo path_via wf have "Subobjs P C (Cs@\<^sub>pCs')" and "last (Cs@\<^sub>pCs') = C'" apply(auto intro:Subobjs_appendPath simp:path_via_def) apply(drule_tac Cs="Cs'" in Subobjs_nonempty) by(rule sym[OF appendPath_last]) with T h v' have ?thesis by auto } with casts typeof wf typeof leq show ?thesis by(cases v,auto elim:casts_to.cases split:if_split_asm) qed theorem assumes wf:"wwf_prog P" shows red_preserves_hconf: "P,E \ \e,(h,l)\ \ \e',(h',l')\ \ (\T. \ P,E,h \ e : T; P \ h \ \ \ P \ h' \)" and reds_preserves_hconf: "P,E \ \es,(h,l)\ [\] \es',(h',l')\ \ (\Ts. \ P,E,h \ es [:] Ts; P \ h \ \ \ P \ h' \)" proof (induct rule:red_reds_inducts) case (RedNew h a h' C E l) have new: "new_Addr h = Some a" and h':"h' = h(a \ (C, Collect (init_obj P C)))" and hconf:"P \ h \" and wt_New:"P,E,h \ new C : T" by fact+ from new have None: "h a = None" by(rule new_Addr_SomeD) with wf have oconf:"P,h \ (C, Collect (init_obj P C)) \" apply (auto simp:oconf_def) apply (rule_tac x="init_class_fieldmap P (last Cs)" in exI) by (fastforce intro:init_obj.intros fconf_init_fields elim: init_obj.cases dest!:Subobj_last_isClass simp:is_class_def)+ thus ?case using h' None by(fast intro: hconf_new[OF hconf]) next case (RedFAss h a D S Cs' F T Cs v v' Ds fs' E l T') let ?fs' = "fs'(F \ v')" let ?S' = "insert (Ds, ?fs') (S - {(Ds, fs')})" have ha:"h a = Some(D,S)" and hconf:"P \ h \" and field:"P \ last Cs' has least F:T via Cs" and casts:"P \ T casts v to v'" and Ds:"Ds = Cs' @\<^sub>p Cs" and S:"(Ds,fs') \ S" and wte:"P,E,h \ ref(a,Cs')\F{Cs} := Val v : T'" by fact+ from wte have "P \ last Cs' has least F:T' via Cs" by (auto split:if_split_asm) with field have eq:"T = T'" by (rule sees_field_fun) with casts wte wf have conf:"P,h \ v' :\ T'" by(auto intro:casts_conf) from hconf ha have oconf:"P,h \ (D,S) \" by (fastforce simp:hconf_def) with S have suboD:"Subobjs P D Ds" by (fastforce simp:oconf_def) from field obtain Bs fs ms where subo:"Subobjs P (last Cs') Cs" and "class": "class P (last Cs) = Some(Bs,fs,ms)" and map:"map_of fs F = Some T" by (auto simp:LeastFieldDecl_def FieldDecls_def) from Ds subo have last:"last Cs = last Ds" by(fastforce dest:Subobjs_nonempty intro:appendPath_last simp:appendPath_last) with "class" have classDs:"class P (last Ds) = Some(Bs,fs,ms)" by simp with S suboD oconf have "P,h \ fs' (:\) map_of fs" apply (auto simp:oconf_def) apply (erule allE) apply (erule_tac x="Ds" in allE) apply (erule_tac x="fs'" in allE) apply clarsimp done with map conf eq have fconf:"P,h \ fs'(F \ v') (:\) map_of fs" by (simp add:fconf_def) from oconf have "\Cs fs'. (Cs,fs') \ S \ Subobjs P D Cs \ (\fs Bs ms. class P (last Cs) = Some (Bs,fs,ms) \ P,h \ fs' (:\) map_of fs)" by(simp add:oconf_def) with suboD classDs fconf have oconf':"\Cs fs'. (Cs,fs') \ ?S' \ Subobjs P D Cs \ (\fs Bs ms. class P (last Cs) = Some (Bs,fs,ms) \ P,h \ fs' (:\) map_of fs)" by auto from oconf have all:"\Cs. Subobjs P D Cs \ (\!fs'. (Cs,fs') \ S)" by(simp add:oconf_def) with S have "\Cs. Subobjs P D Cs \ (\!fs'. (Cs,fs') \ ?S')" by blast with oconf' have oconf':"P,h \ (D,?S') \" by (simp add:oconf_def) with hconf ha show ?case by (rule hconf_upd_obj) next case (CallObj E e h l e' h' l' Copt M es) thus ?case by (cases Copt) auto next case (CallParams E es h l es' h' l' v Copt M) thus ?case by (cases Copt) auto next case (RedCallNull E Copt M vs h l) thus ?case by (cases Copt) auto qed auto theorem assumes wf:"wwf_prog P" shows red_preserves_lconf: "P,E \ \e,(h,l)\ \ \e',(h',l')\ \ (\T. \ P,E,h \ e:T; P,h \ l (:\)\<^sub>w E; P \ E \ \ \ P,h' \ l' (:\)\<^sub>w E)" and reds_preserves_lconf: "P,E \ \es,(h,l)\ [\] \es',(h',l')\ \ (\Ts. \ P,E,h \ es[:]Ts; P,h \ l (:\)\<^sub>w E; P \ E \ \ \ P,h' \ l' (:\)\<^sub>w E)" proof(induct rule:red_reds_inducts) case RedNew thus ?case by(fast intro:lconf_hext red_hext_incr[OF red_reds.RedNew]) next case (RedLAss E V T v v' h l T') have casts:"P \ T casts v to v'" and env:"E V = Some T" and wt:"P,E,h \ V:=Val v : T'" and lconf:"P,h \ l (:\)\<^sub>w E" by fact+ from wt env have eq:"T = T'" by auto with casts wt wf have conf:"P,h \ v' :\ T'" by(auto intro:casts_conf) with lconf env eq show ?case by (simp del:fun_upd_apply)(erule lconf_upd,simp_all) next case RedFAss thus ?case by(auto intro:lconf_hext red_hext_incr[OF red_reds.RedFAss] simp del:fun_upd_apply) next case (BlockRedNone E V T e h l e' h' l' T') have red:"P,E(V \ T) \ \e,(h, l(V := None))\ \ \e',(h', l')\" and IH: "\T''. \ P,E(V \ T),h \ e : T''; P,h \ l(V:=None) (:\)\<^sub>w E(V \ T); envconf P (E(V \ T)) \ \ P,h' \ l' (:\)\<^sub>w E(V \ T)" and lconf: "P,h \ l (:\)\<^sub>w E" and wte: "P,E,h \ {V:T; e} : T'" and envconf:"envconf P E" by fact+ from lconf_hext[OF lconf red_hext_incr[OF red]] have lconf':"P,h' \ l (:\)\<^sub>w E" . from wte have wte':"P,E(V\T),h \ e : T'" and type:"is_type P T" by (auto elim:WTrt.cases) from envconf type have envconf':"envconf P (E(V \ T))" by(auto simp:envconf_def) from lconf have "P,h \ (l(V := None)) (:\)\<^sub>w E(V\T)" by (simp add:lconf_def fun_upd_apply) from IH[OF wte' this envconf'] have "P,h' \ l' (:\)\<^sub>w E(V\T)" . with lconf' show ?case by (fastforce simp:lconf_def fun_upd_apply split:if_split_asm) next case (BlockRedSome E V T e h l e' h' l' v T') have red:"P,E(V \ T) \ \e,(h, l(V := None))\ \ \e',(h', l')\" and IH: "\T''. \ P,E(V \ T),h \ e : T''; P,h \ l(V:=None) (:\)\<^sub>w E(V \ T); envconf P (E(V \ T)) \ \ P,h' \ l' (:\)\<^sub>w E(V \ T)" and lconf: "P,h \ l (:\)\<^sub>w E" and wte: "P,E,h \ {V:T; e} : T'" and envconf:"envconf P E" by fact+ from lconf_hext[OF lconf red_hext_incr[OF red]] have lconf':"P,h' \ l (:\)\<^sub>w E" . from wte have wte':"P,E(V\T),h \ e : T'" and type:"is_type P T" by (auto elim:WTrt.cases) from envconf type have envconf':"envconf P (E(V \ T))" by(auto simp:envconf_def) from lconf have "P,h \ (l(V := None)) (:\)\<^sub>w E(V\T)" by (simp add:lconf_def fun_upd_apply) from IH[OF wte' this envconf'] have "P,h' \ l' (:\)\<^sub>w E(V\T)" . with lconf' show ?case by (fastforce simp:lconf_def fun_upd_apply split:if_split_asm) next case (InitBlockRed E V T e h l v' e' h' l' v'' v T') have red: "P,E(V \ T) \ \e, (h, l(V\v'))\ \ \e',(h', l')\" and IH: "\T''. \ P,E(V \ T),h \ e : T''; P,h \ l(V \ v') (:\)\<^sub>w E(V \ T); envconf P (E(V \ T)) \ \ P,h' \ l' (:\)\<^sub>w E(V \ T)" and lconf:"P,h \ l (:\)\<^sub>w E" and l':"l' V = Some v''" and wte:"P,E,h \ {V:T; V:=Val v;; e} : T'" and casts:"P \ T casts v to v'" and envconf:"envconf P E" by fact+ from lconf_hext[OF lconf red_hext_incr[OF red]] have lconf':"P,h' \ l (:\)\<^sub>w E" . from wte obtain T'' where wte':"P,E(V\T),h \ e : T'" and wt:"P,E(V \ T),h \ V:=Val v : T''" and type:"is_type P T" by (auto elim:WTrt.cases) from envconf type have envconf':"envconf P (E(V \ T))" by(auto simp:envconf_def) from wt have "T'' = T" by auto with wf casts wt have "P,h \ v' :\ T" by(auto intro:casts_conf) with lconf have "P,h \ l(V \ v') (:\)\<^sub>w E(V\T)" by -(rule lconf_upd2) from IH[OF wte' this envconf'] have "P,h' \ l' (:\)\<^sub>w E(V \ T)" . with lconf' show ?case by (fastforce simp:lconf_def fun_upd_apply split:if_split_asm) next case (CallObj E e h l e' h' l' Copt M es) thus ?case by (cases Copt) auto next case (CallParams E es h l es' h' l' v Copt M) thus ?case by (cases Copt) auto next case (RedCallNull E Copt M vs h l) thus ?case by (cases Copt) auto qed auto text\Preservation of definite assignment more complex and requires a few lemmas first.\ lemma [iff]: "\A. \ length Vs = length Ts; length vs = length Ts\ \ \ (blocks (Vs,Ts,vs,e)) A = \ e (A \ \set Vs\)" apply(induct Vs Ts vs e rule:blocks_old_induct) apply(simp_all add:hyperset_defs) done lemma red_lA_incr: "P,E \ \e,(h,l)\ \ \e',(h',l')\ \ \dom l\ \ \ e \ \dom l'\ \ \ e'" and reds_lA_incr: "P,E \ \es,(h,l)\ [\] \es',(h',l')\ \ \dom l\ \ \s es \ \dom l'\ \ \s es'" apply (induct rule:red_reds_inducts) apply (simp_all del: fun_upd_apply add: hyperset_defs) apply blast apply blast apply blast apply blast apply blast apply blast apply blast apply auto done text\Now preservation of definite assignment.\ lemma assumes wf: "wf_C_prog P" shows red_preserves_defass: "P,E \ \e,(h,l)\ \ \e',(h',l')\ \ \ e \dom l\ \ \ e' \dom l'\" and "P,E \ \es,(h,l)\ [\] \es',(h',l')\ \ \s es \dom l\ \ \s es' \dom l'\" proof (induct rule:red_reds_inducts) case BinOpRed1 thus ?case by (auto elim!: D_mono[OF red_lA_incr]) next case FAssRed1 thus ?case by (auto elim!: D_mono[OF red_lA_incr]) next case CallObj thus ?case by (auto elim!: Ds_mono[OF red_lA_incr]) next case (RedCall h l a C S Cs M Ts' T' pns' body' Ds Ts T pns body Cs' vs bs new_body E) thus ?case apply (auto dest!:select_method_wf_mdecl[OF wf] simp:wf_mdecl_def elim!:D_mono') apply(cases T') apply auto by(rule_tac A="\insert this (set pns)\" in D_mono,clarsimp simp:hyperset_defs, assumption)+ next case RedStaticCall thus ?case apply (auto dest!:has_least_wf_mdecl[OF wf] simp:wf_mdecl_def elim!:D_mono') by(auto simp:hyperset_defs) next case InitBlockRed thus ?case by(auto simp:hyperset_defs elim!:D_mono' simp del:fun_upd_apply) next case BlockRedNone thus ?case by(auto simp:hyperset_defs elim!:D_mono' simp del:fun_upd_apply) next case BlockRedSome thus ?case by(auto simp:hyperset_defs elim!:D_mono' simp del:fun_upd_apply) next case SeqRed thus ?case by (auto elim!: D_mono[OF red_lA_incr]) next case CondRed thus ?case by (auto elim!: D_mono[OF red_lA_incr]) next case RedWhile thus ?case by(auto simp:hyperset_defs elim!:D_mono') next case ListRed1 thus ?case by (auto elim!: Ds_mono[OF red_lA_incr]) qed (auto simp:hyperset_defs) text\Combining conformance of heap and local variables:\ definition sconf :: "prog \ env \ state \ bool" ("_,_ \ _ \" [51,51,51]50) where "P,E \ s \ \ let (h,l) = s in P \ h \ \ P,h \ l (:\)\<^sub>w E \ P \ E \" lemma red_preserves_sconf: "\ P,E \ \e,s\ \ \e',s'\; P,E,hp s \ e : T; P,E \ s \; wwf_prog P\ \ P,E \ s' \" by(fastforce intro:red_preserves_hconf red_preserves_lconf simp add:sconf_def) lemma reds_preserves_sconf: "\ P,E \ \es,s\ [\] \es',s'\; P,E,hp s \ es [:] Ts; P,E \ s \; wwf_prog P\ \ P,E \ s' \" by(fastforce intro:reds_preserves_hconf reds_preserves_lconf simp add:sconf_def) subsection "Subject reduction" lemma wt_blocks: "\E. \ length Vs = length Ts; length vs = length Ts; \T' \ set Ts. is_type P T'\ \ (P,E,h \ blocks(Vs,Ts,vs,e) : T) = (P,E(Vs[\]Ts),h \ e:T \ (\Ts'. map (P \ typeof\<^bsub>h\<^esub>) vs = map Some Ts' \ P \ Ts' [\] Ts))" proof(induct Vs Ts vs e rule:blocks_old_induct) case (5 V Vs T' Ts v vs e) have length:"length (V#Vs) = length (T'#Ts)" "length (v#vs) = length (T'#Ts)" and type:"\S \ set (T'#Ts). is_type P S" and IH:"\E. \length Vs = length Ts; length vs = length Ts; \S \ set Ts. is_type P S\ \ (P,E,h \ blocks (Vs, Ts, vs, e) : T) = (P,E(Vs [\] Ts),h \ e : T \ (\Ts'. map P \ typeof\<^bsub>h\<^esub> vs = map Some Ts' \ P \ Ts' [\] Ts))" by fact+ from type have typeT':"is_type P T'" and type':"\S \ set Ts. is_type P S" by simp_all from length have "length Vs = length Ts" "length vs = length Ts" by simp_all from IH[OF this type'] have eq:"(P,E(V \ T'),h \ blocks (Vs,Ts,vs,e) : T) = - (P,E(V \ T')(Vs [\] Ts),h \ e : T \ + (P,E(V \ T', Vs [\] Ts),h \ e : T \ (\Ts'. map P \ typeof\<^bsub>h\<^esub> vs = map Some Ts' \ P \ Ts' [\] Ts))" . show ?case proof(rule iffI) assume "P,E,h \ blocks (V#Vs,T'#Ts,v#vs,e) : T" then have wt:"P,E(V \ T'),h \ V:=Val v : T'" and blocks:"P,E(V \ T'),h \ blocks (Vs,Ts,vs,e) : T" by auto - from blocks eq obtain Ts' where wte:"P,E(V \ T')(Vs [\] Ts),h \ e : T" + from blocks eq obtain Ts' where wte:"P,E(V \ T', Vs [\] Ts),h \ e : T" and typeof:"map P \ typeof\<^bsub>h\<^esub> vs = map Some Ts'" and subs:"P \ Ts' [\] Ts" by auto from wt obtain T'' where "P \ typeof\<^bsub>h\<^esub> v = Some T''" and "P \ T'' \ T'" by auto with wte typeof subs show "P,E(V # Vs [\] T' # Ts),h \ e : T \ (\Ts'. map P \ typeof\<^bsub>h\<^esub> (v # vs) = map Some Ts' \ P \ Ts' [\] (T' # Ts))" by auto next assume "P,E(V # Vs [\] T' # Ts),h \ e : T \ (\Ts'. map P \ typeof\<^bsub>h\<^esub> (v # vs) = map Some Ts' \ P \ Ts' [\] (T' # Ts))" then obtain Ts' where wte:"P,E(V # Vs [\] T' # Ts),h \ e : T" and typeof:"map P \ typeof\<^bsub>h\<^esub> (v # vs) = map Some Ts'" and subs:"P \ Ts' [\] (T'#Ts)" by auto from subs obtain U Us where Ts':"Ts' = U#Us" by(cases Ts') auto with wte typeof subs eq have blocks:"P,E(V \ T'),h \ blocks (Vs,Ts,vs,e) : T" by auto from Ts' typeof subs have "P \ typeof\<^bsub>h\<^esub> v = Some U" and "P \ U \ T'" by (auto simp:fun_of_def) hence wtval:"P,E(V \ T'),h \ V:=Val v : T'" by auto with blocks typeT' show "P,E,h \ blocks (V#Vs,T'#Ts,v#vs,e) : T" by auto qed qed auto theorem assumes wf: "wf_C_prog P" shows subject_reduction2: "P,E \ \e,(h,l)\ \ \e',(h',l')\ \ (\T. \ P,E \ (h,l) \; P,E,h \ e : T \ \ P,E,h' \ e' :\<^bsub>NT\<^esub> T)" and subjects_reduction2: "P,E \ \es,(h,l)\ [\] \es',(h',l')\ \ (\Ts.\ P,E \ (h,l) \; P,E,h \ es [:] Ts \ \ types_conf P E h' es' Ts)" proof (induct rule:red_reds_inducts) case (RedNew h a h' C E l) have new:"new_Addr h = Some a" and h':"h' = h(a \ (C, Collect (init_obj P C)))" and wt:"P,E,h \ new C : T" by fact+ from wt have eq:"T = Class C" and "class": "is_class P C" by auto from "class" have subo:"Subobjs P C [C]" by(rule Subobjs_Base) from h' have "h' a = Some(C, Collect (init_obj P C))" by(simp add:map_upd_Some_unfold) with subo have "P,E,h' \ ref(a,[C]) : Class C" by auto with eq show ?case by auto next case (RedNewFail h E C l) have sconf:"P,E \ (h, l) \" by fact from wf have "is_class P OutOfMemory" by (fastforce intro:is_class_xcpt wf_prog_wwf_prog) hence "preallocated h \ P \ typeof\<^bsub>h\<^esub> (Ref (addr_of_sys_xcpt OutOfMemory,[OutOfMemory])) = Some(Class OutOfMemory)" by (auto elim: preallocatedE dest!:preallocatedD Subobjs_Base) with sconf have "P,E,h \ THROW OutOfMemory : T" by(auto simp:sconf_def hconf_def) thus ?case by (fastforce intro:wt_same_type_typeconf) next case (StaticCastRed E e h l e' h' l' C) have wt:"P,E,h \ \C\e : T" and IH:"\T'. \P,E \ (h,l) \; P,E,h \ e : T'\ \ P,E,h' \ e' :\<^bsub>NT\<^esub> T'" and sconf:"P,E \ (h, l) \" by fact+ from wt obtain T' where wte:"P,E,h \ e : T'" and isref:"is_refT T'" and "class": "is_class P C" and T:"T = Class C" by auto from isref have "P,E,h' \ \C\e' : Class C" proof(rule refTE) assume "T' = NT" with IH[OF sconf wte] isref "class" show ?thesis by auto next fix D assume "T' = Class D" with IH[OF sconf wte] isref "class" show ?thesis by auto qed with T show ?case by (fastforce intro:wt_same_type_typeconf) next case RedStaticCastNull thus ?case by (auto elim:WTrt.cases) next case (RedStaticUpCast Cs C Cs' Ds E a h l) have wt:"P,E,h \ \C\ref (a,Cs) : T" and path_via:"P \ Path last Cs to C via Cs'" and Ds:"Ds = Cs @\<^sub>p Cs'" by fact+ from wt have typeof:"P \ typeof\<^bsub>h\<^esub> (Ref(a,Cs)) = Some(Class(last Cs))" and "class": "is_class P C" and T:"T = Class C" by auto from typeof obtain D S where h:"h a = Some(D,S)" and subo:"Subobjs P D Cs" by (auto dest:typeof_Class_Subo split:if_split_asm) from path_via subo wf Ds have "Subobjs P D Ds" and last:"last Ds = C" by(auto intro!:Subobjs_appendPath appendPath_last[THEN sym] Subobjs_nonempty simp:path_via_def) with h have "P,E,h \ ref (a,Ds) : Class C" by auto with T show ?case by (fastforce intro:wt_same_type_typeconf) next case (RedStaticDownCast E C a Cs Cs' h l) have "P,E,h \ \C\ref (a,Cs@[C]@Cs') : T" by fact hence typeof:"P \ typeof\<^bsub>h\<^esub> (Ref(a,Cs@[C]@Cs')) = Some(Class(last(Cs@[C]@Cs')))" and "class": "is_class P C" and T:"T = Class C" by auto from typeof obtain D S where h:"h a = Some(D,S)" and subo:"Subobjs P D (Cs@[C]@Cs')" by (auto dest:typeof_Class_Subo split:if_split_asm) from subo have "Subobjs P D (Cs@[C])" by(fastforce intro:appendSubobj) with h have "P,E,h \ ref (a,Cs@[C]) : Class C" by auto with T show ?case by (fastforce intro:wt_same_type_typeconf) next case (RedStaticCastFail C Cs E a h l) have sconf:"P,E \ (h, l) \" by fact from wf have "is_class P ClassCast" by (fastforce intro:is_class_xcpt wf_prog_wwf_prog) hence "preallocated h \ P \ typeof\<^bsub>h\<^esub> (Ref (addr_of_sys_xcpt ClassCast,[ClassCast])) = Some(Class ClassCast)" by (auto elim: preallocatedE dest!:preallocatedD Subobjs_Base) with sconf have "P,E,h \ THROW ClassCast : T" by(auto simp:sconf_def hconf_def) thus ?case by (fastforce intro:wt_same_type_typeconf) next case (DynCastRed E e h l e' h' l' C) have wt:"P,E,h \ Cast C e : T" and IH:"\T'. \P,E \ (h,l) \; P,E,h \ e : T'\ \ P,E,h' \ e' :\<^bsub>NT\<^esub> T'" and sconf:"P,E \ (h,l) \" by fact+ from wt obtain T' where wte:"P,E,h \ e : T'" and isref:"is_refT T'" and "class": "is_class P C" and T:"T = Class C" by auto from isref have "P,E,h' \ Cast C e' : Class C" proof(rule refTE) assume "T' = NT" with IH[OF sconf wte] isref "class" show ?thesis by auto next fix D assume "T' = Class D" with IH[OF sconf wte] isref "class" show ?thesis by auto qed with T show ?case by (fastforce intro:wt_same_type_typeconf) next case RedDynCastNull thus ?case by (auto elim:WTrt.cases) next case (RedDynCast h l a D S C Cs' E Cs) have wt:"P,E,h \ Cast C (ref (a,Cs)) : T" and path_via:"P \ Path D to C via Cs'" and hp:"hp (h,l) a = Some(D,S)" by fact+ from wt have typeof:"P \ typeof\<^bsub>h\<^esub> (Ref(a,Cs)) = Some(Class(last Cs))" and "class": "is_class P C" and T:"T = Class C" by auto from typeof hp have subo:"Subobjs P D Cs" by (auto dest:typeof_Class_Subo split:if_split_asm) from path_via subo have "Subobjs P D Cs'" and last:"last Cs' = C" by (auto simp:path_via_def) with hp have "P,E,h \ ref (a,Cs') : Class C" by auto with T show ?case by (fastforce intro:wt_same_type_typeconf) next case (RedStaticUpDynCast Cs C Cs' Ds E a h l) have wt:"P,E,h \ Cast C (ref (a,Cs)) : T" and path_via:"P \ Path last Cs to C via Cs'" and Ds:"Ds = Cs @\<^sub>p Cs'" by fact+ from wt have typeof:"P \ typeof\<^bsub>h\<^esub> (Ref(a,Cs)) = Some(Class(last Cs))" and "class": "is_class P C" and T:"T = Class C" by auto from typeof obtain D S where h:"h a = Some(D,S)" and subo:"Subobjs P D Cs" by (auto dest:typeof_Class_Subo split:if_split_asm) from path_via subo wf Ds have "Subobjs P D Ds" and last:"last Ds = C" by(auto intro!:Subobjs_appendPath appendPath_last[THEN sym] Subobjs_nonempty simp:path_via_def) with h have "P,E,h \ ref (a,Ds) : Class C" by auto with T show ?case by (fastforce intro:wt_same_type_typeconf) next case (RedStaticDownDynCast E C a Cs Cs' h l) have "P,E,h \ Cast C (ref (a,Cs@[C]@Cs')) : T" by fact hence typeof:"P \ typeof\<^bsub>h\<^esub> (Ref(a,Cs@[C]@Cs')) = Some(Class(last(Cs@[C]@Cs')))" and "class": "is_class P C" and T:"T = Class C" by auto from typeof obtain D S where h:"h a = Some(D,S)" and subo:"Subobjs P D (Cs@[C]@Cs')" by (auto dest:typeof_Class_Subo split:if_split_asm) from subo have "Subobjs P D (Cs@[C])" by(fastforce intro:appendSubobj) with h have "P,E,h \ ref (a,Cs@[C]) : Class C" by auto with T show ?case by (fastforce intro:wt_same_type_typeconf) next case RedDynCastFail thus ?case by fastforce next case (BinOpRed1 E e h l e' h' l' bop e\<^sub>2) have red:"P,E \ \e,(h, l)\ \ \e',(h', l')\" and wt:"P,E,h \ e \bop\ e\<^sub>2 : T" and IH:"\T'. \P,E \ (h,l) \; P,E,h \ e : T'\ \ P,E,h' \ e' :\<^bsub>NT\<^esub> T'" and sconf:"P,E \ (h,l) \" by fact+ from wt obtain T\<^sub>1 T\<^sub>2 where wte:"P,E,h \ e : T\<^sub>1" and wte2:"P,E,h \ e\<^sub>2 : T\<^sub>2" and binop:"case bop of Eq \ T = Boolean | Add \ T\<^sub>1 = Integer \ T\<^sub>2 = Integer \ T = Integer" by auto from WTrt_hext_mono[OF wte2 red_hext_incr[OF red]] have wte2':"P,E,h' \ e\<^sub>2 : T\<^sub>2" . have "P,E,h' \ e' \bop\ e\<^sub>2 : T" proof (cases bop) assume Eq:"bop = Eq" from IH[OF sconf wte] obtain T' where "P,E,h' \ e' : T'" by (cases "T\<^sub>1") auto with wte2' binop Eq show ?thesis by(cases bop) auto next assume Add:"bop = Add" with binop have Intg:"T\<^sub>1 = Integer" by simp with IH[OF sconf wte] have "P,E,h' \ e' : Integer" by simp with wte2' binop Add show ?thesis by(cases bop) auto qed with binop show ?case by(cases bop) simp_all next case (BinOpRed2 E e h l e' h' l' v\<^sub>1 bop) have red:"P,E \ \e,(h,l)\ \ \e',(h',l')\" and wt:"P,E,h \ Val v\<^sub>1 \bop\ e : T" and IH:"\T'. \P,E \ (h,l) \; P,E,h \ e : T'\ \ P,E,h' \ e' :\<^bsub>NT\<^esub> T'" and sconf:"P,E \ (h,l) \" by fact+ from wt obtain T\<^sub>1 T\<^sub>2 where wtval:"P,E,h \ Val v\<^sub>1 : T\<^sub>1" and wte:"P,E,h \ e : T\<^sub>2" and binop:"case bop of Eq \ T = Boolean | Add \ T\<^sub>1 = Integer \ T\<^sub>2 = Integer \ T = Integer" by auto from WTrt_hext_mono[OF wtval red_hext_incr[OF red]] have wtval':"P,E,h' \ Val v\<^sub>1 : T\<^sub>1" . have "P,E,h' \ Val v\<^sub>1 \bop\ e' : T" proof (cases bop) assume Eq:"bop = Eq" from IH[OF sconf wte] obtain T' where "P,E,h' \ e' : T'" by (cases "T\<^sub>2") auto with wtval' binop Eq show ?thesis by(cases bop) auto next assume Add:"bop = Add" with binop have Intg:"T\<^sub>2 = Integer" by simp with IH[OF sconf wte] have "P,E,h' \ e' : Integer" by simp with wtval' binop Add show ?thesis by(cases bop) auto qed with binop show ?case by(cases bop) simp_all next case (RedBinOp bop v\<^sub>1 v\<^sub>2 v E a b) thus ?case proof (cases bop) case Eq thus ?thesis using RedBinOp by auto next case Add thus ?thesis using RedBinOp by auto qed next case (RedVar h l V v E) have l:"lcl (h, l) V = Some v" and sconf:"P,E \ (h, l) \" and wt:"P,E,h \ Var V : T" by fact+ hence conf:"P,h \ v :\ T" by(force simp:sconf_def lconf_def) show ?case proof(cases "\C. T \ Class C") case True with conf have "P \ typeof\<^bsub>h\<^esub> v = Some T" by(cases T) auto hence "P,E,h \ Val v : T" by auto thus ?thesis by(rule wt_same_type_typeconf) next case False then obtain C where T:"T = Class C" by auto with conf have "P \ typeof\<^bsub>h\<^esub> v = Some(Class C) \ P \ typeof\<^bsub>h\<^esub> v = Some NT" by simp with T show ?thesis by simp qed next case (LAssRed E e h l e' h' l' V) have wt:"P,E,h \ V:=e : T" and sconf:"P,E \ (h, l) \" and IH:"\T'. \P,E \ (h, l) \; P,E,h \ e : T'\ \ P,E,h' \ e' :\<^bsub>NT\<^esub> T'" by fact+ from wt obtain T' where wte:"P,E,h \ e : T'" and env:"E V = Some T" and sub:"P \ T' \ T" by auto from sconf env have "is_type P T" by(auto simp:sconf_def envconf_def) from sub this wf show ?case proof(rule subE) assume eq:"T' = T" and notclass:"\C. T' \ Class C" with IH[OF sconf wte] have "P,E,h' \ e' : T" by(cases T) auto with eq env have "P,E,h' \ V:=e' : T" by auto with eq show ?thesis by(cases T) auto next fix C D assume T':"T' = Class C" and T:"T = Class D" and path_unique:"P \ Path C to D unique" with IH[OF sconf wte] have "P,E,h' \ e' : Class C \ P,E,h' \ e' : NT" by simp hence "P,E,h' \ V:=e' : T" proof(rule disjE) assume "P,E,h' \ e' : Class C" with env T' sub show ?thesis by (fastforce intro:WTrtLAss) next assume "P,E,h' \ e' : NT" with env T show ?thesis by (fastforce intro:WTrtLAss) qed with T show ?thesis by(cases T) auto next fix C assume T':"T' = NT" and T:"T = Class C" with IH[OF sconf wte] have "P,E,h' \ e' : NT" by simp with env T show ?thesis by (fastforce intro:WTrtLAss) qed next case (RedLAss E V T v v' h l T') have env:"E V = Some T" and casts:"P \ T casts v to v'" and sconf:"P,E \ (h, l) \" and wt:"P,E,h \ V:=Val v : T'" by fact+ show ?case proof(cases "\C. T \ Class C") case True with casts wt env show ?thesis by(cases T',auto elim!:casts_to.cases) next case False then obtain C where "T = Class C" by auto with casts wt env wf show ?thesis by(auto elim!:casts_to.cases, auto intro!:sym[OF appendPath_last] Subobjs_nonempty split:if_split_asm simp:path_via_def,drule_tac Cs="Cs" in Subobjs_appendPath,auto) qed next case (FAccRed E e h l e' h' l' F Cs) have red:"P,E \ \e,(h,l)\ \ \e',(h',l')\" and wt:"P,E,h \ e\F{Cs} : T" and IH:"\T'. \P,E \ (h,l) \; P,E,h \ e : T'\ \ P,E,h' \ e' :\<^bsub>NT\<^esub> T'" and sconf:"P,E \ (h,l) \" by fact+ from wt have "P,E,h' \ e'\F{Cs} : T" proof(rule WTrt_elim_cases) fix C assume wte: "P,E,h \ e : Class C" and field:"P \ C has least F:T via Cs" and notemptyCs:"Cs \ []" from field have "class": "is_class P C" by (fastforce intro:Subobjs_isClass simp add:LeastFieldDecl_def FieldDecls_def) from IH[OF sconf wte] have "P,E,h' \ e' : NT \ P,E,h' \ e' : Class C" by auto thus ?thesis proof(rule disjE) assume "P,E,h' \ e' : NT" thus ?thesis by (fastforce intro!:WTrtFAccNT) next assume wte':"P,E,h' \ e' : Class C" from wte' notemptyCs field show ?thesis by(rule WTrtFAcc) qed next assume wte: "P,E,h \ e : NT" from IH[OF sconf wte] have "P,E,h' \ e' : NT" by auto thus ?thesis by (rule WTrtFAccNT) qed thus ?case by(rule wt_same_type_typeconf) next case (RedFAcc h l a D S Ds Cs' Cs fs' F v E) have h:"hp (h,l) a = Some(D,S)" and Ds:"Ds = Cs'@\<^sub>pCs" and S:"(Ds,fs') \ S" and fs':"fs' F = Some v" and sconf:"P,E \ (h,l) \" and wte:"P,E,h \ ref (a,Cs')\F{Cs} : T" by fact+ from wte have field:"P \ last Cs' has least F:T via Cs" and notemptyCs:"Cs \ []" by (auto split:if_split_asm) from h S sconf obtain Bs fs ms where classDs:"class P (last Ds) = Some (Bs,fs,ms)" and fconf:"P,h \ fs' (:\) map_of fs" by (simp add:sconf_def hconf_def oconf_def) blast from field Ds have "last Cs = last Ds" by (fastforce intro!:appendPath_last Subobjs_nonempty simp:LeastFieldDecl_def FieldDecls_def) with field classDs have map:"map_of fs F = Some T" by (simp add:LeastFieldDecl_def FieldDecls_def) with fconf fs' have conf:"P,h \ v :\ T" by (simp add:fconf_def,erule_tac x="F" in allE,fastforce) thus ?case by (cases T) auto next case (RedFAccNull E F Cs h l) have sconf:"P,E \ (h, l) \" by fact from wf have "is_class P NullPointer" by (fastforce intro:is_class_xcpt wf_prog_wwf_prog) hence "preallocated h \ P \ typeof\<^bsub>h\<^esub> (Ref (addr_of_sys_xcpt NullPointer,[NullPointer])) = Some(Class NullPointer)" by (auto elim: preallocatedE dest!:preallocatedD Subobjs_Base) with sconf have "P,E,h \ THROW NullPointer : T" by(auto simp:sconf_def hconf_def) thus ?case by (fastforce intro:wt_same_type_typeconf wf_prog_wwf_prog) next case (FAssRed1 E e h l e' h' l' F Cs e\<^sub>2) have red:"P,E \ \e,(h,l)\ \ \e',(h',l')\" and wt:"P,E,h \ e\F{Cs} := e\<^sub>2 : T" and IH:"\T'. \P,E \ (h,l) \; P,E,h \ e : T'\ \ P,E,h' \ e' :\<^bsub>NT\<^esub> T'" and sconf:"P,E \ (h,l) \" by fact+ from wt have "P,E,h' \ e'\F{Cs} := e\<^sub>2 : T" proof (rule WTrt_elim_cases) fix C T' assume wte: "P,E,h \ e : Class C" and field:"P \ C has least F:T via Cs" and notemptyCs:"Cs \ []" and wte2:"P,E,h \ e\<^sub>2 : T'" and sub:"P \ T' \ T" have wte2': "P,E,h' \ e\<^sub>2 : T'" by(rule WTrt_hext_mono[OF wte2 red_hext_incr[OF red]]) from IH[OF sconf wte] have "P,E,h' \ e' : Class C \ P,E,h' \ e' : NT" by simp thus ?thesis proof(rule disjE) assume wte':"P,E,h' \ e' : Class C" from wte' notemptyCs field wte2' sub show ?thesis by (rule WTrtFAss) next assume wte':"P,E,h' \ e' : NT" from wte' wte2' sub show ?thesis by (rule WTrtFAssNT) qed next fix T' assume wte:"P,E,h \ e : NT" and wte2:"P,E,h \ e\<^sub>2 : T'" and sub:"P \ T' \ T" have wte2': "P,E,h' \ e\<^sub>2 : T'" by(rule WTrt_hext_mono[OF wte2 red_hext_incr[OF red]]) from IH[OF sconf wte] have wte':"P,E,h' \ e' : NT" by simp from wte' wte2' sub show ?thesis by (rule WTrtFAssNT) qed thus ?case by(rule wt_same_type_typeconf) next case (FAssRed2 E e h l e' h' l' v F Cs) have red:"P,E \ \e,(h,l)\ \ \e',(h',l')\" and wt:"P,E,h \ Val v\F{Cs} := e : T" and IH:"\T'. \P,E \ (h,l) \; P,E,h \ e : T'\ \ P,E,h' \ e' :\<^bsub>NT\<^esub> T'" and sconf:"P,E \ (h,l) \" by fact+ from wt have "P,E,h' \ Val v\F{Cs}:=e' : T" proof (rule WTrt_elim_cases) fix C T' assume wtval:"P,E,h \ Val v : Class C" and field:"P \ C has least F:T via Cs" and notemptyCs:"Cs \ []" and wte:"P,E,h \ e : T'" and sub:"P \ T' \ T" have wtval':"P,E,h' \ Val v : Class C" by(rule WTrt_hext_mono[OF wtval red_hext_incr[OF red]]) from field wf have type:"is_type P T" by(rule least_field_is_type) from sub type wf show ?thesis proof(rule subE) assume "T' = T" and notclass:"\C. T' \ Class C" from IH[OF sconf wte] notclass have wte':"P,E,h' \ e' : T'" by(cases T') auto from wtval' notemptyCs field wte' sub show ?thesis by(rule WTrtFAss) next fix C' D assume T':"T' = Class C'" and T:"T = Class D" and path_unique:"P \ Path C' to D unique" from IH[OF sconf wte] T' have "P,E,h' \ e' : Class C' \ P,E,h' \ e' : NT" by simp thus ?thesis proof(rule disjE) assume wte':"P,E,h' \ e' : Class C'" from wtval' notemptyCs field wte' sub T' show ?thesis by (fastforce intro: WTrtFAss) next assume wte':"P,E,h' \ e' : NT" from wtval' notemptyCs field wte' sub T show ?thesis by (fastforce intro: WTrtFAss) qed next fix C' assume T':"T' = NT" and T:"T = Class C'" from IH[OF sconf wte] T' have wte':"P,E,h' \ e' : NT" by simp from wtval' notemptyCs field wte' sub T show ?thesis by (fastforce intro: WTrtFAss) qed next fix T' assume wtval:"P,E,h \ Val v : NT" and wte:"P,E,h \ e : T'" and sub:"P \ T' \ T" have wtval':"P,E,h' \ Val v : NT" by(rule WTrt_hext_mono[OF wtval red_hext_incr[OF red]]) from IH[OF sconf wte] sub obtain T'' where wte':"P,E,h' \ e' : T''" and sub':"P \ T'' \ T" by (cases T',auto,cases T,auto) from wtval' wte' sub' show ?thesis by(rule WTrtFAssNT) qed thus ?case by(rule wt_same_type_typeconf) next case (RedFAss h a D S Cs' F T Cs v v' Ds fs E l T') let ?fs' = "fs(F \ v')" let ?S' = "insert (Ds, ?fs') (S - {(Ds, fs)})" let ?h' = "h(a \ (D,?S'))" have h:"h a = Some(D,S)" and casts:"P \ T casts v to v'" and field:"P \ last Cs' has least F:T via Cs" and wt:"P,E,h \ ref (a,Cs')\F{Cs} := Val v : T'" by fact+ from wt wf have type:"is_type P T'" by (auto dest:least_field_is_type split:if_split_asm) from wt field obtain T'' where wtval:"P,E,h \ Val v : T''" and eq:"T = T'" and leq:"P \ T'' \ T'" by (auto dest:sees_field_fun split:if_split_asm) from casts eq wtval show ?case proof(induct rule:casts_to.induct) case (casts_prim T\<^sub>0 w) have "T\<^sub>0 = T'" and "\C. T\<^sub>0 \ Class C" and wtval':"P,E,h \ Val w : T''" by fact+ with leq have "T' = T''" by(cases T',auto) with wtval' have "P,E,h \ Val w : T'" by simp with h have "P,E,(h(a\(D,insert(Ds,fs(F \ w))(S-{(Ds,fs)})))) \ Val w : T'" by(cases w,auto split:if_split_asm) thus "P,E,(h(a\(D,insert(Ds,fs(F \ w))(S-{(Ds,fs)})))) \ (Val w) :\<^bsub>NT\<^esub> T'" by(rule wt_same_type_typeconf) next case (casts_null C'') have T':"Class C'' = T'" by fact have "P,E,(h(a\(D,insert(Ds,fs(F \ Null))(S-{(Ds,fs)})))) \ null : NT" by simp with sym[OF T'] show "P,E,(h(a\(D,insert(Ds,fs(F \ Null))(S-{(Ds,fs)})))) \ null :\<^bsub>NT\<^esub> T'" by simp next case (casts_ref Xs C'' Xs' Ds'' a') have "Class C'' = T'" and "Ds'' = Xs @\<^sub>p Xs'" and "P \ Path last Xs to C'' via Xs'" and "P,E,h \ ref (a', Xs) : T''" by fact+ with wf have "P,E,h \ ref (a',Ds'') : T'" by (auto intro!:appendPath_last[THEN sym] Subobjs_nonempty split:if_split_asm simp:path_via_def, drule_tac Cs="Xs" in Subobjs_appendPath,auto) with h have "P,E,(h(a\(D,insert(Ds,fs(F \ Ref(a',Ds'')))(S-{(Ds,fs)})))) \ ref (a',Ds'') : T'" by auto thus "P,E,(h(a\(D,insert(Ds,fs(F \ Ref(a',Ds'')))(S-{(Ds,fs)})))) \ ref (a',Ds'') :\<^bsub>NT\<^esub> T'" by(rule wt_same_type_typeconf) qed next case (RedFAssNull E F Cs v h l) have sconf:"P,E \ (h, l) \" by fact from wf have "is_class P NullPointer" by (fastforce intro:is_class_xcpt wf_prog_wwf_prog) hence "preallocated h \ P \ typeof\<^bsub>h\<^esub> (Ref (addr_of_sys_xcpt NullPointer,[NullPointer])) = Some(Class NullPointer)" by (auto elim: preallocatedE dest!:preallocatedD Subobjs_Base) with sconf have "P,E,h \ THROW NullPointer : T" by(auto simp:sconf_def hconf_def) thus ?case by (fastforce intro:wt_same_type_typeconf wf_prog_wwf_prog) next case (CallObj E e h l e' h' l' Copt M es) have red: "P,E \ \e,(h,l)\ \ \e',(h',l')\" and IH: "\T'. \P,E \ (h,l) \; P,E,h \ e : T'\ \ P,E,h' \ e' :\<^bsub>NT\<^esub> T'" and sconf: "P,E \ (h,l) \" and wt: "P,E,h \ Call e Copt M es : T" by fact+ from wt have "P,E,h' \ Call e' Copt M es : T" proof(cases Copt) case None with wt have "P,E,h \ e\M(es) : T" by simp hence "P,E,h' \ e'\M(es) : T" proof(rule WTrt_elim_cases) fix C Cs Ts Ts' m assume wte:"P,E,h \ e : Class C" and "method":"P \ C has least M = (Ts, T, m) via Cs" and wtes:"P,E,h \ es [:] Ts'" and subs: "P \ Ts' [\] Ts" from IH[OF sconf wte] have "P,E,h' \ e' : NT \ P,E,h' \ e' : Class C" by auto thus ?thesis proof(rule disjE) assume wte':"P,E,h' \ e' : NT" have "P,E,h' \ es [:] Ts'" by(rule WTrts_hext_mono[OF wtes red_hext_incr[OF red]]) with wte' show ?thesis by(rule WTrtCallNT) next assume wte':"P,E,h' \ e' : Class C" have wtes':"P,E,h' \ es [:] Ts'" by(rule WTrts_hext_mono[OF wtes red_hext_incr[OF red]]) from wte' "method" wtes' subs show ?thesis by(rule WTrtCall) qed next fix Ts assume wte:"P,E,h \ e : NT" and wtes:"P,E,h \ es [:] Ts" from IH[OF sconf wte] have wte':"P,E,h' \ e' : NT" by simp have "P,E,h' \ es [:] Ts" by(rule WTrts_hext_mono[OF wtes red_hext_incr[OF red]]) with wte' show ?thesis by(rule WTrtCallNT) qed with None show ?thesis by simp next case (Some C) with wt have "P,E,h \ e\(C::)M(es) : T" by simp hence "P,E,h' \ e'\(C::)M(es) : T" proof(rule WTrt_elim_cases) fix C' Cs Ts Ts' m assume wte:"P,E,h \ e : Class C'" and path_unique:"P \ Path C' to C unique" and "method":"P \ C has least M = (Ts, T, m) via Cs" and wtes:"P,E,h \ es [:] Ts'" and subs: "P \ Ts' [\] Ts" from IH[OF sconf wte] have "P,E,h' \ e' : NT \ P,E,h' \ e' : Class C'" by auto thus ?thesis proof(rule disjE) assume wte':"P,E,h' \ e' : NT" have "P,E,h' \ es [:] Ts'" by(rule WTrts_hext_mono[OF wtes red_hext_incr[OF red]]) with wte' show ?thesis by(rule WTrtCallNT) next assume wte':"P,E,h' \ e' : Class C'" have wtes':"P,E,h' \ es [:] Ts'" by(rule WTrts_hext_mono[OF wtes red_hext_incr[OF red]]) from wte' path_unique "method" wtes' subs show ?thesis by(rule WTrtStaticCall) qed next fix Ts assume wte:"P,E,h \ e : NT" and wtes:"P,E,h \ es [:] Ts" from IH[OF sconf wte] have wte':"P,E,h' \ e' : NT" by simp have "P,E,h' \ es [:] Ts" by(rule WTrts_hext_mono[OF wtes red_hext_incr[OF red]]) with wte' show ?thesis by(rule WTrtCallNT) qed with Some show ?thesis by simp qed thus ?case by (rule wt_same_type_typeconf) next case (CallParams E es h l es' h' l' v Copt M) have reds: "P,E \ \es,(h,l)\ [\] \es',(h',l')\" and IH: "\Ts. \P,E \ (h,l) \; P,E,h \ es [:] Ts\ \ types_conf P E h' es' Ts" and sconf: "P,E \ (h,l) \" and wt: "P,E,h \ Call (Val v) Copt M es : T" by fact+ from wt have "P,E,h' \ Call (Val v) Copt M es' : T" proof(cases Copt) case None with wt have "P,E,h \ (Val v)\M(es) : T" by simp hence "P,E,h' \ Val v\M(es') : T" proof (rule WTrt_elim_cases) fix C Cs Ts Ts' m assume wte: "P,E,h \ Val v : Class C" and "method":"P \ C has least M = (Ts,T,m) via Cs" and wtes: "P,E,h \ es [:] Ts'" and subs:"P \ Ts' [\] Ts" from wtes have "length es = length Ts'" by(rule WTrts_same_length) with reds have "length es' = length Ts'" by -(drule reds_length,simp) with IH[OF sconf wtes] subs obtain Ts'' where wtes':"P,E,h' \ es' [:] Ts''" and subs':"P \ Ts'' [\] Ts" by(auto dest:types_conf_smaller_types) have wte':"P,E,h' \ Val v : Class C" by(rule WTrt_hext_mono[OF wte reds_hext_incr[OF reds]]) from wte' "method" wtes' subs' show ?thesis by(rule WTrtCall) next fix Ts assume wte:"P,E,h \ Val v : NT" and wtes:"P,E,h \ es [:] Ts" from wtes have "length es = length Ts" by(rule WTrts_same_length) with reds have "length es' = length Ts" by -(drule reds_length,simp) with IH[OF sconf wtes] obtain Ts' where wtes':"P,E,h' \ es' [:] Ts'" and "P \ Ts' [\] Ts" by(auto dest:types_conf_smaller_types) have wte':"P,E,h' \ Val v : NT" by(rule WTrt_hext_mono[OF wte reds_hext_incr[OF reds]]) from wte' wtes' show ?thesis by(rule WTrtCallNT) qed with None show ?thesis by simp next case (Some C) with wt have "P,E,h \ (Val v)\(C::)M(es) : T" by simp hence "P,E,h' \ (Val v)\(C::)M(es') : T" proof(rule WTrt_elim_cases) fix C' Cs Ts Ts' m assume wte:"P,E,h \ Val v : Class C'" and path_unique:"P \ Path C' to C unique" and "method":"P \ C has least M = (Ts,T,m) via Cs" and wtes:"P,E,h \ es [:] Ts'" and subs: "P \ Ts' [\] Ts" from wtes have "length es = length Ts'" by(rule WTrts_same_length) with reds have "length es' = length Ts'" by -(drule reds_length,simp) with IH[OF sconf wtes] subs obtain Ts'' where wtes':"P,E,h' \ es' [:] Ts''" and subs':"P \ Ts'' [\] Ts" by(auto dest:types_conf_smaller_types) have wte':"P,E,h' \ Val v : Class C'" by(rule WTrt_hext_mono[OF wte reds_hext_incr[OF reds]]) from wte' path_unique "method" wtes' subs' show ?thesis by(rule WTrtStaticCall) next fix Ts assume wte:"P,E,h \ Val v : NT" and wtes:"P,E,h \ es [:] Ts" from wtes have "length es = length Ts" by(rule WTrts_same_length) with reds have "length es' = length Ts" by -(drule reds_length,simp) with IH[OF sconf wtes] obtain Ts' where wtes':"P,E,h' \ es' [:] Ts'" and "P \ Ts' [\] Ts" by(auto dest:types_conf_smaller_types) have wte':"P,E,h' \ Val v : NT" by(rule WTrt_hext_mono[OF wte reds_hext_incr[OF reds]]) from wte' wtes' show ?thesis by(rule WTrtCallNT) qed with Some show ?thesis by simp qed thus ?case by (rule wt_same_type_typeconf) next case (RedCall h l a C S Cs M Ts' T' pns' body' Ds Ts T pns body Cs' vs bs new_body E T'') have hp:"hp (h,l) a = Some(C,S)" and "method":"P \ last Cs has least M = (Ts',T',pns',body') via Ds" and select:"P \ (C,Cs@\<^sub>pDs) selects M = (Ts,T,pns,body) via Cs'" and length1:"length vs = length pns" and length2:"length Ts = length pns" and bs:"bs = blocks(this#pns,Class(last Cs')#Ts,Ref(a,Cs')#vs,body)" and body_case:"new_body = (case T' of Class D \ \D\bs | _ \ bs)" and wt:"P,E,h \ ref (a,Cs)\M(map Val vs) : T''" by fact+ from wt hp "method" wf obtain Ts'' where wtref:"P,E,h \ ref (a,Cs) : Class (last Cs)" and eq:"T'' = T'" and wtes:"P,E,h \ map Val vs [:] Ts''" and subs: "P \ Ts'' [\] Ts'" by(auto dest:wf_sees_method_fun split:if_split_asm) from select wf have "is_class P (last Cs')" by(induct rule:SelectMethodDef.induct, auto intro:Subobj_last_isClass simp:FinalOverriderMethodDef_def OverriderMethodDefs_def MinimalMethodDefs_def LeastMethodDef_def MethodDefs_def) with select_method_wf_mdecl[OF wf select] have length_pns:"length (this#pns) = length (Class(last Cs')#Ts)" and notNT:"T \ NT" and type:"\T\set (Class(last Cs')#Ts). is_type P T" and wtabody:"P,[this\Class(last Cs'),pns[\]Ts] \ body :: T" by(auto simp:wf_mdecl_def) from wtes hp select have map:"map (P \ typeof\<^bsub>h\<^esub>) (Ref(a,Cs')#vs) = map Some (Class(last Cs')#Ts'')" by(auto elim:SelectMethodDef.cases split:if_split_asm simp:FinalOverriderMethodDef_def OverriderMethodDefs_def MinimalMethodDefs_def LeastMethodDef_def MethodDefs_def) from wtref hp have "P \ Path C to (last Cs) via Cs" by (auto simp:path_via_def split:if_split_asm) with select "method" wf have "Ts' = Ts \ P \ T \ T'" by -(rule select_least_methods_subtypes,simp_all) hence eqs:"Ts' = Ts" and sub:"P \ T \ T'" by auto from wf wtabody have "P,Map.empty(this\Class(last Cs'),pns[\]Ts),h \ body : T" by -(rule WT_implies_WTrt,simp_all) hence wtbody:"P,E(this#pns [\] Class (last Cs')#Ts),h \ body : T" by(rule WTrt_env_mono) simp from wtes have "length vs = length Ts''" by (fastforce dest:WTrts_same_length) with eqs subs have length_vs:"length (Ref(a,Cs')#vs) = length (Class(last Cs')#Ts)" by (simp add:list_all2_iff) from subs eqs have "P \ (Class(last Cs')#Ts'') [\] (Class(last Cs')#Ts)" by (simp add:fun_of_def) with wt_blocks[OF length_pns length_vs type] wtbody map eq have blocks:"P,E,h \ blocks(this#pns,Class(last Cs')#Ts,Ref(a,Cs')#vs,body) : T" by auto have "P,E,h \ new_body : T'" proof(cases "\C. T' \ Class C") case True with sub notNT have "T = T'" by (cases T') auto with blocks True body_case bs show ?thesis by(cases T') auto next case False then obtain D where T':"T' = Class D" by auto with "method" sub wf have "class": "is_class P D" by (auto elim!:widen.cases dest:least_method_is_type intro:Subobj_last_isClass simp:path_unique_def) with blocks T' body_case bs "class" sub show ?thesis by(cases T',auto,cases T,auto) qed with eq show ?case by(fastforce intro:wt_same_type_typeconf) next case (RedStaticCall Cs C Cs'' M Ts T pns body Cs' Ds vs E a h l T') have "method":"P \ C has least M = (Ts, T, pns, body) via Cs'" and length1:"length vs = length pns" and length2:"length Ts = length pns" and path_unique:"P \ Path last Cs to C unique" and path_via:"P \ Path last Cs to C via Cs''" and Ds:"Ds = (Cs @\<^sub>p Cs'') @\<^sub>p Cs'" and wt:"P,E,h \ ref (a,Cs)\(C::)M(map Val vs) : T'" by fact+ from wt "method" wf obtain Ts' where wtref:"P,E,h \ ref (a,Cs) : Class (last Cs)" and wtes:"P,E,h \ map Val vs [:] Ts'" and subs:"P \ Ts' [\] Ts" and TeqT':"T = T'" by(auto dest:wf_sees_method_fun split:if_split_asm) from wtref obtain D S where hp:"h a = Some(D,S)" and subo:"Subobjs P D Cs" by (auto split:if_split_asm) from length1 length2 have length_vs: "length (Ref(a,Ds)#vs) = length (Class (last Ds)#Ts)" by simp from length2 have length_pns:"length (this#pns) = length (Class (last Ds)#Ts)" by simp from "method" have "Cs' \ []" by (fastforce intro!:Subobjs_nonempty simp add:LeastMethodDef_def MethodDefs_def) with Ds have last:"last Cs' = last Ds" by (fastforce dest:appendPath_last) with "method" have "is_class P (last Ds)" by(auto simp:LeastMethodDef_def MethodDefs_def is_class_def) with last has_least_wf_mdecl[OF wf "method"] have wtabody: "P,[this#pns [\] Class (last Ds)#Ts] \ body :: T" and type:"\T\set (Class(last Ds)#Ts). is_type P T" by(auto simp:wf_mdecl_def) from path_via have suboCs'':"Subobjs P (last Cs) Cs''" and lastCs'':"last Cs'' = C" by (auto simp add:path_via_def) with subo wf have subo':"Subobjs P D (Cs@\<^sub>pCs'')" by(fastforce intro: Subobjs_appendPath) from lastCs'' suboCs'' have lastC:"C = last(Cs@\<^sub>pCs'')" by (fastforce dest:Subobjs_nonempty intro:appendPath_last) from "method" have "Subobjs P C Cs'" by (auto simp:LeastMethodDef_def MethodDefs_def) with subo' wf lastC have "Subobjs P D ((Cs @\<^sub>p Cs'') @\<^sub>p Cs')" by (fastforce intro:Subobjs_appendPath) with Ds have suboDs:"Subobjs P D Ds" by simp from wtabody have "P,Map.empty(this#pns [\] Class (last Ds)#Ts),h \ body : T" by(rule WT_implies_WTrt) hence "P,E(this#pns [\] Class (last Ds)#Ts),h \ body : T" by(rule WTrt_env_mono) simp hence "P,E,h \ blocks(this#pns, Class (last Ds)#Ts, Ref(a,Ds)#vs, body) : T" using wtes subs wt_blocks[OF length_pns length_vs type] hp suboDs by(auto simp add:rel_list_all2_Cons2) with TeqT' show ?case by(fastforce intro:wt_same_type_typeconf) next case (RedCallNull E Copt M vs h l) have sconf:"P,E \ (h, l) \" by fact from wf have "is_class P NullPointer" by (fastforce intro:is_class_xcpt wf_prog_wwf_prog) hence "preallocated h \ P \ typeof\<^bsub>h\<^esub> (Ref (addr_of_sys_xcpt NullPointer,[NullPointer])) = Some(Class NullPointer)" by (auto elim: preallocatedE dest!:preallocatedD Subobjs_Base) with sconf have "P,E,h \ THROW NullPointer : T" by(auto simp:sconf_def hconf_def) thus ?case by (fastforce intro:wt_same_type_typeconf) next case (BlockRedNone E V T e h l e' h' l' T') have IH:"\T'. \P,E(V \ T) \ (h, l(V := None)) \; P,E(V \ T),h \ e : T'\ \ P,E(V \ T),h' \ e' :\<^bsub>NT\<^esub> T'" and sconf:"P,E \ (h, l) \" and wt:"P,E,h \ {V:T; e} : T'" by fact+ from wt have type:"is_type P T" and wte:"P,E(V\T),h \ e : T'" by auto from sconf type have "P,E(V \ T) \ (h, l(V := None)) \" by (auto simp:sconf_def lconf_def envconf_def) from IH[OF this wte] type show ?case by (cases T') auto next case (BlockRedSome E V T e h l e' h' l' v T') have red:"P,E(V \ T) \ \e,(h, l(V := None))\ \ \e',(h', l')\" and IH:"\T'. \P,E(V \ T) \ (h, l(V := None)) \; P,E(V \ T),h \ e : T'\ \ P,E(V \ T),h' \ e' :\<^bsub>NT\<^esub> T'" and Some:"l' V = Some v" and sconf:"P,E \ (h, l) \" and wt:"P,E,h \ {V:T; e} : T'" by fact+ from wt have wte:"P,E(V\T),h \ e : T'" and type:"is_type P T" by auto with sconf wf red type have "P,h' \ l' (:\)\<^sub>w E(V \ T)" by -(auto simp:sconf_def,rule red_preserves_lconf, auto intro:wf_prog_wwf_prog simp:envconf_def lconf_def) hence conf:"P,h' \ v :\ T" using Some by(auto simp:lconf_def,erule_tac x="V" in allE,clarsimp) have wtval:"P,E(V \ T),h' \ V:=Val v : T" proof(cases T) case Void with conf show ?thesis by auto next case Boolean with conf show ?thesis by auto next case Integer with conf show ?thesis by auto next case NT with conf show ?thesis by auto next case (Class C) with conf have "P,E(V \ T),h' \ Val v : T \ P,E(V \ T),h' \ Val v : NT" by auto with Class show ?thesis by auto qed from sconf type have "P,E(V \ T) \ (h, l(V := None)) \" by (auto simp:sconf_def lconf_def envconf_def) from IH[OF this wte] wtval type show ?case by(cases T') auto next case (InitBlockRed E V T e h l v' e' h' l' v'' v T') have red:"P,E(V \ T) \ \e,(h, l(V \ v'))\ \ \e',(h', l')\" and IH:"\T'. \P,E(V \ T) \ (h, l(V \ v')) \; P,E(V \ T),h \ e : T'\ \ P,E(V \ T),h' \ e' :\<^bsub>NT\<^esub> T'" and Some:"l' V = Some v''" and casts:"P \ T casts v to v'" and sconf:"P,E \ (h, l) \" and wt:"P,E,h \ {V:T := Val v; e} : T'" by fact+ from wt have wte:"P,E(V \ T),h \ e : T'" and wtval:"P,E(V \ T),h \ V:=Val v : T" and type:"is_type P T" by auto from wf casts wtval have "P,h \ v' :\ T" by(fastforce intro!:casts_conf wf_prog_wwf_prog) with sconf have lconf:"P,h \ l(V \ v') (:\)\<^sub>w E(V \ T)" by (fastforce intro!:lconf_upd2 simp:sconf_def) from sconf type have "envconf P (E(V \ T))" by(simp add:sconf_def envconf_def) from red_preserves_lconf[OF wf_prog_wwf_prog[OF wf] red wte lconf this] have "P,h' \ l' (:\)\<^sub>w E(V \ T)" . with Some have "P,h' \ v'' :\ T" by(simp add:lconf_def,erule_tac x="V" in allE,auto) hence wtval':"P,E(V \ T),h' \ V:=Val v'' : T" by(cases T) auto from lconf sconf type have "P,E(V \ T) \ (h, l(V \ v')) \" by(auto simp:sconf_def envconf_def) from IH[OF this wte] wtval' type show ?case by(cases T') auto next case RedBlock thus ?case by (fastforce intro:wt_same_type_typeconf) next case RedInitBlock thus ?case by (fastforce intro:wt_same_type_typeconf) next case (SeqRed E e h l e' h' l' e\<^sub>2 T) have red:"P,E \ \e,(h, l)\ \ \e',(h', l')\" and IH:"\T'. \P,E \ (h, l) \; P,E,h \ e : T'\ \ P,E,h' \ e' :\<^bsub>NT\<^esub> T'" and sconf:"P,E \ (h, l) \" and wt:"P,E,h \ e;; e\<^sub>2 : T" by fact+ from wt obtain T' where wte:"P,E,h \ e : T'" and wte2:"P,E,h \ e\<^sub>2 : T" by auto from WTrt_hext_mono[OF wte2 red_hext_incr[OF red]] have wte2':"P,E,h' \ e\<^sub>2 : T" . from IH[OF sconf wte] obtain T'' where "P,E,h' \ e' : T''" by(cases T') auto with wte2' have "P,E,h' \ e';; e\<^sub>2 : T" by auto thus ?case by(rule wt_same_type_typeconf) next case RedSeq thus ?case by (fastforce intro:wt_same_type_typeconf) next case (CondRed E e h l e' h' l' e\<^sub>1 e\<^sub>2) have red:"P,E \ \e,(h, l)\ \ \e',(h', l')\" and IH: "\T. \P,E \ (h,l) \; P,E,h \ e : T\ \ P,E,h' \ e' :\<^bsub>NT\<^esub> T" and wt:"P,E,h \ if (e) e\<^sub>1 else e\<^sub>2 : T" and sconf:"P,E \ (h,l) \" by fact+ from wt have wte:"P,E,h \ e : Boolean" and wte1:"P,E,h \ e\<^sub>1 : T" and wte2:"P,E,h \ e\<^sub>2 : T" by auto from IH[OF sconf wte] have wte':"P,E,h' \ e' : Boolean" by auto from wte' WTrt_hext_mono[OF wte1 red_hext_incr[OF red]] WTrt_hext_mono[OF wte2 red_hext_incr[OF red]] have "P,E,h' \ if (e') e\<^sub>1 else e\<^sub>2 : T" by (rule WTrtCond) thus ?case by(rule wt_same_type_typeconf) next case RedCondT thus ?case by (fastforce intro: wt_same_type_typeconf) next case RedCondF thus ?case by (fastforce intro: wt_same_type_typeconf) next case RedWhile thus ?case by (fastforce intro: wt_same_type_typeconf) next case (ThrowRed E e h l e' h' l' T) have IH:"\T. \P,E \ (h, l) \; P,E,h \ e : T\ \ P,E,h' \ e' :\<^bsub>NT\<^esub> T" and sconf:"P,E \ (h, l) \" and wt:"P,E,h \ throw e : T" by fact+ from wt obtain T' where wte:"P,E,h \ e : T'" and ref:"is_refT T'" by auto from ref have "P,E,h' \ throw e' : T" proof(rule refTE) assume T':"T' = NT" with wte have "P,E,h \ e : NT" by simp from IH[OF sconf this] ref T' show ?thesis by auto next fix C assume T':"T' = Class C" with wte have "P,E,h \ e : Class C" by simp from IH[OF sconf this] have "P,E,h' \ e' : Class C \ P,E,h' \ e' : NT" by simp thus ?thesis proof(rule disjE) assume wte':"P,E,h' \ e' : Class C" have "is_refT (Class C)" by simp with wte' show ?thesis by auto next assume wte':"P,E,h' \ e' : NT" have "is_refT NT" by simp with wte' show ?thesis by auto qed qed thus ?case by (rule wt_same_type_typeconf) next case (RedThrowNull E h l) have sconf:"P,E \ (h, l) \" by fact from wf have "is_class P NullPointer" by (fastforce intro:is_class_xcpt wf_prog_wwf_prog) hence "preallocated h \ P \ typeof\<^bsub>h\<^esub> (Ref (addr_of_sys_xcpt NullPointer,[NullPointer])) = Some(Class NullPointer)" by (auto elim: preallocatedE dest!:preallocatedD Subobjs_Base) with sconf have "P,E,h \ THROW NullPointer : T" by(auto simp:sconf_def hconf_def) thus ?case by (fastforce intro:wt_same_type_typeconf wf_prog_wwf_prog) next case (ListRed1 E e h l e' h' l' es Ts) have red:"P,E \ \e,(h, l)\ \ \e',(h', l')\" and IH:"\T. \P,E \ (h, l) \; P,E,h \ e : T\ \ P,E,h' \ e' :\<^bsub>NT\<^esub> T" and sconf:"P,E \ (h, l) \" and wt:"P,E,h \ e # es [:] Ts" by fact+ from wt obtain U Us where Ts:"Ts = U#Us" by(cases Ts) auto with wt have wte:"P,E,h \ e : U" and wtes:"P,E,h \ es [:] Us" by simp_all from WTrts_hext_mono[OF wtes red_hext_incr[OF red]] have wtes':"P,E,h' \ es [:] Us" . hence "length es = length Us" by (rule WTrts_same_length) with wtes' have "types_conf P E h' es Us" by (fastforce intro:wts_same_types_typesconf) with IH[OF sconf wte] Ts show ?case by simp next case (ListRed2 E es h l es' h' l' v Ts) have reds:"P,E \ \es,(h, l)\ [\] \es',(h', l')\" and IH:"\Ts. \P,E \ (h, l) \; P,E,h \ es [:] Ts\ \ types_conf P E h' es' Ts" and sconf:"P,E \ (h, l) \" and wt:"P,E,h \ Val v#es [:] Ts" by fact+ from wt obtain U Us where Ts:"Ts = U#Us" by(cases Ts) auto with wt have wtval:"P,E,h \ Val v : U" and wtes:"P,E,h \ es [:] Us" by simp_all from WTrt_hext_mono[OF wtval reds_hext_incr[OF reds]] have "P,E,h' \ Val v : U" . hence "P,E,h' \ (Val v) :\<^bsub>NT\<^esub> U" by(rule wt_same_type_typeconf) with IH[OF sconf wtes] Ts show ?case by simp next case (CallThrowObj E h l Copt M es h' l') thus ?case by(cases Copt)(auto intro:wt_same_type_typeconf) next case (CallThrowParams es vs h l es' E v Copt M h' l') thus ?case by(cases Copt)(auto intro:wt_same_type_typeconf) qed (fastforce intro:wt_same_type_typeconf)+ corollary subject_reduction: "\ wf_C_prog P; P,E \ \e,s\ \ \e',s'\; P,E \ s \; P,E,hp s \ e:T \ \ P,E,(hp s') \ e' :\<^bsub>NT\<^esub> T" by(cases s, cases s', fastforce dest:subject_reduction2) corollary subjects_reduction: "\ wf_C_prog P; P,E \ \es,s\ [\] \es',s'\; P,E \ s \; P,E,hp s \ es[:]Ts \ \ types_conf P E (hp s') es' Ts" by(cases s, cases s', fastforce dest:subjects_reduction2) subsection \Lifting to \\*\\ text\Now all these preservation lemmas are first lifted to the transitive closure \dots\ lemma step_preserves_sconf: assumes wf: "wf_C_prog P" and step: "P,E \ \e,s\ \* \e',s'\" shows "\T. \ P,E,hp s \ e : T; P,E \ s \ \ \ P,E \ s' \" using step proof (induct rule:converse_rtrancl_induct2) case refl show ?case by fact next case step thus ?case using wf apply simp apply (frule subject_reduction[OF wf]) apply (rule step.prems) apply (rule step.prems) apply (cases T) apply (auto dest:red_preserves_sconf intro:wf_prog_wwf_prog) done qed lemma steps_preserves_sconf: assumes wf: "wf_C_prog P" and step: "P,E \ \es,s\ [\]* \es',s'\" shows "\Ts. \ P,E,hp s \ es [:] Ts; P,E \ s \ \ \ P,E \ s' \" using step proof (induct rule:converse_rtrancl_induct2) case refl show ?case by fact next case (step es s es'' s'' Ts) have Reds:"((es, s), es'', s'') \ Reds P E" and reds:"P,E \ \es'',s''\ [\]* \es',s'\" and wtes:"P,E,hp s \ es [:] Ts" and sconf:"P,E \ s \" and IH:"\Ts. \P,E,hp s'' \ es'' [:] Ts; P,E \ s'' \\ \ P,E \ s' \" by fact+ from Reds have reds1:"P,E \ \es,s\ [\] \es'',s''\" by simp from subjects_reduction[OF wf this sconf wtes] have type:"types_conf P E (hp s'') es'' Ts" . from reds1 wtes sconf wf have sconf':"P,E \ s'' \" by(fastforce intro:wf_prog_wwf_prog reds_preserves_sconf) from type have "\Ts'. P,E,hp s'' \ es'' [:] Ts'" proof (induct Ts arbitrary: es'') fix esi assume "types_conf P E (hp s'') esi []" thus "\Ts'. P,E,hp s'' \ esi [:] Ts'" proof(induct esi) case Nil thus "\Ts'. P,E,hp s'' \ [] [:] Ts'" by simp next fix ex esx assume "types_conf P E (hp s'') (ex#esx) []" thus "\Ts'. P,E,hp s'' \ ex#esx [:] Ts'" by simp qed next fix T' Ts' esi assume type':"types_conf P E (hp s'') esi (T'#Ts')" and IH:"\es''. types_conf P E (hp s'') es'' Ts' \ \Ts''. P,E,hp s'' \ es'' [:] Ts''" from type' show "\Ts'. P,E,hp s'' \ esi [:] Ts'" proof(induct esi) case Nil thus "\Ts'. P,E,hp s'' \ [] [:] Ts'" by simp next fix ex esx assume "types_conf P E (hp s'') (ex#esx) (T'#Ts')" hence type':"P,E,hp s'' \ ex :\<^bsub>NT\<^esub> T'" and types':"types_conf P E (hp s'') esx Ts'" by simp_all from type' obtain Tx where type'':"P,E,hp s'' \ ex : Tx" by(cases T') auto from IH[OF types'] obtain Tsx where "P,E,hp s'' \ esx [:] Tsx" by auto with type'' show "\Ts'. P,E,hp s'' \ ex#esx [:] Ts'" by auto qed qed then obtain Ts' where "P,E,hp s'' \ es'' [:] Ts'" by blast from IH[OF this sconf'] show ?case . qed lemma step_preserves_defass: assumes wf: "wf_C_prog P" and step: "P,E \ \e,s\ \* \e',s'\" shows "\ e \dom(lcl s)\ \ \ e' \dom(lcl s')\" using step proof (induct rule:converse_rtrancl_induct2) case refl thus ?case . next case (step e s e' s') thus ?case by(cases s,cases s')(auto dest:red_preserves_defass[OF wf]) qed lemma step_preserves_type: assumes wf: "wf_C_prog P" and step: "P,E \ \e,s\ \* \e',s'\" shows "\T. \ P,E \ s \; P,E,hp s \ e:T \ \ P,E,(hp s') \ e' :\<^bsub>NT\<^esub> T" using step proof (induct rule:converse_rtrancl_induct2) case refl thus ?case by -(rule wt_same_type_typeconf) next case (step e s e'' s'' T) thus ?case using wf apply simp apply (frule subject_reduction[OF wf]) apply (auto dest!:red_preserves_sconf intro:wf_prog_wwf_prog) apply(cases T) apply fastforce+ done qed text\predicate to show the same lemma for lists\ fun conformable :: "ty list \ ty list \ bool" where "conformable [] [] \ True" | "conformable (T''#Ts'') (T'#Ts') \ (T'' = T' \ (\C. T'' = NT \ T' = Class C)) \ conformable Ts'' Ts'" | "conformable _ _ \ False" lemma types_conf_conf_types_conf: "\types_conf P E h es Ts; conformable Ts Ts'\ \ types_conf P E h es Ts'" proof (induct Ts arbitrary: Ts' es) case Nil thus ?case by (cases Ts') (auto split: if_split_asm) next case (Cons T'' Ts'') have type:"types_conf P E h es (T''#Ts'')" and conf:"conformable (T''#Ts'') Ts'" and IH:"\Ts' es. \types_conf P E h es Ts''; conformable Ts'' Ts'\ \ types_conf P E h es Ts'" by fact+ from type obtain e' es' where es:"es = e'#es'" by (cases es) auto with type have type':"P,E,h \ e' :\<^bsub>NT\<^esub> T''" and types': "types_conf P E h es' Ts''" by simp_all from conf obtain U Us where Ts': "Ts' = U#Us" by (cases Ts') auto with conf have disj:"T'' = U \ (\C. T'' = NT \ U = Class C)" and conf':"conformable Ts'' Us" by simp_all from type' disj have "P,E,h \ e' :\<^bsub>NT\<^esub> U" by auto with IH[OF types' conf'] Ts' es show ?case by simp qed lemma types_conf_Wtrt_conf: "types_conf P E h es Ts \ \Ts'. P,E,h \ es [:] Ts' \ conformable Ts' Ts" proof (induct Ts arbitrary: es) case Nil thus ?case by (cases es) (auto split:if_split_asm) next case (Cons T'' Ts'') have type:"types_conf P E h es (T''#Ts'')" and IH:"\es. types_conf P E h es Ts'' \ \Ts'. P,E,h \ es [:] Ts' \ conformable Ts' Ts''" by fact+ from type obtain e' es' where es:"es = e'#es'" by (cases es) auto with type have type':"P,E,h \ e' :\<^bsub>NT\<^esub> T''" and types': "types_conf P E h es' Ts''" by simp_all from type' obtain T' where "P,E,h \ e' : T'" and "T' = T'' \ (\C. T' = NT \ T'' = Class C)" by(cases T'') auto with IH[OF types'] es show ?case by(auto,rule_tac x="T''#Ts'" in exI,simp,rule_tac x="NT#Ts'" in exI,simp) qed lemma steps_preserves_types: assumes wf: "wf_C_prog P" and steps: "P,E \ \es,s\ [\]* \es',s'\" shows "\Ts. \ P,E \ s \; P,E,hp s \ es [:] Ts\ \ types_conf P E (hp s') es' Ts" using steps proof (induct rule:converse_rtrancl_induct2) case refl thus ?case by -(rule wts_same_types_typesconf) next case (step es s es'' s'' Ts) have Reds:"((es, s), es'', s'') \ Reds P E" and steps:"P,E \ \es'',s''\ [\]* \es',s'\" and sconf:"P,E \ s \" and wtes:"P,E,hp s \ es [:] Ts" and IH:"\Ts. \P,E \ s'' \; P,E,hp s'' \ es'' [:] Ts \ \ types_conf P E (hp s') es' Ts" by fact+ from Reds have step:"P,E \ \es,s\ [\] \es'',s''\" by simp with wtes sconf wf have sconf':"P,E \ s'' \" by(auto intro:reds_preserves_sconf wf_prog_wwf_prog) from wtes have "length es = length Ts" by(fastforce dest:WTrts_same_length) from step sconf wtes have type': "types_conf P E (hp s'') es'' Ts" by (rule subjects_reduction[OF wf]) then obtain Ts' where wtes'':"P,E,hp s'' \ es'' [:] Ts'" and conf:"conformable Ts' Ts" by (auto dest:types_conf_Wtrt_conf) from IH[OF sconf' wtes''] have "types_conf P E (hp s') es' Ts'" . with conf show ?case by(fastforce intro:types_conf_conf_types_conf) qed subsection \Lifting to \\\\ text\\dots and now to the big step semantics, just for fun.\ lemma eval_preserves_sconf: "\ wf_C_prog P; P,E \ \e,s\ \ \e',s'\; P,E \ e::T; P,E \ s \ \ \ P,E \ s' \" by(blast intro:step_preserves_sconf big_by_small WT_implies_WTrt wf_prog_wwf_prog) lemma evals_preserves_sconf: "\ wf_C_prog P; P,E \ \es,s\ [\] \es',s'\; P,E \ es [::] Ts; P,E \ s \ \ \ P,E \ s' \" by(blast intro:steps_preserves_sconf bigs_by_smalls WTs_implies_WTrts wf_prog_wwf_prog) lemma eval_preserves_type: assumes wf: "wf_C_prog P" shows "\ P,E \ \e,s\ \ \e',s'\; P,E \ s \; P,E \ e::T \ \ P,E,(hp s') \ e' :\<^bsub>NT\<^esub> T" using wf by (auto dest!:big_by_small[OF wf_prog_wwf_prog[OF wf]] WT_implies_WTrt intro:wf_prog_wwf_prog dest!:step_preserves_type[OF wf]) lemma evals_preserves_types: assumes wf: "wf_C_prog P" shows "\ P,E \ \es,s\ [\] \es',s'\; P,E \ s \; P,E \ es [::] Ts \ \ types_conf P E (hp s') es' Ts" using wf by (auto dest!:bigs_by_smalls[OF wf_prog_wwf_prog[OF wf]] WTs_implies_WTrts intro:wf_prog_wwf_prog dest!:steps_preserves_types[OF wf]) subsection \The final polish\ text\The above preservation lemmas are now combined and packed nicely.\ definition wf_config :: "prog \ env \ state \ expr \ ty \ bool" ("_,_,_ \ _ : _ \" [51,0,0,0,0]50) where "P,E,s \ e:T \ \ P,E \ s \ \ P,E,hp s \ e : T" theorem Subject_reduction: assumes wf: "wf_C_prog P" shows "P,E \ \e,s\ \ \e',s'\ \ P,E,s \ e : T \ \ P,E,(hp s') \ e' :\<^bsub>NT\<^esub> T" using wf by (force elim!:red_preserves_sconf intro:wf_prog_wwf_prog dest:subject_reduction[OF wf] simp:wf_config_def) theorem Subject_reductions: assumes wf: "wf_C_prog P" and reds: "P,E \ \e,s\ \* \e',s'\" shows "\T. P,E,s \ e : T \ \ P,E,(hp s') \ e' :\<^bsub>NT\<^esub> T" using reds proof (induct rule:converse_rtrancl_induct2) case refl thus ?case by (fastforce intro:wt_same_type_typeconf simp:wf_config_def) next case (step e s e'' s'' T) have Red:"((e, s), e'', s'') \ Red P E" and IH:"\T. P,E,s'' \ e'' : T \ \ P,E,(hp s') \ e' :\<^bsub>NT\<^esub> T" and wte:"P,E,s \ e : T \" by fact+ from Red have red:"P,E \ \e,s\ \ \e'',s''\" by simp from red_preserves_sconf[OF red] wte wf have sconf:"P,E \ s'' \" by(fastforce dest:wf_prog_wwf_prog simp:wf_config_def) from wf red wte have type_conf:"P,E,(hp s'') \ e'' :\<^bsub>NT\<^esub> T" by(rule Subject_reduction) show ?case proof(cases T) case Void with type_conf have "P,E,hp s'' \ e'' : T" by simp with sconf have "P,E,s'' \ e'' : T \" by(simp add:wf_config_def) from IH[OF this] show ?thesis . next case Boolean with type_conf have "P,E,hp s'' \ e'' : T" by simp with sconf have "P,E,s'' \ e'' : T \" by(simp add:wf_config_def) from IH[OF this] show ?thesis . next case Integer with type_conf have "P,E,hp s'' \ e'' : T" by simp with sconf have "P,E,s'' \ e'' : T \" by(simp add:wf_config_def) from IH[OF this] show ?thesis . next case NT with type_conf have "P,E,hp s'' \ e'' : T" by simp with sconf have "P,E,s'' \ e'' : T \" by(simp add:wf_config_def) from IH[OF this] show ?thesis . next case (Class C) with type_conf have "P,E,hp s'' \ e'' : T \ P,E,hp s'' \ e'' : NT" by simp thus ?thesis proof(rule disjE) assume "P,E,hp s'' \ e'' : T" with sconf have "P,E,s'' \ e'' : T \" by(simp add:wf_config_def) from IH[OF this] show ?thesis . next assume "P,E,hp s'' \ e'' : NT" with sconf have "P,E,s'' \ e'' : NT \" by(simp add:wf_config_def) from IH[OF this] have "P,E,hp s' \ e' : NT" by simp with Class show ?thesis by simp qed qed qed corollary Progress: assumes wf: "wf_C_prog P" shows "\ P,E,s \ e : T \; \ e \dom(lcl s)\; \ final e \ \ \e' s'. P,E \ \e,s\ \ \e',s'\" using progress[OF wf_prog_wwf_prog[OF wf]] by(auto simp:wf_config_def sconf_def) corollary TypeSafety: fixes s s' :: state assumes wf:"wf_C_prog P" and sconf:"P,E \ s \" and wte:"P,E \ e :: T" and D:"\ e \dom(lcl s)\" and step:"P,E \ \e,s\ \* \e',s'\" and nored:"\(\e'' s''. P,E \ \e',s'\ \ \e'',s''\)" shows "(\v. e' = Val v \ P,hp s' \ v :\ T) \ (\r. e' = Throw r \ the_addr (Ref r) \ dom(hp s'))" proof - from sconf wte wf have wf_config:"P,E,s \ e : T \" by(fastforce intro:WT_implies_WTrt simp:wf_config_def) with wf step have type_conf:"P,E,(hp s') \ e' :\<^bsub>NT\<^esub> T" by(rule Subject_reductions) from step_preserves_sconf[OF wf step wte[THEN WT_implies_WTrt] sconf] wf have sconf':"P,E \ s' \" by simp from wf step D have D':"\ e' \dom(lcl s')\" by(rule step_preserves_defass) show ?thesis proof(cases T) case Void with type_conf have wte':"P,E,hp s' \ e' : T" by simp with sconf' have wf_config':"P,E,s' \ e' : T \" by(simp add:wf_config_def) { assume "\ final e'" from Progress[OF wf wf_config' D' this] nored have False by simp } hence "final e'" by fast with wte' show ?thesis by(auto simp:final_def) next case Boolean with type_conf have wte':"P,E,hp s' \ e' : T" by simp with sconf' have wf_config':"P,E,s' \ e' : T \" by(simp add:wf_config_def) { assume "\ final e'" from Progress[OF wf wf_config' D' this] nored have False by simp } hence "final e'" by fast with wte' show ?thesis by(auto simp:final_def) next case Integer with type_conf have wte':"P,E,hp s' \ e' : T" by simp with sconf' have wf_config':"P,E,s' \ e' : T \" by(simp add:wf_config_def) { assume "\ final e'" from Progress[OF wf wf_config' D' this] nored have False by simp } hence "final e'" by fast with wte' show ?thesis by(auto simp:final_def) next case NT with type_conf have wte':"P,E,hp s' \ e' : T" by simp with sconf' have wf_config':"P,E,s' \ e' : T \" by(simp add:wf_config_def) { assume "\ final e'" from Progress[OF wf wf_config' D' this] nored have False by simp } hence "final e'" by fast with wte' show ?thesis by(auto simp:final_def) next case (Class C) with type_conf have wte':"P,E,hp s' \ e' : T \ P,E,hp s' \ e' : NT" by simp thus ?thesis proof(rule disjE) assume wte':"P,E,hp s' \ e' : T" with sconf' have wf_config':"P,E,s' \ e' : T \" by(simp add:wf_config_def) { assume "\ final e'" from Progress[OF wf wf_config' D' this] nored have False by simp } hence "final e'" by fast with wte' show ?thesis by(auto simp:final_def) next assume wte':"P,E,hp s' \ e' : NT" with sconf' have wf_config':"P,E,s' \ e' : NT \" by(simp add:wf_config_def) { assume "\ final e'" from Progress[OF wf wf_config' D' this] nored have False by simp } hence "final e'" by fast with wte' Class show ?thesis by(auto simp:final_def) qed qed qed - - end diff --git a/thys/Featherweight_OCL/examples/Employee_Model/Analysis/Analysis_UML.thy b/thys/Featherweight_OCL/examples/Employee_Model/Analysis/Analysis_UML.thy --- a/thys/Featherweight_OCL/examples/Employee_Model/Analysis/Analysis_UML.thy +++ b/thys/Featherweight_OCL/examples/Employee_Model/Analysis/Analysis_UML.thy @@ -1,1384 +1,1384 @@ (***************************************************************************** * Featherweight-OCL --- A Formal Semantics for UML-OCL Version OCL 2.5 * for the OMG Standard. * http://www.brucker.ch/projects/hol-testgen/ * * Analysis_UML.thy --- OCL Contracts and an Example. * This file is part of HOL-TestGen. * * Copyright (c) 2012-2015 Université Paris-Saclay, Univ. Paris-Sud, France * 2013-2015 IRT SystemX, France * * All rights reserved. * * Redistribution and use in source and binary forms, with or without * modification, are permitted provided that the following conditions are * met: * * * Redistributions of source code must retain the above copyright * notice, this list of conditions and the following disclaimer. * * * Redistributions in binary form must reproduce the above * copyright notice, this list of conditions and the following * disclaimer in the documentation and/or other materials provided * with the distribution. * * * Neither the name of the copyright holders nor the names of its * contributors may be used to endorse or promote products derived * from this software without specific prior written permission. * * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS * "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT * LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR * A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT * OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, * SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT * LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, * DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY * THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT * (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. ******************************************************************************) chapter\Example: The Employee Analysis Model\ (* UML part *) theory Analysis_UML imports "../../../UML_Main" begin text \\label{ex:employee-analysis:uml}\ section\Introduction\ text\ For certain concepts like classes and class-types, only a generic definition for its resulting semantics can be given. Generic means, there is a function outside HOL that ``compiles'' a concrete, closed-world class diagram into a ``theory'' of this data model, consisting of a bunch of definitions for classes, accessors, method, casts, and tests for actual types, as well as proofs for the fundamental properties of these operations in this concrete data model.\ text\Such generic function or ``compiler'' can be implemented in Isabelle on the ML level. This has been done, for a semantics following the open-world assumption, for UML 2.0 in~\<^cite>\"brucker.ea:extensible:2008-b" and "brucker:interactive:2007"\. In this paper, we follow another approach for UML 2.4: we define the concepts of the compilation informally, and present a concrete example which is verified in Isabelle/HOL.\ subsection\Outlining the Example\ text\We are presenting here an ``analysis-model'' of the (slightly modified) example Figure 7.3, page 20 of the OCL standard~\<^cite>\"omg:ocl:2012"\. Here, analysis model means that associations were really represented as relation on objects on the state---as is intended by the standard---rather by pointers between objects as is done in our ``design model'' \isatagafp (see \autoref{ex:employee-design:uml}). \endisatagafp \isatagannexa (see \url{http://isa-afp.org/entries/Featherweight_OCL.shtml}). \endisatagannexa To be precise, this theory contains the formalization of the data-part covered by the UML class model (see \autoref{fig:person-ana}):\ text\ \begin{figure} \centering\scalebox{.3}{\includegraphics{figures/person.png}}% \caption{A simple UML class model drawn from Figure 7.3, page 20 of~\<^cite>\"omg:ocl:2012"\. \label{fig:person-ana}} \end{figure} \ text\This means that the association (attached to the association class \inlineocl{EmployeeRanking}) with the association ends \inlineocl+boss+ and \inlineocl+employees+ is implemented by the attribute \inlineocl+boss+ and the operation \inlineocl+employees+ (to be discussed in the OCL part captured by the subsequent theory). \ section\Example Data-Universe and its Infrastructure\ text\Ideally, the following is generated automatically from a UML class model.\ text\Our data universe consists in the concrete class diagram just of node's, and implicitly of the class object. Each class implies the existence of a class type defined for the corresponding object representations as follows:\ datatype type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n = mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid (* the oid to the person itself *) "int option" (* the attribute "salary" or null *) datatype type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y = mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid (* the oid to the oclany itself *) "(int option) option" (* the extensions to "person"; used to denote objects of actual type "person" casted to "oclany"; in case of existence of several subclasses of oclany, sums of extensions have to be provided. *) text\Now, we construct a concrete ``universe of OclAny types'' by injection into a sum type containing the class types. This type of OclAny will be used as instance for all respective type-variables.\ datatype \ = in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n | in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y text\Having fixed the object universe, we can introduce type synonyms that exactly correspond to OCL types. Again, we exploit that our representation of OCL is a ``shallow embedding'' with a one-to-one correspondance of OCL-types to types of the meta-language HOL.\ type_synonym Boolean = " \ Boolean" type_synonym Integer = " \ Integer" type_synonym Void = " \ Void" type_synonym OclAny = "(\, type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y option option) val" type_synonym Person = "(\, type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n option option) val" type_synonym Set_Integer = "(\, int option option) Set" type_synonym Set_Person = "(\, type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n option option) Set" text\Just a little check:\ typ "Boolean" text\To reuse key-elements of the library like referential equality, we have to show that the object universe belongs to the type class ``oclany,'' \ie, each class type has to provide a function @{term oid_of} yielding the object id (oid) of the object.\ instantiation type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: object begin definition oid_of_type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def: "oid_of x = (case x of mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid _ \ oid)" instance .. end instantiation type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y :: object begin definition oid_of_type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_def: "oid_of x = (case x of mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid _ \ oid)" instance .. end instantiation \ :: object begin definition oid_of_\_def: "oid_of x = (case x of in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person \ oid_of person | in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oclany \ oid_of oclany)" instance .. end section\Instantiation of the Generic Strict Equality\ text\We instantiate the referential equality on \Person\ and \OclAny\\ overloading StrictRefEq \ "StrictRefEq :: [Person,Person] \ Boolean" begin definition StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n : "(x::Person) \ y \ StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t x y" end overloading StrictRefEq \ "StrictRefEq :: [OclAny,OclAny] \ Boolean" begin definition StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y : "(x::OclAny) \ y \ StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t x y" end lemmas cps23 = cp_StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t[of "x::Person" "y::Person" "\", simplified StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n[symmetric]] cp_intro(9) [of "P::Person \Person""Q::Person \Person", simplified StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n[symmetric] ] StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def [of "x::Person" "y::Person", simplified StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n[symmetric]] StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_defargs [of _ "x::Person" "y::Person", simplified StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n[symmetric]] StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_strict1 [of "x::Person", simplified StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n[symmetric]] StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_strict2 [of "x::Person", simplified StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n[symmetric]] for x y \ P Q text\For each Class \emph{C}, we will have a casting operation \inlineocl{.oclAsType($C$)}, a test on the actual type \inlineocl{.oclIsTypeOf($C$)} as well as its relaxed form \inlineocl{.oclIsKindOf($C$)} (corresponding exactly to Java's \verb+instanceof+-operator. \ text\Thus, since we have two class-types in our concrete class hierarchy, we have two operations to declare and to provide two overloading definitions for the two static types. \ section\OclAsType\ subsection\Definition\ consts OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y :: "'\ \ OclAny" ("(_) .oclAsType'(OclAny')") consts OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: "'\ \ Person" ("(_) .oclAsType'(Person')") definition "OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\ = (\u. \case u of in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y a \ a | in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid a) \ mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid \a\\)" lemma OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\_some: "OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\ x \ None" by(simp add: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\_def) overloading OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y \ "OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y :: OclAny \ OclAny" begin definition OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny: "(X::OclAny) .oclAsType(OclAny) \ X" end overloading OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y \ "OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y :: Person \ OclAny" begin definition OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person: "(X::Person) .oclAsType(OclAny) \ (\\. case X \ of \ \ invalid \ | \\\ \ null \ | \\mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid a \\ \ \\ (mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid \a\) \\)" end definition "OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\ = (\u. case u of in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n p \ \p\ | in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid \a\) \ \mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid a\ | _ \ None)" overloading OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n \ "OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: OclAny \ Person" begin definition OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny: "(X::OclAny) .oclAsType(Person) \ (\\. case X \ of \ \ invalid \ | \\\ \ null \ | \\mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid \ \\ \ invalid \ \ \down-cast exception\ | \\mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid \a\ \\ \ \\mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid a\\)" end overloading OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n \ "OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: Person \ Person" begin definition OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person: "(X::Person) .oclAsType(Person) \ X " (* to avoid identity for null ? *) end text_raw\\isatagafp\ lemmas [simp] = OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person subsection\Context Passing\ lemma cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Person: "cp P \ cp(\X. (P (X::Person)::Person) .oclAsType(OclAny))" by(rule cpI1, simp_all add: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) lemma cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_OclAny: "cp P \ cp(\X. (P (X::OclAny)::OclAny) .oclAsType(OclAny))" by(rule cpI1, simp_all add: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny) lemma cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_Person: "cp P \ cp(\X. (P (X::Person)::Person) .oclAsType(Person))" by(rule cpI1, simp_all add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person) lemma cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_OclAny: "cp P \ cp(\X. (P (X::OclAny)::OclAny) .oclAsType(Person))" by(rule cpI1, simp_all add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny) lemma cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_OclAny: "cp P \ cp(\X. (P (X::Person)::OclAny) .oclAsType(OclAny))" by(rule cpI1, simp_all add: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny) lemma cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Person: "cp P \ cp(\X. (P (X::OclAny)::Person) .oclAsType(OclAny))" by(rule cpI1, simp_all add: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) lemma cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_OclAny: "cp P \ cp(\X. (P (X::Person)::OclAny) .oclAsType(Person))" by(rule cpI1, simp_all add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny) lemma cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Person: "cp P \ cp(\X. (P (X::OclAny)::Person) .oclAsType(Person))" by(rule cpI1, simp_all add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person) lemmas [simp] = cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Person cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_OclAny cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_Person cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_OclAny cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_OclAny cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Person cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_OclAny cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Person text_raw\\endisatagafp\ subsection\Execution with Invalid or Null as Argument\ lemma OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_strict : "(invalid::OclAny) .oclAsType(OclAny) = invalid" by(simp) lemma OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_nullstrict : "(null::OclAny) .oclAsType(OclAny) = null" by(simp) lemma OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_strict[simp] : "(invalid::Person) .oclAsType(OclAny) = invalid" by(rule ext, simp add: bot_option_def invalid_def OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) lemma OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_nullstrict[simp] : "(null::Person) .oclAsType(OclAny) = null" by(rule ext, simp add: null_fun_def null_option_def bot_option_def OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) lemma OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_strict[simp] : "(invalid::OclAny) .oclAsType(Person) = invalid" by(rule ext, simp add: bot_option_def invalid_def OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny) lemma OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_nullstrict[simp] : "(null::OclAny) .oclAsType(Person) = null" by(rule ext, simp add: null_fun_def null_option_def bot_option_def OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny) lemma OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_strict : "(invalid::Person) .oclAsType(Person) = invalid" by(simp) lemma OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_nullstrict : "(null::Person) .oclAsType(Person) = null" by(simp) section\OclIsTypeOf\ subsection\Definition\ consts OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y :: "'\ \ Boolean" ("(_).oclIsTypeOf'(OclAny')") consts OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: "'\ \ Boolean" ("(_).oclIsTypeOf'(Person')") overloading OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y \ "OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y :: OclAny \ Boolean" begin definition OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny: "(X::OclAny) .oclIsTypeOf(OclAny) \ (\\. case X \ of \ \ invalid \ | \\\ \ true \ \ \invalid ??\ | \\mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid \ \\ \ true \ | \\mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid \_\ \\ \ false \)" end lemma OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny': "(X::OclAny) .oclIsTypeOf(OclAny) = (\ \. if \ \ \ X then (case X \ of \\\ \ true \ \ \invalid ??\ | \\mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid \ \\ \ true \ | \\mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid \_\ \\ \ false \) else invalid \)" apply(rule ext, simp add: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny) by(case_tac "\ \ \ X", auto simp: foundation18' bot_option_def) interpretation OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny : profile_mono_schemeV "OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y::OclAny \ Boolean" "\ X. (case X of \None\ \ \\True\\ \ \invalid ??\ | \\mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid None \\ \ \\True\\ | \\mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid \_\ \\ \ \\False\\)" apply(unfold_locales, simp add: atomize_eq, rule ext) by(auto simp: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny' OclValid_def true_def false_def split: option.split type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y.split) overloading OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y \ "OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y :: Person \ Boolean" begin definition OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person: "(X::Person) .oclIsTypeOf(OclAny) \ (\\. case X \ of \ \ invalid \ | \\\ \ true \ \ \invalid ??\ | \\ _ \\ \ false \) \ \must have actual type \Person\ otherwise\" end overloading OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n \ "OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: OclAny \ Boolean" begin definition OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny: "(X::OclAny) .oclIsTypeOf(Person) \ (\\. case X \ of \ \ invalid \ | \\\ \ true \ | \\mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid \ \\ \ false \ | \\mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid \_\ \\ \ true \)" end overloading OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n \ "OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: Person \ Boolean" begin definition OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person: "(X::Person) .oclIsTypeOf(Person) \ (\\. case X \ of \ \ invalid \ | _ \ true \)" (* for (* \\ _ \\ \ true \ *) : must have actual type Node otherwise *) end text_raw\\isatagafp\ subsection\Context Passing\ lemma cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Person: "cp P \ cp(\X.(P(X::Person)::Person).oclIsTypeOf(OclAny))" by(rule cpI1, simp_all add: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) lemma cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_OclAny: "cp P \ cp(\X.(P(X::OclAny)::OclAny).oclIsTypeOf(OclAny))" by(rule cpI1, simp_all add: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny) lemma cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_Person: "cp P \ cp(\X.(P(X::Person)::Person).oclIsTypeOf(Person))" by(rule cpI1, simp_all add: OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person) lemma cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_OclAny: "cp P \ cp(\X.(P(X::OclAny)::OclAny).oclIsTypeOf(Person))" by(rule cpI1, simp_all add: OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny) lemma cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_OclAny: "cp P \ cp(\X.(P(X::Person)::OclAny).oclIsTypeOf(OclAny))" by(rule cpI1, simp_all add: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny) lemma cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Person: "cp P \ cp(\X.(P(X::OclAny)::Person).oclIsTypeOf(OclAny))" by(rule cpI1, simp_all add: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) lemma cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_OclAny: "cp P \ cp(\X.(P(X::Person)::OclAny).oclIsTypeOf(Person))" by(rule cpI1, simp_all add: OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny) lemma cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Person: "cp P \ cp(\X.(P(X::OclAny)::Person).oclIsTypeOf(Person))" by(rule cpI1, simp_all add: OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person) lemmas [simp] = cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Person cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_OclAny cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_Person cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_OclAny cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_OclAny cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Person cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_OclAny cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Person text_raw\\endisatagafp\ subsection\Execution with Invalid or Null as Argument\ lemma OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_strict1[simp]: "(invalid::OclAny) .oclIsTypeOf(OclAny) = invalid" by(rule ext, simp add: null_fun_def null_option_def bot_option_def null_def invalid_def OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny) lemma OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_strict2[simp]: "(null::OclAny) .oclIsTypeOf(OclAny) = true" by(rule ext, simp add: null_fun_def null_option_def bot_option_def null_def invalid_def OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny) lemma OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_strict1[simp]: "(invalid::Person) .oclIsTypeOf(OclAny) = invalid" by(rule ext, simp add: null_fun_def null_option_def bot_option_def null_def invalid_def OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) lemma OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_strict2[simp]: "(null::Person) .oclIsTypeOf(OclAny) = true" by(rule ext, simp add: null_fun_def null_option_def bot_option_def null_def invalid_def OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) lemma OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_strict1[simp]: "(invalid::OclAny) .oclIsTypeOf(Person) = invalid" by(rule ext, simp add: null_fun_def null_option_def bot_option_def null_def invalid_def OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny) lemma OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_strict2[simp]: "(null::OclAny) .oclIsTypeOf(Person) = true" by(rule ext, simp add: null_fun_def null_option_def bot_option_def null_def invalid_def OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny) lemma OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_strict1[simp]: "(invalid::Person) .oclIsTypeOf(Person) = invalid" by(rule ext, simp add: null_fun_def null_option_def bot_option_def null_def invalid_def OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person) lemma OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_strict2[simp]: "(null::Person) .oclIsTypeOf(Person) = true" by(rule ext, simp add: null_fun_def null_option_def bot_option_def null_def invalid_def OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person) subsection\Up Down Casting\ lemma actualType_larger_staticType: assumes isdef: "\ \ (\ X)" shows "\ \ (X::Person) .oclIsTypeOf(OclAny) \ false" using isdef by(auto simp : null_option_def bot_option_def OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person foundation22 foundation16) lemma down_cast_type: assumes isOclAny: "\ \ (X::OclAny) .oclIsTypeOf(OclAny)" and non_null: "\ \ (\ X)" shows "\ \ (X .oclAsType(Person)) \ invalid" using isOclAny non_null apply(auto simp : bot_fun_def null_fun_def null_option_def bot_option_def null_def invalid_def OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny foundation22 foundation16 split: option.split type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y.split type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n.split) by(simp add: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny OclValid_def false_def true_def) lemma down_cast_type': assumes isOclAny: "\ \ (X::OclAny) .oclIsTypeOf(OclAny)" and non_null: "\ \ (\ X)" shows "\ \ not (\ (X .oclAsType(Person)))" by(rule foundation15[THEN iffD1], simp add: down_cast_type[OF assms]) lemma up_down_cast : assumes isdef: "\ \ (\ X)" shows "\ \ ((X::Person) .oclAsType(OclAny) .oclAsType(Person) \ X)" using isdef by(auto simp : null_fun_def null_option_def bot_option_def null_def invalid_def OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny foundation22 foundation16 split: option.split type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n.split) lemma up_down_cast_Person_OclAny_Person [simp]: shows "((X::Person) .oclAsType(OclAny) .oclAsType(Person) = X)" apply(rule ext, rename_tac \) apply(rule foundation22[THEN iffD1]) apply(case_tac "\ \ (\ X)", simp add: up_down_cast) apply(simp add: defined_split, elim disjE) apply(erule StrongEq_L_subst2_rev, simp, simp)+ done lemma up_down_cast_Person_OclAny_Person': assumes "\ \ \ X" shows "\ \ (((X :: Person) .oclAsType(OclAny) .oclAsType(Person)) \ X)" apply(simp only: up_down_cast_Person_OclAny_Person StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) by(rule StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_sym, simp add: assms) lemma up_down_cast_Person_OclAny_Person'': assumes "\ \ \ (X :: Person)" shows "\ \ (X .oclIsTypeOf(Person) implies (X .oclAsType(OclAny) .oclAsType(Person)) \ X)" apply(simp add: OclValid_def) apply(subst cp_OclImplies) apply(simp add: StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_sym[OF assms, simplified OclValid_def]) apply(subst cp_OclImplies[symmetric]) by simp section\OclIsKindOf\ subsection\Definition\ consts OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y :: "'\ \ Boolean" ("(_).oclIsKindOf'(OclAny')") consts OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: "'\ \ Boolean" ("(_).oclIsKindOf'(Person')") overloading OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y \ "OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y :: OclAny \ Boolean" begin definition OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny: "(X::OclAny) .oclIsKindOf(OclAny) \ (\\. case X \ of \ \ invalid \ | _ \ true \)" end overloading OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y \ "OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y :: Person \ Boolean" begin definition OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person: "(X::Person) .oclIsKindOf(OclAny) \ (\\. case X \ of \ \ invalid \ | _\ true \)" (* for (* \\mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n e oid _ \\ \ true \ *) : must have actual type Person otherwise *) end overloading OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n \ "OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: OclAny \ Boolean" begin definition OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny: "(X::OclAny) .oclIsKindOf(Person) \ (\\. case X \ of \ \ invalid \ | \\\ \ true \ | \\mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid \ \\ \ false \ | \\mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid \_\ \\ \ true \)" end overloading OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n \ "OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: Person \ Boolean" begin definition OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person: "(X::Person) .oclIsKindOf(Person) \ (\\. case X \ of \ \ invalid \ | _ \ true \)" end text_raw\\isatagafp\ subsection\Context Passing\ lemma cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Person: "cp P \ cp(\X.(P(X::Person)::Person).oclIsKindOf(OclAny))" by(rule cpI1, simp_all add: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) lemma cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_OclAny: "cp P \ cp(\X.(P(X::OclAny)::OclAny).oclIsKindOf(OclAny))" by(rule cpI1, simp_all add: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny) lemma cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_Person: "cp P \ cp(\X.(P(X::Person)::Person).oclIsKindOf(Person))" by(rule cpI1, simp_all add: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person) lemma cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_OclAny: "cp P \ cp(\X.(P(X::OclAny)::OclAny).oclIsKindOf(Person))" by(rule cpI1, simp_all add: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny) lemma cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_OclAny: "cp P \ cp(\X.(P(X::Person)::OclAny).oclIsKindOf(OclAny))" by(rule cpI1, simp_all add: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny) lemma cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Person: "cp P \ cp(\X.(P(X::OclAny)::Person).oclIsKindOf(OclAny))" by(rule cpI1, simp_all add: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) lemma cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_OclAny: "cp P \ cp(\X.(P(X::Person)::OclAny).oclIsKindOf(Person))" by(rule cpI1, simp_all add: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny) lemma cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Person: "cp P \ cp(\X.(P(X::OclAny)::Person).oclIsKindOf(Person))" by(rule cpI1, simp_all add: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person) lemmas [simp] = cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Person cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_OclAny cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_Person cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_OclAny cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_OclAny cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Person cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_OclAny cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Person text_raw\\endisatagafp\ subsection\Execution with Invalid or Null as Argument\ lemma OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_strict1[simp] : "(invalid::OclAny) .oclIsKindOf(OclAny) = invalid" by(rule ext, simp add: invalid_def bot_option_def OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny) lemma OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_strict2[simp] : "(null::OclAny) .oclIsKindOf(OclAny) = true" by(rule ext, simp add: null_fun_def null_option_def OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny) lemma OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_strict1[simp] : "(invalid::Person) .oclIsKindOf(OclAny) = invalid" by(rule ext, simp add: bot_option_def invalid_def OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) lemma OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_strict2[simp] : "(null::Person) .oclIsKindOf(OclAny) = true" by(rule ext, simp add: null_fun_def null_option_def bot_option_def OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) lemma OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_strict1[simp]: "(invalid::OclAny) .oclIsKindOf(Person) = invalid" by(rule ext, simp add: null_fun_def null_option_def bot_option_def null_def invalid_def OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny) lemma OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_strict2[simp]: "(null::OclAny) .oclIsKindOf(Person) = true" by(rule ext, simp add: null_fun_def null_option_def bot_option_def null_def invalid_def OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny) lemma OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_strict1[simp]: "(invalid::Person) .oclIsKindOf(Person) = invalid" by(rule ext, simp add: null_fun_def null_option_def bot_option_def null_def invalid_def OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person) lemma OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_strict2[simp]: "(null::Person) .oclIsKindOf(Person) = true" by(rule ext, simp add: null_fun_def null_option_def bot_option_def null_def invalid_def OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person) subsection\Up Down Casting\ lemma actualKind_larger_staticKind: assumes isdef: "\ \ (\ X)" shows "\ \ ((X::Person) .oclIsKindOf(OclAny) \ true)" using isdef by(auto simp : bot_option_def OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person foundation22 foundation16) lemma down_cast_kind: assumes isOclAny: "\ (\ \ ((X::OclAny).oclIsKindOf(Person)))" and non_null: "\ \ (\ X)" shows "\ \ ((X .oclAsType(Person)) \ invalid)" using isOclAny non_null apply(auto simp : bot_fun_def null_fun_def null_option_def bot_option_def null_def invalid_def OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny foundation22 foundation16 split: option.split type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y.split type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n.split) by(simp add: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny OclValid_def false_def true_def) section\OclAllInstances\ text\To denote OCL-types occurring in OCL expressions syntactically---as, for example, as ``argument'' of \inlineisar{oclAllInstances()}---we use the inverses of the injection functions into the object universes; we show that this is sufficient ``characterization.''\ definition "Person \ OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\" definition "OclAny \ OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\" lemmas [simp] = Person_def OclAny_def lemma OclAllInstances_generic\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_exec: "OclAllInstances_generic pre_post OclAny = (\\. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\ Some ` OclAny ` ran (heap (pre_post \)) \\)" proof - let ?S1 = "\\. OclAny ` ran (heap (pre_post \))" let ?S2 = "\\. ?S1 \ - {None}" have B : "\\. ?S2 \ \ ?S1 \" by auto have C : "\\. ?S1 \ \ ?S2 \" by(auto simp: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\_some) show ?thesis by(insert equalityI[OF B C], simp) qed lemma OclAllInstances_at_post\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_exec: "OclAny .allInstances() = (\\. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\ Some ` OclAny ` ran (heap (snd \)) \\)" unfolding OclAllInstances_at_post_def by(rule OclAllInstances_generic\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_exec) lemma OclAllInstances_at_pre\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_exec: "OclAny .allInstances@pre() = (\\. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\ Some ` OclAny ` ran (heap (fst \)) \\) " unfolding OclAllInstances_at_pre_def by(rule OclAllInstances_generic\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_exec) subsection\OclIsTypeOf\ lemma OclAny_allInstances_generic_oclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y1: assumes [simp]: "\x. pre_post (x, x) = x" shows "\\. (\ \ ((OclAllInstances_generic pre_post OclAny)->forAll\<^sub>S\<^sub>e\<^sub>t(X|X .oclIsTypeOf(OclAny))))" apply(rule_tac x = \\<^sub>0 in exI, simp add: \\<^sub>0_def OclValid_def del: OclAllInstances_generic_def) apply(simp only: assms UML_Set.OclForall_def refl if_True OclAllInstances_generic_defined[simplified OclValid_def]) apply(simp only: OclAllInstances_generic_def) apply(subst (1 2 3) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp add: bot_option_def) by(simp add: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny) lemma OclAny_allInstances_at_post_oclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y1: "\\. (\ \ (OclAny .allInstances()->forAll\<^sub>S\<^sub>e\<^sub>t(X|X .oclIsTypeOf(OclAny))))" unfolding OclAllInstances_at_post_def by(rule OclAny_allInstances_generic_oclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y1, simp) lemma OclAny_allInstances_at_pre_oclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y1: "\\. (\ \ (OclAny .allInstances@pre()->forAll\<^sub>S\<^sub>e\<^sub>t(X|X .oclIsTypeOf(OclAny))))" unfolding OclAllInstances_at_pre_def by(rule OclAny_allInstances_generic_oclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y1, simp) lemma OclAny_allInstances_generic_oclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y2: assumes [simp]: "\x. pre_post (x, x) = x" shows "\\. (\ \ not ((OclAllInstances_generic pre_post OclAny)->forAll\<^sub>S\<^sub>e\<^sub>t(X|X .oclIsTypeOf(OclAny))))" proof - fix oid a let ?t0 = "\heap = Map.empty(oid \ in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid \a\)), assocs = Map.empty\" show ?thesis apply(rule_tac x = "(?t0, ?t0)" in exI, simp add: OclValid_def del: OclAllInstances_generic_def) apply(simp only: UML_Set.OclForall_def refl if_True OclAllInstances_generic_defined[simplified OclValid_def]) apply(simp only: OclAllInstances_generic_def OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\_def) apply(subst (1 2 3) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp add: bot_option_def) by(simp add: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny OclNot_def OclAny_def) qed lemma OclAny_allInstances_at_post_oclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y2: "\\. (\ \ not (OclAny .allInstances()->forAll\<^sub>S\<^sub>e\<^sub>t(X|X .oclIsTypeOf(OclAny))))" unfolding OclAllInstances_at_post_def by(rule OclAny_allInstances_generic_oclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y2, simp) lemma OclAny_allInstances_at_pre_oclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y2: "\\. (\ \ not (OclAny .allInstances@pre()->forAll\<^sub>S\<^sub>e\<^sub>t(X|X .oclIsTypeOf(OclAny))))" unfolding OclAllInstances_at_pre_def by(rule OclAny_allInstances_generic_oclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y2, simp) lemma Person_allInstances_generic_oclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n: "\ \ ((OclAllInstances_generic pre_post Person)->forAll\<^sub>S\<^sub>e\<^sub>t(X|X .oclIsTypeOf(Person)))" apply(simp add: OclValid_def del: OclAllInstances_generic_def) apply(simp only: UML_Set.OclForall_def refl if_True OclAllInstances_generic_defined[simplified OclValid_def]) apply(simp only: OclAllInstances_generic_def) apply(subst (1 2 3) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp add: bot_option_def) by(simp add: OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person) lemma Person_allInstances_at_post_oclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n: "\ \ (Person .allInstances()->forAll\<^sub>S\<^sub>e\<^sub>t(X|X .oclIsTypeOf(Person)))" unfolding OclAllInstances_at_post_def by(rule Person_allInstances_generic_oclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) lemma Person_allInstances_at_pre_oclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n: "\ \ (Person .allInstances@pre()->forAll\<^sub>S\<^sub>e\<^sub>t(X|X .oclIsTypeOf(Person)))" unfolding OclAllInstances_at_pre_def by(rule Person_allInstances_generic_oclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) subsection\OclIsKindOf\ lemma OclAny_allInstances_generic_oclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y: "\ \ ((OclAllInstances_generic pre_post OclAny)->forAll\<^sub>S\<^sub>e\<^sub>t(X|X .oclIsKindOf(OclAny)))" apply(simp add: OclValid_def del: OclAllInstances_generic_def) apply(simp only: UML_Set.OclForall_def refl if_True OclAllInstances_generic_defined[simplified OclValid_def]) apply(simp only: OclAllInstances_generic_def) apply(subst (1 2 3) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp add: bot_option_def) by(simp add: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny) lemma OclAny_allInstances_at_post_oclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y: "\ \ (OclAny .allInstances()->forAll\<^sub>S\<^sub>e\<^sub>t(X|X .oclIsKindOf(OclAny)))" unfolding OclAllInstances_at_post_def by(rule OclAny_allInstances_generic_oclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y) lemma OclAny_allInstances_at_pre_oclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y: "\ \ (OclAny .allInstances@pre()->forAll\<^sub>S\<^sub>e\<^sub>t(X|X .oclIsKindOf(OclAny)))" unfolding OclAllInstances_at_pre_def by(rule OclAny_allInstances_generic_oclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y) lemma Person_allInstances_generic_oclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y: "\ \ ((OclAllInstances_generic pre_post Person)->forAll\<^sub>S\<^sub>e\<^sub>t(X|X .oclIsKindOf(OclAny)))" apply(simp add: OclValid_def del: OclAllInstances_generic_def) apply(simp only: UML_Set.OclForall_def refl if_True OclAllInstances_generic_defined[simplified OclValid_def]) apply(simp only: OclAllInstances_generic_def) apply(subst (1 2 3) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp add: bot_option_def) by(simp add: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) lemma Person_allInstances_at_post_oclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y: "\ \ (Person .allInstances()->forAll\<^sub>S\<^sub>e\<^sub>t(X|X .oclIsKindOf(OclAny)))" unfolding OclAllInstances_at_post_def by(rule Person_allInstances_generic_oclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y) lemma Person_allInstances_at_pre_oclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y: "\ \ (Person .allInstances@pre()->forAll\<^sub>S\<^sub>e\<^sub>t(X|X .oclIsKindOf(OclAny)))" unfolding OclAllInstances_at_pre_def by(rule Person_allInstances_generic_oclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y) lemma Person_allInstances_generic_oclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n: "\ \ ((OclAllInstances_generic pre_post Person)->forAll\<^sub>S\<^sub>e\<^sub>t(X|X .oclIsKindOf(Person)))" apply(simp add: OclValid_def del: OclAllInstances_generic_def) apply(simp only: UML_Set.OclForall_def refl if_True OclAllInstances_generic_defined[simplified OclValid_def]) apply(simp only: OclAllInstances_generic_def) apply(subst (1 2 3) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp add: bot_option_def) by(simp add: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person) lemma Person_allInstances_at_post_oclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n: "\ \ (Person .allInstances()->forAll\<^sub>S\<^sub>e\<^sub>t(X|X .oclIsKindOf(Person)))" unfolding OclAllInstances_at_post_def by(rule Person_allInstances_generic_oclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) lemma Person_allInstances_at_pre_oclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n: "\ \ (Person .allInstances@pre()->forAll\<^sub>S\<^sub>e\<^sub>t(X|X .oclIsKindOf(Person)))" unfolding OclAllInstances_at_pre_def by(rule Person_allInstances_generic_oclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) section\The Accessors (any, boss, salary)\ text\\label{sec:eam-accessors}\ text\Should be generated entirely from a class-diagram.\ subsection\Definition (of the association Employee-Boss)\ text\We start with a oid for the association; this oid can be used in presence of association classes to represent the association inside an object, pretty much similar to the \inlineisar+Design_UML+, where we stored an \verb+oid+ inside the class as ``pointer.''\ definition oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\ ::"oid" where "oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\ = 10" text\From there on, we can already define an empty state which must contain for $\mathit{oid}_{Person}\mathcal{BOSS}$ the empty relation (encoded as association list, since there are associations with a Sequence-like structure).\ definition eval_extract :: "('\,('a::object) option option) val \ (oid \ ('\,'c::null) val) \ ('\,'c::null) val" where "eval_extract X f = (\ \. case X \ of \ \ invalid \ \ \exception propagation\ | \ \ \ \ invalid \ \ \dereferencing null pointer\ | \\ obj \\ \ f (oid_of obj) \)" definition "choose\<^sub>2_1 = fst" definition "choose\<^sub>2_2 = snd" definition "List_flatten = (\l. (foldl ((\acc. (\l. (foldl ((\acc. (\l. (Cons (l) (acc))))) (acc) ((rev (l))))))) (Nil) ((rev (l)))))" definition "deref_assocs\<^sub>2" :: "('\ state \ '\ state \ '\ state) \ (oid list list \ oid list \ oid list) \ oid \ (oid list \ ('\,'f)val) \ oid \ ('\, 'f::null)val" where "deref_assocs\<^sub>2 pre_post to_from assoc_oid f oid = (\\. case (assocs (pre_post \)) assoc_oid of \ S \ \ f (List_flatten (map (choose\<^sub>2_2 \ to_from) (filter (\ p. List.member (choose\<^sub>2_1 (to_from p)) oid) S))) \ | _ \ invalid \)" text\The \pre_post\-parameter is configured with \fst\ or \snd\, the \to_from\-parameter either with the identity @{term id} or the following combinator \switch\:\ definition "switch\<^sub>2_1 = (\[x,y]\ (x,y))" definition "switch\<^sub>2_2 = (\[x,y]\ (y,x))" definition "switch\<^sub>3_1 = (\[x,y,z]\ (x,y))" definition "switch\<^sub>3_2 = (\[x,y,z]\ (x,z))" definition "switch\<^sub>3_3 = (\[x,y,z]\ (y,x))" definition "switch\<^sub>3_4 = (\[x,y,z]\ (y,z))" definition "switch\<^sub>3_5 = (\[x,y,z]\ (z,x))" definition "switch\<^sub>3_6 = (\[x,y,z]\ (z,y))" definition deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: "(\ state \ \ state \ \ state) \ (type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n \ (\, 'c::null)val) \ oid \ (\, 'c::null)val" where "deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n fst_snd f oid = (\\. case (heap (fst_snd \)) oid of \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n obj \ \ f obj \ | _ \ invalid \)" definition deref_oid\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y :: "(\ state \ \ state \ \ state) \ (type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y \ (\, 'c::null)val) \ oid \ (\, 'c::null)val" where "deref_oid\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y fst_snd f oid = (\\. case (heap (fst_snd \)) oid of \ in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y obj \ \ f obj \ | _ \ invalid \)" text\pointer undefined in state or not referencing a type conform object representation\ definition "select\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\\ f = (\ X. case X of (mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y _ \) \ null | (mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y _ \any\) \ f (\x _. \\x\\) any)" definition "select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\ f = select_object mtSet UML_Set.OclIncluding UML_Set.OclANY (f (\x _. \\x\\))" definition "select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\\\ f = (\ X. case X of (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n _ \) \ null | (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n _ \salary\) \ f (\x _. \\x\\) salary)" definition "deref_assocs\<^sub>2\\\\ fst_snd f = (\ mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid _ \ deref_assocs\<^sub>2 fst_snd switch\<^sub>2_1 oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\ f oid)" definition "in_pre_state = fst" definition "in_post_state = snd" definition "reconst_basetype = (\ convert x. convert x)" definition dot\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\\ :: "OclAny \ _" ("(1(_).any)" 50) where "(X).any = eval_extract X (deref_oid\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y in_post_state (select\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\\ reconst_basetype))" definition dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\ :: "Person \ Person" ("(1(_).boss)" 50) where "(X).boss = eval_extract X (deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n in_post_state (deref_assocs\<^sub>2\\\\ in_post_state (select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\ (deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n in_post_state))))" definition dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\\\ :: "Person \ Integer" ("(1(_).salary)" 50) where "(X).salary = eval_extract X (deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n in_post_state (select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\\\ reconst_basetype))" definition dot\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\\_at_pre :: "OclAny \ _" ("(1(_).any@pre)" 50) where "(X).any@pre = eval_extract X (deref_oid\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y in_pre_state (select\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\\ reconst_basetype))" definition dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\_at_pre:: "Person \ Person" ("(1(_).boss@pre)" 50) where "(X).boss@pre = eval_extract X (deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n in_pre_state (deref_assocs\<^sub>2\\\\ in_pre_state (select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\ (deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n in_pre_state))))" definition dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\\\_at_pre:: "Person \ Integer" ("(1(_).salary@pre)" 50) where "(X).salary@pre = eval_extract X (deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n in_pre_state (select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\\\ reconst_basetype))" lemmas dot_accessor = dot\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\\_def dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\_def dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\\\_def dot\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\\_at_pre_def dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\_at_pre_def dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\\\_at_pre_def subsection\Context Passing\ lemmas [simp] = eval_extract_def lemma cp_dot\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\\: "((X).any) \ = ((\_. X \).any) \" by (simp add: dot_accessor) lemma cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\: "((X).boss) \ = ((\_. X \).boss) \" by (simp add: dot_accessor) lemma cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\\\: "((X).salary) \ = ((\_. X \).salary) \" by (simp add: dot_accessor) lemma cp_dot\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\\_at_pre: "((X).any@pre) \ = ((\_. X \).any@pre) \" by (simp add: dot_accessor) lemma cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\_at_pre: "((X).boss@pre) \ = ((\_. X \).boss@pre) \" by (simp add: dot_accessor) lemma cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\\\_at_pre: "((X).salary@pre) \ = ((\_. X \).salary@pre) \" by (simp add: dot_accessor) lemmas cp_dot\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\\_I [simp, intro!]= cp_dot\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\\[THEN allI[THEN allI], of "\ X _. X" "\ _ \. \", THEN cpI1] lemmas cp_dot\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\\_at_pre_I [simp, intro!]= cp_dot\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\\_at_pre[THEN allI[THEN allI], of "\ X _. X" "\ _ \. \", THEN cpI1] lemmas cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\_I [simp, intro!]= cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\[THEN allI[THEN allI], of "\ X _. X" "\ _ \. \", THEN cpI1] lemmas cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\_at_pre_I [simp, intro!]= cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\_at_pre[THEN allI[THEN allI], of "\ X _. X" "\ _ \. \", THEN cpI1] lemmas cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\\\_I [simp, intro!]= cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\\\[THEN allI[THEN allI], of "\ X _. X" "\ _ \. \", THEN cpI1] lemmas cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\\\_at_pre_I [simp, intro!]= cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\\\_at_pre[THEN allI[THEN allI], of "\ X _. X" "\ _ \. \", THEN cpI1] subsection\Execution with Invalid or Null as Argument\ lemma dot\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\\_nullstrict [simp]: "(null).any = invalid" by(rule ext, simp add: dot_accessor null_fun_def null_option_def bot_option_def null_def invalid_def) lemma dot\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\\_at_pre_nullstrict [simp] : "(null).any@pre = invalid" by(rule ext, simp add: dot_accessor null_fun_def null_option_def bot_option_def null_def invalid_def) lemma dot\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\\_strict [simp] : "(invalid).any = invalid" by(rule ext, simp add: dot_accessor null_fun_def null_option_def bot_option_def null_def invalid_def) lemma dot\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\\_at_pre_strict [simp] : "(invalid).any@pre = invalid" by(rule ext, simp add: dot_accessor null_fun_def null_option_def bot_option_def null_def invalid_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\_nullstrict [simp]: "(null).boss = invalid" by(rule ext, simp add: dot_accessor null_fun_def null_option_def bot_option_def null_def invalid_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\_at_pre_nullstrict [simp] : "(null).boss@pre = invalid" by(rule ext, simp add: dot_accessor null_fun_def null_option_def bot_option_def null_def invalid_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\_strict [simp] : "(invalid).boss = invalid" by(rule ext, simp add: dot_accessor null_fun_def null_option_def bot_option_def null_def invalid_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\_at_pre_strict [simp] : "(invalid).boss@pre = invalid" by(rule ext, simp add: dot_accessor null_fun_def null_option_def bot_option_def null_def invalid_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\\\_nullstrict [simp]: "(null).salary = invalid" by(rule ext, simp add: dot_accessor null_fun_def null_option_def bot_option_def null_def invalid_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\\\_at_pre_nullstrict [simp] : "(null).salary@pre = invalid" by(rule ext, simp add: dot_accessor null_fun_def null_option_def bot_option_def null_def invalid_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\\\_strict [simp] : "(invalid).salary = invalid" by(rule ext, simp add: dot_accessor null_fun_def null_option_def bot_option_def null_def invalid_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\\\_at_pre_strict [simp] : "(invalid).salary@pre = invalid" by(rule ext, simp add: dot_accessor null_fun_def null_option_def bot_option_def null_def invalid_def) subsection\Representation in States\ lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\_def_mono:"\ \ \(X .boss) \ \ \ \(X)" apply(case_tac "\ \ (X \ invalid)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .boss)))" and \ = "\" and x = "X" and y = "invalid"], simp add: foundation16') apply(case_tac "\ \ (X \ null)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .boss)))" and \ = "\" and x = "X" and y = "null"], simp add: foundation16') by(simp add: defined_split) lemma repr_boss: assumes A : "\ \ \(x .boss)" shows "is_represented_in_state in_post_state (x .boss) Person \" apply(insert A[simplified foundation16] A[THEN dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\_def_mono, simplified foundation16]) unfolding is_represented_in_state_def dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\_def eval_extract_def select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\_def in_post_state_def oops lemma repr_bossX : assumes A: "\ \ \(x .boss)" shows "\ \ ((Person .allInstances()) ->includes\<^sub>S\<^sub>e\<^sub>t(x .boss))" oops section\A Little Infra-structure on Example States\ text\ The example we are defining in this section comes from the figure~\ref{fig:eam1_system-states}. \begin{figure} \includegraphics[width=\textwidth]{figures/pre-post.pdf} \caption{(a) pre-state $\sigma_1$ and (b) post-state $\sigma_1'$.} \label{fig:eam1_system-states} \end{figure} \ text_raw\\isatagafp\ definition OclInt1000 ("\\\\") where "OclInt1000 = (\ _ . \\1000\\)" definition OclInt1200 ("\\\\") where "OclInt1200 = (\ _ . \\1200\\)" definition OclInt1300 ("\\\\") where "OclInt1300 = (\ _ . \\1300\\)" definition OclInt1800 ("\\\\") where "OclInt1800 = (\ _ . \\1800\\)" definition OclInt2600 ("\\\\") where "OclInt2600 = (\ _ . \\2600\\)" definition OclInt2900 ("\\\\") where "OclInt2900 = (\ _ . \\2900\\)" definition OclInt3200 ("\\\\") where "OclInt3200 = (\ _ . \\3200\\)" definition OclInt3500 ("\\\\") where "OclInt3500 = (\ _ . \\3500\\)" definition "oid0 \ 0" definition "oid1 \ 1" definition "oid2 \ 2" definition "oid3 \ 3" definition "oid4 \ 4" definition "oid5 \ 5" definition "oid6 \ 6" definition "oid7 \ 7" definition "oid8 \ 8" definition "person1 \ mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid0 \1300\" definition "person2 \ mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid1 \1800\" definition "person3 \ mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid2 None" definition "person4 \ mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid3 \2900\" definition "person5 \ mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid4 \3500\" definition "person6 \ mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid5 \2500\" definition "person7 \ mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid6 \\3200\\" definition "person8 \ mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid7 None" definition "person9 \ mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid8 \0\" text_raw\\endisatagafp\ definition - "\\<^sub>1 \ \ heap = Map.empty(oid0 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid0 \1000\)) - (oid1 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid1 \1200\)) + "\\<^sub>1 \ \ heap = Map.empty(oid0 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid0 \1000\), + oid1 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid1 \1200\), \<^cancel>\oid2\ - (oid3 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid3 \2600\)) - (oid4 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person5) - (oid5 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid5 \2300\)) + oid3 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid3 \2600\), + oid4 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person5, + oid5 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid5 \2300\), \<^cancel>\oid6\ \<^cancel>\oid7\ - (oid8 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person9), + oid8 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person9), assocs = Map.empty(oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\ \ [[[oid0],[oid1]],[[oid3],[oid4]],[[oid5],[oid3]]]) \" definition - "\\<^sub>1' \ \ heap = Map.empty(oid0 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person1) - (oid1 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person2) - (oid2 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person3) - (oid3 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person4) + "\\<^sub>1' \ \ heap = Map.empty(oid0 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person1, + oid1 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person2, + oid2 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person3, + oid3 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person4, \<^cancel>\oid4\ - (oid5 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person6) - (oid6 \ in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y person7) - (oid7 \ in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y person8) - (oid8 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person9), + oid5 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person6, + oid6 \ in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y person7, + oid7 \ in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y person8, + oid8 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person9), assocs = Map.empty(oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\ \ [[[oid0],[oid1]],[[oid1],[oid1]],[[oid5],[oid6]],[[oid6],[oid6]]]) \" definition "\\<^sub>0 \ \ heap = Map.empty, assocs = Map.empty \" lemma basic_\_wff: "WFF(\\<^sub>1,\\<^sub>1')" by(auto simp: WFF_def \\<^sub>1_def \\<^sub>1'_def oid0_def oid1_def oid2_def oid3_def oid4_def oid5_def oid6_def oid7_def oid8_def oid_of_\_def oid_of_type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def oid_of_type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_def person1_def person2_def person3_def person4_def person5_def person6_def person7_def person8_def person9_def) lemma [simp,code_unfold]: "dom (heap \\<^sub>1) = {oid0,oid1\<^cancel>\,oid2\,oid3,oid4,oid5\<^cancel>\,oid6,oid7\,oid8}" by(auto simp: \\<^sub>1_def) lemma [simp,code_unfold]: "dom (heap \\<^sub>1') = {oid0,oid1,oid2,oid3\<^cancel>\,oid4\,oid5,oid6,oid7,oid8}" by(auto simp: \\<^sub>1'_def) text_raw\\isatagafp\ definition "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 :: Person \ \ _ .\\ person1 \\" definition "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 :: Person \ \ _ .\\ person2 \\" definition "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 :: Person \ \ _ .\\ person3 \\" definition "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4 :: Person \ \ _ .\\ person4 \\" definition "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5 :: Person \ \ _ .\\ person5 \\" definition "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 :: Person \ \ _ .\\ person6 \\" definition "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 :: OclAny \ \ _ .\\ person7 \\" definition "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8 :: OclAny \ \ _ .\\ person8 \\" definition "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 :: Person \ \ _ .\\ person9 \\" lemma [code_unfold]: "((x::Person) \ y) = StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t x y" by(simp only: StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) lemma [code_unfold]: "((x::OclAny) \ y) = StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t x y" by(simp only: StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y) lemmas [simp,code_unfold] = OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person text_raw\\endisatagafp\ Assert "\s\<^sub>p\<^sub>r\<^sub>e . (s\<^sub>p\<^sub>r\<^sub>e,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .salary <> \\\\)" Assert "\s\<^sub>p\<^sub>r\<^sub>e . (s\<^sub>p\<^sub>r\<^sub>e,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .salary \ \\\\)" Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .salary@pre \ \\\\)" Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .salary@pre <> \\\\)" (*Assert "\s\<^sub>p\<^sub>r\<^sub>e . (s\<^sub>p\<^sub>r\<^sub>e,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .boss <> X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1)" Assert "\s\<^sub>p\<^sub>r\<^sub>e . (s\<^sub>p\<^sub>r\<^sub>e,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .boss .salary \ \\\\)" Assert "\s\<^sub>p\<^sub>r\<^sub>e . (s\<^sub>p\<^sub>r\<^sub>e,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .boss .boss <> X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1)" Assert "\s\<^sub>p\<^sub>r\<^sub>e . (s\<^sub>p\<^sub>r\<^sub>e,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .boss .boss \ X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2)" Assert " (\\<^sub>1,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .boss@pre .salary \ \\\\)" Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .boss@pre .salary@pre \ \\\\)" Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .boss@pre .salary@pre <> \\\\)" Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .boss@pre \ X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2)" Assert " (\\<^sub>1,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .boss@pre .boss \ X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2)" Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .boss@pre .boss@pre \ null)" Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ not(\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .boss@pre .boss@pre .boss@pre))" *) lemma " (\\<^sub>1,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .oclIsMaintained())" by(simp add: OclValid_def OclIsMaintained_def \\<^sub>1_def \\<^sub>1'_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1_def person1_def oid0_def oid1_def oid2_def oid3_def oid4_def oid5_def oid6_def oid_of_option_def oid_of_type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def) lemma "\s\<^sub>p\<^sub>r\<^sub>e s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (s\<^sub>p\<^sub>r\<^sub>e,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ ((X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .oclAsType(OclAny) .oclAsType(Person)) \ X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1)" by(rule up_down_cast_Person_OclAny_Person', simp add: X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1_def) Assert "\s\<^sub>p\<^sub>r\<^sub>e s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (s\<^sub>p\<^sub>r\<^sub>e,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .oclIsTypeOf(Person))" Assert "\s\<^sub>p\<^sub>r\<^sub>e s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (s\<^sub>p\<^sub>r\<^sub>e,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ not(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .oclIsTypeOf(OclAny))" Assert "\s\<^sub>p\<^sub>r\<^sub>e s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (s\<^sub>p\<^sub>r\<^sub>e,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .oclIsKindOf(Person))" Assert "\s\<^sub>p\<^sub>r\<^sub>e s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (s\<^sub>p\<^sub>r\<^sub>e,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .oclIsKindOf(OclAny))" Assert "\s\<^sub>p\<^sub>r\<^sub>e s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (s\<^sub>p\<^sub>r\<^sub>e,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ not(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .oclAsType(OclAny) .oclIsTypeOf(OclAny))" Assert "\s\<^sub>p\<^sub>r\<^sub>e . (s\<^sub>p\<^sub>r\<^sub>e,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 .salary \ \\\\)" Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 .salary@pre \ \\\\)" (*Assert "\s\<^sub>p\<^sub>r\<^sub>e . (s\<^sub>p\<^sub>r\<^sub>e,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 .boss \ X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2)" Assert " (\\<^sub>1,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 .boss .salary@pre \ \\\\)" Assert " (\\<^sub>1,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 .boss .boss@pre \ null)" Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 .boss@pre \ null)" Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 .boss@pre <> X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2)" Assert " (\\<^sub>1,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 .boss@pre <> (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 .boss))" Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ not(\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 .boss@pre .boss))" Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ not(\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 .boss@pre .salary@pre))" *) lemma " (\\<^sub>1,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 .oclIsMaintained())" by(simp add: OclValid_def OclIsMaintained_def \\<^sub>1_def \\<^sub>1'_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2_def person2_def oid0_def oid1_def oid2_def oid3_def oid4_def oid5_def oid6_def oid_of_option_def oid_of_type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def) Assert "\s\<^sub>p\<^sub>r\<^sub>e . (s\<^sub>p\<^sub>r\<^sub>e,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 .salary \ null)" Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ not(\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 .salary@pre))" (*Assert "\s\<^sub>p\<^sub>r\<^sub>e . (s\<^sub>p\<^sub>r\<^sub>e,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 .boss \ null)" Assert "\s\<^sub>p\<^sub>r\<^sub>e . (s\<^sub>p\<^sub>r\<^sub>e,\\<^sub>1') \ not(\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 .boss .salary))" Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ not(\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 .boss@pre))" *)lemma " (\\<^sub>1,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 .oclIsNew())" by(simp add: OclValid_def OclIsNew_def \\<^sub>1_def \\<^sub>1'_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3_def person3_def oid0_def oid1_def oid2_def oid3_def oid4_def oid5_def oid6_def oid8_def oid_of_option_def oid_of_type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def) (*Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4 .boss@pre \ X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5)" Assert " (\\<^sub>1,\\<^sub>1') \ not(\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4 .boss@pre .salary))" Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4 .boss@pre .salary@pre \ \\\\)" *) lemma " (\\<^sub>1,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4 .oclIsMaintained())" by(simp add: OclValid_def OclIsMaintained_def \\<^sub>1_def \\<^sub>1'_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4_def person4_def oid0_def oid1_def oid2_def oid3_def oid4_def oid5_def oid6_def oid_of_option_def oid_of_type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def) Assert "\s\<^sub>p\<^sub>r\<^sub>e . (s\<^sub>p\<^sub>r\<^sub>e,\\<^sub>1') \ not(\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5 .salary))" Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5 .salary@pre \ \\\\)" (*Assert "\s\<^sub>p\<^sub>r\<^sub>e . (s\<^sub>p\<^sub>r\<^sub>e,\\<^sub>1') \ not(\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5 .boss))" *) lemma " (\\<^sub>1,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5 .oclIsDeleted())" by(simp add: OclNot_def OclValid_def OclIsDeleted_def \\<^sub>1_def \\<^sub>1'_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5_def person5_def oid0_def oid1_def oid2_def oid3_def oid4_def oid5_def oid6_def oid7_def oid8_def oid_of_option_def oid_of_type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def) (* (* access to an oclany object not yet supported *) Assert " (\\<^sub>1,\\<^sub>1') \ ((X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 .boss .salary) \ \\\\ )"*) (*Assert "\s\<^sub>p\<^sub>r\<^sub>e . (s\<^sub>p\<^sub>r\<^sub>e,\\<^sub>1') \ not(\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 .boss .salary@pre))" Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 .boss@pre \ X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4)" Assert " (\\<^sub>1,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 .boss@pre .salary \ \\\\)" Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 .boss@pre .salary@pre \ \\\\)" Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 .boss@pre .boss@pre \ X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5)" *) lemma " (\\<^sub>1,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 .oclIsMaintained())" by(simp add: OclValid_def OclIsMaintained_def \\<^sub>1_def \\<^sub>1'_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6_def person6_def oid0_def oid1_def oid2_def oid3_def oid4_def oid5_def oid6_def oid_of_option_def oid_of_type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def) (* (* access to an oclany object not yet supported *) Assert " (\\<^sub>1,\\<^sub>1') \ ((X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 .oclAsType(Person) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 .boss)))" *) (* (* access to an oclany object not yet supported *) Assert " (\\<^sub>1,\\<^sub>1') \ ((X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 .oclAsType(Person) .boss) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 .oclAsType(Person)) )" *) (* (* access to an oclany object not yet supported *) Assert " (\\<^sub>1,\\<^sub>1') \ ((X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 .oclAsType(Person) .boss .salary) \ \\\\ )" *) Assert "\s\<^sub>p\<^sub>r\<^sub>e s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (s\<^sub>p\<^sub>r\<^sub>e,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ \(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 .oclAsType(Person))" (*Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ not(\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 .oclAsType(Person) .boss@pre))" *) lemma "\s\<^sub>p\<^sub>r\<^sub>e s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (s\<^sub>p\<^sub>r\<^sub>e,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ ((X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 .oclAsType(Person) .oclAsType(OclAny) .oclAsType(Person)) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 .oclAsType(Person)))" by(rule up_down_cast_Person_OclAny_Person', simp add: X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7_def OclValid_def valid_def person7_def) lemma " (\\<^sub>1,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 .oclIsNew())" by(simp add: OclValid_def OclIsNew_def \\<^sub>1_def \\<^sub>1'_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7_def person7_def oid0_def oid1_def oid2_def oid3_def oid4_def oid5_def oid6_def oid8_def oid_of_option_def oid_of_type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_def) Assert "\s\<^sub>p\<^sub>r\<^sub>e s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (s\<^sub>p\<^sub>r\<^sub>e,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8 <> X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7)" Assert "\s\<^sub>p\<^sub>r\<^sub>e s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (s\<^sub>p\<^sub>r\<^sub>e,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ not(\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8 .oclAsType(Person)))" Assert "\s\<^sub>p\<^sub>r\<^sub>e s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (s\<^sub>p\<^sub>r\<^sub>e,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8 .oclIsTypeOf(OclAny))" Assert "\s\<^sub>p\<^sub>r\<^sub>e s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (s\<^sub>p\<^sub>r\<^sub>e,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ not(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8 .oclIsTypeOf(Person))" Assert "\s\<^sub>p\<^sub>r\<^sub>e s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (s\<^sub>p\<^sub>r\<^sub>e,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ not(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8 .oclIsKindOf(Person))" Assert "\s\<^sub>p\<^sub>r\<^sub>e s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (s\<^sub>p\<^sub>r\<^sub>e,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8 .oclIsKindOf(OclAny))" lemma \_modifiedonly: "(\\<^sub>1,\\<^sub>1') \ (Set{ X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .oclAsType(OclAny) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 .oclAsType(OclAny) \<^cancel>\, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 .oclAsType(OclAny)\ , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4 .oclAsType(OclAny) \<^cancel>\, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5 .oclAsType(OclAny)\ , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 .oclAsType(OclAny) \<^cancel>\, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 .oclAsType(OclAny)\ \<^cancel>\, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8 .oclAsType(OclAny)\ \<^cancel>\, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 .oclAsType(OclAny)\}->oclIsModifiedOnly())" apply(simp add: OclIsModifiedOnly_def OclValid_def oid0_def oid1_def oid2_def oid3_def oid4_def oid5_def oid6_def oid7_def oid8_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9_def person1_def person2_def person3_def person4_def person5_def person6_def person7_def person8_def person9_def image_def) apply(simp add: OclIncluding_rep_set mtSet_rep_set null_option_def bot_option_def) apply(simp add: oid_of_option_def oid_of_type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_def, clarsimp) apply(simp add: \\<^sub>1_def \\<^sub>1'_def oid0_def oid1_def oid2_def oid3_def oid4_def oid5_def oid6_def oid7_def oid8_def) done lemma "(\\<^sub>1,\\<^sub>1') \ ((X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 @pre (\x. \OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\ x\)) \ X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9)" by(simp add: OclSelf_at_pre_def \\<^sub>1_def oid_of_option_def oid_of_type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9_def person9_def oid8_def OclValid_def StrongEq_def OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\_def) lemma "(\\<^sub>1,\\<^sub>1') \ ((X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 @post (\x. \OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\ x\)) \ X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9)" by(simp add: OclSelf_at_post_def \\<^sub>1'_def oid_of_option_def oid_of_type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9_def person9_def oid8_def OclValid_def StrongEq_def OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\_def) lemma "(\\<^sub>1,\\<^sub>1') \ (((X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 .oclAsType(OclAny)) @pre (\x. \OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\ x\)) \ ((X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 .oclAsType(OclAny)) @post (\x. \OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\ x\)))" proof - have including4 : "\a b c d \. Set{\\. \\a\\, \\. \\b\\, \\. \\c\\, \\. \\d\\} \ = Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\ {\\a\\, \\b\\, \\c\\, \\d\\} \\" apply(subst abs_rep_simp'[symmetric], simp) apply(simp add: OclIncluding_rep_set mtSet_rep_set) by(rule arg_cong[of _ _ "\x. (Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e(\\ x \\))"], auto) have excluding1: "\S a b c d e \. (\_. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\ {\\a\\, \\b\\, \\c\\, \\d\\} \\)->excluding\<^sub>S\<^sub>e\<^sub>t(\\. \\e\\) \ = Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\ {\\a\\, \\b\\, \\c\\, \\d\\} - {\\e\\} \\" apply(simp add: UML_Set.OclExcluding_def) apply(simp add: defined_def OclValid_def false_def true_def bot_fun_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_fun_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def) apply(rule conjI) apply(rule impI, subst (asm) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject) apply( simp add: bot_option_def)+ apply(rule conjI) apply(rule impI, subst (asm) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject) apply( simp add: bot_option_def null_option_def)+ apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp add: bot_option_def, simp) done show ?thesis apply(rule framing[where X = "Set{ X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .oclAsType(OclAny) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 .oclAsType(OclAny) \<^cancel>\, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 .oclAsType(OclAny)\ , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4 .oclAsType(OclAny) \<^cancel>\, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5 .oclAsType(OclAny)\ , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 .oclAsType(OclAny) \<^cancel>\, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 .oclAsType(OclAny)\ \<^cancel>\, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8 .oclAsType(OclAny)\ \<^cancel>\, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 .oclAsType(OclAny)\}"]) apply(cut_tac \_modifiedonly) apply(simp only: OclValid_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9_def person1_def person2_def person3_def person4_def person5_def person6_def person7_def person8_def person9_def OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) apply(subst cp_OclIsModifiedOnly, subst UML_Set.OclExcluding.cp0, subst (asm) cp_OclIsModifiedOnly, simp add: including4 excluding1) apply(simp only: X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9_def person1_def person2_def person3_def person4_def person5_def person6_def person7_def person8_def person9_def) apply(simp add: OclIncluding_rep_set mtSet_rep_set oid0_def oid1_def oid2_def oid3_def oid4_def oid5_def oid6_def oid7_def oid8_def) apply(simp add: StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def oid_of_option_def oid_of_type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_def OclNot_def OclValid_def null_option_def bot_option_def) done qed lemma perm_\\<^sub>1' : "\\<^sub>1' = \ heap = Map.empty - (oid8 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person9) - (oid7 \ in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y person8) - (oid6 \ in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y person7) - (oid5 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person6) + (oid8 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person9, + oid7 \ in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y person8, + oid6 \ in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y person7, + oid5 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person6, \<^cancel>\oid4\ - (oid3 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person4) - (oid2 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person3) - (oid1 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person2) - (oid0 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person1) + oid3 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person4, + oid2 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person3, + oid1 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person2, + oid0 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person1) , assocs = assocs \\<^sub>1' \" proof - note P = fun_upd_twist show ?thesis apply(simp add: \\<^sub>1'_def oid0_def oid1_def oid2_def oid3_def oid4_def oid5_def oid6_def oid7_def oid8_def) apply(subst (1) P, simp) apply(subst (2) P, simp) apply(subst (1) P, simp) apply(subst (3) P, simp) apply(subst (2) P, simp) apply(subst (1) P, simp) apply(subst (4) P, simp) apply(subst (3) P, simp) apply(subst (2) P, simp) apply(subst (1) P, simp) apply(subst (5) P, simp) apply(subst (4) P, simp) apply(subst (3) P, simp) apply(subst (2) P, simp) apply(subst (1) P, simp) apply(subst (6) P, simp) apply(subst (5) P, simp) apply(subst (4) P, simp) apply(subst (3) P, simp) apply(subst (2) P, simp) apply(subst (1) P, simp) apply(subst (7) P, simp) apply(subst (6) P, simp) apply(subst (5) P, simp) apply(subst (4) P, simp) apply(subst (3) P, simp) apply(subst (2) P, simp) apply(subst (1) P, simp) by(simp) qed declare const_ss [simp] lemma "\\\<^sub>1. (\\<^sub>1,\\<^sub>1') \ (Person .allInstances() \ Set{ X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^cancel>\, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 .oclAsType(Person)\<^cancel>\, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8\, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 })" apply(subst perm_\\<^sub>1') apply(simp only: oid0_def oid1_def oid2_def oid3_def oid4_def oid5_def oid6_def oid7_def oid8_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9_def person7_def) apply(subst state_update_vs_allInstances_at_post_tc, simp, simp add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\_def, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp, simp, simp, rule OclIncluding_cong, simp, simp) apply(subst state_update_vs_allInstances_at_post_tc, simp, simp add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\_def, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp, simp, simp, rule OclIncluding_cong, simp, simp) apply(subst state_update_vs_allInstances_at_post_tc, simp, simp add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\_def, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp, simp, simp, rule OclIncluding_cong, simp, simp) apply(subst state_update_vs_allInstances_at_post_tc, simp, simp add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\_def, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp, simp, simp, rule OclIncluding_cong, simp, simp) apply(subst state_update_vs_allInstances_at_post_tc, simp, simp add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\_def, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp, simp, simp, rule OclIncluding_cong, simp, simp) apply(subst state_update_vs_allInstances_at_post_tc, simp, simp add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\_def, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp, simp, simp, rule OclIncluding_cong, simp, simp) apply(subst state_update_vs_allInstances_at_post_ntc, simp, simp add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\_def person8_def, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp, simp, simp) apply(subst state_update_vs_allInstances_at_post_tc, simp, simp add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\_def, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp, simp, simp, rule OclIncluding_cong, simp, simp) apply(rule state_update_vs_allInstances_at_post_empty) by(simp_all add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\_def) lemma "\\\<^sub>1. (\\<^sub>1,\\<^sub>1') \ (OclAny .allInstances() \ Set{ X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .oclAsType(OclAny), X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 .oclAsType(OclAny), X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 .oclAsType(OclAny), X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4 .oclAsType(OclAny) \<^cancel>\, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 .oclAsType(OclAny), X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 .oclAsType(OclAny) })" apply(subst perm_\\<^sub>1') apply(simp only: oid0_def oid1_def oid2_def oid3_def oid4_def oid5_def oid6_def oid7_def oid8_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9_def person1_def person2_def person3_def person4_def person5_def person6_def person9_def) apply(subst state_update_vs_allInstances_at_post_tc, simp, simp add: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\_def, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp, simp, simp, rule OclIncluding_cong, simp, simp)+ apply(rule state_update_vs_allInstances_at_post_empty) by(simp_all add: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\_def) end diff --git a/thys/Featherweight_OCL/examples/Employee_Model/Design/Design_UML.thy b/thys/Featherweight_OCL/examples/Employee_Model/Design/Design_UML.thy --- a/thys/Featherweight_OCL/examples/Employee_Model/Design/Design_UML.thy +++ b/thys/Featherweight_OCL/examples/Employee_Model/Design/Design_UML.thy @@ -1,1337 +1,1337 @@ (***************************************************************************** * Featherweight-OCL --- A Formal Semantics for UML-OCL Version OCL 2.5 * for the OMG Standard. * http://www.brucker.ch/projects/hol-testgen/ * * Design_UML.thy --- OCL Contracts and an Example. * This file is part of HOL-TestGen. * * Copyright (c) 2012-2015 Université Paris-Saclay, Univ. Paris-Sud, France * 2013-2015 IRT SystemX, France * * All rights reserved. * * Redistribution and use in source and binary forms, with or without * modification, are permitted provided that the following conditions are * met: * * * Redistributions of source code must retain the above copyright * notice, this list of conditions and the following disclaimer. * * * Redistributions in binary form must reproduce the above * copyright notice, this list of conditions and the following * disclaimer in the documentation and/or other materials provided * with the distribution. * * * Neither the name of the copyright holders nor the names of its * contributors may be used to endorse or promote products derived * from this software without specific prior written permission. * * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS * "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT * LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR * A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT * OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, * SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT * LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, * DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY * THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT * (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. ******************************************************************************) chapter\Example: The Employee Design Model\ (* UML part *) theory Design_UML imports "../../../UML_Main" begin text \\label{ex:employee-design:uml}\ section\Introduction\ text\ For certain concepts like classes and class-types, only a generic definition for its resulting semantics can be given. Generic means, there is a function outside HOL that ``compiles'' a concrete, closed-world class diagram into a ``theory'' of this data model, consisting of a bunch of definitions for classes, accessors, method, casts, and tests for actual types, as well as proofs for the fundamental properties of these operations in this concrete data model.\ text\Such generic function or ``compiler'' can be implemented in Isabelle on the ML level. This has been done, for a semantics following the open-world assumption, for UML 2.0 in~\<^cite>\"brucker.ea:extensible:2008-b" and "brucker:interactive:2007"\. In this paper, we follow another approach for UML 2.4: we define the concepts of the compilation informally, and present a concrete example which is verified in Isabelle/HOL.\ subsection\Outlining the Example\ text\We are presenting here a ``design-model'' of the (slightly modified) example Figure 7.3, page 20 of the OCL standard~\<^cite>\"omg:ocl:2012"\. To be precise, this theory contains the formalization of the data-part covered by the UML class model (see \autoref{fig:person}):\ text\ \begin{figure} \centering\scalebox{.3}{\includegraphics{figures/person.png}}% \caption{A simple UML class model drawn from Figure 7.3, page 20 of~\<^cite>\"omg:ocl:2012"\. \label{fig:person}} \end{figure} \ text\This means that the association (attached to the association class \inlineocl{EmployeeRanking}) with the association ends \inlineocl+boss+ and \inlineocl+employees+ is implemented by the attribute \inlineocl+boss+ and the operation \inlineocl+employees+ (to be discussed in the OCL part captured by the subsequent theory). \ section\Example Data-Universe and its Infrastructure\ text\Ideally, the following is generated automatically from a UML class model.\ text\Our data universe consists in the concrete class diagram just of node's, and implicitly of the class object. Each class implies the existence of a class type defined for the corresponding object representations as follows:\ datatype type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n = mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid (* the oid to the person itself *) "int option" (* the attribute "salary" or null *) "oid option" (* the attribute "boss" or null *) datatype type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y = mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid (* the oid to the oclany itself *) "(int option \ oid option) option" (* the extensions to "person"; used to denote objects of actual type "person" casted to "oclany"; in case of existence of several subclasses of oclany, sums of extensions have to be provided. *) text\Now, we construct a concrete ``universe of OclAny types'' by injection into a sum type containing the class types. This type of OclAny will be used as instance for all respective type-variables.\ datatype \ = in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n | in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y text\Having fixed the object universe, we can introduce type synonyms that exactly correspond to OCL types. Again, we exploit that our representation of OCL is a ``shallow embedding'' with a one-to-one correspondance of OCL-types to types of the meta-language HOL.\ type_synonym Boolean = " \ Boolean" type_synonym Integer = " \ Integer" type_synonym Void = " \ Void" type_synonym OclAny = "(\, type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y option option) val" type_synonym Person = "(\, type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n option option) val" type_synonym Set_Integer = "(\, int option option) Set" type_synonym Set_Person = "(\, type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n option option) Set" text\Just a little check:\ typ "Boolean" text\To reuse key-elements of the library like referential equality, we have to show that the object universe belongs to the type class ``oclany,'' \ie, each class type has to provide a function @{term oid_of} yielding the object id (oid) of the object.\ instantiation type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: object begin definition oid_of_type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def: "oid_of x = (case x of mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid _ _ \ oid)" instance .. end instantiation type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y :: object begin definition oid_of_type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_def: "oid_of x = (case x of mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid _ \ oid)" instance .. end instantiation \ :: object begin definition oid_of_\_def: "oid_of x = (case x of in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person \ oid_of person | in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oclany \ oid_of oclany)" instance .. end section\Instantiation of the Generic Strict Equality\ text\We instantiate the referential equality on \Person\ and \OclAny\\ overloading StrictRefEq \ "StrictRefEq :: [Person,Person] \ Boolean" begin definition StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n : "(x::Person) \ y \ StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t x y" end overloading StrictRefEq \ "StrictRefEq :: [OclAny,OclAny] \ Boolean" begin definition StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y : "(x::OclAny) \ y \ StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t x y" end lemmas cps23 = cp_StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t[of "x::Person" "y::Person" "\", simplified StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n[symmetric]] cp_intro(9) [of "P::Person \Person""Q::Person \Person", simplified StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n[symmetric] ] StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def [of "x::Person" "y::Person", simplified StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n[symmetric]] StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_defargs [of _ "x::Person" "y::Person", simplified StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n[symmetric]] StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_strict1 [of "x::Person", simplified StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n[symmetric]] StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_strict2 [of "x::Person", simplified StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n[symmetric]] for x y \ P Q text\For each Class \emph{C}, we will have a casting operation \inlineocl{.oclAsType($C$)}, a test on the actual type \inlineocl{.oclIsTypeOf($C$)} as well as its relaxed form \inlineocl{.oclIsKindOf($C$)} (corresponding exactly to Java's \verb+instanceof+-operator. \ text\Thus, since we have two class-types in our concrete class hierarchy, we have two operations to declare and to provide two overloading definitions for the two static types. \ section\OclAsType\ subsection\Definition\ consts OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y :: "'\ \ OclAny" ("(_) .oclAsType'(OclAny')") consts OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: "'\ \ Person" ("(_) .oclAsType'(Person')") definition "OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\ = (\u. \case u of in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y a \ a | in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid a b) \ mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid \(a,b)\\)" lemma OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\_some: "OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\ x \ None" by(simp add: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\_def) overloading OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y \ "OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y :: OclAny \ OclAny" begin definition OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny: "(X::OclAny) .oclAsType(OclAny) \ X" end overloading OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y \ "OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y :: Person \ OclAny" begin definition OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person: "(X::Person) .oclAsType(OclAny) \ (\\. case X \ of \ \ invalid \ | \\\ \ null \ | \\mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid a b \\ \ \\ (mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid \(a,b)\) \\)" end definition "OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\ = (\u. case u of in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n p \ \p\ | in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid \(a,b)\) \ \mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid a b\ | _ \ None)" overloading OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n \ "OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: OclAny \ Person" begin definition OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny: "(X::OclAny) .oclAsType(Person) \ (\\. case X \ of \ \ invalid \ | \\\ \ null \ | \\mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid \ \\ \ invalid \ \ \down-cast exception\ | \\mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid \(a,b)\ \\ \ \\mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid a b \\)" end overloading OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n \ "OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: Person \ Person" begin definition OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person: "(X::Person) .oclAsType(Person) \ X " (* to avoid identity for null ? *) end text_raw\\isatagafp\ lemmas [simp] = OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person subsection\Context Passing\ lemma cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Person: "cp P \ cp(\X. (P (X::Person)::Person) .oclAsType(OclAny))" by(rule cpI1, simp_all add: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) lemma cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_OclAny: "cp P \ cp(\X. (P (X::OclAny)::OclAny) .oclAsType(OclAny))" by(rule cpI1, simp_all add: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny) lemma cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_Person: "cp P \ cp(\X. (P (X::Person)::Person) .oclAsType(Person))" by(rule cpI1, simp_all add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person) lemma cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_OclAny: "cp P \ cp(\X. (P (X::OclAny)::OclAny) .oclAsType(Person))" by(rule cpI1, simp_all add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny) lemma cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_OclAny: "cp P \ cp(\X. (P (X::Person)::OclAny) .oclAsType(OclAny))" by(rule cpI1, simp_all add: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny) lemma cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Person: "cp P \ cp(\X. (P (X::OclAny)::Person) .oclAsType(OclAny))" by(rule cpI1, simp_all add: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) lemma cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_OclAny: "cp P \ cp(\X. (P (X::Person)::OclAny) .oclAsType(Person))" by(rule cpI1, simp_all add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny) lemma cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Person: "cp P \ cp(\X. (P (X::OclAny)::Person) .oclAsType(Person))" by(rule cpI1, simp_all add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person) lemmas [simp] = cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Person cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_OclAny cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_Person cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_OclAny cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_OclAny cp_OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Person cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_OclAny cp_OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Person text_raw\\endisatagafp\ subsection\Execution with Invalid or Null as Argument\ lemma OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_strict : "(invalid::OclAny) .oclAsType(OclAny) = invalid" by(simp) lemma OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_nullstrict : "(null::OclAny) .oclAsType(OclAny) = null" by(simp) lemma OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_strict[simp] : "(invalid::Person) .oclAsType(OclAny) = invalid" by(rule ext, simp add: bot_option_def invalid_def OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) lemma OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_nullstrict[simp] : "(null::Person) .oclAsType(OclAny) = null" by(rule ext, simp add: null_fun_def null_option_def bot_option_def OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) lemma OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_strict[simp] : "(invalid::OclAny) .oclAsType(Person) = invalid" by(rule ext, simp add: bot_option_def invalid_def OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny) lemma OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_nullstrict[simp] : "(null::OclAny) .oclAsType(Person) = null" by(rule ext, simp add: null_fun_def null_option_def bot_option_def OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny) lemma OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_strict : "(invalid::Person) .oclAsType(Person) = invalid" by(simp) lemma OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_nullstrict : "(null::Person) .oclAsType(Person) = null" by(simp) section\OclIsTypeOf\ subsection\Definition\ consts OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y :: "'\ \ Boolean" ("(_).oclIsTypeOf'(OclAny')") consts OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: "'\ \ Boolean" ("(_).oclIsTypeOf'(Person')") overloading OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y \ "OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y :: OclAny \ Boolean" begin definition OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny: "(X::OclAny) .oclIsTypeOf(OclAny) \ (\\. case X \ of \ \ invalid \ | \\\ \ true \ \ \invalid ??\ | \\mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid \ \\ \ true \ | \\mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid \_\ \\ \ false \)" end lemma OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny': "(X::OclAny) .oclIsTypeOf(OclAny) = (\ \. if \ \ \ X then (case X \ of \\\ \ true \ \ \invalid ??\ | \\mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid \ \\ \ true \ | \\mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid \_\ \\ \ false \) else invalid \)" apply(rule ext, simp add: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny) by(case_tac "\ \ \ X", auto simp: foundation18' bot_option_def) interpretation OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny : profile_mono_schemeV "OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y::OclAny \ Boolean" "\ X. (case X of \None\ \ \\True\\ \ \invalid ??\ | \\mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid None \\ \ \\True\\ | \\mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid \_\ \\ \ \\False\\)" apply(unfold_locales, simp add: atomize_eq, rule ext) by(auto simp: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny' OclValid_def true_def false_def split: option.split type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y.split) overloading OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y \ "OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y :: Person \ Boolean" begin definition OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person: "(X::Person) .oclIsTypeOf(OclAny) \ (\\. case X \ of \ \ invalid \ | \\\ \ true \ \ \invalid ??\ | \\ _ \\ \ false \) \ \must have actual type \Person\ otherwise\" end overloading OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n \ "OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: OclAny \ Boolean" begin definition OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny: "(X::OclAny) .oclIsTypeOf(Person) \ (\\. case X \ of \ \ invalid \ | \\\ \ true \ | \\mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid \ \\ \ false \ | \\mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid \_\ \\ \ true \)" end overloading OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n \ "OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: Person \ Boolean" begin definition OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person: "(X::Person) .oclIsTypeOf(Person) \ (\\. case X \ of \ \ invalid \ | _ \ true \)" (* for (* \\ _ \\ \ true \ *) : must have actual type Node otherwise *) end text_raw\\isatagafp\ subsection\Context Passing\ lemma cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Person: "cp P \ cp(\X.(P(X::Person)::Person).oclIsTypeOf(OclAny))" by(rule cpI1, simp_all add: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) lemma cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_OclAny: "cp P \ cp(\X.(P(X::OclAny)::OclAny).oclIsTypeOf(OclAny))" by(rule cpI1, simp_all add: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny) lemma cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_Person: "cp P \ cp(\X.(P(X::Person)::Person).oclIsTypeOf(Person))" by(rule cpI1, simp_all add: OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person) lemma cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_OclAny: "cp P \ cp(\X.(P(X::OclAny)::OclAny).oclIsTypeOf(Person))" by(rule cpI1, simp_all add: OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny) lemma cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_OclAny: "cp P \ cp(\X.(P(X::Person)::OclAny).oclIsTypeOf(OclAny))" by(rule cpI1, simp_all add: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny) lemma cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Person: "cp P \ cp(\X.(P(X::OclAny)::Person).oclIsTypeOf(OclAny))" by(rule cpI1, simp_all add: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) lemma cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_OclAny: "cp P \ cp(\X.(P(X::Person)::OclAny).oclIsTypeOf(Person))" by(rule cpI1, simp_all add: OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny) lemma cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Person: "cp P \ cp(\X.(P(X::OclAny)::Person).oclIsTypeOf(Person))" by(rule cpI1, simp_all add: OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person) lemmas [simp] = cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Person cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_OclAny cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_Person cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_OclAny cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_OclAny cp_OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Person cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_OclAny cp_OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Person text_raw\\endisatagafp\ subsection\Execution with Invalid or Null as Argument\ lemma OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_strict1[simp]: "(invalid::OclAny) .oclIsTypeOf(OclAny) = invalid" by(rule ext, simp add: null_fun_def null_option_def bot_option_def null_def invalid_def OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny) lemma OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_strict2[simp]: "(null::OclAny) .oclIsTypeOf(OclAny) = true" by(rule ext, simp add: null_fun_def null_option_def bot_option_def null_def invalid_def OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny) lemma OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_strict1[simp]: "(invalid::Person) .oclIsTypeOf(OclAny) = invalid" by(rule ext, simp add: null_fun_def null_option_def bot_option_def null_def invalid_def OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) lemma OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_strict2[simp]: "(null::Person) .oclIsTypeOf(OclAny) = true" by(rule ext, simp add: null_fun_def null_option_def bot_option_def null_def invalid_def OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) lemma OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_strict1[simp]: "(invalid::OclAny) .oclIsTypeOf(Person) = invalid" by(rule ext, simp add: null_fun_def null_option_def bot_option_def null_def invalid_def OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny) lemma OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_strict2[simp]: "(null::OclAny) .oclIsTypeOf(Person) = true" by(rule ext, simp add: null_fun_def null_option_def bot_option_def null_def invalid_def OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny) lemma OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_strict1[simp]: "(invalid::Person) .oclIsTypeOf(Person) = invalid" by(rule ext, simp add: null_fun_def null_option_def bot_option_def null_def invalid_def OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person) lemma OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_strict2[simp]: "(null::Person) .oclIsTypeOf(Person) = true" by(rule ext, simp add: null_fun_def null_option_def bot_option_def null_def invalid_def OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person) subsection\Up Down Casting\ lemma actualType_larger_staticType: assumes isdef: "\ \ (\ X)" shows "\ \ (X::Person) .oclIsTypeOf(OclAny) \ false" using isdef by(auto simp : null_option_def bot_option_def OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person foundation22 foundation16) lemma down_cast_type: assumes isOclAny: "\ \ (X::OclAny) .oclIsTypeOf(OclAny)" and non_null: "\ \ (\ X)" shows "\ \ (X .oclAsType(Person)) \ invalid" using isOclAny non_null apply(auto simp : bot_fun_def null_fun_def null_option_def bot_option_def null_def invalid_def OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny foundation22 foundation16 split: option.split type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y.split type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n.split) by(simp add: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny OclValid_def false_def true_def) lemma down_cast_type': assumes isOclAny: "\ \ (X::OclAny) .oclIsTypeOf(OclAny)" and non_null: "\ \ (\ X)" shows "\ \ not (\ (X .oclAsType(Person)))" by(rule foundation15[THEN iffD1], simp add: down_cast_type[OF assms]) lemma up_down_cast : assumes isdef: "\ \ (\ X)" shows "\ \ ((X::Person) .oclAsType(OclAny) .oclAsType(Person) \ X)" using isdef by(auto simp : null_fun_def null_option_def bot_option_def null_def invalid_def OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny foundation22 foundation16 split: option.split type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n.split) lemma up_down_cast_Person_OclAny_Person [simp]: shows "((X::Person) .oclAsType(OclAny) .oclAsType(Person) = X)" apply(rule ext, rename_tac \) apply(rule foundation22[THEN iffD1]) apply(case_tac "\ \ (\ X)", simp add: up_down_cast) apply(simp add: defined_split, elim disjE) apply(erule StrongEq_L_subst2_rev, simp, simp)+ done lemma up_down_cast_Person_OclAny_Person': assumes "\ \ \ X" shows "\ \ (((X :: Person) .oclAsType(OclAny) .oclAsType(Person)) \ X)" apply(simp only: up_down_cast_Person_OclAny_Person StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) by(rule StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_sym, simp add: assms) lemma up_down_cast_Person_OclAny_Person'': assumes "\ \ \ (X :: Person)" shows "\ \ (X .oclIsTypeOf(Person) implies (X .oclAsType(OclAny) .oclAsType(Person)) \ X)" apply(simp add: OclValid_def) apply(subst cp_OclImplies) apply(simp add: StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_sym[OF assms, simplified OclValid_def]) apply(subst cp_OclImplies[symmetric]) by simp section\OclIsKindOf\ subsection\Definition\ consts OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y :: "'\ \ Boolean" ("(_).oclIsKindOf'(OclAny')") consts OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: "'\ \ Boolean" ("(_).oclIsKindOf'(Person')") overloading OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y \ "OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y :: OclAny \ Boolean" begin definition OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny: "(X::OclAny) .oclIsKindOf(OclAny) \ (\\. case X \ of \ \ invalid \ | _ \ true \)" end overloading OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y \ "OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y :: Person \ Boolean" begin definition OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person: "(X::Person) .oclIsKindOf(OclAny) \ (\\. case X \ of \ \ invalid \ | _\ true \)" (* for (* \\mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n e oid _ \\ \ true \ *) : must have actual type Person otherwise *) end overloading OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n \ "OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: OclAny \ Boolean" begin definition OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny: "(X::OclAny) .oclIsKindOf(Person) \ (\\. case X \ of \ \ invalid \ | \\\ \ true \ | \\mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid \ \\ \ false \ | \\mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid \_\ \\ \ true \)" end overloading OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n \ "OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: Person \ Boolean" begin definition OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person: "(X::Person) .oclIsKindOf(Person) \ (\\. case X \ of \ \ invalid \ | _ \ true \)" end text_raw\\isatagafp\ subsection\Context Passing\ lemma cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Person: "cp P \ cp(\X.(P(X::Person)::Person).oclIsKindOf(OclAny))" by(rule cpI1, simp_all add: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) lemma cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_OclAny: "cp P \ cp(\X.(P(X::OclAny)::OclAny).oclIsKindOf(OclAny))" by(rule cpI1, simp_all add: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny) lemma cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_Person: "cp P \ cp(\X.(P(X::Person)::Person).oclIsKindOf(Person))" by(rule cpI1, simp_all add: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person) lemma cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_OclAny: "cp P \ cp(\X.(P(X::OclAny)::OclAny).oclIsKindOf(Person))" by(rule cpI1, simp_all add: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny) lemma cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_OclAny: "cp P \ cp(\X.(P(X::Person)::OclAny).oclIsKindOf(OclAny))" by(rule cpI1, simp_all add: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny) lemma cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Person: "cp P \ cp(\X.(P(X::OclAny)::Person).oclIsKindOf(OclAny))" by(rule cpI1, simp_all add: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) lemma cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_OclAny: "cp P \ cp(\X.(P(X::Person)::OclAny).oclIsKindOf(Person))" by(rule cpI1, simp_all add: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny) lemma cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Person: "cp P \ cp(\X.(P(X::OclAny)::Person).oclIsKindOf(Person))" by(rule cpI1, simp_all add: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person) lemmas [simp] = cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_Person cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_OclAny cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_Person cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_OclAny cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_OclAny cp_OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_Person cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_OclAny cp_OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_Person text_raw\\endisatagafp\ subsection\Execution with Invalid or Null as Argument\ lemma OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_strict1[simp] : "(invalid::OclAny) .oclIsKindOf(OclAny) = invalid" by(rule ext, simp add: invalid_def bot_option_def OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny) lemma OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny_strict2[simp] : "(null::OclAny) .oclIsKindOf(OclAny) = true" by(rule ext, simp add: null_fun_def null_option_def OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny) lemma OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_strict1[simp] : "(invalid::Person) .oclIsKindOf(OclAny) = invalid" by(rule ext, simp add: bot_option_def invalid_def OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) lemma OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person_strict2[simp] : "(null::Person) .oclIsKindOf(OclAny) = true" by(rule ext, simp add: null_fun_def null_option_def bot_option_def OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) lemma OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_strict1[simp]: "(invalid::OclAny) .oclIsKindOf(Person) = invalid" by(rule ext, simp add: null_fun_def null_option_def bot_option_def null_def invalid_def OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny) lemma OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny_strict2[simp]: "(null::OclAny) .oclIsKindOf(Person) = true" by(rule ext, simp add: null_fun_def null_option_def bot_option_def null_def invalid_def OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny) lemma OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_strict1[simp]: "(invalid::Person) .oclIsKindOf(Person) = invalid" by(rule ext, simp add: null_fun_def null_option_def bot_option_def null_def invalid_def OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person) lemma OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person_strict2[simp]: "(null::Person) .oclIsKindOf(Person) = true" by(rule ext, simp add: null_fun_def null_option_def bot_option_def null_def invalid_def OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person) subsection\Up Down Casting\ lemma actualKind_larger_staticKind: assumes isdef: "\ \ (\ X)" shows "\ \ ((X::Person) .oclIsKindOf(OclAny) \ true)" using isdef by(auto simp : bot_option_def OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person foundation22 foundation16) lemma down_cast_kind: assumes isOclAny: "\ (\ \ ((X::OclAny).oclIsKindOf(Person)))" and non_null: "\ \ (\ X)" shows "\ \ ((X .oclAsType(Person)) \ invalid)" using isOclAny non_null apply(auto simp : bot_fun_def null_fun_def null_option_def bot_option_def null_def invalid_def OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny foundation22 foundation16 split: option.split type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y.split type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n.split) by(simp add: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny OclValid_def false_def true_def) section\OclAllInstances\ text\To denote OCL-types occurring in OCL expressions syntactically---as, for example, as ``argument'' of \inlineisar{oclAllInstances()}---we use the inverses of the injection functions into the object universes; we show that this is sufficient ``characterization.''\ definition "Person \ OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\" definition "OclAny \ OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\" lemmas [simp] = Person_def OclAny_def lemma OclAllInstances_generic\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_exec: "OclAllInstances_generic pre_post OclAny = (\\. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\ Some ` OclAny ` ran (heap (pre_post \)) \\)" proof - let ?S1 = "\\. OclAny ` ran (heap (pre_post \))" let ?S2 = "\\. ?S1 \ - {None}" have B : "\\. ?S2 \ \ ?S1 \" by auto have C : "\\. ?S1 \ \ ?S2 \" by(auto simp: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\_some) show ?thesis by(insert equalityI[OF B C], simp) qed lemma OclAllInstances_at_post\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_exec: "OclAny .allInstances() = (\\. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\ Some ` OclAny ` ran (heap (snd \)) \\)" unfolding OclAllInstances_at_post_def by(rule OclAllInstances_generic\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_exec) lemma OclAllInstances_at_pre\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_exec: "OclAny .allInstances@pre() = (\\. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\ Some ` OclAny ` ran (heap (fst \)) \\) " unfolding OclAllInstances_at_pre_def by(rule OclAllInstances_generic\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_exec) subsection\OclIsTypeOf\ lemma OclAny_allInstances_generic_oclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y1: assumes [simp]: "\x. pre_post (x, x) = x" shows "\\. (\ \ ((OclAllInstances_generic pre_post OclAny)->forAll\<^sub>S\<^sub>e\<^sub>t(X|X .oclIsTypeOf(OclAny))))" apply(rule_tac x = \\<^sub>0 in exI, simp add: \\<^sub>0_def OclValid_def del: OclAllInstances_generic_def) apply(simp only: assms UML_Set.OclForall_def refl if_True OclAllInstances_generic_defined[simplified OclValid_def]) apply(simp only: OclAllInstances_generic_def) apply(subst (1 2 3) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp add: bot_option_def) by(simp add: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny) lemma OclAny_allInstances_at_post_oclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y1: "\\. (\ \ (OclAny .allInstances()->forAll\<^sub>S\<^sub>e\<^sub>t(X|X .oclIsTypeOf(OclAny))))" unfolding OclAllInstances_at_post_def by(rule OclAny_allInstances_generic_oclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y1, simp) lemma OclAny_allInstances_at_pre_oclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y1: "\\. (\ \ (OclAny .allInstances@pre()->forAll\<^sub>S\<^sub>e\<^sub>t(X|X .oclIsTypeOf(OclAny))))" unfolding OclAllInstances_at_pre_def by(rule OclAny_allInstances_generic_oclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y1, simp) lemma OclAny_allInstances_generic_oclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y2: assumes [simp]: "\x. pre_post (x, x) = x" shows "\\. (\ \ not ((OclAllInstances_generic pre_post OclAny)->forAll\<^sub>S\<^sub>e\<^sub>t(X|X .oclIsTypeOf(OclAny))))" proof - fix oid a let ?t0 = "\heap = Map.empty(oid \ in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y (mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid \a\)), assocs = Map.empty\" show ?thesis apply(rule_tac x = "(?t0, ?t0)" in exI, simp add: OclValid_def del: OclAllInstances_generic_def) apply(simp only: UML_Set.OclForall_def refl if_True OclAllInstances_generic_defined[simplified OclValid_def]) apply(simp only: OclAllInstances_generic_def OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\_def) apply(subst (1 2 3) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp add: bot_option_def) by(simp add: OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny OclNot_def OclAny_def) qed lemma OclAny_allInstances_at_post_oclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y2: "\\. (\ \ not (OclAny .allInstances()->forAll\<^sub>S\<^sub>e\<^sub>t(X|X .oclIsTypeOf(OclAny))))" unfolding OclAllInstances_at_post_def by(rule OclAny_allInstances_generic_oclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y2, simp) lemma OclAny_allInstances_at_pre_oclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y2: "\\. (\ \ not (OclAny .allInstances@pre()->forAll\<^sub>S\<^sub>e\<^sub>t(X|X .oclIsTypeOf(OclAny))))" unfolding OclAllInstances_at_pre_def by(rule OclAny_allInstances_generic_oclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y2, simp) lemma Person_allInstances_generic_oclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n: "\ \ ((OclAllInstances_generic pre_post Person)->forAll\<^sub>S\<^sub>e\<^sub>t(X|X .oclIsTypeOf(Person)))" apply(simp add: OclValid_def del: OclAllInstances_generic_def) apply(simp only: UML_Set.OclForall_def refl if_True OclAllInstances_generic_defined[simplified OclValid_def]) apply(simp only: OclAllInstances_generic_def) apply(subst (1 2 3) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp add: bot_option_def) by(simp add: OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person) lemma Person_allInstances_at_post_oclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n: "\ \ (Person .allInstances()->forAll\<^sub>S\<^sub>e\<^sub>t(X|X .oclIsTypeOf(Person)))" unfolding OclAllInstances_at_post_def by(rule Person_allInstances_generic_oclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) lemma Person_allInstances_at_pre_oclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n: "\ \ (Person .allInstances@pre()->forAll\<^sub>S\<^sub>e\<^sub>t(X|X .oclIsTypeOf(Person)))" unfolding OclAllInstances_at_pre_def by(rule Person_allInstances_generic_oclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) subsection\OclIsKindOf\ lemma OclAny_allInstances_generic_oclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y: "\ \ ((OclAllInstances_generic pre_post OclAny)->forAll\<^sub>S\<^sub>e\<^sub>t(X|X .oclIsKindOf(OclAny)))" apply(simp add: OclValid_def del: OclAllInstances_generic_def) apply(simp only: UML_Set.OclForall_def refl if_True OclAllInstances_generic_defined[simplified OclValid_def]) apply(simp only: OclAllInstances_generic_def) apply(subst (1 2 3) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp add: bot_option_def) by(simp add: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny) lemma OclAny_allInstances_at_post_oclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y: "\ \ (OclAny .allInstances()->forAll\<^sub>S\<^sub>e\<^sub>t(X|X .oclIsKindOf(OclAny)))" unfolding OclAllInstances_at_post_def by(rule OclAny_allInstances_generic_oclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y) lemma OclAny_allInstances_at_pre_oclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y: "\ \ (OclAny .allInstances@pre()->forAll\<^sub>S\<^sub>e\<^sub>t(X|X .oclIsKindOf(OclAny)))" unfolding OclAllInstances_at_pre_def by(rule OclAny_allInstances_generic_oclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y) lemma Person_allInstances_generic_oclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y: "\ \ ((OclAllInstances_generic pre_post Person)->forAll\<^sub>S\<^sub>e\<^sub>t(X|X .oclIsKindOf(OclAny)))" apply(simp add: OclValid_def del: OclAllInstances_generic_def) apply(simp only: UML_Set.OclForall_def refl if_True OclAllInstances_generic_defined[simplified OclValid_def]) apply(simp only: OclAllInstances_generic_def) apply(subst (1 2 3) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp add: bot_option_def) by(simp add: OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) lemma Person_allInstances_at_post_oclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y: "\ \ (Person .allInstances()->forAll\<^sub>S\<^sub>e\<^sub>t(X|X .oclIsKindOf(OclAny)))" unfolding OclAllInstances_at_post_def by(rule Person_allInstances_generic_oclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y) lemma Person_allInstances_at_pre_oclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y: "\ \ (Person .allInstances@pre()->forAll\<^sub>S\<^sub>e\<^sub>t(X|X .oclIsKindOf(OclAny)))" unfolding OclAllInstances_at_pre_def by(rule Person_allInstances_generic_oclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y) lemma Person_allInstances_generic_oclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n: "\ \ ((OclAllInstances_generic pre_post Person)->forAll\<^sub>S\<^sub>e\<^sub>t(X|X .oclIsKindOf(Person)))" apply(simp add: OclValid_def del: OclAllInstances_generic_def) apply(simp only: UML_Set.OclForall_def refl if_True OclAllInstances_generic_defined[simplified OclValid_def]) apply(simp only: OclAllInstances_generic_def) apply(subst (1 2 3) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp add: bot_option_def) by(simp add: OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person) lemma Person_allInstances_at_post_oclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n: "\ \ (Person .allInstances()->forAll\<^sub>S\<^sub>e\<^sub>t(X|X .oclIsKindOf(Person)))" unfolding OclAllInstances_at_post_def by(rule Person_allInstances_generic_oclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) lemma Person_allInstances_at_pre_oclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n: "\ \ (Person .allInstances@pre()->forAll\<^sub>S\<^sub>e\<^sub>t(X|X .oclIsKindOf(Person)))" unfolding OclAllInstances_at_pre_def by(rule Person_allInstances_generic_oclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) section\The Accessors (any, boss, salary)\ text\\label{sec:edm-accessors}\ text\Should be generated entirely from a class-diagram.\ subsection\Definition\ definition eval_extract :: "('\,('a::object) option option) val \ (oid \ ('\,'c::null) val) \ ('\,'c::null) val" where "eval_extract X f = (\ \. case X \ of \ \ invalid \ \ \exception propagation\ | \ \ \ \ invalid \ \ \dereferencing null pointer\ | \\ obj \\ \ f (oid_of obj) \)" definition deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n :: "(\ state \ \ state \ \ state) \ (type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n \ (\, 'c::null)val) \ oid \ (\, 'c::null)val" where "deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n fst_snd f oid = (\\. case (heap (fst_snd \)) oid of \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n obj \ \ f obj \ | _ \ invalid \)" definition deref_oid\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y :: "(\ state \ \ state \ \ state) \ (type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y \ (\, 'c::null)val) \ oid \ (\, 'c::null)val" where "deref_oid\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y fst_snd f oid = (\\. case (heap (fst_snd \)) oid of \ in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y obj \ \ f obj \ | _ \ invalid \)" text\pointer undefined in state or not referencing a type conform object representation\ definition "select\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\\ f = (\ X. case X of (mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y _ \) \ null | (mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y _ \any\) \ f (\x _. \\x\\) any)" definition "select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\ f = (\ X. case X of (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n _ _ \) \ null \ \object contains null pointer\ | (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n _ _ \boss\) \ f (\x _. \\x\\) boss)" definition "select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\\\ f = (\ X. case X of (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n _ \ _) \ null | (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n _ \salary\ _) \ f (\x _. \\x\\) salary)" definition "in_pre_state = fst" definition "in_post_state = snd" definition "reconst_basetype = (\ convert x. convert x)" definition dot\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\\ :: "OclAny \ _" ("(1(_).any)" 50) where "(X).any = eval_extract X (deref_oid\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y in_post_state (select\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\\ reconst_basetype))" definition dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\ :: "Person \ Person" ("(1(_).boss)" 50) where "(X).boss = eval_extract X (deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n in_post_state (select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\ (deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n in_post_state)))" definition dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\\\ :: "Person \ Integer" ("(1(_).salary)" 50) where "(X).salary = eval_extract X (deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n in_post_state (select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\\\ reconst_basetype))" definition dot\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\\_at_pre :: "OclAny \ _" ("(1(_).any@pre)" 50) where "(X).any@pre = eval_extract X (deref_oid\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y in_pre_state (select\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\\ reconst_basetype))" definition dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\_at_pre:: "Person \ Person" ("(1(_).boss@pre)" 50) where "(X).boss@pre = eval_extract X (deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n in_pre_state (select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\ (deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n in_pre_state)))" definition dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\\\_at_pre:: "Person \ Integer" ("(1(_).salary@pre)" 50) where "(X).salary@pre = eval_extract X (deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n in_pre_state (select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\\\ reconst_basetype))" lemmas dot_accessor = dot\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\\_def dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\_def dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\\\_def dot\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\\_at_pre_def dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\_at_pre_def dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\\\_at_pre_def subsection\Context Passing\ lemmas [simp] = eval_extract_def lemma cp_dot\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\\: "((X).any) \ = ((\_. X \).any) \" by (simp add: dot_accessor) lemma cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\: "((X).boss) \ = ((\_. X \).boss) \" by (simp add: dot_accessor) lemma cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\\\: "((X).salary) \ = ((\_. X \).salary) \" by (simp add: dot_accessor) lemma cp_dot\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\\_at_pre: "((X).any@pre) \ = ((\_. X \).any@pre) \" by (simp add: dot_accessor) lemma cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\_at_pre: "((X).boss@pre) \ = ((\_. X \).boss@pre) \" by (simp add: dot_accessor) lemma cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\\\_at_pre: "((X).salary@pre) \ = ((\_. X \).salary@pre) \" by (simp add: dot_accessor) lemmas cp_dot\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\\_I [simp, intro!]= cp_dot\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\\[THEN allI[THEN allI], of "\ X _. X" "\ _ \. \", THEN cpI1] lemmas cp_dot\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\\_at_pre_I [simp, intro!]= cp_dot\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\\_at_pre[THEN allI[THEN allI], of "\ X _. X" "\ _ \. \", THEN cpI1] lemmas cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\_I [simp, intro!]= cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\[THEN allI[THEN allI], of "\ X _. X" "\ _ \. \", THEN cpI1] lemmas cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\_at_pre_I [simp, intro!]= cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\_at_pre[THEN allI[THEN allI], of "\ X _. X" "\ _ \. \", THEN cpI1] lemmas cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\\\_I [simp, intro!]= cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\\\[THEN allI[THEN allI], of "\ X _. X" "\ _ \. \", THEN cpI1] lemmas cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\\\_at_pre_I [simp, intro!]= cp_dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\\\_at_pre[THEN allI[THEN allI], of "\ X _. X" "\ _ \. \", THEN cpI1] subsection\Execution with Invalid or Null as Argument\ lemma dot\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\\_nullstrict [simp]: "(null).any = invalid" by(rule ext, simp add: dot_accessor null_fun_def null_option_def bot_option_def null_def invalid_def) lemma dot\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\\_at_pre_nullstrict [simp] : "(null).any@pre = invalid" by(rule ext, simp add: dot_accessor null_fun_def null_option_def bot_option_def null_def invalid_def) lemma dot\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\\_strict [simp] : "(invalid).any = invalid" by(rule ext, simp add: dot_accessor null_fun_def null_option_def bot_option_def null_def invalid_def) lemma dot\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y\\\_at_pre_strict [simp] : "(invalid).any@pre = invalid" by(rule ext, simp add: dot_accessor null_fun_def null_option_def bot_option_def null_def invalid_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\_nullstrict [simp]: "(null).boss = invalid" by(rule ext, simp add: dot_accessor null_fun_def null_option_def bot_option_def null_def invalid_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\_at_pre_nullstrict [simp] : "(null).boss@pre = invalid" by(rule ext, simp add: dot_accessor null_fun_def null_option_def bot_option_def null_def invalid_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\_strict [simp] : "(invalid).boss = invalid" by(rule ext, simp add: dot_accessor null_fun_def null_option_def bot_option_def null_def invalid_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\_at_pre_strict [simp] : "(invalid).boss@pre = invalid" by(rule ext, simp add: dot_accessor null_fun_def null_option_def bot_option_def null_def invalid_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\\\_nullstrict [simp]: "(null).salary = invalid" by(rule ext, simp add: dot_accessor null_fun_def null_option_def bot_option_def null_def invalid_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\\\_at_pre_nullstrict [simp] : "(null).salary@pre = invalid" by(rule ext, simp add: dot_accessor null_fun_def null_option_def bot_option_def null_def invalid_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\\\_strict [simp] : "(invalid).salary = invalid" by(rule ext, simp add: dot_accessor null_fun_def null_option_def bot_option_def null_def invalid_def) lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\\\_at_pre_strict [simp] : "(invalid).salary@pre = invalid" by(rule ext, simp add: dot_accessor null_fun_def null_option_def bot_option_def null_def invalid_def) subsection\Representation in States\ lemma dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\_def_mono:"\ \ \(X .boss) \ \ \ \(X)" apply(case_tac "\ \ (X \ invalid)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .boss)))" and \ = "\" and x = "X" and y = "invalid"], simp add: foundation16') apply(case_tac "\ \ (X \ null)", insert StrongEq_L_subst2[where P = "(\x. (\ (x .boss)))" and \ = "\" and x = "X" and y = "null"], simp add: foundation16') by(simp add: defined_split) lemma repr_boss: assumes A : "\ \ \(x .boss)" shows "is_represented_in_state in_post_state (x .boss) Person \" apply(insert A[simplified foundation16] A[THEN dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\_def_mono, simplified foundation16]) unfolding is_represented_in_state_def dot\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\_def eval_extract_def select\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\\\_def in_post_state_def by(auto simp: deref_oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def bot_fun_def bot_option_def null_option_def null_fun_def invalid_def OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\_def image_def ran_def split: type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n.split option.split \.split) lemma repr_bossX : assumes A: "\ \ \(x .boss)" shows "\ \ ((Person .allInstances()) ->includes\<^sub>S\<^sub>e\<^sub>t(x .boss))" proof - have B : "\S f. (x .boss) \ \ (Some ` f ` S) \ (x .boss) \ \ (Some ` (f ` S - {None}))" apply(auto simp: image_def ran_def, metis) by(insert A[simplified foundation16], simp add: null_option_def bot_option_def) show ?thesis apply(insert repr_boss[OF A] OclAllInstances_at_post_defined[where H = Person and \ = \]) unfolding is_represented_in_state_def OclValid_def OclAllInstances_at_post_def OclAllInstances_generic_def OclIncludes_def in_post_state_def apply(simp add: A[THEN foundation20, simplified OclValid_def]) apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp, metis bot_option_def option.distinct(1)) by(simp add: image_comp B true_def) qed section\A Little Infra-structure on Example States\ text\ The example we are defining in this section comes from the figure~\ref{fig:edm1_system-states}. \begin{figure} \includegraphics[width=\textwidth]{figures/pre-post.pdf} \caption{(a) pre-state $\sigma_1$ and (b) post-state $\sigma_1'$.} \label{fig:edm1_system-states} \end{figure} \ text_raw\\isatagafp\ definition OclInt1000 ("\\\\") where "OclInt1000 = (\ _ . \\1000\\)" definition OclInt1200 ("\\\\") where "OclInt1200 = (\ _ . \\1200\\)" definition OclInt1300 ("\\\\") where "OclInt1300 = (\ _ . \\1300\\)" definition OclInt1800 ("\\\\") where "OclInt1800 = (\ _ . \\1800\\)" definition OclInt2600 ("\\\\") where "OclInt2600 = (\ _ . \\2600\\)" definition OclInt2900 ("\\\\") where "OclInt2900 = (\ _ . \\2900\\)" definition OclInt3200 ("\\\\") where "OclInt3200 = (\ _ . \\3200\\)" definition OclInt3500 ("\\\\") where "OclInt3500 = (\ _ . \\3500\\)" definition "oid0 \ 0" definition "oid1 \ 1" definition "oid2 \ 2" definition "oid3 \ 3" definition "oid4 \ 4" definition "oid5 \ 5" definition "oid6 \ 6" definition "oid7 \ 7" definition "oid8 \ 8" definition "person1 \ mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid0 \1300\ \oid1\" definition "person2 \ mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid1 \1800\ \oid1\" definition "person3 \ mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid2 None None" definition "person4 \ mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid3 \2900\ None" definition "person5 \ mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid4 \3500\ None" definition "person6 \ mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid5 \2500\ \oid6\" definition "person7 \ mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid6 \(\3200\, \oid6\)\" definition "person8 \ mk\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y oid7 None" definition "person9 \ mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid8 \0\ None" definition - "\\<^sub>1 \ \ heap = Map.empty(oid0 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid0 \1000\ \oid1\)) - (oid1 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid1 \1200\ None)) + "\\<^sub>1 \ \ heap = Map.empty(oid0 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid0 \1000\ \oid1\), + oid1 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid1 \1200\ None), \<^cancel>\oid2\ - (oid3 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid3 \2600\ \oid4\)) - (oid4 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person5) - (oid5 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid5 \2300\ \oid3\)) + oid3 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid3 \2600\ \oid4\), + oid4 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person5, + oid5 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n oid5 \2300\ \oid3\), \<^cancel>\oid6\ \<^cancel>\oid7\ - (oid8 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person9), + oid8 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person9), assocs = Map.empty \" definition - "\\<^sub>1' \ \ heap = Map.empty(oid0 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person1) - (oid1 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person2) - (oid2 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person3) - (oid3 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person4) + "\\<^sub>1' \ \ heap = Map.empty(oid0 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person1, + oid1 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person2, + oid2 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person3, + oid3 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person4, \<^cancel>\oid4\ - (oid5 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person6) - (oid6 \ in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y person7) - (oid7 \ in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y person8) - (oid8 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person9), + oid5 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person6, + oid6 \ in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y person7, + oid7 \ in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y person8, + oid8 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person9), assocs = Map.empty \" definition "\\<^sub>0 \ \ heap = Map.empty, assocs = Map.empty \" lemma basic_\_wff: "WFF(\\<^sub>1,\\<^sub>1')" by(auto simp: WFF_def \\<^sub>1_def \\<^sub>1'_def oid0_def oid1_def oid2_def oid3_def oid4_def oid5_def oid6_def oid7_def oid8_def oid_of_\_def oid_of_type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def oid_of_type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_def person1_def person2_def person3_def person4_def person5_def person6_def person7_def person8_def person9_def) lemma [simp,code_unfold]: "dom (heap \\<^sub>1) = {oid0,oid1\<^cancel>\,oid2\,oid3,oid4,oid5\<^cancel>\,oid6,oid7\,oid8}" by(auto simp: \\<^sub>1_def) lemma [simp,code_unfold]: "dom (heap \\<^sub>1') = {oid0,oid1,oid2,oid3\<^cancel>\,oid4\,oid5,oid6,oid7,oid8}" by(auto simp: \\<^sub>1'_def) text_raw\\isatagafp\ definition "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 :: Person \ \ _ .\\ person1 \\" definition "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 :: Person \ \ _ .\\ person2 \\" definition "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 :: Person \ \ _ .\\ person3 \\" definition "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4 :: Person \ \ _ .\\ person4 \\" definition "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5 :: Person \ \ _ .\\ person5 \\" definition "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 :: Person \ \ _ .\\ person6 \\" definition "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 :: OclAny \ \ _ .\\ person7 \\" definition "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8 :: OclAny \ \ _ .\\ person8 \\" definition "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 :: Person \ \ _ .\\ person9 \\" lemma [code_unfold]: "((x::Person) \ y) = StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t x y" by(simp only: StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) lemma [code_unfold]: "((x::OclAny) \ y) = StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t x y" by(simp only: StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y) lemmas [simp,code_unfold] = OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny OclIsTypeOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny OclIsTypeOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_OclAny OclIsKindOf\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_OclAny OclIsKindOf\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_Person text_raw\\endisatagafp\ Assert "\s\<^sub>p\<^sub>r\<^sub>e . (s\<^sub>p\<^sub>r\<^sub>e,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .salary <> \\\\)" Assert "\s\<^sub>p\<^sub>r\<^sub>e . (s\<^sub>p\<^sub>r\<^sub>e,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .salary \ \\\\)" Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .salary@pre \ \\\\)" Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .salary@pre <> \\\\)" Assert "\s\<^sub>p\<^sub>r\<^sub>e . (s\<^sub>p\<^sub>r\<^sub>e,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .boss <> X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1)" Assert "\s\<^sub>p\<^sub>r\<^sub>e . (s\<^sub>p\<^sub>r\<^sub>e,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .boss .salary \ \\\\)" Assert "\s\<^sub>p\<^sub>r\<^sub>e . (s\<^sub>p\<^sub>r\<^sub>e,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .boss .boss <> X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1)" Assert "\s\<^sub>p\<^sub>r\<^sub>e . (s\<^sub>p\<^sub>r\<^sub>e,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .boss .boss \ X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2)" Assert " (\\<^sub>1,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .boss@pre .salary \ \\\\)" Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .boss@pre .salary@pre \ \\\\)" Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .boss@pre .salary@pre <> \\\\)" Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .boss@pre \ X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2)" Assert " (\\<^sub>1,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .boss@pre .boss \ X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2)" Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .boss@pre .boss@pre \ null)" Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ not(\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .boss@pre .boss@pre .boss@pre))" lemma " (\\<^sub>1,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .oclIsMaintained())" by(simp add: OclValid_def OclIsMaintained_def \\<^sub>1_def \\<^sub>1'_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1_def person1_def oid0_def oid1_def oid2_def oid3_def oid4_def oid5_def oid6_def oid_of_option_def oid_of_type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def) lemma "\s\<^sub>p\<^sub>r\<^sub>e s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (s\<^sub>p\<^sub>r\<^sub>e,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ ((X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .oclAsType(OclAny) .oclAsType(Person)) \ X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1)" by(rule up_down_cast_Person_OclAny_Person', simp add: X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1_def) Assert "\s\<^sub>p\<^sub>r\<^sub>e s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (s\<^sub>p\<^sub>r\<^sub>e,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .oclIsTypeOf(Person))" Assert "\s\<^sub>p\<^sub>r\<^sub>e s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (s\<^sub>p\<^sub>r\<^sub>e,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ not(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .oclIsTypeOf(OclAny))" Assert "\s\<^sub>p\<^sub>r\<^sub>e s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (s\<^sub>p\<^sub>r\<^sub>e,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .oclIsKindOf(Person))" Assert "\s\<^sub>p\<^sub>r\<^sub>e s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (s\<^sub>p\<^sub>r\<^sub>e,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .oclIsKindOf(OclAny))" Assert "\s\<^sub>p\<^sub>r\<^sub>e s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (s\<^sub>p\<^sub>r\<^sub>e,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ not(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .oclAsType(OclAny) .oclIsTypeOf(OclAny))" Assert "\s\<^sub>p\<^sub>r\<^sub>e . (s\<^sub>p\<^sub>r\<^sub>e,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 .salary \ \\\\)" Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 .salary@pre \ \\\\)" Assert "\s\<^sub>p\<^sub>r\<^sub>e . (s\<^sub>p\<^sub>r\<^sub>e,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 .boss \ X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2)" Assert " (\\<^sub>1,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 .boss .salary@pre \ \\\\)" Assert " (\\<^sub>1,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 .boss .boss@pre \ null)" Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 .boss@pre \ null)" Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 .boss@pre <> X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2)" Assert " (\\<^sub>1,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 .boss@pre <> (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 .boss))" Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ not(\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 .boss@pre .boss))" Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ not(\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 .boss@pre .salary@pre))" lemma " (\\<^sub>1,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 .oclIsMaintained())" by(simp add: OclValid_def OclIsMaintained_def \\<^sub>1_def \\<^sub>1'_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2_def person2_def oid0_def oid1_def oid2_def oid3_def oid4_def oid5_def oid6_def oid_of_option_def oid_of_type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def) Assert "\s\<^sub>p\<^sub>r\<^sub>e . (s\<^sub>p\<^sub>r\<^sub>e,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 .salary \ null)" Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ not(\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 .salary@pre))" Assert "\s\<^sub>p\<^sub>r\<^sub>e . (s\<^sub>p\<^sub>r\<^sub>e,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 .boss \ null)" Assert "\s\<^sub>p\<^sub>r\<^sub>e . (s\<^sub>p\<^sub>r\<^sub>e,\\<^sub>1') \ not(\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 .boss .salary))" Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ not(\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 .boss@pre))" lemma " (\\<^sub>1,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 .oclIsNew())" by(simp add: OclValid_def OclIsNew_def \\<^sub>1_def \\<^sub>1'_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3_def person3_def oid0_def oid1_def oid2_def oid3_def oid4_def oid5_def oid6_def oid8_def oid_of_option_def oid_of_type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def) Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4 .boss@pre \ X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5)" Assert " (\\<^sub>1,\\<^sub>1') \ not(\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4 .boss@pre .salary))" Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4 .boss@pre .salary@pre \ \\\\)" lemma " (\\<^sub>1,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4 .oclIsMaintained())" by(simp add: OclValid_def OclIsMaintained_def \\<^sub>1_def \\<^sub>1'_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4_def person4_def oid0_def oid1_def oid2_def oid3_def oid4_def oid5_def oid6_def oid_of_option_def oid_of_type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def) Assert "\s\<^sub>p\<^sub>r\<^sub>e . (s\<^sub>p\<^sub>r\<^sub>e,\\<^sub>1') \ not(\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5 .salary))" Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5 .salary@pre \ \\\\)" Assert "\s\<^sub>p\<^sub>r\<^sub>e . (s\<^sub>p\<^sub>r\<^sub>e,\\<^sub>1') \ not(\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5 .boss))" lemma " (\\<^sub>1,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5 .oclIsDeleted())" by(simp add: OclNot_def OclValid_def OclIsDeleted_def \\<^sub>1_def \\<^sub>1'_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5_def person5_def oid0_def oid1_def oid2_def oid3_def oid4_def oid5_def oid6_def oid7_def oid8_def oid_of_option_def oid_of_type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def) (* (* access to an oclany object not yet supported *) Assert " (\\<^sub>1,\\<^sub>1') \ ((X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 .boss .salary) \ \\\\ )"*) Assert "\s\<^sub>p\<^sub>r\<^sub>e . (s\<^sub>p\<^sub>r\<^sub>e,\\<^sub>1') \ not(\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 .boss .salary@pre))" Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 .boss@pre \ X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4)" Assert " (\\<^sub>1,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 .boss@pre .salary \ \\\\)" Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 .boss@pre .salary@pre \ \\\\)" Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 .boss@pre .boss@pre \ X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5)" lemma " (\\<^sub>1,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 .oclIsMaintained())" by(simp add: OclValid_def OclIsMaintained_def \\<^sub>1_def \\<^sub>1'_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6_def person6_def oid0_def oid1_def oid2_def oid3_def oid4_def oid5_def oid6_def oid_of_option_def oid_of_type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def) (* (* access to an oclany object not yet supported *) Assert " (\\<^sub>1,\\<^sub>1') \ ((X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 .oclAsType(Person) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 .boss)))" *) (* (* access to an oclany object not yet supported *) Assert " (\\<^sub>1,\\<^sub>1') \ ((X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 .oclAsType(Person) .boss) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 .oclAsType(Person)) )" *) (* (* access to an oclany object not yet supported *) Assert " (\\<^sub>1,\\<^sub>1') \ ((X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 .oclAsType(Person) .boss .salary) \ \\\\ )" *) Assert "\s\<^sub>p\<^sub>r\<^sub>e s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (s\<^sub>p\<^sub>r\<^sub>e,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ \(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 .oclAsType(Person))" Assert "\ s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (\\<^sub>1,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ not(\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 .oclAsType(Person) .boss@pre))" lemma "\s\<^sub>p\<^sub>r\<^sub>e s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (s\<^sub>p\<^sub>r\<^sub>e,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ ((X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 .oclAsType(Person) .oclAsType(OclAny) .oclAsType(Person)) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 .oclAsType(Person)))" by(rule up_down_cast_Person_OclAny_Person', simp add: X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7_def OclValid_def valid_def person7_def) lemma " (\\<^sub>1,\\<^sub>1') \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 .oclIsNew())" by(simp add: OclValid_def OclIsNew_def \\<^sub>1_def \\<^sub>1'_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7_def person7_def oid0_def oid1_def oid2_def oid3_def oid4_def oid5_def oid6_def oid8_def oid_of_option_def oid_of_type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_def) Assert "\s\<^sub>p\<^sub>r\<^sub>e s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (s\<^sub>p\<^sub>r\<^sub>e,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8 <> X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7)" Assert "\s\<^sub>p\<^sub>r\<^sub>e s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (s\<^sub>p\<^sub>r\<^sub>e,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ not(\(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8 .oclAsType(Person)))" Assert "\s\<^sub>p\<^sub>r\<^sub>e s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (s\<^sub>p\<^sub>r\<^sub>e,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8 .oclIsTypeOf(OclAny))" Assert "\s\<^sub>p\<^sub>r\<^sub>e s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (s\<^sub>p\<^sub>r\<^sub>e,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ not(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8 .oclIsTypeOf(Person))" Assert "\s\<^sub>p\<^sub>r\<^sub>e s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (s\<^sub>p\<^sub>r\<^sub>e,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ not(X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8 .oclIsKindOf(Person))" Assert "\s\<^sub>p\<^sub>r\<^sub>e s\<^sub>p\<^sub>o\<^sub>s\<^sub>t. (s\<^sub>p\<^sub>r\<^sub>e,s\<^sub>p\<^sub>o\<^sub>s\<^sub>t) \ (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8 .oclIsKindOf(OclAny))" lemma \_modifiedonly: "(\\<^sub>1,\\<^sub>1') \ (Set{ X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .oclAsType(OclAny) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 .oclAsType(OclAny) \<^cancel>\, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 .oclAsType(OclAny)\ , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4 .oclAsType(OclAny) \<^cancel>\, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5 .oclAsType(OclAny)\ , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 .oclAsType(OclAny) \<^cancel>\, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 .oclAsType(OclAny)\ \<^cancel>\, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8 .oclAsType(OclAny)\ \<^cancel>\, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 .oclAsType(OclAny)\}->oclIsModifiedOnly())" apply(simp add: OclIsModifiedOnly_def OclValid_def oid0_def oid1_def oid2_def oid3_def oid4_def oid5_def oid6_def oid7_def oid8_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9_def person1_def person2_def person3_def person4_def person5_def person6_def person7_def person8_def person9_def image_def) apply(simp add: OclIncluding_rep_set mtSet_rep_set null_option_def bot_option_def) apply(simp add: oid_of_option_def oid_of_type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_def, clarsimp) apply(simp add: \\<^sub>1_def \\<^sub>1'_def oid0_def oid1_def oid2_def oid3_def oid4_def oid5_def oid6_def oid7_def oid8_def) done lemma "(\\<^sub>1,\\<^sub>1') \ ((X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 @pre (\x. \OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\ x\)) \ X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9)" by(simp add: OclSelf_at_pre_def \\<^sub>1_def oid_of_option_def oid_of_type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9_def person9_def oid8_def OclValid_def StrongEq_def OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\_def) lemma "(\\<^sub>1,\\<^sub>1') \ ((X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 @post (\x. \OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\ x\)) \ X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9)" by(simp add: OclSelf_at_post_def \\<^sub>1'_def oid_of_option_def oid_of_type\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9_def person9_def oid8_def OclValid_def StrongEq_def OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\_def) lemma "(\\<^sub>1,\\<^sub>1') \ (((X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 .oclAsType(OclAny)) @pre (\x. \OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\ x\)) \ ((X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 .oclAsType(OclAny)) @post (\x. \OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\ x\)))" proof - have including4 : "\a b c d \. Set{\\. \\a\\, \\. \\b\\, \\. \\c\\, \\. \\d\\} \ = Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\ {\\a\\, \\b\\, \\c\\, \\d\\} \\" apply(subst abs_rep_simp'[symmetric], simp) apply(simp add: OclIncluding_rep_set mtSet_rep_set) by(rule arg_cong[of _ _ "\x. (Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e(\\ x \\))"], auto) have excluding1: "\S a b c d e \. (\_. Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\ {\\a\\, \\b\\, \\c\\, \\d\\} \\)->excluding\<^sub>S\<^sub>e\<^sub>t(\\. \\e\\) \ = Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e \\ {\\a\\, \\b\\, \\c\\, \\d\\} - {\\e\\} \\" apply(simp add: UML_Set.OclExcluding_def) apply(simp add: defined_def OclValid_def false_def true_def bot_fun_def bot_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def null_fun_def null_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_def) apply(rule conjI) apply(rule impI, subst (asm) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject) apply( simp add: bot_option_def)+ apply(rule conjI) apply(rule impI, subst (asm) Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inject) apply( simp add: bot_option_def null_option_def)+ apply(subst Abs_Set\<^sub>b\<^sub>a\<^sub>s\<^sub>e_inverse, simp add: bot_option_def, simp) done show ?thesis apply(rule framing[where X = "Set{ X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .oclAsType(OclAny) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 .oclAsType(OclAny) \<^cancel>\, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 .oclAsType(OclAny)\ , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4 .oclAsType(OclAny) \<^cancel>\, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5 .oclAsType(OclAny)\ , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 .oclAsType(OclAny) \<^cancel>\, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 .oclAsType(OclAny)\ \<^cancel>\, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8 .oclAsType(OclAny)\ \<^cancel>\, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 .oclAsType(OclAny)\}"]) apply(cut_tac \_modifiedonly) apply(simp only: OclValid_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9_def person1_def person2_def person3_def person4_def person5_def person6_def person7_def person8_def person9_def OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_Person) apply(subst cp_OclIsModifiedOnly, subst UML_Set.OclExcluding.cp0, subst (asm) cp_OclIsModifiedOnly, simp add: including4 excluding1) apply(simp only: X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9_def person1_def person2_def person3_def person4_def person5_def person6_def person7_def person8_def person9_def) apply(simp add: OclIncluding_rep_set mtSet_rep_set oid0_def oid1_def oid2_def oid3_def oid4_def oid5_def oid6_def oid7_def oid8_def) apply(simp add: StrictRefEq\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def oid_of_option_def oid_of_type\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_def OclNot_def OclValid_def null_option_def bot_option_def) done qed lemma perm_\\<^sub>1' : "\\<^sub>1' = \ heap = Map.empty - (oid8 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person9) - (oid7 \ in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y person8) - (oid6 \ in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y person7) - (oid5 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person6) + (oid8 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person9, + oid7 \ in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y person8, + oid6 \ in\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y person7, + oid5 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person6, \<^cancel>\oid4\ - (oid3 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person4) - (oid2 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person3) - (oid1 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person2) - (oid0 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person1) + oid3 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person4, + oid2 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person3, + oid1 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person2, + oid0 \ in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n person1) , assocs = assocs \\<^sub>1' \" proof - note P = fun_upd_twist show ?thesis apply(simp add: \\<^sub>1'_def oid0_def oid1_def oid2_def oid3_def oid4_def oid5_def oid6_def oid7_def oid8_def) apply(subst (1) P, simp) apply(subst (2) P, simp) apply(subst (1) P, simp) apply(subst (3) P, simp) apply(subst (2) P, simp) apply(subst (1) P, simp) apply(subst (4) P, simp) apply(subst (3) P, simp) apply(subst (2) P, simp) apply(subst (1) P, simp) apply(subst (5) P, simp) apply(subst (4) P, simp) apply(subst (3) P, simp) apply(subst (2) P, simp) apply(subst (1) P, simp) apply(subst (6) P, simp) apply(subst (5) P, simp) apply(subst (4) P, simp) apply(subst (3) P, simp) apply(subst (2) P, simp) apply(subst (1) P, simp) apply(subst (7) P, simp) apply(subst (6) P, simp) apply(subst (5) P, simp) apply(subst (4) P, simp) apply(subst (3) P, simp) apply(subst (2) P, simp) apply(subst (1) P, simp) by(simp) qed declare const_ss [simp] lemma "\\\<^sub>1. (\\<^sub>1,\\<^sub>1') \ (Person .allInstances() \ Set{ X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4\<^cancel>\, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7 .oclAsType(Person)\<^cancel>\, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8\, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 })" apply(subst perm_\\<^sub>1') apply(simp only: oid0_def oid1_def oid2_def oid3_def oid4_def oid5_def oid6_def oid7_def oid8_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9_def person7_def) apply(subst state_update_vs_allInstances_at_post_tc, simp, simp add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\_def, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp, simp, simp, rule OclIncluding_cong, simp, simp) apply(subst state_update_vs_allInstances_at_post_tc, simp, simp add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\_def, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp, simp, simp, rule OclIncluding_cong, simp, simp) apply(subst state_update_vs_allInstances_at_post_tc, simp, simp add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\_def, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp, simp, simp, rule OclIncluding_cong, simp, simp) apply(subst state_update_vs_allInstances_at_post_tc, simp, simp add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\_def, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp, simp, simp, rule OclIncluding_cong, simp, simp) apply(subst state_update_vs_allInstances_at_post_tc, simp, simp add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\_def, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp, simp, simp, rule OclIncluding_cong, simp, simp) apply(subst state_update_vs_allInstances_at_post_tc, simp, simp add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\_def, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp, simp, simp, rule OclIncluding_cong, simp, simp) apply(subst state_update_vs_allInstances_at_post_ntc, simp, simp add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\_def person8_def, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp, simp, simp) apply(subst state_update_vs_allInstances_at_post_tc, simp, simp add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\_def, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp, simp, simp, rule OclIncluding_cong, simp, simp) apply(rule state_update_vs_allInstances_at_post_empty) by(simp_all add: OclAsType\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\_def) lemma "\\\<^sub>1. (\\<^sub>1,\\<^sub>1') \ (OclAny .allInstances() \ Set{ X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 .oclAsType(OclAny), X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 .oclAsType(OclAny), X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 .oclAsType(OclAny), X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4 .oclAsType(OclAny) \<^cancel>\, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5\, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6 .oclAsType(OclAny), X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9 .oclAsType(OclAny) })" apply(subst perm_\\<^sub>1') apply(simp only: oid0_def oid1_def oid2_def oid3_def oid4_def oid5_def oid6_def oid7_def oid8_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n4_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n5_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n6_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n7_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n8_def X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n9_def person1_def person2_def person3_def person4_def person5_def person6_def person9_def) apply(subst state_update_vs_allInstances_at_post_tc, simp, simp add: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\_def, simp, rule const_StrictRefEq\<^sub>S\<^sub>e\<^sub>t_including, simp, simp, simp, rule OclIncluding_cong, simp, simp)+ apply(rule state_update_vs_allInstances_at_post_empty) by(simp_all add: OclAsType\<^sub>O\<^sub>c\<^sub>l\<^sub>A\<^sub>n\<^sub>y_\_def) end diff --git a/thys/Hoare_Time/Big_StepT_Partial.thy b/thys/Hoare_Time/Big_StepT_Partial.thy --- a/thys/Hoare_Time/Big_StepT_Partial.thy +++ b/thys/Hoare_Time/Big_StepT_Partial.thy @@ -1,500 +1,500 @@ subsection "Big step Semantics on partial states" theory Big_StepT_Partial imports Partial_Evaluation Big_StepT "SepLogAdd/Sep_Algebra_Add" "HOL-Eisbach.Eisbach" begin type_synonym lvname = string type_synonym pstate_t = "partstate * nat" type_synonym assnp = "partstate \ bool" type_synonym assn2 = "pstate_t \ bool" subsubsection \helper functions\ paragraph \restrict\ definition restrict where "restrict S s = (%x. if x:S then Some (s x) else None)" lemma restrictI: "\x\S. s1 x = s2 x \ restrict S s1 = restrict S s2" unfolding restrict_def by fastforce lemma restrictE: "restrict S s1 = restrict S s2 \ s1 = s2 on S" unfolding restrict_def by (meson option.inject) lemma dom_restrict[simp]: "dom (restrict S s) = S" unfolding restrict_def using domIff by fastforce lemma restrict_less_part: "restrict S t \ part t" unfolding restrict_def map_le_substate_conv[symmetric] map_le_def part_def apply auto by (metis option.simps(3)) paragraph \Heap helper functions\ fun lmaps_to_expr :: "aexp \ val \ assn2" where "lmaps_to_expr a v = (%(s,c). dom s = vars a \ paval a s = v \ c = 0)" fun lmaps_to_expr_x :: "vname \ aexp \ val \ assn2" where "lmaps_to_expr_x x a v = (%(s,c). dom s = vars a \ {x} \ paval a s = v \ c = 0)" lemma subState: "x \ y \ v \ dom x \ x v = y v" unfolding map_le_substate_conv[symmetric] map_le_def by blast lemma fixes ps:: partstate and s::state assumes "vars a \ dom ps" "ps \ part s" shows emb_update: "emb [x \ paval a ps] s = (emb ps s) (x := aval a (emb ps s))" using assms unfolding emb_def apply auto apply (rule ext) apply(case_tac "v=x") apply(simp add: paval_aval) apply(simp) unfolding part_def apply(case_tac "v \ dom ps") using subState apply fastforce by (simp add: domIff) lemma paval_aval2: "vars a \ dom ps \ ps \ part s \ paval a ps = aval a s" apply(induct a) using subState unfolding part_def apply auto by fastforce lemma fixes ps:: partstate and s::state assumes "vars a \ dom ps" "ps \ part s" shows emb_update2: "emb (ps(x \ paval a ps)) s = (emb ps s)(x := aval a (emb ps s))" using assms unfolding emb_def apply auto apply (rule ext) apply(case_tac "v=x") apply(simp add: paval_aval) by (simp) subsubsection "Big step Semantics on partial states" inductive big_step_t_part :: "com \ partstate \ nat \ partstate \ bool" ("_ \\<^sub>A _ \ _" 55) where Skip: "(SKIP,s) \\<^sub>A Suc 0 \ s" | Assign: "\ vars a \ {x} \ dom ps; paval a ps = v ; ps' = ps(x \ v) \ \ (x ::= a,ps) \\<^sub>A Suc 0 \ ps'" | Seq: "\ (c1,s1) \\<^sub>A x \ s2; (c2,s2) \\<^sub>A y \ s3 ; z=x+y \ \ (c1;;c2, s1) \\<^sub>A z \ s3" | IfTrue: "\ vars b \ dom ps ; dom ps' = dom ps ; pbval b ps; (c1,ps) \\<^sub>A x \ ps'; y=x+1 \ \ (IF b THEN c1 ELSE c2, ps) \\<^sub>A y \ ps'" | IfFalse: "\ vars b \ dom ps ; dom ps' = dom ps ; \ pbval b ps; (c2,ps) \\<^sub>A x \ ps'; y=x+1 \ \ (IF b THEN c1 ELSE c2, ps) \\<^sub>A y \ ps'" | WhileFalse: "\ vars b \ dom s; \ pbval b s \ \ (WHILE b DO c,s) \\<^sub>A Suc 0 \ s" | WhileTrue: "\ pbval b s1; vars b \ dom s1; (c,s1) \\<^sub>A x \ s2; (WHILE b DO c, s2) \\<^sub>A y \ s3; 1+x+y=z \ \ (WHILE b DO c, s1) \\<^sub>A z \ s3" declare big_step_t_part.intros [intro] inductive_cases Skip_tE3[elim!]: "(SKIP,s) \\<^sub>A x \ t" thm Skip_tE3 inductive_cases Assign_tE3[elim!]: "(x ::= a,s) \\<^sub>A p \ t" thm Assign_tE3 inductive_cases Seq_tE3[elim!]: "(c1;;c2,s1) \\<^sub>A p \ s3" thm Seq_tE3 inductive_cases If_tE3[elim!]: "(IF b THEN c1 ELSE c2,s) \\<^sub>A x \ t" thm If_tE3 inductive_cases While_tE3[elim]: "(WHILE b DO c,s) \\<^sub>A x \ t" thm While_tE3 lemmas big_step_t_part_induct = big_step_t_part.induct[split_format(complete)] lemma big_step_t3_post_dom_conv: "(c,ps) \\<^sub>A t \ ps' \ dom ps' = dom ps" apply(induct rule: big_step_t_part_induct) apply (auto simp: sep_disj_fun_def plus_fun_def) apply metis done lemma add_update_distrib: "ps1 x1 = Some y \ ps1 ## ps2 \ vars x2 \ dom ps1 \ ps1(x1 \ paval x2 ps1) + ps2 = (ps1 + ps2)(x1 \ paval x2 ps1)" apply (rule ext) apply (auto simp: sep_disj_fun_def plus_fun_def) by (metis disjoint_iff_not_equal domI domain_conv) lemma paval_extend: "ps1 ## ps2 \ vars a \ dom ps1 \ paval a (ps1 + ps2) = paval a ps1" apply(induct a) apply (auto simp: sep_disj_fun_def domain_conv) by (metis domI map_add_comm map_add_dom_app_simps(1) option.sel plus_fun_conv) lemma pbval_extend: "ps1 ## ps2 \ vars b \ dom ps1 \ pbval b (ps1 + ps2) = pbval b ps1" apply(induct b) by (auto simp: paval_extend) lemma Framer: "(C, ps1) \\<^sub>A m \ ps1' \ ps1 ## ps2 \ (C, ps1 + ps2) \\<^sub>A m \ ps1'+ps2" proof (induct rule: big_step_t_part_induct) case (Skip s) then show ?case by (auto simp: big_step_t_part.Skip) next case (Assign a x ps v ps') show ?case apply(rule big_step_t_part.Assign) using Assign apply (auto simp: plus_fun_def) apply(rule ext) apply(case_tac "xa=x") subgoal apply auto subgoal using paval_extend[unfolded plus_fun_def] by auto unfolding sep_disj_fun_def by (metis disjoint_iff_not_equal domI domain_conv) subgoal by auto done next case (IfTrue b ps ps' c1 x y c2) then show ?case apply (auto ) apply(subst big_step_t_part.IfTrue) apply (auto simp: pbval_extend) subgoal by (auto simp: plus_fun_def) subgoal by (auto simp: plus_fun_def) subgoal by (auto simp: plus_fun_def) done next case (IfFalse b ps ps' c2 x y c1) then show ?case apply (auto ) apply(subst big_step_t_part.IfFalse) apply (auto simp: pbval_extend) subgoal by (auto simp: plus_fun_def) subgoal by (auto simp: plus_fun_def) subgoal by (auto simp: plus_fun_def) done next case (WhileFalse b s c) then show ?case apply(subst big_step_t_part.WhileFalse) subgoal by (auto simp: plus_fun_def) subgoal by (auto simp: pbval_extend) by auto next case (WhileTrue b s1 c x s2 y s3 z) from big_step_t3_post_dom_conv[OF WhileTrue(3)] have "dom s2 = dom s1" by auto with WhileTrue(8) have "s2 ## ps2" unfolding sep_disj_fun_def domain_conv by auto with WhileTrue show ?case apply auto apply(subst big_step_t_part.WhileTrue) subgoal by (auto simp: pbval_extend) subgoal by (auto simp: plus_fun_def) apply (auto) done next case (Seq c1 s1 x s2 c2 y s3 z) from big_step_t3_post_dom_conv[OF Seq(1)] have "dom s2 = dom s1" by auto with Seq(6) have "s2 ## ps2" unfolding sep_disj_fun_def domain_conv by auto with Seq show ?case apply (subst big_step_t_part.Seq) by auto qed lemma Framer2: "(C, ps1) \\<^sub>A m \ ps1' \ ps1 ## ps2 \ ps = ps1 + ps2 \ ps' = ps1'+ps2 \ (C, ps) \\<^sub>A m \ ps' " using Framer by auto (* connection to bigstep2 *) subsubsection \Relation to BigStep Semantic on full states\ lemma paval_aval_part: "paval a (part s) = aval a s" apply(induct a) by (auto simp: part_def) lemma pbval_bval_part: "pbval b (part s) = bval b s" apply(induct b) by (auto simp: paval_aval_part) -lemma part_paval_aval: "part (s(x := aval a s)) = part s(x \ paval a (part s))" +lemma part_paval_aval: "part (s(x := aval a s)) = (part s)(x \ paval a (part s))" apply(rule ext) apply(case_tac "xa=x") unfolding part_def apply auto by (metis (full_types) domIff map_le_def map_le_substate_conv option.distinct(1) part_def paval_aval2 subsetI) lemma full_to_part: "(C, s) \ m \ s' \ (C, part s) \\<^sub>A m \ part s' " apply(induct rule: big_step_t_induct) using Skip apply simp apply (subst Assign) using part_paval_aval apply(simp_all add: ) apply(rule Seq) apply auto apply(rule IfTrue) apply (auto simp: pbval_bval_part) apply(rule IfFalse) apply (auto simp: pbval_bval_part) apply(rule WhileFalse) apply (auto simp: pbval_bval_part) apply(rule WhileTrue) apply (auto simp: pbval_bval_part) done lemma part_to_full': "(C, ps) \\<^sub>A m \ ps' \ (C, emb ps s) \ m \ emb ps' s" proof (induct rule: big_step_t_part_induct) case (Assign a x ps v ps') have z: "paval a ps = aval a (emb ps s)" apply(rule paval_aval_vars) using Assign(1) by auto have g :"emb ps' s = (emb ps s)(x:=aval a (emb ps s) )" apply(simp only: Assign z[symmetric]) unfolding emb_def by auto show ?case apply(simp only: g) by(rule big_step_t.Assign) qed (auto simp: pbval_bval_vars[symmetric]) lemma part_to_full: "(C, part s) \\<^sub>A m \ part s' \ (C, s) \ m \ s'" proof - assume "(C, part s) \\<^sub>A m \ part s'" then have "(C, emb (part s) s) \ m \ emb (part s') s" by (rule part_to_full') then show "(C, s) \ m \ s'" by auto qed lemma part_full_equiv: "(C, s) \ m \ s' \ (C, part s) \\<^sub>A m \ part s'" using part_to_full full_to_part by metis subsubsection \more properties\ lemma big_step_t3_gt0: "(C, ps) \\<^sub>A x \ ps' \ x > 0" apply(induct rule: big_step_t_part_induct) apply auto done lemma big_step_t3_same: "(C, ps) \\<^sub>A m \ ps' ==> ps = ps' on UNIV - lvars C" apply(induct rule: big_step_t_part_induct) by (auto simp: sep_disj_fun_def plus_fun_def) lemma avalDirekt3_correct: " (x ::= N v, ps) \\<^sub>A m \ ps' \ paval' a ps = Some v \ (x ::= a, ps) \\<^sub>A m \ ps'" apply(auto) apply(subst Assign) by (auto simp: paval_paval' paval'dom) subsection \Partial State\ (* partialstate and nat is a separation algebra ! *) lemma fixes h :: "(vname \ val option) * nat" shows "(P ** Q ** H) h = (Q ** H ** P) h" by (simp add: sep_conj_ac) lemma separate_othogonal_commuted': assumes "\ps n. P (ps,n) \ ps = 0" "\ps n. Q (ps,n) \ n = 0" shows "(P ** Q) s \ P (0,snd s) \ Q (fst s,0)" using assms unfolding sep_conj_def by force lemma separate_othogonal_commuted: assumes "\ps n. P (ps,n) \ ps = 0" "\ps n. Q (ps,n) \ n = 0" shows "(P ** Q) (ps,n) \ P (0,n) \ Q (ps,0)" using assms unfolding sep_conj_def by force lemma separate_othogonal: assumes "\ps n. P (ps,n) \ n = 0" "\ps n. Q (ps,n) \ ps = 0" shows "(P ** Q) (ps,n) \ P (ps,0) \ Q (0,n)" using assms unfolding sep_conj_def by force lemma assumes " ((\(s, n). P (s, n) \ vars b \ dom s) \* (\(s, c). s = 0 \ c = Suc 0)) (ps, n)" shows "\ n'. P (ps, n') \ vars b \ dom ps \ n = Suc n'" proof - from assms obtain x y where " x ## y" and "(ps, n) = x + y" and 2: "(case x of (s, n) \ P (s, n) \ vars b \ dom s)" and "(case y of (s, c) \ s = 0 \ c = Suc 0)" unfolding sep_conj_def by blast then have "y = (0, Suc 0)" and f: "fst x = ps" and n: "n = snd x + Suc 0" by auto with 2 have "P (ps, snd x) \ vars b \ dom ps \ n = Suc (snd x)" by auto then show ?thesis by simp qed subsection \Dollar and Pointsto\ definition dollar :: "nat \ assn2" ("$") where "dollar q = (%(s,c). s = 0 \ c=q)" lemma sep_reorder_dollar_aux: "NO_MATCH ($X) A \ ($B ** A) = (A ** $B)" "($X ** $Y) = $(X+Y)" apply (auto simp: sep_simplify) unfolding dollar_def sep_conj_def sep_disj_prod_def sep_disj_nat_def by auto lemmas sep_reorder_dollar = sep_conj_assoc sep_reorder_dollar_aux lemma stardiff: assumes "(P \* $m) (ps, n)" shows P: "P (ps, n - m)" and "m\n" using assms unfolding sep_conj_def dollar_def by auto lemma [simp]: "(Q ** $0) = Q" unfolding dollar_def sep_conj_def sep_disj_prod_def sep_disj_nat_def by auto definition embP :: "(partstate \ bool) \ partstate \ nat \ bool" where "embP P = (%(s,n). P s \ n = 0)" lemma orthogonal_split: assumes "(embP Q \* $ n) = (embP P \* $ m)" shows "(Q = P \ n = m) \ Q = (\s. False) \ P = (\s. False)" using assms unfolding embP_def dollar_def apply (auto intro!: ext) unfolding sep_conj_def apply auto unfolding sep_disj_prod_def plus_prod_def apply (metis fst_conv snd_conv)+ done (* how to set up case rules *) lemma F: assumes "(embP Q \* $ n) = (embP P \* $ m)" obtains (blub) "Q = P" and "n = m" | (da) "Q = (\s. False)" and "P = (\s. False)" using assms orthogonal_split by auto lemma T: assumes "(embP Q \* $ n) = (embP P \* $ m)" obtains (blub) x::nat where "Q = P" and "n = m" and "x=x" | (da) "Q = (\s. False)" and "P = (\s. False)" using assms orthogonal_split by auto definition pointsto :: "vname \ val \ assn2" ("_ \ _" [56,51] 56) where "v \ n = (%(s,c). s = [v \ n] \ c=0)" (* If you don't mind syntax ambiguity: *) notation pred_ex (binder "\" 10) definition maps_to_ex :: "vname \ assn2" ("_ \ -" [56] 56) where "x \ - \ \y. x \ y" fun lmaps_to_ex :: "vname set \ assn2" where "lmaps_to_ex xs = (%(s,c). dom s = xs \ c = 0)" lemma "(x \ -) (s,n) \ x \ dom s" unfolding maps_to_ex_def pointsto_def by auto fun lmaps_to_axpr :: "bexp \ bool \ assnp" where "lmaps_to_axpr b bv = (%ps. vars b \ dom ps \ pbval b ps = bv )" definition lmaps_to_axpr' :: "bexp \ bool \ assnp" where "lmaps_to_axpr' b bv = lmaps_to_axpr b bv" subsection \Frame Inference\ definition Frame where "Frame P Q F \ \s. (P imp (Q ** F)) s" definition Frame' where "Frame' P P' Q F \ \s. (( P' ** P) imp (Q ** F)) s" definition cnv where "cnv x y == x = y" lemma cnv_I: "cnv x x" unfolding cnv_def by simp lemma Frame'_conv: "Frame P Q F = Frame' (P ** \) \ (Q ** \) F" unfolding Frame_def Frame'_def apply auto done lemma Frame'I: "Frame' (P ** \) \ (Q ** \) F \ cnv F F' \ Frame P Q F'" unfolding Frame_def Frame'_def cnv_def apply auto done lemma FrameD: assumes "Frame P Q F" " P s " shows "(F ** Q) s" using assms unfolding Frame_def by (auto simp: sep_conj_commute) lemma Frame'_match: "Frame' (P ** P') \ Q F \ Frame' (x \ v ** P) P' (x \ v ** Q) F" unfolding Frame_def Frame'_def apply (auto simp: sep_conj_ac) by (metis (no_types, opaque_lifting) prod.collapse sep.mult_assoc sep_conj_impl1) lemma R: assumes "\s. (A imp B) s" shows "((A ** $n) imp (B ** $n)) s" proof (safe) assume "(A \* $ n) s" then obtain h1 h2 where A: "A h1" and n: "$n h2" and disj: "h1 ## h2" "s = h1+h2" unfolding sep_conj_def by blast from assms A have B: "B h1" by auto show "(B ** $n) s" using B n disj unfolding sep_conj_def by blast qed lemma Frame'_matchdollar: assumes "Frame' (P ** P' ** $(n-m)) \ Q F" and nm: "n\m" shows "Frame' ($n ** P) P' ($m ** Q) F" using assms(1) unfolding Frame_def Frame'_def apply (auto simp: sep_conj_ac) proof (goal_cases) case (1 a b) have g: "((P \* P' \* $ n) imp (F \* Q \* $ m)) (a, b) \ (((P \* P' \* $(n-m)) ** $m) imp ((F \* Q) \* $ m)) (a, b)" by(simp add: nm sep_reorder_dollar) have "((P \* P' \* $ n) imp (F \* Q \* $ m)) (a, b)" apply(subst g) apply(rule R) using 1(1) by auto then have "(P \* P' \* $ n) (a, b) \ (F \* Q \* $ m) (a, b)" by blast then show ?case using 1(2) by auto qed lemma Frame'_nomatch: " Frame' P (p ** P') (x \ v ** Q) F\ Frame' (p ** P) P' (x \ v ** Q) F" unfolding Frame'_def by (auto simp: sep_conj_ac) lemma Frame'_nomatchempty: " Frame' P P' (x \ v ** Q) F\ Frame' (\ ** P) P' (x \ v ** Q) F" unfolding Frame'_def by (auto simp: sep_conj_ac) (* this will only be reached after a Frame'_match move, where P' is \ *) lemma Frame'_end: " Frame' P \ \ P" unfolding Frame'_def by (auto simp: sep_conj_ac) schematic_goal "Frame (x \ v1 \* y \ v2) (x \ ?v) ?F" apply(rule Frame'I) apply(simp only: sep_conj_assoc) apply(rule Frame'_match) apply(rule Frame'_end) apply(simp only: sep_conj_ac sep_conj_empty' sep_conj_empty) apply(rule cnv_I) done schematic_goal "Frame (x \ v1 \* y \ v2) (y \ ?v) ?F" apply(rule Frame'I) apply(simp only: sep_conj_assoc) apply(rule Frame'_end Frame'_match Frame'_nomatchempty Frame'_nomatch; (simp only: sep_conj_assoc)?)+ apply(simp only: sep_conj_ac sep_conj_empty' sep_conj_empty) apply(rule cnv_I) done method frame_inference_init = (rule Frame'I, (simp only: sep_conj_assoc)?) method frame_inference_solve = (rule Frame'_matchdollar Frame'_end Frame'_match Frame'_nomatchempty Frame'_nomatch; (simp only: sep_conj_assoc)?)+ method frame_inference_cleanup = ( (simp only: sep_conj_ac sep_conj_empty' sep_conj_empty)?; rule cnv_I) method frame_inference = (frame_inference_init, (frame_inference_solve; fail), (frame_inference_cleanup; fail)) method frame_inference_debug = (frame_inference_init, frame_inference_solve) subsubsection \tests\ schematic_goal "Frame (x \ v1 \* y \ v2) (y \ ?v) ?F" by frame_inference schematic_goal "Frame (x \ v1 ** P ** \ ** y \ v2 \* z \ v2 ** Q) (z \ ?v ** y \ ?v2) ?F" by frame_inference (* with dollar *) schematic_goal " 1 \ v \ Frame ($ (2 * v) \* ''x'' \ int v) ($ 1 \* ''x'' \ ?d) ?F" apply(rule Frame'I) apply(simp only: sep_conj_assoc) apply(rule Frame'_matchdollar Frame'_end Frame'_match Frame'_nomatchempty Frame'_nomatch; (simp only: sep_conj_assoc)?)+ apply (simp only: sep_conj_ac sep_conj_empty' sep_conj_empty)? apply (rule cnv_I) done schematic_goal " 0 < v \ Frame ($ (2 * v) \* ''x'' \ int v) ($ 1 \* ''x'' \ ?d) ?F" apply frame_inference done subsection \Expression evaluation\ definition symeval where "symeval P e v \ (\s n. P (s,n) \ paval' e s = Some v)" definition symevalb where "symevalb P e v \ (\s n. P (s,n) \ pbval' e s = Some v)" lemma symeval_c: "symeval P (N v) v" unfolding symeval_def apply auto done lemma symeval_v: assumes "Frame P (x \ v) F" shows "symeval P (V x) v" unfolding symeval_def apply auto apply (drule FrameD[OF assms]) unfolding sep_conj_def pointsto_def apply (auto simp: plus_fun_conv) done lemma symeval_plus: assumes "symeval P e1 v1" "symeval P e2 v2" shows "symeval P (Plus e1 e2) (v1+v2)" using assms unfolding symeval_def by auto lemma symevalb_c: "symevalb P (Bc v) v" unfolding symevalb_def apply auto done lemma symevalb_and: assumes "symevalb P e1 v1" "symevalb P e2 v2" shows "symevalb P (And e1 e2) (v1 \ v2)" using assms unfolding symevalb_def by auto lemma symevalb_not: assumes "symevalb P e v" shows "symevalb P (Not e) (\ v)" using assms unfolding symevalb_def by auto lemma symevalb_less: assumes "symeval P e1 v1" "symeval P e2 v2" shows "symevalb P (Less e1 e2) (v1 < v2)" using assms unfolding symevalb_def symeval_def by auto lemmas symeval = symeval_c symeval_v symeval_plus symevalb_c symevalb_and symevalb_not symevalb_less schematic_goal "symevalb ( (x \ v1) ** (y \ v2) ) (Less (Plus (V x) (V y)) (N 5)) ?g" apply(rule symeval | frame_inference)+ done end \ No newline at end of file diff --git a/thys/Isabelle_Meta_Model/toy_example/document_generated/Design_generated.thy b/thys/Isabelle_Meta_Model/toy_example/document_generated/Design_generated.thy --- a/thys/Isabelle_Meta_Model/toy_example/document_generated/Design_generated.thy +++ b/thys/Isabelle_Meta_Model/toy_example/document_generated/Design_generated.thy @@ -1,107 +1,108 @@ theory Design_generated imports "../Toy_Library" "../Toy_Library_Static" "../embedding/Generator_dynamic_sequential" begin (* 1 ************************************ 0 + 1 *) text \ For certain concepts like classes and class-types, only a generic definition for its resulting semantics can be given. Generic means, there is a function outside HOL that ``compiles'' a concrete, closed-world class diagram into a ``theory'' of this data model, consisting of a bunch of definitions for classes, accessors, method, casts, and tests for actual types, as well as proofs for the fundamental properties of these operations in this concrete data model. \ (* 2 ************************************ 1 + 1 *) text \ Our data universe consists in the concrete class diagram just of node's, and implicitly of the class object. Each class implies the existence of a class type defined for the corresponding object representations as follows: \ (* 3 ************************************ 2 + 10 *) datatype ty\\\\<^sub>A\<^sub>t\<^sub>o\<^sub>m = mk\\\\<^sub>A\<^sub>t\<^sub>o\<^sub>m "oid" "oid list option" "int option" "bool option" "nat option" "unit option" datatype ty\<^sub>A\<^sub>t\<^sub>o\<^sub>m = mk\<^sub>A\<^sub>t\<^sub>o\<^sub>m "ty\\\\<^sub>A\<^sub>t\<^sub>o\<^sub>m" "int option" datatype ty\\\\<^sub>M\<^sub>o\<^sub>l\<^sub>e\<^sub>c\<^sub>u\<^sub>l\<^sub>e = mk\\\\<^sub>M\<^sub>o\<^sub>l\<^sub>e\<^sub>c\<^sub>u\<^sub>l\<^sub>e_\<^sub>A\<^sub>t\<^sub>o\<^sub>m "ty\<^sub>A\<^sub>t\<^sub>o\<^sub>m" | mk\\\\<^sub>M\<^sub>o\<^sub>l\<^sub>e\<^sub>c\<^sub>u\<^sub>l\<^sub>e "oid" "oid list option" "int option" "bool option" "nat option" "unit option" datatype ty\<^sub>M\<^sub>o\<^sub>l\<^sub>e\<^sub>c\<^sub>u\<^sub>l\<^sub>e = mk\<^sub>M\<^sub>o\<^sub>l\<^sub>e\<^sub>c\<^sub>u\<^sub>l\<^sub>e "ty\\\\<^sub>M\<^sub>o\<^sub>l\<^sub>e\<^sub>c\<^sub>u\<^sub>l\<^sub>e" datatype ty\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n = mk\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\<^sub>M\<^sub>o\<^sub>l\<^sub>e\<^sub>c\<^sub>u\<^sub>l\<^sub>e "ty\<^sub>M\<^sub>o\<^sub>l\<^sub>e\<^sub>c\<^sub>u\<^sub>l\<^sub>e" | mk\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\<^sub>A\<^sub>t\<^sub>o\<^sub>m "ty\<^sub>A\<^sub>t\<^sub>o\<^sub>m" | mk\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n "oid" "nat option" "unit option" datatype ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n = mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n "ty\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" "oid list option" "int option" "bool option" datatype ty\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y = mk\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" | mk\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\<^sub>M\<^sub>o\<^sub>l\<^sub>e\<^sub>c\<^sub>u\<^sub>l\<^sub>e "ty\<^sub>M\<^sub>o\<^sub>l\<^sub>e\<^sub>c\<^sub>u\<^sub>l\<^sub>e" | mk\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\<^sub>A\<^sub>t\<^sub>o\<^sub>m "ty\<^sub>A\<^sub>t\<^sub>o\<^sub>m" | mk\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y "oid" datatype ty\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y = mk\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y "ty\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y" "nat option" "unit option" datatype ty\\\\<^sub>T\<^sub>o\<^sub>y\<^sub>A\<^sub>n\<^sub>y = mk\\\\<^sub>T\<^sub>o\<^sub>y\<^sub>A\<^sub>n\<^sub>y_\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y "ty\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y" | mk\\\\<^sub>T\<^sub>o\<^sub>y\<^sub>A\<^sub>n\<^sub>y_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" | mk\\\\<^sub>T\<^sub>o\<^sub>y\<^sub>A\<^sub>n\<^sub>y_\<^sub>M\<^sub>o\<^sub>l\<^sub>e\<^sub>c\<^sub>u\<^sub>l\<^sub>e "ty\<^sub>M\<^sub>o\<^sub>l\<^sub>e\<^sub>c\<^sub>u\<^sub>l\<^sub>e" | mk\\\\<^sub>T\<^sub>o\<^sub>y\<^sub>A\<^sub>n\<^sub>y_\<^sub>A\<^sub>t\<^sub>o\<^sub>m "ty\<^sub>A\<^sub>t\<^sub>o\<^sub>m" | mk\\\\<^sub>T\<^sub>o\<^sub>y\<^sub>A\<^sub>n\<^sub>y "oid" datatype ty\<^sub>T\<^sub>o\<^sub>y\<^sub>A\<^sub>n\<^sub>y = mk\<^sub>T\<^sub>o\<^sub>y\<^sub>A\<^sub>n\<^sub>y "ty\\\\<^sub>T\<^sub>o\<^sub>y\<^sub>A\<^sub>n\<^sub>y" (* 4 ************************************ 12 + 1 *) text \ Now, we construct a concrete ``universe of ToyAny types'' by injection into a sum type containing the class types. This type of ToyAny will be used as instance for all respective type-variables. \ (* 5 ************************************ 13 + 1 *) datatype \ = in\<^sub>A\<^sub>t\<^sub>o\<^sub>m "ty\<^sub>A\<^sub>t\<^sub>o\<^sub>m" | in\<^sub>M\<^sub>o\<^sub>l\<^sub>e\<^sub>c\<^sub>u\<^sub>l\<^sub>e "ty\<^sub>M\<^sub>o\<^sub>l\<^sub>e\<^sub>c\<^sub>u\<^sub>l\<^sub>e" | in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" | in\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y "ty\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y" | in\<^sub>T\<^sub>o\<^sub>y\<^sub>A\<^sub>n\<^sub>y "ty\<^sub>T\<^sub>o\<^sub>y\<^sub>A\<^sub>n\<^sub>y" (* 6 ************************************ 14 + 1 *) text \ Having fixed the object universe, we can introduce type synonyms that exactly correspond to Toy types. Again, we exploit that our representation of Toy is a ``shallow embedding'' with a one-to-one correspondance of Toy-types to types of the meta-language HOL. \ (* 7 ************************************ 15 + 5 *) type_synonym Atom = "\\ty\<^sub>A\<^sub>t\<^sub>o\<^sub>m\\<^sub>\\\<^sub>\" type_synonym Molecule = "\\ty\<^sub>M\<^sub>o\<^sub>l\<^sub>e\<^sub>c\<^sub>u\<^sub>l\<^sub>e\\<^sub>\\\<^sub>\" type_synonym Person = "\\ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\<^sub>\\\<^sub>\" type_synonym Galaxy = "\\ty\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y\\<^sub>\\\<^sub>\" type_synonym ToyAny = "\\ty\<^sub>T\<^sub>o\<^sub>y\<^sub>A\<^sub>n\<^sub>y\\<^sub>\\\<^sub>\" (* 8 ************************************ 20 + 3 *) definition "oid\<^sub>A\<^sub>t\<^sub>o\<^sub>m_0___boss = 0" definition "oid\<^sub>M\<^sub>o\<^sub>l\<^sub>e\<^sub>c\<^sub>u\<^sub>l\<^sub>e_0___boss = 0" definition "oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___boss = 0" (* 9 ************************************ 23 + 2 *) definition "switch\<^sub>2_01 = (\ [x0 , x1] \ (x0 , x1))" definition "switch\<^sub>2_10 = (\ [x0 , x1] \ (x1 , x0))" (* 10 ************************************ 25 + 3 *) definition "oid1 = 1" definition "oid2 = 2" definition "inst_assoc1 = (\oid_class to_from oid. ((case (deref_assocs_list ((to_from::oid list list \ oid list \ oid list)) ((oid::oid)) ((drop ((((map_of_list ([(oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___boss , (List.map ((\(x , y). [x , y]) o switch\<^sub>2_01) ([[[oid1] , [oid2]]])))]))) ((oid_class::oid))))))) of Nil \ None | l \ (Some (l)))::oid list option))" (* 11 ************************************ 28 + 0 *) (* 12 ************************************ 28 + 2 *) definition "oid3 = 3" definition "inst_assoc3 = (\oid_class to_from oid. ((case (deref_assocs_list ((to_from::oid list list \ oid list \ oid list)) ((oid::oid)) ((drop ((((map_of_list ([]))) ((oid_class::oid))))))) of Nil \ None | l \ (Some (l)))::oid list option))" (* 13 ************************************ 30 + 0 *) (* 14 ************************************ 30 + 4 *) generation_syntax [ shallow (generation_semantics [ design ]) ] setup \(Generation_mode.update_compiler_config ((K (let open META in Compiler_env_config_ext (true, NONE, Oids ((Code_Numeral.natural_of_integer 0), (Code_Numeral.natural_of_integer 1), (Code_Numeral.natural_of_integer 4)), I ((Code_Numeral.natural_of_integer 0), (Code_Numeral.natural_of_integer 0)), Gen_only_design, SOME (ToyClass ((META.SS_base (META.ST "ToyAny")), nil, uncurry cons (ToyClass ((META.SS_base (META.ST "Galaxy")), uncurry cons (I ((META.SS_base (META.ST "wormhole")), ToyTy_base_unlimitednatural), uncurry cons (I ((META.SS_base (META.ST "is_sound")), ToyTy_base_void), nil)), uncurry cons (ToyClass ((META.SS_base (META.ST "Person")), uncurry cons (I ((META.SS_base (META.ST "boss")), ToyTy_object (ToyTyObj (ToyTyCore (Toy_ty_class_ext ((META.SS_base (META.ST "oid")), (Code_Numeral.natural_of_integer 0), (Code_Numeral.natural_of_integer 2), Toy_ty_class_node_ext ((Code_Numeral.natural_of_integer 0), Toy_multiplicity_ext (uncurry cons (I (Mult_star, NONE), nil), NONE, uncurry cons (Set, nil), ()), (META.SS_base (META.ST "Person")), ()), Toy_ty_class_node_ext ((Code_Numeral.natural_of_integer 1), Toy_multiplicity_ext (uncurry cons (I (Mult_nat ((Code_Numeral.natural_of_integer 0)), SOME (Mult_nat ((Code_Numeral.natural_of_integer 1)))), nil), SOME ((META.SS_base (META.ST "boss"))), uncurry cons (Set, nil), ()), (META.SS_base (META.ST "Person")), ()), ())), nil))), uncurry cons (I ((META.SS_base (META.ST "salary")), ToyTy_base_integer), uncurry cons (I ((META.SS_base (META.ST "is_meta_thinking")), ToyTy_base_boolean), nil))), uncurry cons (ToyClass ((META.SS_base (META.ST "Molecule")), nil, uncurry cons (ToyClass ((META.SS_base (META.ST "Atom")), uncurry cons (I ((META.SS_base (META.ST "size")), ToyTy_base_integer), nil), nil), nil)), nil)), nil)), nil))), uncurry cons (META_instance (ToyInstance (uncurry cons (Toy_instance_single_ext (SOME ((META.SS_base (META.ST "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3"))), SOME ((META.SS_base (META.ST "Person"))), ToyAttrNoCast (uncurry cons (I (NONE, I ((META.SS_base (META.ST "salary")), ShallB_term (ToyDefInteger ((META.SS_base (META.ST "1")))))), nil)), ()), nil))), uncurry cons (META_instance (ToyInstance (uncurry cons (Toy_instance_single_ext (SOME ((META.SS_base (META.ST "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1"))), SOME ((META.SS_base (META.ST "Person"))), ToyAttrNoCast (uncurry cons (I (NONE, I ((META.SS_base (META.ST "salary")), ShallB_term (ToyDefInteger ((META.SS_base (META.ST "1300")))))), uncurry cons (I (NONE, I ((META.SS_base (META.ST "boss")), ShallB_str ((META.SS_base (META.ST "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2"))))), nil))), ()), uncurry cons (Toy_instance_single_ext (SOME ((META.SS_base (META.ST "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2"))), SOME ((META.SS_base (META.ST "Person"))), ToyAttrNoCast (uncurry cons (I (NONE, I ((META.SS_base (META.ST "salary")), ShallB_term (ToyDefInteger ((META.SS_base (META.ST "1800")))))), nil)), ()), nil)))), uncurry cons (META_class_raw (Floor1, Toy_class_raw_ext (ToyTyObj (ToyTyCore_pre ((META.SS_base (META.ST "Person"))), uncurry cons (uncurry cons (ToyTyCore_pre ((META.SS_base (META.ST "Galaxy"))), nil), nil)), uncurry cons (I ((META.SS_base (META.ST "salary")), ToyTy_base_integer), uncurry cons (I ((META.SS_base (META.ST "boss")), ToyTy_object (ToyTyObj (ToyTyCore_pre ((META.SS_base (META.ST "Person"))), nil))), uncurry cons (I ((META.SS_base (META.ST "is_meta_thinking")), ToyTy_base_boolean), nil))), nil, false, ())), uncurry cons (META_class_raw (Floor1, Toy_class_raw_ext (ToyTyObj (ToyTyCore_pre ((META.SS_base (META.ST "Galaxy"))), nil), uncurry cons (I ((META.SS_base (META.ST "wormhole")), ToyTy_base_unlimitednatural), uncurry cons (I ((META.SS_base (META.ST "is_sound")), ToyTy_base_void), nil)), nil, false, ())), uncurry cons (META_class_raw (Floor1, Toy_class_raw_ext (ToyTyObj (ToyTyCore_pre ((META.SS_base (META.ST "Molecule"))), uncurry cons (uncurry cons (ToyTyCore_pre ((META.SS_base (META.ST "Person"))), nil), nil)), nil, nil, false, ())), uncurry cons (META_class_raw (Floor1, Toy_class_raw_ext (ToyTyObj (ToyTyCore_pre ((META.SS_base (META.ST "Atom"))), uncurry cons (uncurry cons (ToyTyCore_pre ((META.SS_base (META.ST "Molecule"))), nil), nil)), uncurry cons (I ((META.SS_base (META.ST "size")), ToyTy_base_integer), nil), nil, false, ())), nil)))))), uncurry cons (I ((META.ST "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3"), I (Toy_instance_single_ext (SOME ((META.SS_base (META.ST "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3"))), SOME ((META.SS_base (META.ST "Person"))), ToyAttrNoCast (uncurry cons (I (NONE, I ((META.SS_base (META.ST "salary")), ShallB_term (ToyDefInteger ((META.SS_base (META.ST "1")))))), nil)), ()), Oids ((Code_Numeral.natural_of_integer 0), (Code_Numeral.natural_of_integer 1), (Code_Numeral.natural_of_integer 3)))), uncurry cons (I ((META.ST "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2"), I (Toy_instance_single_ext (SOME ((META.SS_base (META.ST "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2"))), SOME ((META.SS_base (META.ST "Person"))), ToyAttrNoCast (uncurry cons (I (NONE, I ((META.SS_base (META.ST "salary")), ShallB_term (ToyDefInteger ((META.SS_base (META.ST "1800")))))), nil)), ()), Oids ((Code_Numeral.natural_of_integer 0), (Code_Numeral.natural_of_integer 1), (Code_Numeral.natural_of_integer 2)))), uncurry cons (I ((META.ST "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1"), I (Toy_instance_single_ext (SOME ((META.SS_base (META.ST "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1"))), SOME ((META.SS_base (META.ST "Person"))), ToyAttrNoCast (uncurry cons (I (NONE, I ((META.SS_base (META.ST "salary")), ShallB_term (ToyDefInteger ((META.SS_base (META.ST "1300")))))), uncurry cons (I (NONE, I ((META.SS_base (META.ST "boss")), ShallB_str ((META.SS_base (META.ST "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2"))))), nil))), ()), Oids ((Code_Numeral.natural_of_integer 0), (Code_Numeral.natural_of_integer 1), (Code_Numeral.natural_of_integer 1)))), nil))), nil, false, false, I (nil, nil), nil, I (NONE, false), ()) end))))\ Instance \\<^sub>1_object0 :: Person = [ "salary" = 1000, "boss" = self 1 ] and \\<^sub>1_object1 :: Person = [ "salary" = 1200 ] and \\<^sub>1_object2 :: Person = [ "salary" = 2600, "boss" = X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 ] and \\<^sub>1_object4 :: Person = [ "salary" = 2300, "boss" = self 2 ] +(* Fails because of modified map update syntax priorities State[shallow] \\<^sub>1 = [ \\<^sub>1_object0, \\<^sub>1_object1, \\<^sub>1_object2, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1, \\<^sub>1_object4, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 ] (* 15 ************************************ 34 + 1 *) State[shallow] \\<^sub>1' = [ X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2, X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 ] (* 16 ************************************ 35 + 1 *) PrePost[shallow] \\<^sub>1 \\<^sub>1' - +*) end diff --git a/thys/Isabelle_Meta_Model/toy_example/document_generated/Design_generated_generated.thy b/thys/Isabelle_Meta_Model/toy_example/document_generated/Design_generated_generated.thy --- a/thys/Isabelle_Meta_Model/toy_example/document_generated/Design_generated_generated.thy +++ b/thys/Isabelle_Meta_Model/toy_example/document_generated/Design_generated_generated.thy @@ -1,225 +1,225 @@ theory Design_generated_generated imports "../Toy_Library" "../Toy_Library_Static" begin (* 1 ************************************ 0 + 1 *) text \ For certain concepts like classes and class-types, only a generic definition for its resulting semantics can be given. Generic means, there is a function outside HOL that ``compiles'' a concrete, closed-world class diagram into a ``theory'' of this data model, consisting of a bunch of definitions for classes, accessors, method, casts, and tests for actual types, as well as proofs for the fundamental properties of these operations in this concrete data model. \ (* 2 ************************************ 1 + 1 *) text \ Our data universe consists in the concrete class diagram just of node's, and implicitly of the class object. Each class implies the existence of a class type defined for the corresponding object representations as follows: \ (* 3 ************************************ 2 + 10 *) datatype ty\\\\<^sub>A\<^sub>t\<^sub>o\<^sub>m = mk\\\\<^sub>A\<^sub>t\<^sub>o\<^sub>m "oid" "oid list option" "int option" "bool option" "nat option" "unit option" datatype ty\<^sub>A\<^sub>t\<^sub>o\<^sub>m = mk\<^sub>A\<^sub>t\<^sub>o\<^sub>m "ty\\\\<^sub>A\<^sub>t\<^sub>o\<^sub>m" "int option" datatype ty\\\\<^sub>M\<^sub>o\<^sub>l\<^sub>e\<^sub>c\<^sub>u\<^sub>l\<^sub>e = mk\\\\<^sub>M\<^sub>o\<^sub>l\<^sub>e\<^sub>c\<^sub>u\<^sub>l\<^sub>e_\<^sub>A\<^sub>t\<^sub>o\<^sub>m "ty\<^sub>A\<^sub>t\<^sub>o\<^sub>m" | mk\\\\<^sub>M\<^sub>o\<^sub>l\<^sub>e\<^sub>c\<^sub>u\<^sub>l\<^sub>e "oid" "oid list option" "int option" "bool option" "nat option" "unit option" datatype ty\<^sub>M\<^sub>o\<^sub>l\<^sub>e\<^sub>c\<^sub>u\<^sub>l\<^sub>e = mk\<^sub>M\<^sub>o\<^sub>l\<^sub>e\<^sub>c\<^sub>u\<^sub>l\<^sub>e "ty\\\\<^sub>M\<^sub>o\<^sub>l\<^sub>e\<^sub>c\<^sub>u\<^sub>l\<^sub>e" datatype ty\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n = mk\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\<^sub>M\<^sub>o\<^sub>l\<^sub>e\<^sub>c\<^sub>u\<^sub>l\<^sub>e "ty\<^sub>M\<^sub>o\<^sub>l\<^sub>e\<^sub>c\<^sub>u\<^sub>l\<^sub>e" | mk\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_\<^sub>A\<^sub>t\<^sub>o\<^sub>m "ty\<^sub>A\<^sub>t\<^sub>o\<^sub>m" | mk\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n "oid" "nat option" "unit option" datatype ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n = mk\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n "ty\\\\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" "oid list option" "int option" "bool option" datatype ty\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y = mk\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" | mk\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\<^sub>M\<^sub>o\<^sub>l\<^sub>e\<^sub>c\<^sub>u\<^sub>l\<^sub>e "ty\<^sub>M\<^sub>o\<^sub>l\<^sub>e\<^sub>c\<^sub>u\<^sub>l\<^sub>e" | mk\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y_\<^sub>A\<^sub>t\<^sub>o\<^sub>m "ty\<^sub>A\<^sub>t\<^sub>o\<^sub>m" | mk\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y "oid" datatype ty\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y = mk\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y "ty\\\\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y" "nat option" "unit option" datatype ty\\\\<^sub>T\<^sub>o\<^sub>y\<^sub>A\<^sub>n\<^sub>y = mk\\\\<^sub>T\<^sub>o\<^sub>y\<^sub>A\<^sub>n\<^sub>y_\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y "ty\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y" | mk\\\\<^sub>T\<^sub>o\<^sub>y\<^sub>A\<^sub>n\<^sub>y_\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" | mk\\\\<^sub>T\<^sub>o\<^sub>y\<^sub>A\<^sub>n\<^sub>y_\<^sub>M\<^sub>o\<^sub>l\<^sub>e\<^sub>c\<^sub>u\<^sub>l\<^sub>e "ty\<^sub>M\<^sub>o\<^sub>l\<^sub>e\<^sub>c\<^sub>u\<^sub>l\<^sub>e" | mk\\\\<^sub>T\<^sub>o\<^sub>y\<^sub>A\<^sub>n\<^sub>y_\<^sub>A\<^sub>t\<^sub>o\<^sub>m "ty\<^sub>A\<^sub>t\<^sub>o\<^sub>m" | mk\\\\<^sub>T\<^sub>o\<^sub>y\<^sub>A\<^sub>n\<^sub>y "oid" datatype ty\<^sub>T\<^sub>o\<^sub>y\<^sub>A\<^sub>n\<^sub>y = mk\<^sub>T\<^sub>o\<^sub>y\<^sub>A\<^sub>n\<^sub>y "ty\\\\<^sub>T\<^sub>o\<^sub>y\<^sub>A\<^sub>n\<^sub>y" (* 4 ************************************ 12 + 1 *) text \ Now, we construct a concrete ``universe of ToyAny types'' by injection into a sum type containing the class types. This type of ToyAny will be used as instance for all respective type-variables. \ (* 5 ************************************ 13 + 1 *) datatype \ = in\<^sub>A\<^sub>t\<^sub>o\<^sub>m "ty\<^sub>A\<^sub>t\<^sub>o\<^sub>m" | in\<^sub>M\<^sub>o\<^sub>l\<^sub>e\<^sub>c\<^sub>u\<^sub>l\<^sub>e "ty\<^sub>M\<^sub>o\<^sub>l\<^sub>e\<^sub>c\<^sub>u\<^sub>l\<^sub>e" | in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" | in\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y "ty\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y" | in\<^sub>T\<^sub>o\<^sub>y\<^sub>A\<^sub>n\<^sub>y "ty\<^sub>T\<^sub>o\<^sub>y\<^sub>A\<^sub>n\<^sub>y" (* 6 ************************************ 14 + 1 *) text \ Having fixed the object universe, we can introduce type synonyms that exactly correspond to Toy types. Again, we exploit that our representation of Toy is a ``shallow embedding'' with a one-to-one correspondance of Toy-types to types of the meta-language HOL. \ (* 7 ************************************ 15 + 5 *) type_synonym Atom = "\\ty\<^sub>A\<^sub>t\<^sub>o\<^sub>m\\<^sub>\\\<^sub>\" type_synonym Molecule = "\\ty\<^sub>M\<^sub>o\<^sub>l\<^sub>e\<^sub>c\<^sub>u\<^sub>l\<^sub>e\\<^sub>\\\<^sub>\" type_synonym Person = "\\ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\<^sub>\\\<^sub>\" type_synonym Galaxy = "\\ty\<^sub>G\<^sub>a\<^sub>l\<^sub>a\<^sub>x\<^sub>y\\<^sub>\\\<^sub>\" type_synonym ToyAny = "\\ty\<^sub>T\<^sub>o\<^sub>y\<^sub>A\<^sub>n\<^sub>y\\<^sub>\\\<^sub>\" (* 8 ************************************ 20 + 3 *) definition "oid\<^sub>A\<^sub>t\<^sub>o\<^sub>m_0___boss = 0" definition "oid\<^sub>M\<^sub>o\<^sub>l\<^sub>e\<^sub>c\<^sub>u\<^sub>l\<^sub>e_0___boss = 0" definition "oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___boss = 0" (* 9 ************************************ 23 + 2 *) definition "switch\<^sub>2_01 = (\ [x0 , x1] \ (x0 , x1))" definition "switch\<^sub>2_10 = (\ [x0 , x1] \ (x1 , x0))" (* 10 ************************************ 25 + 3 *) definition "oid1 = 1" definition "oid2 = 2" definition "inst_assoc1 = (\oid_class to_from oid. ((case (deref_assocs_list ((to_from::oid list list \ oid list \ oid list)) ((oid::oid)) ((drop ((((map_of_list ([(oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___boss , (List.map ((\(x , y). [x , y]) o switch\<^sub>2_01) ([[[oid1] , [oid2]]])))]))) ((oid_class::oid))))))) of Nil \ None | l \ (Some (l)))::oid list option))" (* 11 ************************************ 28 + 0 *) (* 12 ************************************ 28 + 2 *) definition "oid3 = 3" definition "inst_assoc3 = (\oid_class to_from oid. ((case (deref_assocs_list ((to_from::oid list list \ oid list \ oid list)) ((oid::oid)) ((drop ((((map_of_list ([]))) ((oid_class::oid))))))) of Nil \ None | l \ (Some (l)))::oid list option))" (* 13 ************************************ 30 + 0 *) (* 14 ************************************ 30 + 5 *) definition "oid4 = 4" definition "oid5 = 5" definition "oid6 = 6" definition "oid7 = 7" definition "inst_assoc4 = (\oid_class to_from oid. ((case (deref_assocs_list ((to_from::oid list list \ oid list \ oid list)) ((oid::oid)) ((drop ((((map_of_list ([(oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___boss , (List.map ((\(x , y). [x , y]) o switch\<^sub>2_01) ([[[oid7] , [oid6]] , [[oid6] , [oid1]] , [[oid4] , [oid5]]])))]))) ((oid_class::oid))))))) of Nil \ None | l \ (Some (l)))::oid list option))" (* 15 ************************************ 35 + 0 *) (* 16 ************************************ 35 + 1 *) locale state_\\<^sub>1 = fixes "oid4" :: "nat" fixes "oid5" :: "nat" fixes "oid6" :: "nat" fixes "oid1" :: "nat" fixes "oid7" :: "nat" fixes "oid2" :: "nat" assumes distinct_oid: "(distinct ([oid4 , oid5 , oid6 , oid1 , oid7 , oid2]))" fixes "\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" :: "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" fixes "\\<^sub>1_object0" :: "\Person" assumes \\<^sub>1_object0_def: "\\<^sub>1_object0 = (\_. \\\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)" fixes "\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" :: "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" fixes "\\<^sub>1_object1" :: "\Person" assumes \\<^sub>1_object1_def: "\\<^sub>1_object1 = (\_. \\\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)" fixes "\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" :: "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" fixes "\\<^sub>1_object2" :: "\Person" assumes \\<^sub>1_object2_def: "\\<^sub>1_object2 = (\_. \\\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" :: "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1" :: "\Person" assumes X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1_def: "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 = (\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)" fixes "\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" :: "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" fixes "\\<^sub>1_object4" :: "\Person" assumes \\<^sub>1_object4_def: "\\<^sub>1_object4 = (\_. \\\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" :: "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2" :: "\Person" assumes X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2_def: "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 = (\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)" begin -definition "\\<^sub>1 = (state.make ((Map.empty (oid4 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))) (oid5 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))) (oid6 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))) (oid1 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))) (oid7 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))) (oid2 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))) ((map_of_list ([(oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___boss , (List.map ((\(x , y). [x , y]) o switch\<^sub>2_01) ([[[oid4] , [oid2]] , [[oid6] , [oid4]] , [[oid1] , [oid6]] , [[oid7] , [oid3]]])))]))))" +definition "\\<^sub>1 = (state.make ((Map.empty (oid4 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)), oid5 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)), oid6 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)), oid1 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)), oid7 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)), oid2 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))) ((map_of_list ([(oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___boss , (List.map ((\(x , y). [x , y]) o switch\<^sub>2_01) ([[[oid4] , [oid2]] , [[oid6] , [oid4]] , [[oid1] , [oid6]] , [[oid7] , [oid3]]])))]))))" -lemma perm_\\<^sub>1 : "\\<^sub>1 = (state.make ((Map.empty (oid2 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))) (oid7 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))) (oid1 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))) (oid6 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))) (oid5 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))) (oid4 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))) ((assocs (\\<^sub>1))))" +lemma perm_\\<^sub>1 : "\\<^sub>1 = (state.make ((Map.empty (oid2 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)), oid7 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)), oid1 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)), oid6 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)), oid5 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)), oid4 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))) ((assocs (\\<^sub>1))))" apply(simp add: \\<^sub>1_def) apply(subst (1) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (2) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (1) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (3) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (2) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (1) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (4) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (3) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (2) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (1) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (5) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (4) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (3) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (2) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (1) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) by(simp) end (* 17 ************************************ 36 + 1 *) locale state_\\<^sub>1' = fixes "oid1" :: "nat" fixes "oid2" :: "nat" fixes "oid3" :: "nat" assumes distinct_oid: "(distinct ([oid1 , oid2 , oid3]))" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" :: "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1" :: "\Person" assumes X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1_def: "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 = (\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" :: "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2" :: "\Person" assumes X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2_def: "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 = (\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" :: "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3" :: "\Person" assumes X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3_def: "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 = (\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)" begin -definition "\\<^sub>1' = (state.make ((Map.empty (oid1 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))) (oid2 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))) (oid3 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))) ((map_of_list ([(oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___boss , (List.map ((\(x , y). [x , y]) o switch\<^sub>2_01) ([[[oid1] , [oid2]]])))]))))" +definition "\\<^sub>1' = (state.make ((Map.empty (oid1 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)), oid2 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)), oid3 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))) ((map_of_list ([(oid\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n_0___boss , (List.map ((\(x , y). [x , y]) o switch\<^sub>2_01) ([[[oid1] , [oid2]]])))]))))" -lemma perm_\\<^sub>1' : "\\<^sub>1' = (state.make ((Map.empty (oid3 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))) (oid2 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))) (oid1 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))) ((assocs (\\<^sub>1'))))" +lemma perm_\\<^sub>1' : "\\<^sub>1' = (state.make ((Map.empty (oid3 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)), oid2 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n)), oid1 \ (in\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n))))) ((assocs (\\<^sub>1'))))" apply(simp add: \\<^sub>1'_def) apply(subst (1) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (2) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) apply(subst (1) fun_upd_twist, metis distinct_oid distinct_length_2_or_more) by(simp) end (* 18 ************************************ 37 + 1 *) locale pre_post_\\<^sub>1_\\<^sub>1' = fixes "oid1" :: "nat" fixes "oid2" :: "nat" fixes "oid3" :: "nat" fixes "oid4" :: "nat" fixes "oid5" :: "nat" fixes "oid6" :: "nat" fixes "oid7" :: "nat" assumes distinct_oid: "(distinct ([oid1 , oid2 , oid3 , oid4 , oid5 , oid6 , oid7]))" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" :: "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1" :: "\Person" assumes X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1_def: "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 = (\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" :: "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2" :: "\Person" assumes X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2_def: "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 = (\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" :: "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" fixes "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3" :: "\Person" assumes X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3_def: "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 = (\_. \\X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)" fixes "\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" :: "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" fixes "\\<^sub>1_object0" :: "\Person" assumes \\<^sub>1_object0_def: "\\<^sub>1_object0 = (\_. \\\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)" fixes "\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" :: "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" fixes "\\<^sub>1_object1" :: "\Person" assumes \\<^sub>1_object1_def: "\\<^sub>1_object1 = (\_. \\\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)" fixes "\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" :: "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" fixes "\\<^sub>1_object2" :: "\Person" assumes \\<^sub>1_object2_def: "\\<^sub>1_object2 = (\_. \\\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)" fixes "\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" :: "ty\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" fixes "\\<^sub>1_object4" :: "\Person" assumes \\<^sub>1_object4_def: "\\<^sub>1_object4 = (\_. \\\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n\\)" assumes \\<^sub>1: "(state_\\<^sub>1 (oid4) (oid5) (oid6) (oid1) (oid7) (oid2) (\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) (\\<^sub>1_object0) (\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) (\\<^sub>1_object1) (\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) (\\<^sub>1_object2) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1) (\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) (\\<^sub>1_object4) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2))" assumes \\<^sub>1': "(state_\\<^sub>1' (oid1) (oid2) (oid3) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n) (X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3))" begin interpretation state_\\<^sub>1: state_\\<^sub>1 "oid4" "oid5" "oid6" "oid1" "oid7" "oid2" "\\<^sub>1_object0\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" "\\<^sub>1_object0" "\\<^sub>1_object1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" "\\<^sub>1_object1" "\\<^sub>1_object2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" "\\<^sub>1_object2" "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1" "\\<^sub>1_object4\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" "\\<^sub>1_object4" "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2" by(rule \\<^sub>1) interpretation state_\\<^sub>1': state_\\<^sub>1' "oid1" "oid2" "oid3" "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1" "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2" "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n" "X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3" by(rule \\<^sub>1') definition "heap_\\<^sub>1 = (heap (state_\\<^sub>1.\\<^sub>1))" definition "heap_\\<^sub>1' = (heap (state_\\<^sub>1'.\\<^sub>1'))" end end diff --git a/thys/Isabelle_Meta_Model/toy_example/generator/Design_shallow.thy b/thys/Isabelle_Meta_Model/toy_example/generator/Design_shallow.thy --- a/thys/Isabelle_Meta_Model/toy_example/generator/Design_shallow.thy +++ b/thys/Isabelle_Meta_Model/toy_example/generator/Design_shallow.thy @@ -1,136 +1,137 @@ (****************************************************************************** * HOL-TOY * * Copyright (c) 2011-2018 Université Paris-Saclay, Univ. Paris-Sud, France * 2013-2017 IRT SystemX, France * 2011-2015 Achim D. Brucker, Germany * 2016-2018 The University of Sheffield, UK * 2016-2017 Nanyang Technological University, Singapore * 2017-2018 Virginia Tech, USA * * All rights reserved. * * Redistribution and use in source and binary forms, with or without * modification, are permitted provided that the following conditions are * met: * * * Redistributions of source code must retain the above copyright * notice, this list of conditions and the following disclaimer. * * * Redistributions in binary form must reproduce the above * copyright notice, this list of conditions and the following * disclaimer in the documentation and/or other materials provided * with the distribution. * * * Neither the name of the copyright holders nor the names of its * contributors may be used to endorse or promote products derived * from this software without specific prior written permission. * * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS * "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT * LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR * A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT * OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, * SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT * LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, * DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY * THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT * (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. ******************************************************************************) section\Example: A Class Model Interactively Executed\ subsection\Introduction\ theory Design_shallow imports "../Toy_Library" "../Toy_Library_Static" "../embedding/Generator_dynamic_sequential" "../../Antiquote_Setup" begin text\ In this example, we configure our package to execute tactic SML code (corresponding to some generated \verb|.thy| file, @{file \Design_deep.thy\} details how to obtain such generated \verb|.thy| file). Since SML code are already compiled (or reflected) and bound with the native Isabelle API in @{theory Isabelle_Meta_Model.Generator_dynamic_sequential}, nothing is generated in this theory. The system only parses arguments given to meta-commands and immediately calls the corresponding compiled functions. The execution time is comparatively similar as if tactics were written by hand, except that the generated SML code potentially inherits all optimizations performed by the raw code generation of Isabelle (if any). \ generation_syntax [ shallow (generation_semantics [ design ]) (*SORRY*) (*no_dirty*) (*, syntax_print*) ] text\ The configuration in @{keyword "shallow"} mode is straightforward: in this mode @{command generation_syntax} basically terminates in $O(1)$. \ subsection\Designing Class Models (I): Basics\ Class Atom < Molecule Attributes size : Integer End End End End Class Molecule < Person Class Galaxy Attributes wormhole : UnlimitedNatural is_sound : Void End! Class Person < Galaxy Attributes salary : Integer boss : Person is_meta_thinking: Boolean Instance X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 :: Person = [ salary = 1300 , boss = X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 ] and X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 :: Person = [ salary = 1800 ] Instance X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 :: Person = [ salary = 1 ] (* Class Big_Bang < Atom (* This will force the creation of a new universe. *) *) subsection\Designing Class Models (II): Jumping to Another Semantic Floor\ +(* Fails because of modified map update syntax priorities State \\<^sub>1 = [ ([ salary = 1000 , boss = self 1 ] :: Person) , ([ salary = 1200 ] :: Person) (* *) , ([ salary = 2600 , boss = self 3 ] :: Person) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 , ([ salary = 2300 , boss = self 2 ] :: Person) (* *) (* *) , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 ] State \\<^sub>1' = [ X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n1 , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n2 , X\<^sub>P\<^sub>e\<^sub>r\<^sub>s\<^sub>o\<^sub>n3 ] PrePost \\<^sub>1 \\<^sub>1' - +*) (* PrePost \\<^sub>1 [ ([ salary = 1000 , boss = self 1 ] :: Person) ] *) subsection\Designing Class Models (III): Interaction with (Pure) Term\ text\ Here in @{keyword "shallow"} mode, the following expression is directly rejected: \\ \verb|Context Person :: content ()| \\ \verb| Post "><"| \ Context[shallow] Person :: content () Post : "a + b = c" end diff --git a/thys/LightweightJava/Lightweight_Java_Definition.thy b/thys/LightweightJava/Lightweight_Java_Definition.thy --- a/thys/LightweightJava/Lightweight_Java_Definition.thy +++ b/thys/LightweightJava/Lightweight_Java_Definition.thy @@ -1,674 +1,674 @@ (* Title: Lightweight Java, the definition Authors: Rok Strnisa , 2006 Matthew Parkinson , 2006 Maintainer: Note: This file should _not_ be modified directly. Please see the accompanying README file. *) (* generated by Ott 0.20.3 from: lj_common.ott lj_base.ott lj.ott *) theory Lightweight_Java_Definition imports Main "HOL-Library.Multiset" begin (** warning: the backend selected ignores the file structure informations *) (** syntax *) type_synonym "j" = "nat" type_synonym "f" = "string" type_synonym "meth" = "string" type_synonym "var" = "string" type_synonym "dcl" = "string" type_synonym "oid" = "nat" datatype "fqn" = fqn_def "dcl" datatype "cl" = cl_object | cl_fqn "fqn" datatype "x" = x_var "var" | x_this datatype "vd" = vd_def "cl" "var" type_synonym "X" = "x list" datatype "ctx" = ctx_def type_synonym "vds" = "vd list" datatype "s" = s_block "s list" | s_ass "var" "x" | s_read "var" "x" "f" | s_write "x" "f" "x" | s_if "x" "x" "s" "s" | s_new "var" "ctx" "cl" | s_call "var" "x" "meth" "X" datatype "meth_sig" = meth_sig_def "cl" "meth" "vds" datatype "meth_body" = meth_body_def "s list" "x" datatype "fd" = fd_def "cl" "f" datatype "meth_def" = meth_def_def "meth_sig" "meth_body" type_synonym "fds" = "fd list" type_synonym "meth_defs" = "meth_def list" datatype "cld" = cld_def "dcl" "cl" "fds" "meth_defs" type_synonym "ctxcld" = "(ctx \ cld)" datatype "ty" = ty_top | ty_def "ctx" "dcl" datatype "v" = v_null | v_oid "oid" type_synonym "clds" = "cld list" type_synonym "ctxclds" = "ctxcld list" type_synonym "fs" = "f list" type_synonym "ty_opt" = "ty option" type_synonym "tys" = "ty list" type_synonym "L" = "x \ v" type_synonym "H" = "oid \ (ty \ (f \ v))" datatype "Exception" = ex_npe type_synonym "P" = "clds" type_synonym "ctxcld_opt" = "ctxcld option" type_synonym "nn" = "nat" type_synonym "ctxclds_opt" = "ctxclds option" type_synonym "fs_opt" = "fs option" type_synonym "meths" = "meth list" datatype "ty_opt_bot" = ty_opt_bot_opt "ty_opt" | ty_opt_bot_bot type_synonym "meth_def_opt" = "meth_def option" type_synonym "ctxmeth_def_opt" = "(ctx \ meth_def) option" datatype "mty" = mty_def "tys" "ty" type_synonym "\" = "x \ ty" type_synonym "v_opt" = "v option" datatype "config" = config_normal "P" "L" "H" "s list" | config_ex "P" "L" "H" "Exception" type_synonym "T" = "x \ x" (** library functions *) lemma [mono]:" (!! x. f x --> g x) ==> list_all (%b. b) (map f foo_list)--> list_all (%b. b) (map g foo_list) " apply(induct_tac foo_list, auto) done lemma [mono]: "case_prod f p = f (fst p) (snd p)" by (simp add: split_def) (** definitions *) (*defns class_name_def *) inductive class_name :: "cld \ dcl \ bool" where (* defn class_name *) class_nameI: "class_name ((cld_def dcl cl fds meth_defs)) (dcl)" (*defns superclass_name_def *) inductive superclass_name :: "cld \ cl \ bool" where (* defn superclass_name *) superclass_nameI: "superclass_name ((cld_def dcl cl fds meth_defs)) (cl)" (*defns class_fields_def *) inductive class_fields :: "cld \ fds \ bool" where (* defn class_fields *) class_fieldsI: "class_fields ((cld_def dcl cl fds meth_defs)) (fds)" (*defns class_methods_def *) inductive class_methods :: "cld \ meth_defs \ bool" where (* defn class_methods *) class_methodsI: "class_methods ((cld_def dcl cl fds meth_defs)) (meth_defs)" (*defns method_name_def *) inductive method_name :: "meth_def \ meth \ bool" where (* defn method_name *) method_nameI: "method_name ((meth_def_def (meth_sig_def cl meth vds) meth_body)) (meth)" (*defns distinct_names_def *) inductive distinct_names :: "P \ bool" where (* defn distinct_names *) dn_defI: "\ P = ((List.map (%((cld_XXX::cld),(dcl_XXX::dcl)).cld_XXX) cld_dcl_list)) ; list_all (\f. f) ((List.map (%((cld_XXX::cld),(dcl_XXX::dcl)).class_name (cld_XXX) (dcl_XXX)) cld_dcl_list)) ; distinct ( ((List.map (%((cld_XXX::cld),(dcl_XXX::dcl)).dcl_XXX) cld_dcl_list)) ) \ \ distinct_names (P)" (*defns find_cld_def *) inductive find_cld :: "P \ ctx \ fqn \ ctxcld_opt \ bool" where (* defn find_cld *) fc_emptyI: "find_cld ( [] ) (ctx) (fqn) ( None )" | fc_cons_trueI: "\ P = ([(cld)] @ cld_list) ; cld = (cld_def dcl cl fds meth_defs) \ \ find_cld (P) (ctx) ((fqn_def dcl)) ( (Some ( ( ctx , cld ) )) )" | fc_cons_falseI: "\ cld = (cld_def dcl' cl fds meth_defs) ; (cl_fqn (fqn_def dcl)) \ (cl_fqn (fqn_def dcl')) ; find_cld ( (cld_list) ) (ctx) ((fqn_def dcl)) (ctxcld_opt)\ \ find_cld ( ([(cld)] @ cld_list) ) (ctx) ((fqn_def dcl)) (ctxcld_opt)" (*defns find_type_def *) inductive find_type :: "P \ ctx \ cl \ ty_opt \ bool" where (* defn find_type *) ft_objI: "find_type (P) (ctx) (cl_object) ( (Some ( ty_top )) )" | ft_nullI: "\find_cld (P) (ctx) (fqn) ( None )\ \ find_type (P) (ctx) ((cl_fqn fqn)) ( None )" | ft_dclI: "\find_cld (P) (ctx) ((fqn_def dcl)) ( (Some ( ( ctx' , cld ) )) )\ \ find_type (P) (ctx) ((cl_fqn (fqn_def dcl))) ( (Some ( (ty_def ctx' dcl) )) )" (*defns path_length_def *) inductive path_length :: "P \ ctx \ cl \ nn \ bool" where (* defn path_length *) pl_objI: "path_length (P) (ctx) (cl_object) ( 0 )" | pl_fqnI: "\find_cld (P) (ctx) (fqn) ( (Some ( ( ctx' , cld ) )) ) ; superclass_name (cld) (cl) ; path_length (P) (ctx') (cl) (nn)\ \ path_length (P) (ctx) ((cl_fqn fqn)) ( ( nn + 1 ) )" (*defns acyclic_clds_def *) inductive acyclic_clds :: "P \ bool" where (* defn acyclic_clds *) ac_defI: "\ \ ctx fqn . ( ( (\ ctx' cld . find_cld (P) (ctx) (fqn) ( (Some ( ( ctx' , cld ) )) ) ) ) \ (\ nn . path_length (P) (ctx) ((cl_fqn fqn)) (nn) ) ) \ \ acyclic_clds (P)" (*defns find_path_rec_def *) inductive find_path_rec :: "P \ ctx \ cl \ ctxclds \ ctxclds_opt \ bool" where (* defn find_path_rec *) fpr_objI: "find_path_rec (P) (ctx) (cl_object) (ctxclds) ( Some ( ctxclds ) )" | fpr_nullI: "\ ( \ ( acyclic_clds (P) ) ) \ find_cld (P) (ctx) (fqn) ( None ) \ \ find_path_rec (P) (ctx) ((cl_fqn fqn)) (ctxclds) ( None )" | fpr_fqnI: "\ acyclic_clds (P) \ find_cld (P) (ctx) (fqn) ( (Some ( ( ctx' , cld ) )) ) ; superclass_name (cld) (cl) ; find_path_rec (P) (ctx') (cl) ( ctxclds @[ ( ctx' , cld ) ] ) (ctxclds_opt)\ \ find_path_rec (P) (ctx) ((cl_fqn fqn)) (ctxclds) (ctxclds_opt)" (*defns find_path_def *) inductive find_path :: "P \ ctx \ cl \ ctxclds_opt \ bool" where (* defn find_path *) fp_defI: "\find_path_rec (P) (ctx) (cl) ( [] ) (ctxclds_opt)\ \ find_path (P) (ctx) (cl) (ctxclds_opt)" (*defns find_path_ty_def *) inductive find_path_ty :: "P \ ty \ ctxclds_opt \ bool" where (* defn find_path_ty *) fpty_objI: "find_path_ty (P) (ty_top) ( Some ( [] ) )" | fpty_dclI: "\find_path (P) (ctx) ((cl_fqn (fqn_def dcl))) (ctxclds_opt)\ \ find_path_ty (P) ((ty_def ctx dcl)) (ctxclds_opt)" (*defns fields_in_path_def *) inductive fields_in_path :: "ctxclds \ fs \ bool" where (* defn fields_in_path *) fip_emptyI: "fields_in_path ( [] ) ( [] )" | fip_consI: "\class_fields (cld) ( ((List.map (%((cl_XXX::cl),(f_XXX::f)).(fd_def cl_XXX f_XXX)) cl_f_list)) ) ; fields_in_path ( (ctxcld_list) ) (fs) ; fs' = ( ((List.map (%((cl_XXX::cl),(f_XXX::f)).f_XXX) cl_f_list)) @ fs ) \ \ fields_in_path ( ([( ( ctx , cld ) )] @ ctxcld_list) ) (fs')" (*defns fields_def *) inductive fields :: "P \ ty \ fs_opt \ bool" where (* defn fields *) fields_noneI: "\find_path_ty (P) (ty) ( None )\ \ fields (P) (ty) ( None )" | fields_someI: "\find_path_ty (P) (ty) ( Some ( ctxclds ) ) ; fields_in_path (ctxclds) (fs)\ \ fields (P) (ty) ( Some ( fs ) )" (*defns methods_in_path_def *) inductive methods_in_path :: "clds \ meths \ bool" where (* defn methods_in_path *) mip_emptyI: "methods_in_path ( [] ) ( [] )" | mip_consI: "\class_methods (cld) ( ((List.map (%((meth_def_XXX::meth_def),(cl_XXX::cl),(meth_XXX::meth),(vds_XXX::vds),(meth_body_XXX::meth_body)).meth_def_XXX) meth_def_cl_meth_vds_meth_body_list)) ) ; list_all (\f. f) ((List.map (%((meth_def_XXX::meth_def),(cl_XXX::cl),(meth_XXX::meth),(vds_XXX::vds),(meth_body_XXX::meth_body)). meth_def_XXX = (meth_def_def (meth_sig_def cl_XXX meth_XXX vds_XXX) meth_body_XXX) ) meth_def_cl_meth_vds_meth_body_list)) ; methods_in_path ( (cld_list) ) (meths') ; meths = ( ((List.map (%((meth_def_XXX::meth_def),(cl_XXX::cl),(meth_XXX::meth),(vds_XXX::vds),(meth_body_XXX::meth_body)).meth_XXX) meth_def_cl_meth_vds_meth_body_list)) @ meths' ) \ \ methods_in_path ( ([(cld)] @ cld_list) ) (meths)" (*defns methods_def *) inductive methods :: "P \ ty \ meths \ bool" where (* defn methods *) methods_methodsI: "\find_path_ty (P) (ty) ( Some ( ((List.map (%((ctx_XXX::ctx),(cld_XXX::cld)). ( ctx_XXX , cld_XXX ) ) ctx_cld_list)) ) ) ; methods_in_path ( ((List.map (%((ctx_XXX::ctx),(cld_XXX::cld)).cld_XXX) ctx_cld_list)) ) (meths)\ \ methods (P) (ty) (meths)" (*defns ftype_in_fds_def *) inductive ftype_in_fds :: "P \ ctx \ fds \ f \ ty_opt_bot \ bool" where (* defn ftype_in_fds *) ftif_emptyI: "ftype_in_fds (P) (ctx) ( [] ) (f) ((ty_opt_bot_opt None ))" | ftif_cons_botI: "\find_type (P) (ctx) (cl) ( None )\ \ ftype_in_fds (P) (ctx) ( ([((fd_def cl f))] @ fd_list) ) (f) (ty_opt_bot_bot)" | ftif_cons_trueI: "\find_type (P) (ctx) (cl) ( (Some ( ty )) )\ \ ftype_in_fds (P) (ctx) ( ([((fd_def cl f))] @ fd_list) ) (f) ((ty_opt_bot_opt (Some ( ty )) ))" | ftif_cons_falseI: "\ f \ f' ; ftype_in_fds (P) (ctx) ( (fd_list) ) (f') (ty_opt_bot)\ \ ftype_in_fds (P) (ctx) ( ([((fd_def cl f))] @ fd_list) ) (f') (ty_opt_bot)" (*defns ftype_in_path_def *) inductive ftype_in_path :: "P \ ctxclds \ f \ ty_opt \ bool" where (* defn ftype_in_path *) ftip_emptyI: "ftype_in_path (P) ( [] ) (f) ( None )" | ftip_cons_botI: "\class_fields (cld) (fds) ; ftype_in_fds (P) (ctx) (fds) (f) (ty_opt_bot_bot)\ \ ftype_in_path (P) ( ([( ( ctx , cld ) )] @ ctxcld_list) ) (f) ( None )" | ftip_cons_trueI: "\class_fields (cld) (fds) ; ftype_in_fds (P) (ctx) (fds) (f) ((ty_opt_bot_opt (Some ( ty )) ))\ \ ftype_in_path (P) ( ([( ( ctx , cld ) )] @ ctxcld_list) ) (f) ( (Some ( ty )) )" | ftip_cons_falseI: "\class_fields (cld) (fds) ; ftype_in_fds (P) (ctx) (fds) (f) ((ty_opt_bot_opt None )) ; ftype_in_path (P) ( (ctxcld_list) ) (f) (ty_opt)\ \ ftype_in_path (P) ( ([( ( ctx , cld ) )] @ ctxcld_list) ) (f) (ty_opt)" (*defns ftype_def *) inductive ftype :: "P \ ty \ f \ ty \ bool" where (* defn ftype *) ftypeI: "\find_path_ty (P) (ty) ( Some ( ctxclds ) ) ; ftype_in_path (P) (ctxclds) (f) ( (Some ( ty' )) )\ \ ftype (P) (ty) (f) (ty')" (*defns find_meth_def_in_list_def *) inductive find_meth_def_in_list :: "meth_defs \ meth \ meth_def_opt \ bool" where (* defn find_meth_def_in_list *) fmdil_emptyI: "find_meth_def_in_list ( [] ) (meth) ( None )" | fmdil_cons_trueI: "\ meth_def = (meth_def_def (meth_sig_def cl meth vds) meth_body) \ \ find_meth_def_in_list ( ([(meth_def)] @ meth_def_list) ) (meth) ( Some ( meth_def ) )" | fmdil_cons_falseI: "\ meth_def = (meth_def_def (meth_sig_def cl meth' vds) meth_body) ; meth \ meth' ; find_meth_def_in_list ( (meth_def_list) ) (meth) (meth_def_opt)\ \ find_meth_def_in_list ( ([(meth_def)] @ meth_def_list) ) (meth) (meth_def_opt)" (*defns find_meth_def_in_path_def *) inductive find_meth_def_in_path :: "ctxclds \ meth \ ctxmeth_def_opt \ bool" where (* defn find_meth_def_in_path *) fmdip_emptyI: "find_meth_def_in_path ( [] ) (meth) ( (None::ctxmeth_def_opt) )" | fmdip_cons_trueI: "\class_methods (cld) (meth_defs) ; find_meth_def_in_list (meth_defs) (meth) ( Some ( meth_def ) )\ \ find_meth_def_in_path ( ([( ( ctx , cld ) )] @ ctxcld_list) ) (meth) ( (Some ( ctx , meth_def )::ctxmeth_def_opt) )" | fmdip_cons_falseI: "\class_methods (cld) (meth_defs) ; find_meth_def_in_list (meth_defs) (meth) ( None ) ; find_meth_def_in_path ( (ctxcld_list) ) (meth) (ctxmeth_def_opt)\ \ find_meth_def_in_path ( ([( ( ctx , cld ) )] @ ctxcld_list) ) (meth) (ctxmeth_def_opt)" (*defns find_meth_def_def *) inductive find_meth_def :: "P \ ty \ meth \ ctxmeth_def_opt \ bool" where (* defn find_meth_def *) fmd_nullI: "\find_path_ty (P) (ty) ( None )\ \ find_meth_def (P) (ty) (meth) ( (None::ctxmeth_def_opt) )" | fmd_optI: "\find_path_ty (P) (ty) ( Some ( ctxclds ) ) ; find_meth_def_in_path (ctxclds) (meth) (ctxmeth_def_opt)\ \ find_meth_def (P) (ty) (meth) (ctxmeth_def_opt)" (*defns mtype_def *) inductive mtype :: "P \ ty \ meth \ mty \ bool" where (* defn mtype *) mtypeI: "\find_meth_def (P) (ty) (meth) ( (Some ( ctx , meth_def )::ctxmeth_def_opt) ) ; meth_def = (meth_def_def (meth_sig_def cl meth ((List.map (%((cl_XXX::cl),(var_XXX::var),(ty_XXX::ty)).(vd_def cl_XXX var_XXX)) cl_var_ty_list)) ) meth_body) ; find_type (P) (ctx) (cl) ( (Some ( ty' )) ) ; list_all (\f. f) ((List.map (%((cl_XXX::cl),(var_XXX::var),(ty_XXX::ty)).find_type (P) (ctx) (cl_XXX) ( (Some ( ty_XXX )) )) cl_var_ty_list)) ; mty = (mty_def ((List.map (%((cl_XXX::cl),(var_XXX::var),(ty_XXX::ty)).ty_XXX) cl_var_ty_list)) ty') \ \ mtype (P) (ty) (meth) (mty)" (*defns sty_one_def *) inductive sty_one :: "P \ ty \ ty \ bool" where (* defn one *) sty_objI: "\find_path_ty (P) (ty) ( Some ( ctxclds ) )\ \ sty_one (P) (ty) (ty_top)" | sty_dclI: "\find_path_ty (P) (ty) ( Some ( ((List.map (%((ctx_XXX::ctx),(cld_XXX::cld),(dcl_XXX::dcl)). ( ctx_XXX , cld_XXX ) ) ctx_cld_dcl_list)) ) ) ; list_all (\f. f) ((List.map (%((ctx_XXX::ctx),(cld_XXX::cld),(dcl_XXX::dcl)).class_name (cld_XXX) (dcl_XXX)) ctx_cld_dcl_list)) ; ( ctx' , dcl' ) \ set ((List.map (%((ctx_XXX::ctx),(cld_XXX::cld),(dcl_XXX::dcl)).(ctx_XXX,dcl_XXX)) ctx_cld_dcl_list)) \ \ sty_one (P) (ty) ((ty_def ctx' dcl'))" (*defns sty_many_def *) inductive sty_many :: "P \ tys \ tys \ bool" where (* defn many *) sty_manyI: "\ tys = ((List.map (%((ty_XXX::ty),(ty_'::ty)).ty_XXX) ty_ty'_list)) ; tys' = ((List.map (%((ty_XXX::ty),(ty_'::ty)).ty_') ty_ty'_list)) ; list_all (\f. f) ((List.map (%((ty_XXX::ty),(ty_'::ty)).sty_one (P) (ty_XXX) (ty_')) ty_ty'_list)) \ \ sty_many (P) (tys) (tys')" (*defns sty_option_def *) inductive sty_option :: "P \ ty_opt \ ty_opt \ bool" where (* defn option *) sty_optionI: "\ ty_opt = (Some ( ty )) ; ty_opt' = (Some ( ty' )) ; sty_one (P) (ty) (ty')\ \ sty_option (P) (ty_opt) (ty_opt')" (*defns well_formedness *) inductive wf_object :: "P \ H \ v_opt \ ty_opt \ bool" and wf_varstate :: "P \ \ \ H \ L \ bool" and wf_heap :: "P \ H \ bool" and wf_config :: "\ \ config \ bool" and wf_stmt :: "P \ \ \ s \ bool" and wf_meth :: "P \ ty \ meth_def \ bool" and wf_class_common :: "P \ ctx \ dcl \ cl \ fds \ meth_defs \ bool" and wf_class :: "P \ cld \ bool" and wf_program :: "P \ bool" where (* defn object *) wf_nullI: "\ ty_opt = (Some ( ty )) \ \ wf_object (P) (H) ( Some v_null ) (ty_opt)" | wf_objectI: "\sty_option (P) ( (case H oid of None \ None | Some tyfs \ Some (fst tyfs)) ) (ty_opt)\ \ wf_object (P) (H) ( Some (v_oid oid) ) (ty_opt)" | (* defn varstate *) wf_varstateI: "\ finite (dom ( L )) ; \ x \ dom \ . wf_object (P) (H) ( L x ) ( \ x ) \ \ wf_varstate (P) (\) (H) (L)" | (* defn heap *) wf_heapI: "\ finite (dom ( H )) ; \ oid \ dom H . ( \ ty . (case H oid of None \ None | Some tyfs \ Some (fst tyfs)) = (Some ( ty )) \ (\ fs . fields (P) (ty) ( Some ( fs ) ) \ (\ f \ set fs . \ ty' . ( ftype (P) (ty) (f) (ty') \ wf_object (P) (H) ( (case H oid of None \ None | Some tyfs \ (snd tyfs) f ) ) ( (Some ( ty' )) ) ) ) ) ) \ \ wf_heap (P) (H)" | (* defn config *) wf_all_exI: "\wf_program (P) ; wf_heap (P) (H) ; wf_varstate (P) (\) (H) (L)\ \ wf_config (\) ((config_ex P L H Exception))" | wf_allI: "\wf_program (P) ; wf_heap (P) (H) ; wf_varstate (P) (\) (H) (L) ; list_all (\f. f) ((List.map (%(s_XXX::s).wf_stmt (P) (\) (s_XXX)) s_list)) \ \ wf_config (\) ((config_normal P L H (s_list)))" | (* defn stmt *) wf_blockI: "\ list_all (\f. f) ((List.map (%(s_XXX::s).wf_stmt (P) (\) (s_XXX)) s_list)) \ \ wf_stmt (P) (\) ((s_block (s_list)))" | wf_var_assignI: "\sty_option (P) ( \ x ) ( \ (x_var var) )\ \ wf_stmt (P) (\) ((s_ass var x))" | wf_field_readI: "\ \ x = (Some ( ty )) ; ftype (P) (ty) (f) (ty') ; sty_option (P) ( (Some ( ty' )) ) ( \ (x_var var) )\ \ wf_stmt (P) (\) ((s_read var x f))" | wf_field_writeI: "\ \ x = (Some ( ty )) ; ftype (P) (ty) (f) (ty') ; sty_option (P) ( \ y ) ( (Some ( ty' )) )\ \ wf_stmt (P) (\) ((s_write x f y))" | wf_ifI: "\ sty_option (P) ( \ x ) ( \ y ) \ sty_option (P) ( \ y ) ( \ x ) ; wf_stmt (P) (\) (s1) ; wf_stmt (P) (\) (s2)\ \ wf_stmt (P) (\) ((s_if x y s1 s2))" | wf_newI: "\find_type (P) (ctx) (cl) ( (Some ( ty )) ) ; sty_option (P) ( (Some ( ty )) ) ( \ (x_var var) )\ \ wf_stmt (P) (\) ((s_new var ctx cl))" | wf_mcallI: "\ Y = ((List.map (%((y_XXX::x),(ty_XXX::ty)).y_XXX) y_ty_list)) ; \ x = (Some ( ty )) ; mtype (P) (ty) (meth) ((mty_def ((List.map (%((y_XXX::x),(ty_XXX::ty)).ty_XXX) y_ty_list)) ty')) ; list_all (\f. f) ((List.map (%((y_XXX::x),(ty_XXX::ty)).sty_option (P) ( \ y_XXX ) ( (Some ( ty_XXX )) )) y_ty_list)) ; sty_option (P) ( (Some ( ty' )) ) ( \ (x_var var) )\ \ wf_stmt (P) (\) ((s_call var x meth Y))" | (* defn meth *) wf_methodI: "\ distinct ( ((List.map (%((cl_XXX::cl),(var_XXX::var),(ty_XXX::ty)).var_XXX) cl_var_ty_list)) ) ; list_all (\f. f) ((List.map (%((cl_XXX::cl),(var_XXX::var),(ty_XXX::ty)).find_type (P) (ctx) (cl_XXX) ( (Some ( ty_XXX )) )) cl_var_ty_list)) ; \ = ( (map_of ( ((List.map (%((cl_XXX::cl),(var_XXX::var),(ty_XXX::ty)).((x_var var_XXX),ty_XXX)) cl_var_ty_list)) )) ( x_this \ (ty_def ctx dcl) )) ; list_all (\f. f) ((List.map (%(s_XXX::s).wf_stmt (P) (\) (s_XXX)) s_list)) ; find_type (P) (ctx) (cl) ( (Some ( ty )) ) ; sty_option (P) ( \ y ) ( (Some ( ty )) )\ \ wf_meth (P) ((ty_def ctx dcl)) ((meth_def_def (meth_sig_def cl meth ((List.map (%((cl_XXX::cl),(var_XXX::var),(ty_XXX::ty)).(vd_def cl_XXX var_XXX)) cl_var_ty_list)) ) (meth_body_def (s_list) y)))" | (* defn class_common *) wf_class_commonI: "\find_type (P) (ctx) (cl) ( (Some ( ty )) ) ; (ty_def ctx dcl) \ ty ; distinct ( ((List.map (%((cl_XXX::cl),(f_XXX::f),(ty_XXX::ty)).f_XXX) cl_f_ty_list)) ) ; fields (P) (ty) ( Some ( fs ) ) ; (set ((List.map (%((cl_XXX::cl),(f_XXX::f),(ty_XXX::ty)).f_XXX) cl_f_ty_list)) ) \ (set fs ) = {} ; list_all (\f. f) ((List.map (%((cl_XXX::cl),(f_XXX::f),(ty_XXX::ty)).find_type (P) (ctx) (cl_XXX) ( (Some ( ty_XXX )) )) cl_f_ty_list)) ; list_all (\f. f) ((List.map (%((meth_def_XXX::meth_def),(meth_XXX::meth)).wf_meth (P) ((ty_def ctx dcl)) (meth_def_XXX)) meth_def_meth_list)) ; list_all (\f. f) ((List.map (%((meth_def_XXX::meth_def),(meth_XXX::meth)).method_name (meth_def_XXX) (meth_XXX)) meth_def_meth_list)) ; distinct ( ((List.map (%((meth_def_XXX::meth_def),(meth_XXX::meth)).meth_XXX) meth_def_meth_list)) ) ; methods (P) (ty) ( ((List.map (%((meth_'::meth),(mty_XXX::mty),(mty_'::mty)).meth_') meth'_mty_mty'_list)) ) ; list_all (\f. f) ((List.map (%((meth_'::meth),(mty_XXX::mty),(mty_'::mty)).mtype (P) ((ty_def ctx dcl)) (meth_') (mty_XXX)) meth'_mty_mty'_list)) ; list_all (\f. f) ((List.map (%((meth_'::meth),(mty_XXX::mty),(mty_'::mty)).mtype (P) (ty) (meth_') (mty_')) meth'_mty_mty'_list)) ; list_all (\f. f) ((List.map (%((meth_'::meth),(mty_XXX::mty),(mty_'::mty)). meth_' \ set ((List.map (%((meth_def_XXX::meth_def),(meth_XXX::meth)).meth_XXX) meth_def_meth_list)) \ mty_XXX = mty_' ) meth'_mty_mty'_list)) \ \ wf_class_common (P) (ctx) (dcl) (cl) ( ((List.map (%((cl_XXX::cl),(f_XXX::f),(ty_XXX::ty)).(fd_def cl_XXX f_XXX)) cl_f_ty_list)) ) ( ((List.map (%((meth_def_XXX::meth_def),(meth_XXX::meth)).meth_def_XXX) meth_def_meth_list)) )" | (* defn class *) wf_classI: "\ (cld_def dcl cl fds meth_defs) \ set P ; wf_class_common (P) (ctx_def) (dcl) (cl) (fds) (meth_defs)\ \ wf_class (P) ((cld_def dcl cl fds meth_defs))" | (* defn program *) wf_programI: "\ P = (cld_list) ; distinct_names (P) ; list_all (\f. f) ((List.map (%(cld_XXX::cld).wf_class (P) (cld_XXX)) cld_list)) ; acyclic_clds (P)\ \ wf_program (P)" (*defns var_trans *) inductive tr_s :: "T \ s \ s \ bool" where (* defn tr_s *) tr_s_blockI: "\ list_all (\f. f) ((List.map (%((s_XXX::s),(s_'::s)).tr_s (T) (s_XXX) (s_')) s_s'_list)) \ \ tr_s (T) ((s_block ((List.map (%((s_XXX::s),(s_'::s)).s_XXX) s_s'_list)))) ((s_block ((List.map (%((s_XXX::s),(s_'::s)).s_') s_s'_list))))" | tr_s_var_assignI: "\ (case T (x_var var ) of None \ var | Some x' \ (case x' of x_this \ var | x_var var' \ var')) = var' ; (case T x of None \ x | Some x' \ x') = x' \ \ tr_s (T) ((s_ass var x)) ((s_ass var' x'))" | tr_s_field_readI: "\ (case T (x_var var ) of None \ var | Some x' \ (case x' of x_this \ var | x_var var' \ var')) = var' ; (case T x of None \ x | Some x' \ x') = x' \ \ tr_s (T) ((s_read var x f)) ((s_read var' x' f))" | tr_s_field_writeI: "\ (case T x of None \ x | Some x' \ x') = x' ; (case T y of None \ y | Some x' \ x') = y' \ \ tr_s (T) ((s_write x f y)) ((s_write x' f y'))" | tr_s_ifI: "\ (case T x of None \ x | Some x' \ x') = x' ; (case T y of None \ y | Some x' \ x') = y' ; tr_s (T) (s1) (s1') ; tr_s (T) (s2) (s2')\ \ tr_s (T) ((s_if x y s1 s2)) ((s_if x' y' s1' s2'))" | tr_s_newI: "\ (case T (x_var var ) of None \ var | Some x' \ (case x' of x_this \ var | x_var var' \ var')) = var' \ \ tr_s (T) ((s_new var ctx cl)) ((s_new var' ctx cl))" | tr_s_mcallI: "\ (case T (x_var var ) of None \ var | Some x' \ (case x' of x_this \ var | x_var var' \ var')) = var' ; (case T x of None \ x | Some x' \ x') = x' ; list_all (\f. f) ((List.map (%((y_XXX::x),(y_'::x)). (case T y_XXX of None \ y_XXX | Some x' \ x') = y_' ) y_y'_list)) \ \ tr_s (T) ((s_call var x meth ((List.map (%((y_XXX::x),(y_'::x)).y_XXX) y_y'_list)) )) ((s_call var' x' meth ((List.map (%((y_XXX::x),(y_'::x)).y_') y_y'_list)) ))" (*defns reduction *) inductive r_stmt :: "config \ config \ bool" where (* defn stmt *) r_blockI: "r_stmt ((config_normal P L H ([((s_block (s_list)))] @ s'_list))) ((config_normal P L H (s_list @ s'_list)))" | r_var_assignI: "\ L x = Some v \ \ r_stmt ((config_normal P L H ([((s_ass var x))] @ s_list))) ((config_normal P ( L ( (x_var var) \ v )) H (s_list)))" | r_field_read_npeI: "\ L x = Some v_null \ \ r_stmt ((config_normal P L H ([((s_read var x f))] @ s_list))) ((config_ex P L H ex_npe))" | r_field_readI: "\ L x = Some (v_oid oid) ; (case H oid of None \ None | Some tyfs \ (snd tyfs) f ) = Some v \ \ r_stmt ((config_normal P L H ([((s_read var x f))] @ s_list))) ((config_normal P ( L ( (x_var var) \ v )) H (s_list)))" | r_field_write_npeI: "\ L x = Some v_null \ \ r_stmt ((config_normal P L H ([((s_write x f y))] @ s_list))) ((config_ex P L H ex_npe))" | r_field_writeI: "\ L x = Some (v_oid oid) ; L y = Some v \ \ r_stmt ((config_normal P L H ([((s_write x f y))] @ s_list))) ((config_normal P L (case H oid of None \ arbitrary | Some tyfs \ (( H ( oid \ - (fst tyfs, snd tyfs ( f \ v ))))::H)) (s_list)))" + (fst tyfs, (snd tyfs) ( f \ v ))))::H)) (s_list)))" | r_if_trueI: "\ L x = Some v ; L y = Some w ; v = w \ \ r_stmt ((config_normal P L H ([((s_if x y s1 s2))] @ s'_list))) ((config_normal P L H ([(s1)] @ s'_list)))" | r_if_falseI: "\ L x = Some v ; L y = Some w ; v \ w \ \ r_stmt ((config_normal P L H ([((s_if x y s1 s2))] @ s'_list))) ((config_normal P L H ([(s2)] @ s'_list)))" | r_newI: "\find_type (P) (ctx) (cl) ( (Some ( ty )) ) ; fields (P) (ty) ( Some ( (f_list) ) ) ; oid \ dom H ; H' = (( H ( oid \ ( ty , map_of ((List.map (%(f_XXX::f).(f_XXX,v_null)) f_list)) )))::H) \ \ r_stmt ((config_normal P L H ([((s_new var ctx cl))] @ s_list))) ((config_normal P ( L ( (x_var var) \ (v_oid oid) )) H' (s_list)))" | r_mcall_npeI: "\ L x = Some v_null \ \ r_stmt ((config_normal P L H ([((s_call var x meth (y_list) ))] @ s_list))) ((config_ex P L H ex_npe))" | r_mcallI: "\ L x = Some (v_oid oid) ; (case H oid of None \ None | Some tyfs \ Some (fst tyfs)) = (Some ( ty )) ; find_meth_def (P) (ty) (meth) ( (Some ( ctx , (meth_def_def (meth_sig_def cl meth ((List.map (%((y_XXX::x),(cl_XXX::cl),(var_XXX::var),(var_'::var),(v_XXX::v)).(vd_def cl_XXX var_XXX)) y_cl_var_var'_v_list)) ) (meth_body_def ((List.map (%((s_''::s),(s_'::s)).s_') s''_s'_list)) y)) )::ctxmeth_def_opt) ) ; (set ((List.map (%((y_XXX::x),(cl_XXX::cl),(var_XXX::var),(var_'::var),(v_XXX::v)).(x_var var_')) y_cl_var_var'_v_list)) ) Int (dom L ) = {} ; distinct ( ((List.map (%((y_XXX::x),(cl_XXX::cl),(var_XXX::var),(var_'::var),(v_XXX::v)).var_') y_cl_var_var'_v_list)) ) ; x' \ dom L ; x' \ set ((List.map (%((y_XXX::x),(cl_XXX::cl),(var_XXX::var),(var_'::var),(v_XXX::v)).(x_var var_')) y_cl_var_var'_v_list)) ; list_all (\f. f) ((List.map (%((y_XXX::x),(cl_XXX::cl),(var_XXX::var),(var_'::var),(v_XXX::v)). L y_XXX = Some v_XXX ) y_cl_var_var'_v_list)) ; L' = ( ( L ++ (map_of ( ((List.map (%((y_XXX::x),(cl_XXX::cl),(var_XXX::var),(var_'::var),(v_XXX::v)).((x_var var_'),v_XXX)) y_cl_var_var'_v_list)) ))) ( x' \ (v_oid oid) )) ; T = ( (map_of ( ((List.map (%((y_XXX::x),(cl_XXX::cl),(var_XXX::var),(var_'::var),(v_XXX::v)).((x_var var_XXX),(x_var var_'))) y_cl_var_var'_v_list)) )) ( x_this \ x' )) ; list_all (\f. f) ((List.map (%((s_''::s),(s_'::s)).tr_s (T) (s_') (s_'')) s''_s'_list)) ; (case T y of None \ y | Some x' \ x') = y' \ \ r_stmt ((config_normal P L H ([((s_call var x meth ((List.map (%((y_XXX::x),(cl_XXX::cl),(var_XXX::var),(var_'::var),(v_XXX::v)).y_XXX) y_cl_var_var'_v_list)) ))] @ s_list))) ((config_normal P L' H ((List.map (%((s_''::s),(s_'::s)).s_'') s''_s'_list) @ [((s_ass var y'))] @ s_list)))" end diff --git a/thys/LightweightJava/Lightweight_Java_Proof.thy b/thys/LightweightJava/Lightweight_Java_Proof.thy --- a/thys/LightweightJava/Lightweight_Java_Proof.thy +++ b/thys/LightweightJava/Lightweight_Java_Proof.thy @@ -1,1819 +1,1819 @@ (* Title: Lightweight Java, the proof Authors: Rok Strnisa , 2006 Matthew Parkinson , 2006 Maintainer: *) theory Lightweight_Java_Proof imports Lightweight_Java_Equivalence "HOL-Library.Infinite_Set" begin lemmas wf_intros = wf_object_wf_varstate_wf_heap_wf_config_wf_stmt_wf_meth_wf_class_common_wf_class_wf_program.intros [simplified] lemmas wf_induct = wf_object_wf_varstate_wf_heap_wf_config_wf_stmt_wf_meth_wf_class_common_wf_class_wf_program.induct [simplified] lemmas wf_config_normalI = wf_object_wf_varstate_wf_heap_wf_config_wf_stmt_wf_meth_wf_class_common_wf_class_wf_program.wf_allI [simplified] lemmas wf_objectE = wf_object.cases[simplified] lemmas wf_varstateE = wf_varstate.cases[simplified] lemmas wf_heapE = wf_heap.cases[simplified] lemmas wf_configE = wf_config.cases[simplified] lemmas wf_stmtE = wf_stmt.cases[simplified] lemmas wf_methE = wf_meth.cases[simplified] lemmas wf_class_cE = wf_class_common.cases[simplified] lemmas wf_classE = wf_class.cases[simplified] lemmas wf_programE = wf_program.cases[simplified] lemma wf_subvarstateI: "\wf_varstate P \ H L; \ x' = Some ty; wf_object P H (Some v) (Some ty)\ \ wf_varstate P \ H (L (x' \ v))" apply(erule wf_varstateE) apply(rule wf_varstateI) apply(simp) apply(clarsimp) done lemma finite_super_varstate: "finite (dom ((L ++ map_of (map (\(y, cl, var, var', y). (x_var var', y)) y_cl_var_var'_v_list))(x' \ v_oid oid))) = finite (dom L)" apply(induct y_cl_var_var'_v_list) apply(clarsimp) apply(simp add: dom_def) apply(subgoal_tac "{a. a \ x' \ (\y. L a = Some y)} = {a. a = x' \ (\y. L a = Some y)}") apply(simp add: Collect_disj_eq) apply(force) apply(clarsimp simp add: map_add_def dom_def split: if_split_asm) apply(subgoal_tac "{aa. aa \ x_var a \ aa \ x' \ (\y. case_option (L aa) Some (map_of (map (\(y, cl, var, var', y). (x_var var', y)) y_cl_var_var'_v_list) aa) = Some y)} = {aa. aa = x_var a \ (aa = x' \ (\y. case_option (L aa) Some (map_of (map (\(y, cl, var, var', y). (x_var var', y)) y_cl_var_var'_v_list) aa) = Some y))}") apply(simp add: Collect_disj_eq) apply(force) done lemma fst_map_elem: "(y_k, ty_k, var_k, var'_k, v_k) \ set y_ty_var_var'_v_list \ x_var var'_k \ fst ` (\(y, ty, var, var', v). (x_var var', ty)) ` set y_ty_var_var'_v_list" by (force) lemma map_add_x_var[THEN mp]: "var' \ set (map (\(y, cl, var, var', v). var) y_cl_var_var'_v_list) \ x_var var' \ set (map (\(y, cl, var, var', v). x_var var) y_cl_var_var'_v_list)" by (induct y_cl_var_var'_v_list, auto) lemma map_of_map_and_x[simp]: "\\ x = Some ty_x; L x' = None; \x\set y_cl_var_var'_v_list. (\(y, cl, var, var', v). L (x_var var') = None) x; \x\Map.dom \. wf_object Pa H (L x) (\ x)\ \ (if x = x' then Some ty else (\ ++ map_of (map (\((y, cl, var, var', v), y', ty). (x_var var', ty)) (zip y_cl_var_var'_v_list y_ty_list))) x) = Some ty_x" apply(subgoal_tac "x \ x'") apply(subgoal_tac "\(y, cl, var, var', v) \ set y_cl_var_var'_v_list. x_var var' \ x") apply(case_tac "map_of (map (\((y, cl, var, var', v), y', ty). (x_var var', ty)) (zip y_cl_var_var'_v_list y_ty_list)) x") apply(clarsimp simp add: map_add_def) apply(clarsimp) apply(drule map_of_SomeD) apply(clarsimp) apply(rename_tac y cl var var' v y' ty) apply(drule_tac x = "(y, cl, var, var', v)" in bspec) apply(force simp add: set_zip) apply(drule_tac x = "x_var var'" in bspec, simp add: domI) apply(erule wf_objectE, simp+) apply(force elim: wf_objectE)+ done lemma wf_stmt_in_G': "(L x' = None \ (\x\set y_cl_var_var'_v_list. (\(y, cl, var, var', v). L (x_var var') = None) x) \ (\x\dom \. wf_object P H (L x) (\ x)) \ wf_stmt P \ s \ wf_stmt P ((\ ++ map_of (map (\((y, cl, var, var', v), y', y). (x_var var', y)) (zip y_cl_var_var'_v_list y_ty_list)))(x' \ ty)) s) \ (L x' = None \ (\x\set y_cl_var_var'_v_list. (\(y, cl, var, var', v). L (x_var var') = None) x) \ (\x\dom \. wf_object P H (L x) (\ x)) \ (\s\set ss. wf_stmt P \ s) \ (\s\set ss. wf_stmt P ((\ ++ map_of (map (\((y, cl, var, var', v), y', y). (x_var var', y)) (zip y_cl_var_var'_v_list y_ty_list)))(x' \ ty)) s))" apply(rule tr_s_f_tr_ss_f.induct) apply(clarsimp, erule wf_stmtE, simp_all, simp add: wf_blockI) apply(clarify, erule wf_stmtE, simp_all) apply(erule sty_option.cases, force intro: wf_var_assignI sty_optionI) apply(clarify, erule wf_stmtE, simp_all) apply(erule sty_option.cases, force intro: wf_field_readI sty_optionI) apply(clarify, erule wf_stmtE, simp_all) apply(erule sty_option.cases, force intro: wf_field_writeI sty_optionI) apply(clarify, erule wf_stmtE, simp_all) apply(clarsimp) apply(rule wf_ifI) apply(erule disjE) apply(rule disjI1, erule sty_option.cases, force intro: sty_optionI) apply(rule disjI2, erule sty_option.cases, force intro: sty_optionI) apply(assumption+) apply(clarify, erule wf_stmtE, simp_all) apply(erule sty_option.cases, force intro: wf_newI sty_optionI) apply(clarify, erule wf_stmtE, simp_all) apply(rule wf_mcallI, simp, simp, simp) apply(clarsimp split del: if_split) apply(rename_tac y_k ty_k) apply(drule_tac x = "(y_k, ty_k)" in bspec, simp, clarify) apply(erule_tac ?a3.0 = "Some ty_k" in sty_option.cases) apply(force intro: sty_optionI) apply(clarify) apply(erule sty_option.cases) apply(rule sty_optionI, simp+) done lemma map_three_in_four: "var_k \ set (map (\(y, cl, var, u). var) y_cl_var_var'_v_list) \ (\y cl u. (y, cl, var_k, u) \ set y_cl_var_var'_v_list)" by (induct y_cl_var_var'_v_list, auto) lemma map_var: "\(cl_k, var_k, ty_k) \ set cl_var_ty_list; map (\(y, cl, var, u). var) y_cl_var_var'_v_list = map (\(cl, var, ty). var) cl_var_ty_list\ \ \y_k cl'_k u_k. (y_k, cl'_k, var_k, u_k) \ set y_cl_var_var'_v_list" apply(subgoal_tac "var_k \ set (map (\(y, cl, var, u). var) y_cl_var_var'_v_list)") by (force simp add: map_three_in_four)+ lemma length_one_in_two: "length y_ty_list = length (map fst y_ty_list)" by (induct y_ty_list, auto) lemma length_two_in_two: "length y_ty_list = length (map snd y_ty_list)" by (induct y_ty_list, auto) lemma length_two_in_three: "length cl_var_ty_list = length (map (\(cl, var, ty). var) cl_var_ty_list)" by (induct cl_var_ty_list, auto) lemma length_three_in_three: "length cl_var_ty_list = length (map (\(cl, var, ty). ty) cl_var_ty_list)" by (induct cl_var_ty_list, auto) lemma length_three_in_four: "length y_cl_var_var'_v_list = length (map (\(y, cl, var, u). var) y_cl_var_var'_v_list)" by (induct y_cl_var_var'_v_list, auto) lemma length_one_in_five: "length y_cl_var_var'_v_list = length (map (\(y, cl, var, var', v). y) y_cl_var_var'_v_list)" by (induct y_cl_var_var'_v_list, auto) lemma length_three_in_five: "length y_cl_var_var'_v_list = length (map (\(y, cl, var, var', v). var) y_cl_var_var'_v_list)" by (induct y_cl_var_var'_v_list, auto) lemma map_length_list: "length (map (\((y, y'), yy, ty). (y', ty)) list) = length list" by (induct list, auto) lemma map_length_y: "map (\(y, cl, var, var', v). y) y_cl_var_var'_v_list = map fst y_ty_list \ length y_cl_var_var'_v_list = length y_ty_list" by (simp only: length_one_in_five length_one_in_two) lemma map_length_y': "map fst y_y'_list = map fst y_ty_list \ length y_y'_list = length y_ty_list" by (simp only: length_one_in_two length_two_in_two) lemma map_length_var: "map (\(y, cl, var, var', v). var) y_cl_var_var'_v_list = map (\(cl, var, ty). var) cl_var_ty_list \ length y_cl_var_var'_v_list = length cl_var_ty_list" by (simp only: length_three_in_five length_two_in_three) lemma map_length_ty: "map (\(cl, var, ty). ty) cl_var_ty_list = map snd y_ty_list \ length cl_var_ty_list = length y_ty_list" by (simp only: length_three_in_three length_two_in_two) lemma map_length_var_ty: "\map (\(y, cl, var, u). var) y_cl_var_var'_v_list = map (\(cl, var, ty). var) cl_var_ty_list; map (\(cl, var, ty). ty) cl_var_ty_list = map snd y_ty_list\ \ length y_cl_var_var'_v_list = length y_ty_list" apply(rule_tac s = "length cl_var_ty_list" in trans) apply (simp only: length_three_in_four length_two_in_three) apply(simp only: length_three_in_three length_two_in_two) done lemma fun_eq_fst: "(fst \ (\((y, cl, var, var', v), y', y). (x_var var', y))) = (\((y, cl, var, var', v), y', y). (x_var var'))" by (simp add: fun_eq_iff) lemma fst_zip_eq: "length y_cl_var_var'_v_list = length y_ty_list \ (map fst (map (\((y, cl, var, var', v), y', y). (x_var var', y)) (zip y_cl_var_var'_v_list y_ty_list))) = (map (\(y, cl, var, var', v). x_var var') y_cl_var_var'_v_list)" apply(simp) apply(simp add: fun_eq_fst) apply(rule nth_equalityI) apply(force) apply(clarsimp) apply(frule nth_mem) apply(subgoal_tac "\x. (y_cl_var_var'_v_list ! i) = x") apply(force) apply(rule_tac x = "y_cl_var_var'_v_list ! i" in exI) apply(simp) done lemma distinct_x_var: "distinct (map (\(y, cl, var, var', v). x_var var) y_cl_var_var'_v_list) = distinct (map (\(y, cl, var, var', v). var) y_cl_var_var'_v_list)" apply (induct y_cl_var_var'_v_list) by (force simp add: contrapos_np)+ lemma distinct_x_var': "distinct (map (\(y, cl, var, var', v). x_var var') y_cl_var_var'_v_list) = distinct (map (\(y, cl, var, var', v). var') y_cl_var_var'_v_list)" apply(induct y_cl_var_var'_v_list) by (force simp add: contrapos_np)+ lemma map_fst_two: "map fst (map (\(y, cl, var, var', v). (x_var var, x_var var')) y_cl_var_var'_v_list) = map (\(y, cl, var, var', v). x_var var) y_cl_var_var'_v_list" by (induct y_cl_var_var'_v_list, auto) lemma map_fst_two': "map fst (map (\(y, cl, var, var', y). (x_var var', y)) y_cl_var_var'_v_list) = map (\(y, cl, var, var', y). x_var var') y_cl_var_var'_v_list" by (induct y_cl_var_var'_v_list, auto) lemma map_fst_var': "distinct (map fst (map (\(y, cl, var, var', y). (x_v