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_var var', y)) y_cl_var_var'_v_list)) = distinct (map (\(y, cl, var, var', v). var') y_cl_var_var'_v_list)" by (simp only: map_fst_two' distinct_x_var') lemma distinct_var: "\distinct (map (\(cl, var, ty). var) cl_var_ty_list); map (\(y, cl, var, var', v). var) y_cl_var_var'_v_list = map (\(cl, var, ty). var) cl_var_ty_list\ \ distinct (map fst (map (\(y, cl, var, var', v). (x_var var, x_var var')) y_cl_var_var'_v_list))" by (simp only: map_fst_two distinct_x_var) lemma distinct_var': "\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\ \ distinct (map fst (map (\((y, cl, var, var', v), y', y). (x_var var', y)) (zip y_cl_var_var'_v_list y_ty_list))) = distinct (map (\(y, cl, var, var', v). var') y_cl_var_var'_v_list)" by (simp only: map_length_var_ty fst_zip_eq distinct_x_var') lemma weaken_list_var: "map (\(y, cl, var, var', v). var) y_cl_var_var'_v_list = map (\(y, cl, var, u). var) y_cl_var_var'_v_list" by (induct y_cl_var_var'_v_list, auto) lemma weaken_list_fd: "map (\(y, cl, var, var', v). vd_def cl var) list = map (\(y, cl, var, u). vd_def cl var) list" by (induct list, auto) lemma weaken_list_y: "map (\(y, cl, var, var', v). y) y_cl_var_var'_v_list = map fst y_cl_var_var'_v_list" by (induct y_cl_var_var'_v_list, auto) lemma wf_cld_from_lu: "\wf_program P; find_cld P ctx fqn (Some (ctx', cld'))\ \ wf_class P cld'" apply(clarsimp) apply(drule find_to_mem) apply(erule wf_programE) by (simp add: member_def) lemma meth_def_in_set[rule_format]: "find_meth_def_in_list meth_defs meth (Some meth_def) \ meth_def \ set meth_defs" apply(clarsimp) apply(induct meth_defs) by (auto split: meth_def.splits meth_sig.splits if_split_asm) lemma find_meth_def_in_list_mem[rule_format, simp]: "find_meth_def_in_list_f meth_defs meth = Some meth_def \ meth_def \ set meth_defs" apply(induct meth_defs) apply(auto split: meth_def.splits meth_sig.splits) done lemma find_meth_def_in_path_ex_cld[rule_format]: "find_meth_def_in_path_f ctxclds meth = Some (ctx, meth_def) \ (\cld meth_defs. (ctx, cld) \ set ctxclds \ class_methods_f cld = meth_defs \ meth_def \ set meth_defs)" apply(induct ctxclds) apply(simp) apply(clarsimp) apply(clarsimp split: option.splits) apply(rule_tac x = cld in exI) apply(force) apply(force) done lemma map_ctx_cld_dcl_two[simp]: "ctxclds = map (\(ctx, cld, dcl). (ctx, cld)) (map (\(ctx, cld). (ctx, cld, classname_f cld)) ctxclds)" by (induct ctxclds, auto) lemma clds_in_path_exist: "find_path_f P ctx' fqn = Some path \ (\(ctx, cld) \ set path. cld \ set P)" apply(clarify) apply(drule all_in_path_found) apply(simp) apply(clarsimp) apply(drule find_to_mem) apply(simp add: member_def) done lemma wf_meth_defs_in_wf_class[rule_format]: "meth_def \ set (class_methods_f cld) \ wf_class P cld \ wf_meth P (ty_def ctx_def (class_name_f cld)) meth_def" apply(clarsimp) apply(erule wf_classE) apply(clarsimp simp: class_methods_f_def) apply(erule wf_class_cE) apply(clarsimp simp: class_name_f_def) apply(drule(1)bspec, simp) done lemma map_ctxclds: "ctxclds = map ((\(ctx_XXX, cld_XXX, dcl_XXX). (ctx_XXX, cld_XXX)) \ (\(ctx, cld). (ctx, cld, class_name_f cld))) ctxclds" by (induct ctxclds, auto) lemma wf_method_from_find'[rule_format]: "wf_program P \ find_path_ty_f P (ty_def ctxa dcl) = Some ctxclds \ find_meth_def_in_path_f ctxclds meth = Some (ctx, meth_def) \ (\dcla. is_sty_one P (ty_def ctxa dcl) (ty_def ctx dcla) = Some True \ wf_meth P (ty_def ctx_def dcla) meth_def)" apply(clarsimp) apply(drule find_meth_def_in_path_ex_cld) apply(clarsimp) apply(rule_tac x = "class_name_f cld" in exI) apply(rule) apply(rule_tac ctx_cld_dcl_list = "map (\(ctx, cld). (ctx, cld, class_name_f cld)) ctxclds" in sty_dclI[simplified]) apply(simp add: map_ctxclds) apply(clarsimp) apply(force) apply(drule clds_in_path_exist) apply(simp) apply(erule wf_programE) apply(drule_tac x = "(ctx, cld)" in bspec, simp) apply(drule_tac x = cld in bspec, simp) apply(simp add: wf_meth_defs_in_wf_class) done (* fix the ctx problem here\ *) lemma wf_method_from_find: "\wf_program P; find_meth_def P ty_x_d meth (Some (ctx, meth_def))\ \ \dcl. sty_one P ty_x_d (ty_def ctx dcl) \ wf_meth P (ty_def ctx_def dcl) meth_def" apply(erule find_meth_def.cases, clarsimp+) apply(induct ty_x_d) apply(simp) apply(rename_tac ctx' dcl' ctxclds) apply(cut_tac P = P and ctxa = ctx' and dcl = dcl' and ctxclds = ctxclds and meth = meth and ctx = ctx and meth_def = meth_def in wf_method_from_find') apply(simp) apply(simp) done lemma find_type_lift_opts[rule_format]: "(\x\set cl_var_ty_list. (\(cl_k, var_k, ty_k). find_type_f P ctx cl_k = Some ty_k) x) \ lift_opts (map (case_vd (\clk vark. find_type_f P ctx clk)) (map (\(cl_k, var_k, ty_k). vd_def cl_k var_k) cl_var_ty_list)) = Some (map (\(cl, var, ty). ty) cl_var_ty_list)" by (induct cl_var_ty_list, auto) lemma find_md_in_pre_path[rule_format]: "find_meth_def_in_path_f path m = Some ctxmd \ (\path'. find_meth_def_in_path_f (path @ path') m = Some ctxmd)" by (induct path, auto split: option.splits) lemma lift_opts_map: "\x\set cl_var_ty_list. (\(cl_k, var_k, ty_k). find_type_f P ctx cl_k = Some ty_k) x \ lift_opts (map (case_vd (\clk vark. find_type_f P ctx clk)) (map (\(cl_k, var_k, ty_k). vd_def cl_k var_k) cl_var_ty_list)) = Some (map (\(cl, var, ty). ty) cl_var_ty_list)" by (induct cl_var_ty_list, auto) lemma m_in_list[rule_format]: "(\x\set meth_def_meth_list. case_prod (\meth_def. (=) (method_name_f meth_def)) x) \ find_meth_def_in_list_f (map fst meth_def_meth_list) m = Some md \ m \ snd ` set meth_def_meth_list" by (induct meth_def_meth_list, auto simp add: method_name_f_def split: meth_def.splits meth_sig.splits) lemma m_image_eq[rule_format]: "meth_def_def (meth_sig_def cl m list2) meth_body \ set meth_defs \ m \ case_meth_def (\meth_sig meth_body. case meth_sig of meth_sig_def cl meth vds \ meth) ` set meth_defs" by (induct meth_defs, auto) lemma fmdip_to_mip[rule_format]: "find_meth_def_in_path_f path m = Some ctxmd \ m \ set (methods_in_path_f (map snd path))" apply(induct path) apply(simp) apply(clarsimp simp add: class_methods_f_def split: option.splits cld.splits) apply(rename_tac aa bb) apply(case_tac aa) apply(rename_tac meth_sig meth_body) apply(case_tac meth_sig) apply(clarsimp) apply(frule find_md_m_match') apply(clarsimp) apply(frule find_meth_def_in_list_mem) apply(frule m_image_eq) apply(assumption) done lemma lift_opts_for_all[rule_format]: "lift_opts (map (case_vd (\clk vark. find_type_f P ctx clk)) (map (\(cl, var, ty). vd_def cl var) cl_var_ty_list)) = None \ (\x\set cl_var_ty_list. (\(cl, var, ty). find_type_f P ctx_def cl = Some ty) x) \ False" apply(induct cl_var_ty_list) apply(simp) apply(clarsimp) apply(case_tac ctx) apply(simp) done lemma mty_preservation'''': "\find_cld_f P ctx (fqn_def dcl') = Some (ctx, cld_def dcl' (cl_fqn (fqn_def dcl'')) fds' mds'); find_cld_f P ctx (fqn_def dcl'') = Some (ctx, cld_def dcl'' cl'' fds'' mds''); find_path_rec_f P ctx cl'' [(ctx, cld_def dcl'' cl'' fds'' mds'')] = Some path; find_meth_def_in_path_f path m = Some (ctx'', meth_def_def (meth_sig_def cl_r m' vds) mb); find_type_f P ctx'' cl_r = Some ty_r; lift_opts (map (case_vd (\clk vark. find_type_f P ctx'' clk)) vds) = Some tys; lift_opts (map (case_vd (\clk vark. find_type_f P ctx' clk)) vds') = Some tys'; find_type_f P ctx' cl_r' = Some ty_r'; find_meth_def_in_path_f ((ctx, cld_def dcl' (cl_fqn (fqn_def dcl'')) fds' mds') # path) m = Some (ctx', meth_def_def (meth_sig_def cl_r' m'' vds') mb'); wf_program P\ \ tys' = tys \ ty_r' = ty_r" using [[hypsubst_thin = true]] apply(clarsimp simp add: class_methods_f_def split: option.splits) apply(drule wf_cld_from_lu) apply(simp) apply(erule wf_classE) apply(erule wf_class_cE) apply(clarsimp) apply(subgoal_tac "m \ snd ` set meth_def_meth_list") apply(clarsimp) apply(rename_tac md_m m) apply(drule_tac x = "(md_m, m)" in bspec, simp)+ apply(clarsimp split: option.splits) apply(case_tac ctx') apply(clarsimp) apply(case_tac md_m) apply(rename_tac meth_sig meth_body) apply(case_tac meth_sig) apply(rename_tac ms mb_m cl_r_m m vds_m) apply(clarsimp simp add: method_name_f_def) apply(erule_tac Q = "tys' = tys \ ty_r' \ ty_r" in contrapos_pp) apply(clarsimp) apply(clarsimp simp add: methods_f_def find_path_f_def superclass_name_f_def split: option.splits if_split_asm) apply(frule fmdip_to_mip) apply(clarsimp) apply(rename_tac m mty mty') apply(drule_tac x = "(m, mty, mty')" in bspec, simp)+ apply(clarsimp) apply(frule_tac f = snd in imageI) apply(simp) apply(thin_tac "m \ snd ` set meth_def_meth_list") apply(clarsimp) apply(clarsimp simp add: mtype_f_def split: option.splits meth_def.splits meth_sig.splits) apply(clarsimp simp add: find_meth_def_f_def find_path_f_def superclass_name_f_def split: option.splits if_split_asm) apply(frule path_append) apply(rename_tac ad) apply(frule_tac path = ad in path_append) apply(clarsimp) apply(frule_tac prefix = "[(ctx_def, cld_def dcl'' cl'' fds'' mds'')]" and suffix = path'' and prefix' = "[(ctx_def, cld_def dcl' (cl_fqn (fqn_def dcl'')) (map (\(cl, f, ty). fd_def cl f) cl_f_ty_list) (map fst meth_def_meth_list)), (ctx_def, cld_def dcl'' cl'' fds'' mds'')]" in fpr_same_suffix[rule_format]) apply(simp) apply(clarsimp simp add: class_methods_f_def) apply(rule m_in_list) apply(simp) done lemma mty_preservation'''[rule_format]: "\suffix dcl. find_path_rec_f P ctx cl prefix = Some (prefix @ suffix) \ wf_program P \ cl = cl_fqn (fqn_def dcl) \ (\(ctx', cld') \ set suffix. (\prefix' suffix'. find_path_rec_f P ctx' (cl_fqn (fqn_def (class_name_f cld'))) prefix' = Some (prefix' @ suffix') \ mtype_f P (ty_def ctx' (class_name_f cld')) m = Some mty \ mtype_f P (ty_def ctx dcl) m = Some mty))" supply [[simproc del: defined_all]] apply(induct_tac P ctx cl prefix rule: find_path_rec_f.induct) apply(simp) apply(clarsimp split: option.splits) apply(frule path_append) apply(clarify) apply(clarsimp) apply(erule disjE) apply(clarify) apply(clarsimp simp add: class_name_f_def split: cld.splits) apply(frule find_cld_name_eq) apply(frule_tac ctx = a in find_cld_name_eq) apply(clarsimp simp add: class_name_f_def split: cld.splits) apply(case_tac b) apply(rename_tac dcl' cl' fds' mds') apply(clarsimp simp add: superclass_name_f_def) apply(case_tac cl') apply(simp) apply(rename_tac fqn) apply(case_tac fqn) apply(rename_tac dcl'') apply(clarsimp) apply(clarsimp split: option.splits) apply(subgoal_tac "\aa b ab bb. (a = ab \ cld_def dcl' (cl_fqn (fqn_def dcl'')) fds' mds' = bb \ aa = ab \ b = bb) \ (\suffix. find_path_rec_f P ab (case bb of cld_def dcl cl fds mds \ cl) (path @ [(ab, bb)]) = Some (path @ (ab, bb) # suffix) \ (\dcl. (case bb of cld_def dcl cl fds mds \ cl) = cl_fqn (fqn_def dcl) \ (\x\set suffix. (\(ctx', cld'). (\prefix' suffix'. case_option None (case_prod (\ctx' cld. find_path_rec_f P ctx' (superclass_name_f cld) (prefix' @ [(ctx', cld)]))) (find_cld_f P ctx' (fqn_def (class_name_f cld'))) = Some (prefix' @ suffix')) \ mtype_f P (ty_def ctx' (class_name_f cld')) m = Some mty \ mtype_f P (ty_def ab dcl) m = Some mty) x)))") defer apply(force) apply(thin_tac "\aa b x y. \a = aa \ cld_def dcl' (cl_fqn (fqn_def dcl'')) fds' mds' = b; x = aa \ y = b\ \ \suffix. find_path_rec_f P aa (case b of cld_def dcl cl fds mds \ cl) (path @ [(aa, b)]) = Some (path @ (aa, b) # suffix) \ (\dcl. (case b of cld_def dcl cl fds mds \ cl) = cl_fqn (fqn_def dcl) \ (\x\set suffix. case x of (ctx', cld') \ (\prefix' suffix'. case_option None (\(ctx', cld). find_path_rec_f P ctx' (superclass_name_f cld) (prefix' @ [(ctx', cld)])) (find_cld_f P ctx' (fqn_def (class_name_f cld'))) = Some (prefix' @ suffix')) \ mtype_f P (ty_def ctx' (class_name_f cld')) m = Some mty \ mtype_f P (ty_def aa dcl) m = Some mty))") apply(clarsimp) apply(drule_tac x = "(aa, ba)" in bspec, simp) apply(clarsimp) apply(subgoal_tac "(\prefix' suffix'. find_path_rec_f P ab (case bb of cld_def dcl cl fds mds \ cl) (prefix' @ [(ab, bb)]) = Some (prefix' @ suffix'))") apply(erule impE) apply(simp add: superclass_name_f_def) apply(thin_tac "\prefix suffix'. P prefix suffix'" for P) apply(thin_tac "mtype_f P (ty_def aa (class_name_f ba)) m = Some mty") apply(frule find_cld_name_eq) apply(clarsimp split: option.splits if_split_asm) apply(clarsimp simp add: mtype_f_def split: option.splits meth_def.splits meth_sig.splits) apply(rule) apply(clarsimp simp add: find_meth_def_f_def) apply(subgoal_tac "\path'. find_path_f P a (cl_fqn (fqn_def dcl')) = Some path'") apply(case_tac b) apply(clarsimp simp add: find_path_f_def superclass_name_f_def split: if_split_asm option.splits) apply(rename_tac meth_body cl list_char1 list_vd ty list_ty list_char2 cla list_fd list_md list_p1 path') apply(frule_tac path = path' in path_append) apply(rename_tac ag ah) apply(frule_tac path = ag in path_append) apply(clarify) apply(frule_tac suffix = path''a and suffix' = path''b and prefix' = "[(ac, cld_def list_char2 cla list_fd list_md)]" in fpr_same_suffix[rule_format], force) apply(simp) apply(clarify) apply(clarsimp simp add: class_methods_f_def split: option.splits) apply(clarsimp simp add: find_path_f_def superclass_name_f_def split: option.splits if_split_asm) apply(frule_tac path' = "(path @ [(a, cld_def dcl' (cl_fqn (fqn_def dcl'')) fds' mds'), (ac, b)])" in path_append) apply(clarsimp) apply(frule_tac suffix = path''a and prefix' = "[(a, cld_def dcl' (cl_fqn (fqn_def dcl'')) fds' mds'), (ac, b)]" in fpr_same_suffix') back apply(simp) apply(simp) apply(clarsimp) apply(rule) apply(clarsimp) apply(frule_tac ty_x_d = "(ty_def a dcl')" in wf_method_from_find[simplified]) apply(simp) apply(clarsimp) apply(erule wf_methE) apply(rename_tac ag ah ai aj ak al am an ao ap aq ar as at au av) apply(case_tac ag) apply(clarsimp) apply(clarsimp) apply(rule) apply(clarsimp) apply(frule_tac ty_x_d = "(ty_def a dcl')" in wf_method_from_find[simplified]) apply(simp) apply(clarsimp) apply(erule wf_methE) apply(clarsimp) apply(cut_tac lift_opts_for_all) apply(assumption) apply(rule) apply(simp) apply(assumption) apply(clarsimp) apply(thin_tac "find_path_rec_f P ab (case bb of cld_def dcl cl fds mds \ cl) (prefix' @ [(ab, bb)]) = Some (prefix' @ suffix')") apply(thin_tac "find_cld_f P aa (fqn_def (class_name_f ba)) = Some (ab, bb)") apply(thin_tac "(aa, ba) \ set path''") apply(case_tac b) apply(frule_tac dcl = dcl'' in find_cld_name_eq) apply(clarsimp simp add: superclass_name_f_def) apply(rename_tac ctx ctx'' mb cl_r m' vds ty_r tys ctx' cl_r' m'' vds' ty_r' tys' mb' dcl'' cl'' fds'' mds'') apply(clarsimp simp add: find_meth_def_f_def find_path_f_def superclass_name_f_def split: option.splits if_split_asm) apply(thin_tac "find_path_rec_f P ctx cl'' (path @ [(ctx, cld_def dcl' (cl_fqn (fqn_def dcl'')) fds' mds'), (ctx, cld_def dcl'' cl'' fds'' mds'')]) = Some (path @ (ctx, cld_def dcl' (cl_fqn (fqn_def dcl'')) fds' mds') # path'')") apply(clarsimp) apply(rename_tac path'' path') apply(frule_tac path = path' in path_append) apply(frule_tac path = path'' in path_append) apply(clarify) apply(frule_tac suffix = path''a and suffix' = path''b and prefix' = "[(ctx, cld_def dcl'' cl'' fds'' mds'')]" in fpr_same_suffix[rule_format], force) apply(clarsimp simp del: find_meth_def_in_path_f.simps) apply(rename_tac path'') apply(frule_tac ctx' = ctx' and m'' = m'' and cl'' = cl'' and fds'' = fds'' and ctx'' = ctx'' and path = "(ctx, cld_def dcl'' cl'' fds'' mds'') # path''" in mty_preservation'''', simp+) apply(rule_tac x = prefix' in exI) apply(clarsimp simp add: superclass_name_f_def) done lemma mty_preservation'': "\wf_program P; find_path_f P ctx (cl_fqn (fqn_def dcl)) = Some path; (ctx', cld') \ set path\ \ \prefix' suffix'. find_path_rec_f P ctx' (cl_fqn (fqn_def (class_name_f cld'))) prefix' = Some (prefix' @ suffix') \ mtype_f P (ty_def ctx' (class_name_f cld')) m = Some mty \ mtype_f P (ty_def ctx dcl) m = Some mty" apply(cut_tac x = "(ctx', cld')" in mty_preservation'''[of _ _ _ "[]", simplified]) apply(unfold find_path_f_def) apply(assumption) apply(simp+) done lemma mty_preservation': "\wf_program P; find_path_f P ctx (cl_fqn (fqn_def dcl)) = Some path; (ctx', cld') \ set path; find_path_f P ctx' (cl_fqn (fqn_def (class_name_f cld'))) = Some path'; mtype_f P (ty_def ctx' (class_name_f cld')) m = Some mty\ \ mtype_f P (ty_def ctx dcl) m = Some mty" apply(cut_tac path = path and suffix' = path' in mty_preservation''[rule_format, of _ _ _ _ _ _ "[]"]) apply(assumption+) apply(rule) apply(simp add: find_path_f_def) apply(assumption+) done lemma lift_opts_for_all_true[rule_format]: "\y_ty_list. (\x\set cl_var_ty_list. (\(cl_k, var_k, ty_k). find_type_f P ctx cl_k = Some ty_k) x) \ lift_opts (map (case_vd (\clk vark. find_type_f P ctx clk)) (map (\(cl_k, var_k, ty_k). vd_def cl_k var_k) cl_var_ty_list)) = Some (map snd y_ty_list) \ map (\(cl, var, ty). ty) cl_var_ty_list = map snd y_ty_list" by (induct cl_var_ty_list, auto split: option.splits) lemma mty_preservation: "\wf_program P; find_meth_def_f P ty_x_d meth = Some (ctx, meth_def_def (meth_sig_def cl_r meth (map (\(cl_k, var_k, ty_k). vd_def cl_k var_k) cl_var_ty_list)) meth_body); find_type_f P ctx cl_r = Some ty_r_d; \x\set cl_var_ty_list. (\(cl_k, var_k, ty_k). find_type_f P ctx cl_k = Some ty_k) x; mtype_f P ty_x_s meth = Some (mty_def (map snd y_ty_list) ty_r_s); is_sty_one P ty_x_d ty_x_s = Some True\ \ ty_r_d = ty_r_s \ map (\(cl, var, ty). ty) cl_var_ty_list = map snd y_ty_list" apply(case_tac "ty_x_d = ty_x_s") apply(clarsimp) apply(clarsimp simp add: mtype_f_def find_type_lift_opts[simplified] split: option.splits) apply(simp add: is_sty_one_def split: option.splits) apply(case_tac ty_x_s) apply(clarsimp simp add: mtype_f_def find_meth_def_f_def) apply(clarsimp) apply(rename_tac path_d ctx_s cld_s) apply(case_tac ty_x_d) apply(clarsimp) apply(clarsimp) apply(rename_tac ctx_d dcl_d) apply(clarsimp simp add: find_meth_def_f_def split: option.splits) apply(subgoal_tac "\path_s. find_path_f P ctx_s (cl_fqn (fqn_def (class_name_f cld_s))) = Some path_s") apply(clarify) apply(frule_tac path = path_d and path' = path_s in mty_preservation') apply(simp+) apply(clarsimp simp add: mtype_f_def split: option.splits meth_def.splits meth_sig.splits) apply(clarsimp simp add: find_meth_def_f_def) apply(rule_tac P = P and ctx = ctx in lift_opts_for_all_true) apply(simp) apply(clarsimp simp add: mtype_f_def find_meth_def_f_def split: option.splits) done lemma map_not_mem: "\(y_k, ty_k, var_k, var'_k, v_k) \ set y_ty_var_var'_v_list; x' \ (\(y, ty, var, var', v). x_var var') ` set y_ty_var_var'_v_list\ \ x_var var'_k \ x'" by (force) lemma map_fst: "map fst (map (\(y, ty, var, var', v). (x_var var', ty)) y_ty_var_var'_v_list) = map (\(y, ty, var, var', v). x_var var') y_ty_var_var'_v_list" by (force) lemma supertype_top: "sty_one P ty ty' \ ty = ty_top \ ty' = ty_top" by (induct rule: sty_one.induct, auto) lemma top_not_subtype[rule_format]: "sty_one P ty_top (ty_def ctx dcl) \ False" by (rule, drule supertype_top, simp) lemma f_in_found_fds: "ftype_in_fds_f P ctx fds f = ty_opt_bot_opt (Some ty_f) \ f \ case_fd (\cl f. f) ` set fds" by (induct fds, auto split: fd.splits) lemma ftype_fields[rule_format]: "ftype_in_path_f P path_s f = Some ty_f \ f \ set (fields_in_path_f path_s)" apply(induct path_s) apply(simp) apply(clarsimp split: option.splits) apply(case_tac b) apply(clarsimp) apply(rename_tac ctx path_s dcl cl fds mds) apply(clarsimp simp add: class_fields_f_def split: ty_opt_bot.splits option.splits) apply(simp add: f_in_found_fds) done lemma no_field_hiding''': "\suffix. find_path_rec_f P ctx cl prefix = Some (prefix @ suffix) \ wf_program P \ (\(ctx, cld) \ set suffix. (\prefix' suffix'. find_path_rec_f P ctx (cl_fqn (fqn_def (class_name_f cld))) prefix' = Some (prefix' @ suffix') \ f \ set (fields_in_path_f suffix') \ f \ set (fields_in_path_f suffix)))" supply [[simproc del: defined_all]] apply(induct_tac P ctx cl prefix rule: find_path_rec_f.induct) apply(simp) apply(clarsimp split: option.splits) apply(subgoal_tac "\aa ba aaa bb. (a = aaa \ b = bb \ aa = aaa \ ba = bb) \ (\suffix. find_path_rec_f P aaa (superclass_name_f bb) (path @ [(aaa, bb)]) = Some (path @ (aaa, bb) # suffix) \ (\x\set suffix. (\(ctx, cld). (\prefix' suffix'. case_option None (case_prod (\ctx' cld. find_path_rec_f P ctx' (superclass_name_f cld) (prefix' @ [(ctx', cld)]))) (find_cld_f P ctx (fqn_def (class_name_f cld))) = Some (prefix' @ suffix') \ f \ set (fields_in_path_f suffix')) \ f \ set (fields_in_path_f suffix)) x))") defer apply(thin_tac _)+ apply(force) apply(thin_tac "\aa ba x y. \a = aa \ b = ba; x = aa \ y = ba\ \ \suffix. find_path_rec_f P aa (superclass_name_f ba) (path @ [(aa, ba)]) = Some (path @ (aa, ba) # suffix) \ (\x\set suffix. case x of (ctx, cld) \ (\prefix' suffix'. case_option None (\(ctx', cld). find_path_rec_f P ctx' (superclass_name_f cld) (prefix' @ [(ctx', cld)])) (find_cld_f P ctx (fqn_def (class_name_f cld))) = Some (prefix' @ suffix') \ f \ set (fields_in_path_f suffix')) \ f \ set (fields_in_path_f suffix))") apply(clarsimp) apply(frule path_append) apply(clarify) apply(erule_tac x = "tl suffix" in allE) apply(clarsimp) apply(erule disjE) apply(clarify) apply(case_tac fqn) apply(clarsimp simp del: fmdip_cons simp add: class_name_f_def split: cld.splits) apply(frule find_cld_name_eq) apply(frule_tac ctx = a in find_cld_name_eq) apply(clarsimp simp del: fmdip_cons simp add: class_name_f_def split: cld.splits) apply (rename_tac list1 cl list2 list3) apply(frule_tac path = "prefix' @ suffix'" in path_append) apply(clarsimp simp del: fmdip_cons) apply(frule_tac prefix = "path @ [(ab, cld_def list1 cl list2 list3)]" and suffix = path'' and prefix' = "prefix' @ [(ab, cld_def list1 cl list2 list3)]" and suffix' = path''a in fpr_same_suffix[rule_format]) apply(clarsimp) apply(simp) apply(drule_tac x = "(aa, ba)" in bspec, simp) apply(force) done lemma no_field_hiding'': "\find_path_f P ctx' cl = Some suffix; wf_program P; (ctx, cld) \ set suffix\ \ \prefix' suffix'. find_path_rec_f P ctx (cl_fqn (fqn_def (class_name_f cld))) prefix' = Some (prefix' @ suffix') \ f \ set (fields_in_path_f suffix') \ f \ set (fields_in_path_f suffix)" apply(cut_tac x = "(ctx, cld)" in no_field_hiding'''[of _ _ _ "[]", rule_format]) apply(simp add: find_path_f_def)+ done lemma no_field_hiding': "\find_path_f P ctx' cl = Some path; wf_program P; (ctx, cld) \ set path; find_path_f P ctx (cl_fqn (fqn_def (class_name_f cld))) = Some path'; f \ set (fields_in_path_f path')\ \ f \ set (fields_in_path_f path)" apply(cut_tac no_field_hiding''[rule_format, of _ _ _ _ _ _ "[]"]) apply(assumption+) apply(unfold find_path_f_def) apply(rule) apply(simp+) done lemma no_field_hiding: "\wf_program P; is_sty_one P ty_x_d ty_x_s = Some True; ftype_f P ty_x_s f = Some ty_f; fields_f P ty_x_d = Some fields_oid\ \ f \ set fields_oid" apply(case_tac ty_x_s) apply(simp add: ftype_f_def) apply(simp add: is_sty_one_def split: option.splits) apply(case_tac ty_x_d) apply(simp) apply(clarsimp) apply(rename_tac ctx_s dcl_s path_d ctx_d dcl_d) apply(simp add: ftype_f_def split: option.splits) apply(clarsimp simp add: fields_f_def) apply(drule ftype_fields) apply(frule no_field_hiding') apply(simp add: ftype_fields class_name_f_def)+ done lemma ftype_pre_path[rule_format]: "ftype_in_path_f P path_s f = Some ty_f \ ftype_in_path_f P (path_s @ path'') f = Some ty_f" by (induct path_s, auto split: option.splits ty_opt_bot.splits) lemma fields_non_intersection[rule_format]: " (\(cl, f, ty). f) ` set cl_f_ty_list \ set (fields_in_path_f path'') = {} \ f \ set (fields_in_path_f path'') \ ftype_in_fds_f P ctx_def (map (\(cl, f, ty). fd_def cl f) cl_f_ty_list) f = ty_opt_bot_opt None" by (induct cl_f_ty_list, auto) lemma f_notin_list[rule_format]: "f \ set (map (\(cl, f, ty). f) cl_f_ty_list) \ ftype_in_fds_f P ctx_def (map (\(cl, f, ty). fd_def cl f) cl_f_ty_list) f = ty_opt_bot_opt None" by (induct cl_f_ty_list, auto) declare find_path_rec_f.simps[simp del] lemma ftype_preservation''''': "\find_path_rec_f P ctx cl (prefix @ [cld']) = Some (prefix @ cld' # path''); find_type_f P ctx_def cl = Some ty; fields_f P ty = Some fs; ftype_in_path_f P path'' f = Some ty_f; (set (map (\(cl, f, ty). f) cl_f_ty_list)) \ (set fs) = {}\ \ ftype_in_fds_f P ctx (map (\(cl, f, ty). fd_def cl f) cl_f_ty_list) f = ty_opt_bot_opt None" apply(clarsimp simp add: fields_f_def split: option.splits) apply(rename_tac path) apply(case_tac cl) apply(simp add: find_path_rec_f.simps) apply(clarsimp split: fqn.splits option.splits) apply(rename_tac dcl ctx'' cld'') apply(case_tac ctx'') apply(case_tac ctx) apply(clarsimp simp add: find_path_f_def) apply(frule_tac prefix = "prefix @ [cld']" and suffix = path'' and prefix' = "[]" in fpr_same_suffix[rule_format]) apply(simp) apply(clarsimp) apply(drule ftype_fields) apply(cut_tac fields_non_intersection) apply(force) apply(force) done declare find_path_rec_f.simps[simp] lemma ftype_preservation'''': "\suffix. find_path_rec_f P ctx cl prefix = Some (prefix @ suffix) \ wf_program P \ (\(ctx, cld) \ set suffix. (\prefix' suffix'. find_path_rec_f P ctx (cl_fqn (fqn_def (class_name_f cld))) prefix' = Some (prefix' @ suffix') \ ftype_in_path_f P suffix' f = Some ty \ ftype_in_path_f P suffix f = Some ty))" supply [[simproc del: defined_all]] apply(induct_tac P ctx cl prefix rule: find_path_rec_f.induct) apply(simp) apply(clarsimp split: option.splits) apply(subgoal_tac "\aa ba aaa bb. (a = aaa \ b = bb \ aa = aaa \ ba = bb) \ (\suffix. find_path_rec_f P aaa (superclass_name_f bb) (path @ [(aaa, bb)]) = Some (path @ (aaa, bb) # suffix) \ (\x\set suffix. (\(ctx, cld). (\prefix' suffix'. case_option None (case_prod (\ctx' cld. find_path_rec_f P ctx' (superclass_name_f cld) (prefix' @ [(ctx', cld)]))) (find_cld_f P ctx (fqn_def (class_name_f cld))) = Some (prefix' @ suffix') \ ftype_in_path_f P suffix' f = Some ty) \ ftype_in_path_f P suffix f = Some ty) x))") defer apply(thin_tac _)+ apply(force) apply(clarsimp) apply(frule path_append) apply(clarify) apply(erule_tac x = "tl suffix" in allE) apply(clarsimp) apply(erule disjE) apply(clarify) apply(case_tac fqn) apply(clarsimp simp add: class_name_f_def split: cld.splits) apply(frule find_cld_name_eq) apply(frule_tac ctx = a in find_cld_name_eq) apply(clarsimp simp add: class_name_f_def split: cld.splits) apply(rename_tac list1 cl list2 list3) apply(frule_tac path = "prefix' @ suffix'" in path_append) apply(clarsimp) apply(frule_tac prefix = "path @ [(ab, cld_def list1 cl list2 list3)]" and suffix = path'' and prefix' = "prefix' @ [(ab, cld_def list1 cl list2 list3)]" and suffix' = path''a in fpr_same_suffix[rule_format]) apply(clarsimp) apply(clarsimp) apply(drule_tac x = "(aa, ba)" in bspec, simp) apply(clarsimp) apply(subgoal_tac "(\prefix' suffix'. find_path_rec_f P ab (superclass_name_f bb) (prefix' @ [(ab, bb)]) = Some (prefix' @ suffix') \ ftype_in_path_f P suffix' f = Some ty)") apply(clarsimp) apply(subgoal_tac "ftype_in_fds_f P a (class_fields_f b) f = ty_opt_bot_opt None") apply(simp) apply(frule_tac cld' = b in wf_cld_from_lu) apply(simp) apply(erule wf_classE) apply(erule wf_class_cE) apply(clarsimp simp add: superclass_name_f_def class_fields_f_def) apply(rule ftype_preservation''''') apply(simp+) apply(force) done lemma ftype_preservation''': "\find_path_f P ctx' cl = Some suffix; wf_program P; (ctx, cld) \ set suffix\ \ \prefix' suffix'. find_path_rec_f P ctx (cl_fqn (fqn_def (class_name_f cld))) prefix' = Some (prefix' @ suffix') \ ftype_in_path_f P suffix' f = Some ty \ ftype_in_path_f P suffix f = Some ty" apply(cut_tac x = "(ctx, cld)" in ftype_preservation''''[of _ _ _ "[]", rule_format]) apply(simp add: find_path_f_def)+ done lemma ftype_preservation'': "\find_path_f P ctx' cl = Some path; wf_program P; (ctx, cld) \ set path; find_path_f P ctx (cl_fqn (fqn_def (class_name_f cld))) = Some path'; ftype_in_path_f P path' f = Some ty\ \ ftype_in_path_f P path f = Some ty" apply(cut_tac ftype_preservation'''[rule_format, of _ _ _ _ _ _ "[]"]) apply(assumption+) apply(unfold find_path_f_def) apply(rule) apply(simp+) done lemma ftype_preservation'[simplified]: "\wf_program P; sty_one P ty_x_d ty_x_s; ftype P ty_x_s f ty_f\ \ ftype P ty_x_d f ty_f" apply(case_tac ty_x_s) apply(simp add: ftype_f_def) apply(simp add: is_sty_one_def split: option.splits) apply(case_tac ty_x_d) apply(simp) apply(clarsimp) apply(rename_tac ctx_s dcl_s path_d ctx_d dcl_d) apply(simp add: ftype_f_def split: option.splits) apply(rename_tac path_s cl_s fds_s mds_s) apply(rule ftype_preservation'') apply(simp add: class_name_f_def)+ done lemma ftype_preservation[simplified]: "\wf_program P; \x\dom \. wf_object P H (L x) (\ x); \ x = Some ty_x_s; ftype P ty_x_s f ty_f_s; L x = Some (v_oid oid_x); H oid_x = Some (ty_x_d, fs_oid); ftype P ty_x_d f ty_f_d\ \ ty_f_s = ty_f_d" apply(drule_tac x = x in bspec, simp add: domI) apply(erule wf_objectE) apply(simp) apply(clarsimp) apply(erule sty_option.cases) apply(clarsimp) apply(simp add: ftype_preservation') done lemma map_f: "\cl_f_ty_list. map (\(cl, f). fd_def cl f) cl_f_list = map (\(cl, f, ty). fd_def cl f) cl_f_ty_list \ map (\(cl, f). f) cl_f_list = map (\(cl, f, ty). f) cl_f_ty_list" by (induct cl_f_list, auto) lemma fields_to_owner[rule_format]: "f \ set (fields_in_path_f path) \ (\cld \ set (map snd path). \cl_f. fd_def cl_f f \ set (class_fields_f cld))" apply(induct path) apply(simp) apply(clarsimp) apply(rename_tac cld path fd) apply(clarsimp simp add: class_fields_f_def split: option.splits cld.splits fd.splits) apply(force) done lemma ftype_in_fds_none[rule_format]: "f \ (\(cl, f, ty). f) ` set cl_f_ty_list \ ftype_in_fds_f P ctx (map (\(cl, f, ty). fd_def cl f) cl_f_ty_list) f = ty_opt_bot_opt None" by (induct cl_f_ty_list, auto) lemma ftype_in_fds_some[rule_format]: "(cl_f, f, ty_f) \ set cl_f_ty_list \ distinct (map (\(cl, f, ty). f) cl_f_ty_list) \ find_type_f P ctx cl_f = Some ty_f \ ftype_in_fds_f P ctx (map (\(cl, f, ty). fd_def cl f) cl_f_ty_list) f = ty_opt_bot_opt (Some ty_f)" apply(induct cl_f_ty_list) apply(simp) apply(clarsimp) apply(safe) apply(force) apply(simp) apply(frule ftype_in_fds_none) apply(force) done lemma no_ty_bot[THEN mp]: "(\x\set cl_f_ty_list. (\(cl, f, ty). find_type_f P ctx_def cl = Some ty) x) \ ftype_in_fds_f P ctx' (map (\(cl, f, ty). fd_def cl f) cl_f_ty_list) f \ ty_opt_bot_bot" apply(induct cl_f_ty_list) apply(simp) apply(clarsimp split: if_split_asm option.splits) apply(case_tac ctx', simp) (* should get rid of ctx dependency here *) done lemma field_has_type'[rule_format]: "(cl_f, f, ty_f) \ set cl_f_ty_list \ (\(ctx, cld) \ set path. wf_class P cld) \ (ctx, cld_def dcl cl (map (\(cl, f, ty). fd_def cl f) cl_f_ty_list) mds) \ set path \ distinct (map (\(cl, f, ty). f) cl_f_ty_list) \ find_type_f P ctx cl_f = Some ty_f \ (\ty'. ftype_in_path_f P path f = Some ty')" apply(induct path) apply(simp) apply(simp) apply(rule) apply(elim conjE) apply(erule disjE) apply(clarsimp) apply(subgoal_tac "ftype_in_fds_f P ctx (class_fields_f (cld_def dcl cl (map (\(cl, f, ty). fd_def cl f) cl_f_ty_list) mds)) f = ty_opt_bot_opt (Some ty_f)") apply(simp) apply(simp add: class_fields_f_def) apply(rule ftype_in_fds_some) apply(force) apply(case_tac a) apply(rename_tac ctx' cld') apply(simp) apply(case_tac "ftype_in_fds_f P ctx' (class_fields_f cld') f") apply(rename_tac ty_opt) apply(case_tac ty_opt) apply(clarsimp) apply(clarsimp) apply(clarsimp) apply(erule wf_classE) apply(erule wf_class_cE) apply(clarsimp simp add: class_fields_f_def) apply(frule no_ty_bot) apply(force) done lemma mem_class_wf: "\wf_program P; cld \ set P\ \ wf_class P cld" apply(erule wf_programE) apply(drule_tac x = cld in bspec, simp) apply(clarsimp) done lemma field_has_type: "\wf_program P; fields P ty_oid (Some f_list); f \ set f_list\ \ \ty. ftype P ty_oid f ty" apply(simp) apply(clarsimp simp add: fields_f_def split: option.splits) apply(rename_tac path) apply(case_tac ty_oid) apply(simp) apply(clarsimp) apply(rename_tac dcl) apply(frule fields_to_owner) apply(clarsimp) apply(rename_tac ctx' cld cl_f) apply(frule clds_in_path_exist) apply(drule_tac x = "(ctx', cld)" in bspec, simp) apply(clarsimp) apply(frule mem_class_wf, simp) apply(erule wf_classE) apply(erule wf_class_cE) apply(clarsimp simp add: class_fields_f_def) apply(rename_tac cl_f ty_f) apply(drule_tac x = "(cl_f, f, ty_f)" in bspec, simp) apply(clarsimp) apply(simp add: ftype_f_def) apply(rule_tac cl_f = cl_f and ty_f = ty_f and cl_f_ty_list = cl_f_ty_list and ctx = ctx' and dcl = dclb and cl = cla and mds = "(map fst meth_def_meth_list)" in field_has_type') apply(simp) apply(case_tac ctx') apply(simp) apply(clarify) apply(drule_tac cl = "cl_fqn (fqn_def dcl)" and ctxcld = "(a, b)" in all_in_path_found) apply(simp) apply(clarsimp) apply(rule wf_cld_from_lu) apply(simp) apply(force) done lemma exists_three_in_five: "var_k \ set (map (\(y, cl, var, var', v). var) y_cl_var_var'_v_list) \ (\y cl var' v. (y, cl, var_k, var', v) \ set y_cl_var_var'_v_list)" by (induct y_cl_var_var'_v_list, auto) lemma map_xx[rule_format]: "(y, cl, var_k, var', v) \ set y_cl_var_var'_v_list \ (x_var var_k, x_var var') \ set (map (\(y, cl, var, var', v). (x_var var, x_var var')) y_cl_var_var'_v_list)" by (induct y_cl_var_var'_v_list, auto) lemma map_of_vd[rule_format]: "\cl_var_ty_list. map (\(y, cl, var, var', v). vd_def cl var) y_cl_var_var'_v_list = map (\(cl, var, ty). vd_def cl var) cl_var_ty_list \ map (\(y, cl, var, var', v). var) y_cl_var_var'_v_list = map (\(cl, var, ty). var) cl_var_ty_list" by (induct y_cl_var_var'_v_list, auto) lemma mem_vd: "vd_def cl_k var_k \ set (map (\(y, cl, var, var', v). vd_def cl var) y_cl_var_var'_v_list) \ var_k \ set (map (\(y, cl, var, var', v). var) y_cl_var_var'_v_list)" by (induct y_cl_var_var'_v_list, auto) lemma exists_var': "\distinct (map (\(cl, var, ty). var) cl_var_ty_list); map (\(y, cl, var, var', v). var) y_cl_var_var'_v_list = map (\(cl, var, ty). var) cl_var_ty_list; map_of (map (\(cl, var, ty). (x_var var, ty)) cl_var_ty_list) var_k = Some ty_k\ \ \var'_k. map_of (map (\(y, cl, var, var', v). (x_var var, x_var var')) y_cl_var_var'_v_list) var_k = Some (x_var var'_k)" apply(drule map_of_SomeD) apply(clarsimp, hypsubst_thin) apply(rename_tac cl_k var_k) apply(subgoal_tac "var_k \ set (map (\(y, cl, var, var', v). var) y_cl_var_var'_v_list)") apply(simp (no_asm_use)) apply(clarify) apply(rename_tac y_k cl_k' var_k var'_k v_k) apply(rule_tac x = var'_k in exI) apply(rule map_of_is_SomeI) apply(drule distinct_var, assumption+) apply(rule map_xx, assumption) by force lemma map_y: "\((y_k, cl_k, var_k, var'_k, v_k), y'_k, ty_k) \ set (zip y_cl_var_var'_v_list y_ty_list); map (\(y, cl, var, var', v). y) y_cl_var_var'_v_list = map fst y_ty_list\ \ (y_k, ty_k) \ set y_ty_list" apply(clarsimp simp add: set_zip) apply(drule_tac f = "\(y, cl, var, var', v). y" in arg_cong) apply(drule_tac f1 = "\(y, cl, var, var', v). y" in nth_map[THEN sym]) apply(frule_tac f = "\(y, ty). y" in arg_cong) apply(case_tac "y_ty_list ! i") apply(force simp add: in_set_conv_nth) done lemma set_zip_var'_v: "\((y_k, cl_k, var_k, var'_k, v_k), y'_k, ty_k) \ set (zip y_cl_var_var'_v_list y_ty_list)\ \ (x_var var'_k, v_k) \ set (map (\(y, cl, var, var', y). (x_var var', y)) y_cl_var_var'_v_list)" apply(subgoal_tac "(y_k, cl_k, var_k, var'_k, v_k) \ set y_cl_var_var'_v_list") by (force simp add: set_zip)+ lemma map_of_map_zip_none: "\map (\(y, cl, var, var', v). y) y_cl_var_var'_v_list = map fst y_ty_list; 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_G = None\ \ map_of (map (\(y, cl, var, var', y). (x_var var', y)) y_cl_var_var'_v_list) x_G = None" apply(drule map_length_y) apply(clarsimp simp add: map_of_eq_None_iff) apply(erule contrapos_np) apply(simp only: set_map[THEN sym] fst_zip_eq) by force lemma map_of_map_zip_some: "\distinct (map (\(y, cl, var, var', v). var') y_cl_var_var'_v_list); map (\(y, cl, var, var', v). y) y_cl_var_var'_v_list = map fst y_ty_list; 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_G = Some ty_k\ \ \v_k. map_of (map (\(y, cl, var, var', y). (x_var var', y)) y_cl_var_var'_v_list) x_G = Some v_k" apply(drule map_of_SomeD) apply(clarsimp) apply(rename_tac y cl var var' v y') apply(subgoal_tac "(y, cl, var, var', v) \ set y_cl_var_var'_v_list") apply(rule_tac x = v in exI) apply(rule map_of_is_SomeI) apply(simp only: map_fst_two' distinct_x_var') apply(force) by (force simp add: set_zip) lemma x'_not_in_list: "\(y_k, cl_k, var_k, var'_k, v_k) \ set y_cl_var_var'_v_list; x' \ (\(y, cl, var, var', v). x_var var') ` set y_cl_var_var'_v_list\ \ x_var var'_k \ x'" by (erule contrapos_nn, force) lemma nth_in_set: "\i < length y_ty_list; map fst y_ty_list ! i = y_k\ \ y_k \ set (map fst y_ty_list)" apply(subgoal_tac "\i set (map fst y_ty_list) \ (\ty. (y_k, ty) \ set y_ty_list)" by (induct y_ty_list, auto) lemma exists_i: (* can probably simplify a lot *) "\distinct (map (\(cl, var, ty). var) cl_var_ty_list); (y_k, cl_ka, var_k, var'_k, v_k) \ set y_cl_var_var'_v_list; (cl_k, var_k, ty_k) \ set cl_var_ty_list; map (\(y, cl, var, var', v). var) y_cl_var_var'_v_list = map (\(cl, var, ty). var) cl_var_ty_list\ \ \i < length cl_var_ty_list. cl_var_ty_list ! i = (cl_k, var_k, ty_k) \ y_cl_var_var'_v_list ! i = (y_k, cl_ka, var_k, var'_k, v_k)" apply(simp only: in_set_conv_nth) apply(erule_tac P = "\i. i < length cl_var_ty_list \ cl_var_ty_list ! i = (cl_k, var_k, ty_k)" in exE) apply(erule conjE) apply(rule_tac x = x in exI) apply(simp) apply(drule_tac f = "\(cl, var, ty). var" in arg_cong) apply(frule_tac f1 = "\(cl, var, ty). var" in nth_map[THEN sym]) apply(drule_tac t = "map (\(cl, var, ty). var) cl_var_ty_list ! x" in sym) apply(drule_tac s = "(\(cl, var, ty). var) (cl_var_ty_list ! x)" in trans) apply(simp) apply(thin_tac "(\(cl, var, ty). var) (cl_var_ty_list ! x) = (\(cl, var, ty). var) (cl_k, var_k, ty_k)") apply(drule sym) apply(simp) apply(erule exE) apply(rename_tac j) apply(erule conjE) apply(simp only: distinct_conv_nth) apply(erule_tac x = x in allE) apply(drule_tac s = "map (\(cl, var, ty). var) cl_var_ty_list" in sym) apply(drule map_length_var) apply(simp) apply(erule_tac x = j in allE) apply(simp) done lemma y_ty_match: "\i < length cl_var_ty_list; cl_var_ty_list ! i = (cl_k, var_k, ty_k); y_cl_var_var'_v_list ! i = (y_k, cl_ka, var_k, var'_k, v_k); map (\(y, cl, var, var', v). y) y_cl_var_var'_v_list = map fst y_ty_list; map (\(cl, var, ty). ty) cl_var_ty_list = map snd y_ty_list\ \ y_ty_list ! i = (y_k, ty_k)" apply(drule_tac f = "\(cl, var, ty). ty" in arg_cong) apply(frule_tac f1 = "\(cl, var, ty). ty" in nth_map[THEN sym]) apply(drule_tac t = "map (\(cl, var, ty). ty) cl_var_ty_list ! i" in sym) apply(frule map_length_y) apply(frule map_length_ty) apply(drule_tac t = "length y_ty_list" in sym) apply(simp) apply(drule_tac f = "\(y, cl, var, var', v). y" in arg_cong) apply(frule_tac f1 = "\(y, cl, var, var', v). y" in nth_map[THEN sym]) apply(drule_tac t = "map (\(y, cl, var, var', v). y) y_cl_var_var'_v_list ! i" in sym) apply(case_tac "y_ty_list ! i") apply(simp) done lemma map_of_ty_k: "\distinct (map (\(cl, var, ty). var) cl_var_ty_list); distinct (map (\(y, cl, var, var', v). var') y_cl_var_var'_v_list); (y_k, cl_ka, var_k, var'_k, v_k) \ set y_cl_var_var'_v_list; (cl_k, var_k, ty_k) \ set cl_var_ty_list; map (\(y, cl, var, var', v). y) y_cl_var_var'_v_list = map fst y_ty_list; map (\(cl, var, ty). ty) cl_var_ty_list = map snd y_ty_list; map (\(y, cl, var, var', v). var) y_cl_var_var'_v_list = map (\(cl, var, ty). var) cl_var_ty_list\ \ 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_var var'_k) = Some ty_k" apply(frule map_length_y) apply(rule map_of_is_SomeI) apply(simp only: fst_zip_eq) apply(simp only: distinct_x_var') apply(clarsimp) apply(frule map_length_var) apply(frule exists_i, simp+) apply(erule exE) apply(drule_tac t = "length y_ty_list" in sym) apply(clarsimp) apply(frule y_ty_match, assumption+) apply(subgoal_tac "zip y_cl_var_var'_v_list y_ty_list ! i = (y_cl_var_var'_v_list ! i, y_ty_list ! i)") apply(drule_tac f = "\((y, cl, var, var', v), y', y). (x_var var', y)" in arg_cong) apply(subgoal_tac "i < length (zip y_cl_var_var'_v_list y_ty_list)") apply(frule_tac f1 = "\((y, cl, var, var', v), y', y). (x_var var', y)" in nth_map[THEN sym]) apply(drule_tac t = "map (\((y, cl, var, var', v), y', y). (x_var var', y)) (zip y_cl_var_var'_v_list y_ty_list) ! i" in sym) apply(drule_tac s = "(\((y, cl, var, var', v), y', y). (x_var var', y)) (zip y_cl_var_var'_v_list y_ty_list ! i)" in trans) apply(simp) apply(subgoal_tac "\j((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), y', y). (x_var var', y)) (zip y_cl_var_var'_v_list y_ty_list) ! j = (x_var var'_k, ty_k)") apply(simp only: in_set_conv_nth[THEN sym]) apply(simp) apply(rule_tac x = i in exI) apply(simp) apply(force) apply(rule nth_zip) apply(simp) apply(simp) done lemma map_of_ty_k': "\distinct (map (\(cl, var, ty). var) cl_var_ty_list); distinct (map (\(y, cl, var, var', v). var') y_cl_var_var'_v_list); x' \ (\(y, cl, var, var', v). x_var var') ` set y_cl_var_var'_v_list; map fst y_cl_var_var'_v_list = map fst y_ty_list; map (\(cl, var, ty). ty) cl_var_ty_list = map snd y_ty_list; map (\(y, cl, var, u). var) y_cl_var_var'_v_list = map (\(cl, var, ty). var) cl_var_ty_list; map_of (map (\(cl, var, y). (x_var var, y)) cl_var_ty_list) (x_var var) = Some ty_var\ \ (if x_var (case_option var (case_x (\var'. var') var) (map_of (map (\(y, cl, var, var', v). (x_var var, x_var var')) y_cl_var_var'_v_list) (x_var var))) = x' then Some ty_x_m else (\ ++ 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_var (case_option var (case_x (\var'. var') var) (map_of (map (\(y, cl, var, var', v). (x_var var, x_var var')) y_cl_var_var'_v_list) (x_var var))))) = Some ty_var" apply(frule exists_var') apply(simp add: weaken_list_var) apply(simp) apply(clarsimp split del: if_split) apply(clarsimp) apply(drule map_of_SomeD)+ apply(clarsimp split del: if_split) apply(frule x'_not_in_list, simp+) apply(frule map_of_ty_k) apply(simp add: weaken_list_y weaken_list_var)+ done lemma var_not_var'_x': "\\ (x_var var) = Some ty_var; L (x_var var) = Some v_var; \x\set y_cl_var_var'_v_list. L (x_var ((\(y, cl, var, var', v). var') x)) = None; L x' = None\ \ (\ ++ map_of (map (\((y, cl, var, var', v), y', y). (x_var var', y)) (zip y_cl_var_var'_v_list y_ty_list)) ++ Map.empty(x' \ ty_x_d)) (x_var var) = Some ty_var" apply(subgoal_tac "x_var var \ x'") apply(simp add: map_add_def) apply(subgoal_tac "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_var var) = None") apply(simp) apply(clarsimp simp add: map_of_eq_None_iff set_zip) apply(rename_tac y_k cl_k var_k v_k y'_k ty_k i) apply(drule_tac x = "(y_k, cl_k, var_k, var, v_k)" in bspec, simp) apply(force) by force lemma same_el: "\distinct (map (\(y, cl, var, var', v). var') y_cl_var_var'_v_list); i < length y_cl_var_var'_v_list; (y_i, cl_i, var_i, var'_i, v_i) = y_cl_var_var'_v_list ! i; (y_j, cl_j, var_j, var'_i, v_j) \ set y_cl_var_var'_v_list\ \ (y_i, cl_i, var_i, var'_i, v_i) = (y_j, cl_j, var_j, var'_i, v_j)" apply(simp only: in_set_conv_nth) apply(clarsimp) apply(rename_tac j) apply(simp only: distinct_conv_nth) apply(erule_tac x = i in allE) apply(simp) apply(erule_tac x = j in allE) apply(simp) apply(drule sym) apply(simp) done lemma same_y: "\map (\(y, cl, var, var', v). y) y_cl_var_var'_v_list = map fst y_ty_list; i < length y_cl_var_var'_v_list; (y'_k, ty_k) = y_ty_list ! i; y_cl_var_var'_v_list ! i = (y_k, cl_k, var_k, var'_k, v_k)\ \ y'_k = y_k" apply(drule_tac f = "\(y, cl, var, var', v). y" in arg_cong) apply(frule_tac f1 = "\(y, cl, var, var', v). y" in nth_map[THEN sym]) apply(drule_tac f = fst in arg_cong) apply(simp) apply(frule map_length_y) apply(simp) done lemma map_map_zip_y': "map fst y_y'_list = map fst y_ty_list \ map snd y_y'_list = map fst (map (\((y, y'), yy, ty). (y', ty)) (zip y_y'_list y_ty_list))" apply(rule nth_equalityI) apply(induct y_y'_list) apply(simp) apply(clarsimp) apply(frule map_length_y') apply(simp) apply(clarsimp) apply(case_tac "y_y'_list ! i") apply(case_tac "y_ty_list ! i") apply(clarsimp) apply(subgoal_tac "i < length (map (\((y, y'), yy, ty). (y', ty)) (zip y_y'_list y_ty_list))") apply(frule_tac f = fst and xs = "map (\((y, y'), yy, ty). (y', ty)) (zip y_y'_list y_ty_list)" in nth_map) apply(simp) apply(subgoal_tac "length (map (\((y, y'), yy, ty). (y', ty)) list) = length list" for list) apply(simp) apply(frule map_length_y') apply(simp) apply(simp only: map_length_list) done lemma map_map_zip_ty: "map fst y_y'_list = map fst y_ty_list \ map snd (map (\((y, y'), yy, ty). (y', ty)) (zip y_y'_list y_ty_list)) = map snd y_ty_list" apply(frule map_length_y') apply(rule nth_equalityI) apply(induct y_ty_list) apply(simp) apply(simp) apply(clarsimp) apply(case_tac "y_y'_list ! i") apply(case_tac "y_ty_list ! i") apply(simp) done lemma map_y_yy: "\i < length y_y'_list; i < length y_ty_list; map fst y_y'_list = map fst y_ty_list; y_y'_list ! i = (y, y'); y_ty_list ! i = (yy, ty)\ \ y = yy" apply(drule_tac f = fst in arg_cong) apply(frule_tac f1 = fst in nth_map[THEN sym]) by simp lemma var_k_of_ty_k: "\distinct (map (\(cl, var, ty). var) cl_var_ty_list); distinct (map (\(y, cl, var, var', v). var') y_cl_var_var'_v_list); x' \ (\(y, cl, var, var', v). x_var var') ` set y_cl_var_var'_v_list; map fst y_cl_var_var'_v_list = map fst y_ty_list; map (\(cl, var, ty). ty) cl_var_ty_list = map snd y_ty_list; map (\(y, cl, var, u). var) y_cl_var_var'_v_list = map (\(cl, var, ty). var) cl_var_ty_list; (if x = x_this then Some ty_x_m else map_of (map (\(cl, var, y). (x_var var, y)) cl_var_ty_list) x) = Some ty_x\ \ (if (case if x = x_this then Some x' else map_of (map (\(y, cl, var, var', v). (x_var var, x_var var')) y_cl_var_var'_v_list) x of None \ x | Some x' \ x') = x' then Some ty_x_m else (\ ++ map_of (map (\((y, cl, var, var', v), y', y). (x_var var', y)) (zip y_cl_var_var'_v_list y_ty_list))) (case if x = x_this then Some x' else map_of (map (\(y, cl, var, var', v). (x_var var, x_var var')) y_cl_var_var'_v_list) x of None \ x | Some x' \ x')) = Some ty_x" apply(clarsimp) apply(rule) apply(rule) apply(simp) apply(clarsimp) apply(frule exists_var') apply(simp add: weaken_list_var) apply(simp) apply(clarsimp split del: if_split) apply(drule map_of_SomeD)+ apply(clarsimp split del: if_split) apply(frule x'_not_in_list, simp+) apply(frule map_of_ty_k) apply(simp add: weaken_list_y weaken_list_var)+ done lemma x_var_not_this: "(if x_var var = x_this then Some x' else Q) = Q" by simp lemma subtype_through_tr: "\distinct (map (\(cl, var, ty). var) cl_var_ty_list); distinct (map (\(y, cl, var, var', v). var') y_cl_var_var'_v_list); x' \ (\(y, cl, var, var', v). x_var var') ` set y_cl_var_var'_v_list; map fst y_cl_var_var'_v_list = map fst y_ty_list; map (\(cl, var, ty). ty) cl_var_ty_list = map snd y_ty_list; map (\(y, cl, var, u). var) y_cl_var_var'_v_list = map (\(cl, var, ty). var) cl_var_ty_list; is_sty_one P ty_x ty_var = Some True; (if x = x_this then Some ty_x_m else map_of (map (\(cl, var, y). (x_var var, y)) cl_var_ty_list) x) = Some ty_x; map_of (map (\(cl, var, y). (x_var var, y)) cl_var_ty_list) (x_var var) = Some ty_var\ \ sty_option P (if (case if x = x_this then Some x' else map_of (map (\(y, cl, var, var', v). (x_var var, x_var var')) y_cl_var_var'_v_list) x of None \ x | Some x' \ x') = x' then Some ty_x_m else (\ ++ map_of (map (\((y, cl, var, var', v), y', y). (x_var var', y)) (zip y_cl_var_var'_v_list y_ty_list))) (case if x = x_this then Some x' else map_of (map (\(y, cl, var, var', v). (x_var var, x_var var')) y_cl_var_var'_v_list) x of None \ x | Some x' \ x')) (if x_var (case_option var (case_x (\var'. var') var) (map_of (map (\(y, cl, var, var', v). (x_var var, x_var var')) y_cl_var_var'_v_list) (x_var var))) = x' then Some ty_x_m else (\ ++ 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_var (case_option var (case_x (\var'. var') var) (if x_var var = x_this then Some x' else map_of (map (\(y, cl, var, var', v). (x_var var, x_var var')) y_cl_var_var'_v_list) (x_var var)))))" apply(rule_tac ty = ty_x and ty' = ty_var in sty_optionI) apply(simp add: var_k_of_ty_k) apply(simp only: x_var_not_this) apply(rule map_of_ty_k') apply(simp+) done lemma subtypes_through_tr: "\distinct (map (\(cl, var, ty). var) cl_var_ty_list); distinct (map (\(y, cl, var, var', v). var') y_cl_var_var'_v_list); x' \ (\(y, cl, var, var', v). x_var var') ` set y_cl_var_var'_v_list; map fst y_cl_var_var'_v_list = map fst y_ty_list; map (\(cl, var, ty). ty) cl_var_ty_list = map snd y_ty_list; map (\(y, cl, var, u). var) y_cl_var_var'_v_list = map (\(cl, var, ty). var) cl_var_ty_list; \x\set y_y'_list. case_prod (\y. (=) (case if y = x_this then Some x' else map_of (map (\(y, cl, var, var', v). (x_var var, x_var var')) y_cl_var_var'_v_list) y of None \ y | Some x' \ x')) x; \x\set y_ty_lista. (\(y, ty). sty_option P (if y = x_this then Some ty_x_m else map_of (map (\(cl, var, y). (x_var var, y)) cl_var_ty_list) y) (Some ty)) x; map fst y_y'_list = map fst y_ty_lista; (y, y') = y_y'_list ! i; (yy, ty) = y_ty_lista ! i; i < length y_y'_list; i < length y_ty_lista\ \ sty_option P (if y' = x' then Some ty_x_m else (\ ++ map_of (map (\((y, cl, var, var', v), y', y). (x_var var', y)) (zip y_cl_var_var'_v_list y_ty_list))) y') (Some ty)" apply(drule_tac x = "(y, y')" in bspec, simp) apply(drule_tac x = "(yy, ty)" in bspec, simp) apply(clarsimp split del: if_split) apply(frule_tac y_ty_list = y_ty_lista in map_y_yy, assumption+) apply(erule sty_option.cases) apply(clarsimp split del: if_split) apply(rule_tac ty = tya in sty_optionI, (simp add: var_k_of_ty_k)+) done lemma map_override_get: "(\ ++ (\u. if u = x' then Some ty_x_d else None)) x' = Some ty_x_d" apply(simp add: map_add_def) done lemma s_induct': "\\ss. \s\set ss. P s \ P (s_block ss); \list x. P (s_ass list x); \list1 x list2. P (s_read list1 x list2); \x1 list x2. P (s_write x1 list x2); \x1 x2 s1 s2. \P s1; P s2\ \ P (s_if x1 x2 s1 s2); \list1 x list2 list3. P (s_call list1 x list2 list3); \list ctx cl. P (s_new list ctx cl)\ \ P s" by (metis s.induct) lemma wf_tr_stmt_in_G': "\s'. distinct (map (\(cl, var, ty). var) cl_var_ty_list) \ distinct (map (\(y, cl, var, var', v). var') y_cl_var_var'_v_list) \ x' \ (\(y, cl, var, var', v). x_var var') ` set y_cl_var_var'_v_list \ map (\(y, cl, var, var', v). y) y_cl_var_var'_v_list = map fst y_ty_list \ map (\(cl, var, ty). ty) cl_var_ty_list = map snd y_ty_list \ map (\(y, cl, var, var', v). var) y_cl_var_var'_v_list = map (\(cl, var, ty). var) cl_var_ty_list \ is_sty_one P ty_x_d ty_x_m = Some True - \ wf_stmt P (map_of (map (\(cl, var, ty). (x_var var, ty)) cl_var_ty_list)(x_this \ ty_x_m)) s' - \ tr_s (map_of (map (\(y, cl, var, var', v). (x_var var, x_var var')) y_cl_var_var'_v_list)(x_this \ x')) s' s'' + \ wf_stmt P ((map_of (map (\(cl, var, ty). (x_var var, ty)) cl_var_ty_list))(x_this \ ty_x_m)) s' + \ tr_s ((map_of (map (\(y, cl, var, var', v). (x_var var, x_var var')) y_cl_var_var'_v_list))(x_this \ x')) s' 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_x_m)) s''" apply(rule s_induct') apply(clarsimp) apply(erule tr_s.cases, simp_all) apply(clarsimp) apply(erule wf_stmtE, simp_all) apply(clarsimp) apply(rule wf_blockI) apply(clarsimp) apply(rename_tac s'_s''_list s' s'') apply(drule_tac x = "(s', s'')" in bspec, simp)+ apply(clarsimp) apply(clarsimp) apply(rename_tac var x s') apply(erule tr_s.cases, simp+) apply(erule wf_stmtE, simp_all) apply(erule sty_option.cases) apply(rule wf_var_assignI) apply(clarsimp split del: if_split) apply(rule_tac ty_x = ty in subtype_through_tr, simp+) apply(clarsimp) apply(rename_tac var x f s') apply(erule tr_s.cases, simp+) apply(erule wf_stmtE, simp_all) apply(erule sty_option.cases) apply(clarsimp split del: if_split) apply(rename_tac f ty_x var x ty_f ty_var) apply(rule wf_field_readI) apply(simp add: var_k_of_ty_k) apply(simp) apply(rule_tac ty' = ty_var in sty_optionI) apply(simp) apply(rule map_of_ty_k', simp+) apply(clarsimp) apply(rename_tac x f y s') apply(erule tr_s.cases, simp+) apply(erule wf_stmtE, simp_all) apply(erule sty_option.cases) apply(clarsimp split del: if_split) apply(rename_tac f ty_x x y ty_y ty_f) apply(rule wf_field_writeI) apply(simp add: var_k_of_ty_k) apply(simp) apply(rule sty_optionI, (simp add: var_k_of_ty_k)+) apply(clarsimp) apply(rename_tac x y s1 s2 s') apply(erule tr_s.cases, simp+) apply(erule wf_stmtE, simp_all) apply(rule wf_ifI) apply(erule disjE) apply(rule disjI1) apply(erule sty_option.cases) apply(clarsimp split del: if_split) apply(rule sty_optionI, (simp add: var_k_of_ty_k split del: if_split)+) apply(rule disjI2) apply(erule sty_option.cases) apply(clarsimp split del: if_split) apply(rule sty_optionI, (simp add: var_k_of_ty_k)+) apply(clarsimp) apply(clarsimp) apply(clarsimp) apply(rename_tac var x meth ys s') apply(erule tr_s.cases, simp+) apply(erule wf_stmtE, simp_all) apply(clarsimp split del: if_split) apply(rule_tac y_ty_list = "map (\((y, y'), (yy, ty)). (y', ty)) (zip y_y'_list y_ty_lista)" in wf_mcallI[simplified]) apply(force simp add: map_map_zip_y'[simplified]) apply(simp add: var_k_of_ty_k) apply(force simp add: map_map_zip_ty[simplified]) apply(clarsimp simp add: set_zip split del: if_split) apply(simp add: subtypes_through_tr) apply(erule sty_option.cases) apply(clarsimp split del: if_split) apply(rule sty_optionI, (simp add: map_of_ty_k')+) apply(clarsimp) apply(rename_tac var ctx cl s') apply(erule tr_s.cases, simp+) apply(erule wf_stmtE, simp_all) apply(clarsimp) apply(rule wf_newI) apply(simp) apply(erule sty_option.cases) apply(clarsimp split del: if_split) apply(rule sty_optionI, (simp add: map_of_ty_k')+) done lemma wf_object_heap: "\wf_object P H v_opt (Some ty'); is_sty_one Pa ty_y_s ty_f = Some True; H oid_x = Some (ty_x_d, fs_oid)\ \ wf_object P (H(oid_x \ (ty_x_d, fs_oid(f \ v)))) v_opt (Some ty')" apply(erule wf_objectE) apply(clarsimp) apply(rule wf_nullI, simp) apply(erule sty_option.cases) apply(clarsimp split: option.splits) apply(rule wf_objectI) apply(force intro: sty_optionI) done lemma not_dom_is_none: "((\(y, cl, var, var', v). x_var var') ` set y_cl_var_var'_v_list \ Map.dom L = {}) \ (\x\set y_cl_var_var'_v_list. (\(y, cl, var, var', v). L (x_var var') = None) x)" by force lemma type_to_val: "\wf_varstate P \ H L; sty_option P (\ x) (\ y)\ \ \v w. L x = Some v \ L y = Some w" apply(erule sty_option.cases) apply(clarsimp) apply(erule wf_varstateE) apply(frule_tac x = x in bspec, simp add: domI) apply(drule_tac x = y in bspec, simp add: domI) apply(elim wf_objectE) by (simp+) lemma find_path_fields[THEN mp]: "find_path_ty_f P ty = Some path \ (\fs. fields_in_path_f path = fs)" by (force) lemma replicate_eq_length: "replicate x c = replicate y c \ x = y" apply(subgoal_tac "length (replicate x c) = length (replicate y c)") apply(thin_tac _) apply(subgoal_tac "length (replicate x c) = x") apply(subgoal_tac "length (replicate y c) = y") by simp_all lemma string_infinite: "infinite (UNIV :: string set)" apply(simp add: infinite_iff_countable_subset) apply(rule_tac x = "\n. replicate n c" in exI) apply(rule linorder_injI) apply(clarify) apply(frule replicate_eq_length) apply(simp) done lemma x_infinite: "infinite (UNIV :: x set)" apply(simp add: infinite_iff_countable_subset) apply(rule_tac x = "\n. if n = 0 then x_this else x_var (replicate n c)" in exI) apply(rule linorder_injI) apply(clarsimp) done lemma fresh_oid: "wf_heap P H \ (\oid. oid \ dom H)" apply(erule wf_heapE) apply(cut_tac infinite_UNIV_nat) apply(frule_tac A = "dom H" in ex_new_if_finite) apply(assumption) apply(simp) done lemma fresh_x: "finite A \ \x::x. x \ A" apply(cut_tac x_infinite) apply(frule_tac A = A in ex_new_if_finite) apply(assumption) apply(simp) done lemma fresh_var_not_in_list: "finite (A::x set) \ \var::var. x_var var \ A \ var \ set list" apply(cut_tac x_infinite) apply(subgoal_tac "finite (insert x_this A)") apply(frule_tac F = "insert x_this A" and G = "x_var ` set list" in finite_UnI) apply(simp) apply(frule_tac A = "insert x_this A \ x_var ` set list" in ex_new_if_finite) apply(assumption) apply(erule exE) apply(clarsimp) apply(case_tac a) apply(force+) done lemma fresh_vars[rule_format]: "finite (A::x set) \ (\list::var list. length list = i \ x_var` set list \ A = {} \ distinct list)" apply(induct i) apply(simp) apply(clarsimp) apply(frule fresh_var_not_in_list) apply(erule exE) apply(rule_tac x = "var#list" in exI) apply(simp) done lemma fresh_x_not_in_vars': "finite A \ \x'. x' \ A \ x' \ (x_var ` set vars')" apply(subgoal_tac "finite (x_var ` set vars')") apply(frule_tac G = "x_var ` set vars'" in finite_UnI) apply(assumption) apply(cut_tac x_infinite) apply(frule_tac A = "A \ x_var ` set vars'" in ex_new_if_finite) apply(simp) apply(simp) apply(simp) done lemma lift_opts_ft_length[rule_format]: "\tys. lift_opts (map (case_vd (\clk vark. find_type_f P ctx_o clk)) vds) = Some tys \ length tys = length vds" by (induct vds, auto split: vd.splits option.splits) lemma fpr_mem_has_path'[rule_format]: "\suffix. find_path_rec_f P ctx cl prefix = Some (prefix @ suffix) \ (\(ctx', cld') \ set suffix. \prefix'. (\suffix'. find_path_rec_f P ctx' (cl_fqn (fqn_def (class_name_f cld'))) prefix' = Some suffix'))" supply [[simproc del: defined_all]] apply(induct_tac P ctx cl prefix rule: find_path_rec_f.induct) apply(simp) apply(clarsimp split: option.splits) apply(subgoal_tac "\aa ba aaa bb. (a = aaa \ b = bb \ aa = aaa \ ba = bb) \ (\suffix. find_path_rec_f P aaa (superclass_name_f bb) (path @ [(aaa, bb)]) = Some (path @ (aaa, bb) # suffix) \ (\x\set suffix. (\(ctx', cld'). \prefix'. \suffix'. case_option None (case_prod (\ctx' cld. find_path_rec_f P ctx' (superclass_name_f cld) (prefix' @ [(ctx', cld)]))) (find_cld_f P ctx' (fqn_def (class_name_f cld'))) = Some suffix') x))") defer apply(thin_tac _)+ apply(force) apply(thin_tac "\aa ba x y. \a = aa \ b = ba; x = aa \ y = ba\ \ \suffix. find_path_rec_f P aa (superclass_name_f ba) (path @ [(aa, ba)]) = Some (path @ (aa, ba) # suffix) \ (\x\set suffix. case x of (ctx', cld') \ \prefix'. \suffix'. case_option None (\(ctx', cld). find_path_rec_f P ctx' (superclass_name_f cld) (prefix' @ [(ctx', cld)])) (find_cld_f P ctx' (fqn_def (class_name_f cld'))) = Some suffix')") apply(clarsimp) apply(frule path_append) apply(clarify) apply(erule_tac x = "tl suffix" in allE) apply(clarsimp) apply(erule disjE) apply(clarify) apply(clarsimp simp add: class_name_f_def split: cld.splits) apply(rename_tac list1 cl list2 list3) apply(case_tac fqn) apply(clarsimp) apply(frule find_cld_name_eq) apply(clarsimp simp add: class_name_f_def split: cld.splits) apply(frule_tac suffix = path'' and prefix' = " prefix' @ [(a, cld_def list1 cl list2 list3)]" in fpr_same_suffix') apply(simp) apply(simp) apply(drule_tac x = "(aa, ba)" in bspec, simp) apply(clarsimp split: option.splits) done lemma fpr_mem_has_path: "\find_path_f P ctx cl = Some suffix; (ctx', cld') \ set suffix\ \ \suffix'. find_path_rec_f P ctx' (cl_fqn (fqn_def (class_name_f cld'))) [] = Some suffix'" apply(cut_tac x = "(ctx', cld')" in fpr_mem_has_path'[of _ _ _ "[]"]) apply(simp add: find_path_f_def)+ done lemma mtype_to_find_md: "\wf_program P; mtype_f P ty_x_s m = Some (mty_def tys ty_r); is_sty_one P ty_x_d ty_x_s = Some True\ \ \ctx cl_r vds ss' y. find_meth_def_f P ty_x_d m = Some (ctx, meth_def_def (meth_sig_def cl_r m vds) (meth_body_def ss' y)) \ length tys = length vds" apply(clarsimp simp add: is_sty_one_def split: option.splits ty.splits) apply(clarsimp simp add: mtype_f_def find_meth_def_f_def) apply(case_tac ty_x_d) apply(simp) apply(clarsimp) apply(rename_tac path_d ctx_s dcl_s ctx_d dcl_d) apply(rename_tac cl_s fds_s mds_s) apply(frule fpr_mem_has_path) apply(simp) apply(clarsimp simp del: find_path_rec_f.simps) apply(frule_tac path = path_d and m = m and mty = "mty_def tys ty_r" and path' = suffix' in mty_preservation') apply(simp add: class_name_f_def find_path_f_def)+ apply(clarsimp simp add: mtype_f_def split: option.splits if_split_asm meth_def.splits meth_sig.splits) apply(rename_tac meth_body a b c d e f g) apply(case_tac meth_body) apply(simp) apply(clarsimp simp add: lift_opts_ft_length) apply(clarsimp simp add: find_meth_def_f_def split: option.splits) apply(frule find_md_m_match[rule_format]) apply(assumption) done lemma exist_lifted_values: "\\x\dom \. wf_object P H (L x) (\ x); \x\set y_ty_list. (\(y, ty). sty_option P (\ y) (Some ty)) x\ \ \vs. lift_opts (map (\(y, ty). L y) y_ty_list) = Some vs" apply(induct y_ty_list) apply(simp) apply(clarsimp) apply(erule sty_option.cases) apply(clarsimp) apply(drule_tac x = a in bspec, simp add: domI) apply(erule wf_objectE) apply(simp) apply(simp) done lemma [iff]: - "length (tr_ss_f (map_of (zip (map (case_vd (\cl. x_var)) vds) (map x_var vars'))(x_this \ x')) ss') = length ss'" + "length (tr_ss_f ((map_of (zip (map (case_vd (\cl. x_var)) vds) (map x_var vars')))(x_this \ x')) ss') = length ss'" by (induct ss', auto) lemma collect_suc_eq_lt[simp]: "{f i |i::nat. i < Suc n} = {f i |i. i = 0} \ {f (Suc i) |i. i < n}" apply(induct n) apply(simp) apply(clarsimp) apply(simp add: less_Suc_eq) apply(simp add: image_Collect[THEN sym]) apply(simp add: Collect_disj_eq) apply(blast) done lemma [iff]: "\y_ty_list vars vars' vs. length y_ty_list = length vds \ length vars = length vds \ length vars' = length vds \ length vs = length vds \ map (\(y, cl, var, var', v). vd_def cl var) (zip (map fst y_ty_list) (zip (map (case_vd (\cl var. cl)) vds) (zip (map (case_vd (\cl var. var)) vds) (zip vars' vs)))) = vds" apply(induct vds) apply(simp) apply(clarsimp simp add: set_zip length_Suc_conv split: vd.splits) done lemma vars'_eq[rule_format]: "\y_ty_list vds vars vs. length y_ty_list = length vars' \ length vds = length vars' \ length vars = length vars' \ length vs = length vars' \ (\(y, cl, var, var', v). x_var var') ` set (zip (map fst y_ty_list) (zip (map (case_vd (\cl var. cl)) vds) (zip vars (zip vars' vs)))) = x_var ` set vars'" apply(induct vars') apply(simp) apply(clarsimp simp add: set_zip length_Suc_conv) done lemma [iff]: "\y_ty_list vds vars vs. length y_ty_list = length vars' \ length vds = length vars' \ length vars = length vars' \ length vs = length vars' \ map (\(y, cl, var, var', v). var') (zip (map fst y_ty_list) (zip (map (case_vd (\cl var. cl)) vds) (zip vars (zip vars' vs)))) = vars'" apply(induct vars') apply(simp) apply(clarsimp simp add: set_zip length_Suc_conv) done lemma [iff]: "\y_ty_list vds vars vs. length y_ty_list = length vars' \ length vds = length vars' \ length vars = length vars' \ length vs = length vars' \ map (\(y, cl, var, var', v). (x_var var, x_var var')) (zip (map fst y_ty_list) (zip (map (case_vd (\cl var. cl)) vds) (zip (map (case_vd (\cl var. var)) vds) (zip vars' vs)))) = zip (map (case_vd (\cl. x_var)) vds) (map x_var vars')" apply(induct vars') apply(simp) apply(clarsimp simp add: set_zip length_Suc_conv split: vd.splits) done lemma [iff]: "\vds vars vars' vs. length vds = length y_ty_list \ length vars = length y_ty_list \ length vars' = length y_ty_list \ length vs = length y_ty_list \ map (\(y, cl, var, var', v). y) (zip (map fst y_ty_list) (zip (map (case_vd (\cl var. cl)) vds) (zip vars (zip vars' vs)))) = map fst y_ty_list" apply(induct y_ty_list) apply(simp) apply(clarsimp simp add: set_zip length_Suc_conv) done lemma lift_opts_mapping: "\vds vars vars' vs. length vds = length y_ty_list \ length vars = length y_ty_list \ length vars' = length y_ty_list \ length vs = length y_ty_list \ lift_opts (map (\(y, ty). L y) y_ty_list) = Some vs \ (\x\set (zip (map fst y_ty_list) (zip (map (case_vd (\cl var. cl)) vds) (zip (map (case_vd (\cl var. var)) vds) (zip vars' vs)))). (\(y, cl, var, var', v). L y = Some v) x)" apply(induct y_ty_list) apply(simp) apply(clarsimp simp add: set_zip length_Suc_conv split: vd.splits option.splits) apply(rename_tac y1 y_ty_list v2 vs cl1 var1 var'1 v1 vds vars var'2 vars' i cl2 var2) apply(erule_tac x = vds in allE) apply(erule_tac x = vars in allE) apply(erule_tac x = vars' in allE) apply(simp) apply(case_tac i) apply(simp) apply(rename_tac j) apply(clarsimp) apply(erule_tac x = "map fst y_ty_list ! j" in allE) apply(clarsimp) apply(case_tac "zip (map (case_vd (\cl var. cl)) vds) (zip (map (case_vd (\cl var. var)) vds) (zip vars' vs)) ! j") apply(rename_tac cl1 var1 var'1 v1) apply(erule_tac x = cl1 in allE) apply(erule_tac x = var1 in allE) apply(erule_tac x = var'1 in allE) apply(erule_tac x = v1 in allE) apply(clarsimp) apply(erule_tac x = j in allE) apply(simp) done lemma length_y_ty_list_vs[rule_format]: "\vs. lift_opts (map (\(y, ty). L y) y_ty_list) = Some vs \ length y_ty_list = length vs" by (induct y_ty_list, auto split: option.splits) lemma translation_eq: - "\x\set (zip (tr_ss_f (map_of (zip (map (case_vd (\cl. x_var)) vds) (map x_var vars'))(x_this \ x')) ss') ss'). - (\(s'', s'). tr_s (map_of (zip (map (case_vd (\cl. x_var)) vds) (map x_var vars'))(x_this \ x')) s' s'') x" + "\x\set (zip (tr_ss_f ((map_of (zip (map (case_vd (\cl. x_var)) vds) (map x_var vars')))(x_this \ x')) ss') ss'). + (\(s'', s'). tr_s ((map_of (zip (map (case_vd (\cl. x_var)) vds) (map x_var vars')))(x_this \ x')) s' s'') x" apply(induct ss') apply(simp add: tr_rel_f_eq)+ done theorem progress: "\wf_config \ (config_normal P L H S); S \ []\ \ \config'. r_stmt (config_normal P L H S) config'" supply [[simproc del: defined_all]] apply(case_tac S) apply(simp) apply(clarsimp) apply(rename_tac s ss) apply(case_tac s) apply(erule_tac[1-7] wf_configE) apply(simp_all) \ \block\ apply(force intro: r_blockI[simplified]) \ \variable assignment\ apply(clarsimp, erule wf_stmtE, simp_all, clarsimp) apply(frule type_to_val, simp) apply(clarify) apply(frule r_var_assignI) apply(force) \ \field read\ apply(clarsimp, erule wf_stmtE, simp_all, clarsimp) apply(erule wf_varstateE) apply(drule_tac x = x in bspec, simp add: domI) apply(erule wf_objectE) apply(clarsimp) apply(frule r_field_read_npeI) apply(force) apply(clarsimp) apply(erule_tac ?a3.0 = "Some ty" in sty_option.cases) apply(clarsimp split: option.splits) apply(erule wf_heapE) apply(drule_tac x = oid in bspec, simp add: domI) apply(clarsimp) apply(rename_tac x oid ty_x_s ty_x_d fields_oid fs) apply(frule no_field_hiding, simp+) apply(drule_tac x = f in bspec, simp) apply(clarsimp) apply(erule wf_objectE) apply(clarsimp, frule_tac H = H in r_field_readI, simp, force)+ \ \field write\ apply(clarsimp, erule wf_stmtE, simp_all, clarsimp) apply(erule wf_varstateE) apply(frule_tac x = x in bspec, simp add: domI) apply(erule wf_objectE) apply(clarsimp) apply(frule r_field_write_npeI) apply(force) apply(clarsimp) apply(erule sty_option.cases) apply(clarsimp) apply(rename_tac ty_y ty_f) apply(drule_tac x = y in bspec, simp add: domI) apply(clarsimp) apply(erule sty_option.cases) apply(clarsimp split: option.splits) apply(erule wf_objectE) apply(clarsimp) apply(frule_tac H = H and y = y in r_field_writeI, simp, force)+ \ \conditional branch\ apply(clarsimp, erule wf_stmtE, simp_all, clarsimp) apply(erule disjE) apply(frule type_to_val, simp, clarify) apply(case_tac "v = w") apply(frule_tac y = y in r_if_trueI, force+) apply(frule_tac y = y in r_if_falseI, force+) apply(frule type_to_val, simp, clarify) apply(case_tac "v = w") apply(frule_tac y = y and v = w in r_if_trueI, force+) apply(frule_tac y = y and v = w in r_if_falseI, force+) \ \object creation\ apply(clarsimp, erule wf_stmtE, simp_all, clarsimp) apply(rename_tac cl ctx ty var) apply(erule sty_option.cases) apply(clarsimp) apply(rename_tac ty_cl ty_var) apply(simp add: is_sty_one_def split: option.splits) apply(rename_tac path) apply(frule find_path_fields) apply(erule exE) apply(frule fresh_oid) apply(erule exE) apply(frule_tac H = H and L = L and var = var and s_list = ss and f_list = fs in r_newI[simplified]) apply(clarsimp simp add: fields_f_def split: option.splits) apply(assumption) apply(simp) apply(force) \ \method call\ apply(clarsimp, erule wf_stmtE, simp_all, clarsimp) apply(rename_tac ss y_ty_list x ty_x_s m ty_r_s var) apply(erule wf_varstateE) apply(frule_tac x = x in bspec, simp add: domI) apply(clarsimp) apply(erule wf_objectE) apply(clarsimp) apply(frule r_mcall_npeI) apply(force) apply(elim sty_option.cases) apply(clarsimp split: option.splits) apply(rename_tac ty_r_s ty_var_s ty_x_s ty_x_d fs_x) apply(frule mtype_to_find_md, simp+) apply(clarsimp) apply(frule_tac A = "dom L" and i = "length vds" in fresh_vars) apply(clarsimp) apply(rename_tac vars') apply(frule exist_lifted_values) apply(simp) apply(clarify) apply(frule_tac vars' = vars' in fresh_x_not_in_vars') apply(erule exE) apply(erule conjE) apply(subgoal_tac "\vars. vars = map (\vd. case vd of vd_def cl var \ var) vds") apply(erule exE) apply(subgoal_tac "length vars = length vds") apply(frule length_y_ty_list_vs) apply(subgoal_tac "\T. T = (map_of (zip (map (\vd. case vd of vd_def cl var \ x_var var) vds) (map x_var vars')))(x_this \ x')") apply(erule exE) apply(frule_tac H = H and P = P and meth = m and ctx = ctx and cl = cl_r and y = y and ty = ty_x_d and y_cl_var_var'_v_list = "zip (map fst y_ty_list) (zip (map (\vd. case vd of vd_def cl var \ cl) vds) (zip vars (zip vars' vs)))" and s''_s'_list = "zip (tr_ss_f T ss') ss'" and var = var and s_list = ss and x' = x' and T = T in r_mcallI[simplified]) apply(force) apply(simp) apply(simp) apply(simp add: vars'_eq) apply(simp) apply(assumption) apply(simp add: vars'_eq) apply(cut_tac L = L and y_ty_list = y_ty_list in lift_opts_mapping) apply(erule_tac x = vds in allE) apply(erule_tac x = vars in allE) apply(erule_tac x = vars' in allE) apply(erule_tac x = vs in allE) apply(simp) apply(simp) apply(simp) apply(simp add: translation_eq) apply(simp) apply(force) by force+ theorem wf_preservation: "\\ config. \wf_config \ config; r_stmt config config'\ \ \\'. \ \\<^sub>m \' \ wf_config \' config'" supply [[simproc del: defined_all]] apply(erule r_stmt.cases) (* s_read_npe, s_write_npe *) (* \' = \ for all cases except for mcall (case 11) *) apply(rule_tac[1-10] x = \ in exI) apply(erule_tac[1-11] wf_configE) apply(simp_all) (* s_if s_block *) apply(clarsimp) apply(erule wf_stmtE) apply(simp_all) apply(force intro: wf_config_normalI) (* s_ass *) apply(clarsimp) apply(erule wf_stmtE) apply(simp_all) apply(rule wf_config_normalI, assumption+) apply(erule sty_option.cases) apply(rule wf_subvarstateI, assumption+, simp) apply(erule wf_varstateE) apply(drule_tac x = x in bspec, simp add: domI) apply(erule wf_objectE) apply(simp add: wf_nullI) apply(clarsimp) apply(rule wf_objectI) apply(erule sty_option.cases, simp) apply(rule sty_optionI, simp+) apply(rule_tac ty' = ty'a in sty_transitiveI, assumption+) (* s_read *) apply(force intro: wf_intros) apply(clarsimp split: option.splits) apply(erule wf_stmtE) apply(simp_all) apply(rule wf_config_normalI, assumption+) apply(clarsimp) apply(rename_tac L oid H v s_list ty_oid fs_oid \ x ty_x P f ty_f var) apply(erule sty_option.cases) apply(rule wf_subvarstateI, assumption, simp) apply(clarsimp) apply(case_tac v) apply(clarify) apply(rule wf_nullI, simp) apply(clarify) apply(rename_tac ty_f ty_var P oid_v) apply(erule wf_heapE) apply(drule_tac x = oid in bspec, simp add: domI) apply(clarsimp) apply(rename_tac H P fs_oid) apply(erule wf_varstateE) apply(frule_tac x = x in bspec, simp add: domI) apply(clarsimp) apply(erule wf_objectE) apply(simp) apply(erule sty_option.cases) apply(clarsimp split: option.splits) apply(rename_tac L \ H oid_x ty_x_d ty_x_s P) apply(frule_tac ty_x_s = ty_x_s in no_field_hiding, simp+) apply(drule_tac x = f in bspec, simp) apply(clarsimp) apply(frule ftype_preservation[simplified], assumption+) apply(clarsimp) apply(rename_tac ty_f) apply(erule wf_objectE) apply(simp) apply(clarsimp) apply(erule sty_option.cases) apply(clarsimp split: option.splits) apply(rename_tac H oid_v ty_f P ty_v fs_v) apply(rule wf_objectI) apply(rule sty_optionI, simp+) apply(rule sty_transitiveI, assumption+) (* s_write *) apply(force intro: wf_intros) apply(clarsimp, hypsubst_thin) apply(erule wf_stmtE) apply(simp_all) apply(clarsimp split: option.splits) apply(rule) apply(erule wf_varstateE) apply(clarsimp) apply(rename_tac \ P H) apply(drule_tac x = xa in bspec, simp add: domI) apply(erule wf_objectE) apply(simp) apply(clarsimp) apply(elim sty_option.cases) apply(simp) apply(erule sty_option.cases) apply(clarsimp) apply(rule wf_config_normalI) apply(assumption) apply(rename_tac L oid_x v H ss \ x ty_x_s f y ty_y_s ty_f P ty_x_d fs_oid) apply(erule wf_heapE) apply(rule wf_heapI) apply(simp) apply(rule ballI) apply(clarsimp simp add: map_add_def) apply(rule) apply(clarsimp) apply(drule_tac x = oid_x in bspec, simp add: domI) apply(clarsimp) apply(drule_tac x = fa in bspec, simp) apply(clarsimp) apply(rule) apply(clarsimp) apply(case_tac v) apply(clarsimp simp add: wf_nullI) apply(clarsimp) apply(rename_tac H P fields_x_d ty_f_d oid_v) apply(erule wf_varstateE) apply(clarsimp) apply(frule ftype_preservation, assumption+) apply(clarsimp) apply(drule_tac x = y in bspec, simp add: domI) apply(clarsimp) apply(erule_tac ?a4.0 = "Some ty_y_s" in wf_objectE) apply(simp) apply(clarsimp) apply(erule sty_option.cases) apply(clarsimp split: option.splits) apply(rule wf_objectI) apply(force intro: sty_optionI sty_transitiveI) apply(clarsimp) apply(case_tac "fs_oid fa") apply(clarsimp) apply(force elim: wf_objectE) apply(case_tac a) apply(force intro: wf_nullI) apply(clarsimp) apply(rename_tac H P fields_x_d f ty_f_d oid_v) apply(erule wf_varstateE) apply(clarsimp) apply(drule_tac x = y in bspec, simp add: domI) apply(clarsimp) apply(erule wf_objectE) apply(simp) apply(clarsimp) apply(rule wf_objectI) apply(erule sty_option.cases) apply(clarsimp split: option.splits) apply(safe) apply(clarsimp) apply(force intro: sty_optionI) apply(force intro: sty_optionI) apply(drule_tac x = oid in bspec) apply(clarsimp) apply(clarsimp) apply(drule_tac x = fa in bspec, simp) apply(clarsimp simp add: wf_object_heap) apply(rename_tac L oid_x v H ss \ x ty_x_s f y ty_y_s ty_f P ty_x_d fs_x) apply(erule wf_varstateE) apply(rule wf_varstateI) apply(simp) apply(clarsimp) apply(rename_tac L \ P H x' y') apply(drule_tac x = x' in bspec, simp add: domI) apply(clarsimp) apply(erule wf_objectE) apply(clarsimp) apply(rule wf_nullI, simp) apply(clarsimp) apply(erule sty_option.cases) apply(clarsimp split: option.splits) apply(rename_tac H oid_x' ty_x'_s P ty_x'_d fs_x') apply(rule wf_objectI) apply(rule sty_optionI) apply(clarsimp) apply(rule conjI) apply(clarsimp) apply(simp) apply(clarsimp) apply(simp) apply(simp) (* s_if *) apply(erule wf_stmtE) apply(simp_all) apply(force intro: wf_config_normalI) apply(erule wf_stmtE) apply(simp_all) apply(force intro: wf_config_normalI) (* s_new *) apply(erule contrapos_np) apply(erule wf_stmtE, simp_all, clarsimp) apply(erule sty_option.cases, clarsimp) apply(rule wf_config_normalI) apply(assumption) apply(rename_tac H L ss ctx cl \ var ty' ty_var P) apply(erule wf_heapE) apply(rule wf_heapI) apply(simp) apply(safe) apply(simp add: map_upd_Some_unfold) apply(split if_split_asm) apply(clarsimp) apply(frule field_has_type, simp+) apply(erule exE) apply(rule_tac x = ty in exI) apply(simp) apply(force intro: wf_nullI) apply(drule_tac x = oida in bspec, simp add: domI) apply(clarsimp split: option.splits) apply(drule_tac x = f in bspec, simp) apply(safe) apply(rule_tac x = ty'a in exI) apply(simp) apply(erule wf_objectE) apply(clarsimp) apply(force intro: wf_nullI) apply(erule sty_option.cases) apply(clarsimp split: option.splits) apply(rule wf_objectI) apply(clarsimp) apply(rule conjI) apply(force) apply(force intro: sty_optionI) apply(rename_tac oid_new H L ss ctx cl \ var ty_new ty_var P) apply(erule wf_varstateE) apply(clarsimp) apply(rule wf_varstateI) apply(simp) apply(rule ballI) apply(drule_tac x = x in bspec, simp) apply(clarsimp) apply(rule conjI) apply(clarsimp) apply(rule wf_objectI) apply(clarsimp) apply(force intro: sty_optionI) apply(rule) apply(erule wf_objectE) apply(clarsimp) apply(force intro: wf_nullI) apply(erule sty_option.cases) apply(clarsimp split: option.splits) apply(rule wf_objectI) apply(clarsimp) apply(rule) apply(rule) apply(force) apply(rule) apply(force intro: sty_optionI) (* s_call *) apply(force intro: wf_all_exI) (* setting the new type environment *) apply(case_tac "H oid", simp) apply(clarsimp) apply(erule wf_stmtE, simp_all, clarsimp) apply(frule wf_method_from_find[simplified]) apply(simp) apply(safe) apply(erule_tac Q = "\\'. \ \\<^sub>m \' \ ~ (wf_config \' (config_normal Pa ((L ++ map_of (map (\(y_XXX, cl_XXX, var_XXX, var_', y). (x_var var_', y)) y_cl_var_var'_v_list))(x' \ v_oid oid)) H (map fst s''_s'_list @ s_ass vara (case if y = x_this then Some x' else map_of (map (\(y_XXX, cl_XXX, var_XXX, var_', v_XXX). (x_var var_XXX, x_var var_')) y_cl_var_var'_v_list) y of None \ y | Some x' \ x') # s_list)))" in contrapos_pp) apply(simp only: not_all) apply(rule_tac x = "\ ++ map_of (map (\((y, cl, var, var', v), (y', ty)). (x_var var', ty)) (zip y_cl_var_var'_v_list y_ty_list)) ++ Map.empty (x' \ ty_def ctx dcl)" in exI) apply(clarsimp split del: if_split) apply(rule) apply(clarsimp simp add: map_le_def split del: if_split) apply(rule sym) apply(rule_tac L = L and Pa = Pa and H = H in map_of_map_and_x) apply(assumption+) apply(simp add: not_dom_is_none) apply(erule wf_varstateE) apply(clarsimp) apply(rename_tac ss ty_x_d fs_x y_ty_list \ x ty_x_s P meth ty_r var dcl) apply(erule sty_option.cases) apply(clarsimp split del: if_split) apply(rename_tac ty_r ty_var P) apply(erule wf_varstateE, clarify) apply(rename_tac L \ P H) apply(frule_tac x = x in bspec, simp add: domI) apply(frule_tac x = "x_var var" in bspec, simp add: domI) apply(clarsimp split del: if_split) apply(erule wf_objectE, simp, clarsimp split del: if_split) apply(rename_tac P H oid) apply(erule sty_option.cases, clarsimp split del: if_split) apply(rename_tac ty_x_d ty_x_s P) apply(subgoal_tac "x_var var \ dom \") apply(clarify) apply(rename_tac v_var) apply(drule not_dom_is_none) apply(rule wf_config_normalI) apply(assumption) apply(assumption) (* varstate *) apply(rule wf_varstateI) apply(simp only: finite_super_varstate) apply(clarsimp) apply(rule) apply(rule wf_objectI) apply(simp) apply(rule sty_optionI, simp+) apply(clarsimp, hypsubst_thin) apply(rule) apply(rule) apply(rule wf_objectI) apply(clarsimp) apply(rule sty_optionI, simp+) apply(rule) apply(erule disjE) apply(clarsimp) apply(drule map_of_SomeD) apply(clarsimp) apply(rename_tac y_k cl_k var_k var'_k v_k y'_k ty_k) apply(frule_tac x = "(y_k, cl_k, var_k, var'_k, v_k)" in bspec) apply(force simp add: set_zip) apply(clarsimp) apply(subgoal_tac "map_of (map (\(y, cl, var, var', v). (x_var var', v)) y_cl_var_var'_v_list) (x_var var'_k) = Some v_k") apply(clarsimp) apply(drule map_y) apply(simp) apply(drule_tac x = "(y_k, ty_k)" in bspec, assumption) apply(clarsimp) apply(erule sty_option.cases) apply(clarsimp) apply(rename_tac ty_y_k ty_k P) apply(drule_tac x = y_k in bspec, simp add: domI) apply(clarsimp) apply(case_tac v_k) apply(clarify, rule wf_nullI, simp) apply(clarsimp) apply(rename_tac oid_y_k) apply(erule_tac ?a4.0 = "Some ty_y_k" in wf_objectE) apply(simp) apply(erule sty_option.cases) apply(clarsimp split: option.splits) apply(rule wf_objectI) apply(clarsimp) apply(rule sty_optionI, simp+) apply(rule_tac ty' = ty' in sty_transitiveI, assumption+) apply(rule map_of_is_SomeI) apply(simp add: map_fst_var'[simplified]) apply(rule set_zip_var'_v, simp) apply(clarify) apply(rename_tac x_G ty_x_G) apply(frule_tac x = x_G in bspec, simp add: domI) apply(case_tac "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_G") apply(frule map_of_map_zip_none, simp+) apply(simp add: map_add_def) apply(frule map_of_map_zip_some, simp+) apply(clarsimp) apply(drule map_of_SomeD) apply(drule map_of_SomeD) apply(clarsimp simp add: set_zip) apply(frule same_el, simp+) apply(drule_tac s = "(aa, ab, ac, ai, b)" in sym) apply(clarsimp) apply(rename_tac y_k cl_k var_k v_k y'_k ty_k var'_k i) apply(frule_tac y_k = y_k and cl_k = cl_k and var_k = var_k and var'_k = var'_k and v_k = v_k in same_y, simp+) apply(clarsimp) apply(drule_tac x = "(y_k, ty_k)" in bspec, simp) apply(clarsimp) apply(rename_tac y_k ty_k) apply(erule sty_option.cases) apply(clarsimp) apply(drule_tac x = y_k in bspec, simp add: domI) apply(clarsimp) apply(drule_tac x = "(y_k, cl_k, var_k, var'_k, v_k)" in bspec) apply(drule_tac t = "(y_k, cl_k, var_k, var'_k, v_k)" in sym) apply(simp) apply(clarsimp) apply(erule_tac ?a4.0 = "Some ty" in wf_objectE) apply(clarsimp) apply(rule wf_nullI, simp) apply(erule sty_option.cases) apply(clarsimp split: option.splits) apply(rule wf_objectI) apply(clarsimp) apply(rule sty_optionI, simp+) apply(rule_tac ty' = ty'a in sty_transitiveI, assumption+) (* assignment *) apply(clarsimp split del: if_split) apply(rule) apply(rule wf_var_assignI) apply(frule wf_method_from_find) apply(simp) apply(erule exE) apply(erule wf_methE) apply(case_tac "ya = x_this") apply(clarsimp split del: if_split) apply(rule_tac ty' = ty_var in sty_optionI) apply(simp) apply(simp) apply(case_tac ctx, clarify) apply(frule_tac ty_r_d = ty in mty_preservation, assumption+) apply(clarify) apply(erule sty_option.cases) apply(clarsimp) apply(rule_tac ty' = ty' in sty_transitiveI, assumption+) (* y != this *) apply(erule sty_option.cases) apply(clarsimp split del: if_split) apply(frule_tac var_k = ya and ty_k = tya in exists_var') apply(drule map_of_vd) apply(assumption+) apply(erule exE) apply(clarsimp) apply(drule_tac y = "x_var var'_k" in map_of_SomeD) apply(clarsimp split del: if_split) apply(rename_tac y_k cl_k var_k var'_k v_k) apply(frule x'_not_in_list, assumption+) apply(clarsimp) apply(drule map_of_SomeD) apply(clarsimp split del: if_split) apply(rename_tac cl_k' var_k ty_k) apply(case_tac ctx, clarify) apply(frule mty_preservation, assumption+) apply(clarsimp split del: if_split) apply(drule map_of_vd) apply(frule map_of_ty_k, assumption+) apply(rule_tac ty = ty_k and ty' = ty_var in sty_optionI) apply(simp add: map_add_def) apply(simp) apply(simp) apply(rule_tac ty' = ty_r in sty_transitiveI) apply(assumption+) (* wf of translated statements *) apply(clarsimp) apply(erule disjE) apply(erule wf_methE) apply(clarsimp) apply(frule map_of_vd[rule_format]) apply(case_tac ctx, clarify) apply(frule mty_preservation, simp+) apply(clarsimp) apply(rename_tac P dcl cl y meth s'' s') apply(drule_tac x = "(s'', s')" in bspec, simp)+ apply(clarsimp) apply(cut_tac cl_var_ty_list = cl_var_ty_list and y_cl_var_var'_v_list = y_cl_var_var'_v_list and y_ty_list = y_ty_list and P = P and ty_x_d = ty_x_d and ty_x_m = "ty_def ctx_def dcl" and x' = x' and s'' = s'' and \ = \ in wf_tr_stmt_in_G') apply(clarsimp) (* wf of non-translated statements *) apply(cut_tac L = L and x' = x' and y_cl_var_var'_v_list = y_cl_var_var'_v_list and \ = \ and P = P and H = H and y_ty_list = y_ty_list and ty = "ty_def ctx dcl" and ss = ss in wf_stmt_in_G') apply(clarsimp) apply(erule wf_objectE) apply(force) apply(force) done end diff --git a/thys/Locally-Nameless-Sigma/Sigma/Sigma.thy b/thys/Locally-Nameless-Sigma/Sigma/Sigma.thy --- a/thys/Locally-Nameless-Sigma/Sigma/Sigma.thy +++ b/thys/Locally-Nameless-Sigma/Sigma/Sigma.thy @@ -1,2634 +1,2634 @@ section \Locally Nameless representation of basic Sigma calculus enriched with formal parameter\ theory Sigma imports "../preliminary/FMap" begin subsection \Infrastructure for the finite maps\ axiomatization max_label :: nat where LabelAvail: "max_label > 10" definition "Label = {n :: nat. n \ max_label}" typedef Label = Label unfolding Label_def by auto lemmas finite_Label_set = Finite_Set.finite_Collect_le_nat[of max_label] lemma Univ_abs_label: "(UNIV :: (Label set)) = Abs_Label ` {n :: nat. n \ max_label}" proof - have "\x :: Label. x : Abs_Label ` {n :: nat. n \ max_label}" proof fix x :: Label have "Rep_Label x \ {n. n \ max_label}" by (fold Label_def, rule Rep_Label) hence "Abs_Label (Rep_Label x) \ Abs_Label ` {n. n \ max_label}" by (rule imageI) thus "x \ Abs_Label ` {n. n \ max_label}" by (simp add: Rep_Label_inverse) qed thus ?thesis by force qed lemma finite_Label: "finite (UNIV :: (Label set))" by (simp add: Univ_abs_label finite_Label_set) instance Label :: finite proof show "finite (UNIV :: (Label set))" by (rule finite_Label) qed consts Lsuc :: "(Label set) \ Label \ Label" Lmin :: "(Label set) \ Label" Lmax :: "(Label set) \ Label" definition Llt :: "[Label, Label] \ bool" (infixl "<" 50) where "Llt a b == Rep_Label a < Rep_Label b" definition Lle :: "[Label, Label] \ bool" (infixl "\" 50) where "Lle a b == Rep_Label a \ Rep_Label b" definition Ltake_eq :: "[Label set, (Label \ 'a), (Label \ 'a)] \ bool" where "Ltake_eq L f g == \l\L. f l = g l" lemma Ltake_eq_all: fixes f g assumes "dom f = dom g" and "Ltake_eq (dom f) f g" shows "f = g" proof fix x from assms show "f x = g x" unfolding Ltake_eq_def by (cases "x \ dom f", auto) qed lemma Ltake_eq_dom: fixes L :: "Label set" and f :: "Label -~> 'a" assumes "L \ dom f" and "card L = card (dom f)" shows "L = (dom f)" proof (auto) fix x :: Label assume "x \ L" with in_mono[OF assms(1)] show "\y. f x = Some y" by blast next fix x y assume "f x = Some y" from card_seteq[OF finite_dom_fmap[of f] \L \ dom f\] \card L = card (dom f)\ have "L = dom f" by simp with \f x = Some y\ show "x \ L" by force qed subsection \Object-terms in Locally Nameless representation notation, beta-reduction and substitution\ datatype type = Object "Label -~> (type \ type)" datatype bVariable = Self nat | Param nat type_synonym fVariable = string (* each binder introduces 2 variables: self and parameter *) (* thus to each deBruijn index (nat) correspond 2 variables of the same depth *) subsubsection \Enriched Sigma datatype of objects\ datatype sterm = Bvar bVariable (* bound variable -- as deBruijn index *) | Fvar fVariable (* free variable *) | Obj "Label -~> sterm" type (* An object maps labels to terms and has a type *) | Call sterm Label sterm (* Call a l b calls method l on object a with param b *) | Upd sterm Label sterm (* Upd a l b updates method l of object a with body b *) datatype_compat sterm primrec applyPropOnOption:: "(sterm \ bool) \ sterm option \ bool" where f1: "applyPropOnOption P None = True" | f2: "applyPropOnOption P (Some t) = P t" lemma sterm_induct[case_names Bvar Fvar Obj Call Upd empty insert]: fixes t :: sterm and P1 :: "sterm \ bool" and f :: "Label -~> sterm" and P3 :: "(Label -~> sterm) \ bool" assumes "\b. P1 (Bvar b)" and "\x. P1 (Fvar x)" and a_obj: "\f T. P3 f \ P1 (Obj f T)" and "\t1 l t2. \ P1 t1; P1 t2 \ \ P1 (Call t1 l t2)" and "\t1 l t2. \ P1 t1; P1 t2 \ \ P1 (Upd t1 l t2)" and "P3 Map.empty" and a_f: "\t1 f l . \ l \ dom f; P1 t1; P3 f \ \ (P3 (f(l \ t1)))" shows "P1 t \ P3 f" proof - have "\t (f::Label -~> sterm) l. P1 t \ (applyPropOnOption P1) (f l)" proof (intro strip) fix t :: sterm and f' :: "Label -~> sterm" and l :: Label define foobar where "foobar = f' l" from assms show "P1 t \ applyPropOnOption P1 foobar" proof (induct_tac t and foobar rule: compat_sterm.induct compat_sterm_option.induct, auto) fix f :: "Label -~> sterm" and T :: type assume "\x. applyPropOnOption P1 (f x)" with a_f \P3 Map.empty\ have "P3 f" proof (induct f rule: fmap_induct, simp) case (insert F x z) note P1F' = \\xa. applyPropOnOption P1 ((F(x \ z)) xa)\ and predP3 = \\ \l f t1. \l \ dom f; P1 t1; P3 f\ \ P3 (f(l \ t1)); P3 Map.empty; \x. applyPropOnOption P1 (F x)\ \ P3 F\ and P3F' = \\l f t f. \ l \ dom f; P1 t; P3 f \ \ P3 (f(l \ t))\ have "\xa. applyPropOnOption P1 (F xa)" proof - fix xa :: Label show "applyPropOnOption P1 (F xa)" proof (cases "xa = x") case True with P1F' \x \ dom F\ have "F xa = None" by force thus ?thesis by simp next case False hence eq: "F xa = (F(x \ z)) xa" by auto from P1F'[of xa] show "applyPropOnOption P1 (F xa)" by (simp only: ssubst[OF eq]) qed qed from a_f predP3[OF _ \P3 Map.empty\ this] have P3F: "P3 F" by simp from P1F'[of x] have "applyPropOnOption P1 (Some z)" by auto hence "P1 z" by simp from a_f[OF \x \ dom F\ this P3F] show ?case by assumption qed with a_obj show "P1 (Obj f T)" by simp qed qed with assms show ?thesis proof (auto) assume "\l f t1. \l \ dom f; P3 f\ \ P3 (f(l \ t1))" with \P3 Map.empty\ show "P3 f" by (rule fmap_induct) qed qed lemma ball_tsp_P3: fixes P1 :: "sterm \ bool" and P2 :: "sterm \ fVariable \ fVariable \ bool" and P3 :: "sterm \ bool" and f :: "Label -~> sterm" assumes "\t. \ P1 t; \s p. s \ L \ p \ L \ s \ p \ P2 t s p \ \ P3 t" and "\l\dom f. P1 (the(f l))" and "\l\dom f. \s p. s \ L \ p \ L \ s \ p \ P2 (the(f l)) s p" shows "\l\dom f. P3 (the(f l))" proof (intro strip) fix l assume "l \ dom f" with assms(2) have "P1 (the(f l))" by blast moreover from assms(3) \l \ dom f\ have "\s p. s \ L \ p \ L \ s \ p \ P2 (the(f l)) s p" by blast ultimately show "P3 (the(f l))" using assms(1) by simp qed lemma ball_tt'sp_P3: fixes P1 :: "sterm \ sterm \ bool" and P2 :: "sterm \ sterm \ fVariable \ fVariable \ bool" and P3 :: "sterm \ sterm \ bool" and f :: "Label -~> sterm" and f' :: "Label -~> sterm" assumes "\t t'. \ P1 t t'; \s p. s \ L \ p \ L \ s \ p \ P2 t t' s p \ \ P3 t t'" and "dom f = dom f'" and "\l\dom f. P1 (the(f l)) (the(f' l))" and "\l\dom f. \s p. s \ L \ p \ L \ s \ p \ P2 (the(f l)) (the(f' l)) s p" shows "\l\dom f'. P3 (the(f l)) (the(f' l))" proof (intro strip) fix l assume "l \ dom f'" with assms(2-3) have "P1 (the(f l)) (the(f' l))" by blast moreover from assms(2,4) \l \ dom f'\ have "\s p. s \ L \ p \ L \ s \ p \ P2 (the (f l)) (the(f' l)) s p" by blast ultimately show "P3 (the(f l)) (the(f' l))" using assms(1)[of "the(f l)" "the(f' l)"] by simp qed subsubsection \Free variables\ primrec FV :: "sterm \ fVariable set" and FVoption :: "sterm option \ fVariable set" where FV_Bvar : "FV(Bvar b) = {}" | FV_Fvar : "FV(Fvar x) = {x}" | FV_Call : "FV(Call t l a) = FV t \ FV a" | FV_Upd : "FV(Upd t l s) = FV t \ FV s" | FV_Obj : "FV(Obj f T) = (\ l \ dom f. FVoption(f l))" | FV_None : "FVoption None = {}" | FV_Some : "FVoption (Some t) = FV t" definition closed :: "sterm \ bool" where "closed t \ FV t = {}" (* finiteness of FV *) lemma finite_FV_FVoption: "finite (FV t) \ finite (FVoption s)" by(induct _ t _ s rule: compat_sterm_sterm_option.induct, simp_all) (* for infiniteness of string sets see List.infinite_UNIV_listI *) (* NOTE: ex_new_if_finite, infinite_UNIV_listI and finite_FV offer an easy way to proof exists-fresh-statements *) lemma finite_FV[simp]: "finite (FV t)" by (simp add: finite_FV_FVoption) lemma FV_and_cofinite: "\ \x. x \ L \ P x; finite L \ \ \L'. (finite L' \ FV t \ L' \ (\ x. x \ L' \ P x))" by (rule_tac x = "L \ FV t" in exI, auto) lemma exFresh_s_p_cof: fixes L :: "fVariable set" assumes "finite L" shows "\s p. s \ L \ p \ L \ s \ p" proof - from assms have "\s. s \ L" by (simp only: ex_new_if_finite[OF infinite_UNIV_listI]) then obtain s where "s \ L" .. moreover from \finite L\ have "finite (L \ {s})" by simp hence "\p. p \ L \ {s}" by (simp only: ex_new_if_finite[OF infinite_UNIV_listI]) then obtain p where "p \ L \ {s}" .. ultimately show ?thesis by blast qed lemma FV_option_lem: "\l\dom f. FV (the(f l)) = FVoption (f l)" by auto subsubsection \Term opening\ primrec sopen :: "[nat, sterm, sterm, sterm] \ sterm" ("{_ \ [_,_]} _" [0, 0, 0, 300] 300) and sopen_option :: "[nat, sterm, sterm, sterm option] \ sterm option" where sopen_Bvar: "{k \ [s,p]}(Bvar b) = (case b of (Self i) \ (if (k = i) then s else (Bvar b)) | (Param i) \ (if (k = i) then p else (Bvar b)))" | sopen_Fvar: "{k \ [s,p]}(Fvar x) = Fvar x" | sopen_Call: "{k \ [s,p]}(Call t l a) = Call ({k \ [s,p]}t) l ({k \ [s,p]}a)" | sopen_Upd : "{k \ [s,p]}(Upd t l u) = Upd ({k \ [s,p]}t) l ({(Suc k) \ [s,p]}u)" | sopen_Obj : "{k \ [s,p]}(Obj f T) = Obj (\l. sopen_option (Suc k) s p (f l)) T" | sopen_None: "sopen_option k s p None = None" | sopen_Some: "sopen_option k s p (Some t) = Some ({k \ [s,p]}t)" definition openz :: "[sterm, sterm, sterm] \ sterm" ("(_)\<^bsup>[_,_]\<^esup>" [50, 0, 0] 50) where "t\<^bsup>[s,p]\<^esup> = {0 \ [s,p]}t" lemma sopen_eq_Fvar: fixes n s p t x assumes "{n \ [Fvar s,Fvar p]} t = Fvar x" shows "(t = Fvar x) \ (x = s \ t = (Bvar (Self n))) \ (x = p \ t = (Bvar (Param n)))" proof (cases t) case Obj with assms show ?thesis by simp next case Call with assms show ?thesis by simp next case Upd with assms show ?thesis by simp next case (Fvar y) with assms show ?thesis by (cases "y = x", simp_all) next case (Bvar b) thus ?thesis proof (cases b) case (Self k) with assms Bvar show ?thesis by (cases "k = n") simp_all next case (Param k) with assms Bvar show ?thesis by (cases "k = n") simp_all qed qed lemma sopen_eq_Fvar': assumes "{n \ [Fvar s,Fvar p]} t = Fvar x" and "x \ s" and "x \ p" shows "t = Fvar x" proof - from sopen_eq_Fvar[OF assms(1)] \x \ s\ \x \ p\ show ?thesis by auto qed lemma sopen_eq_Bvar: fixes n s p t b assumes "{n \ [Fvar s,Fvar p]} t = Bvar b" shows "t = Bvar b" proof (cases t) case Fvar with assms show ?thesis by (simp add: openz_def) next case Obj with assms show ?thesis by (simp add: openz_def) next case Call with assms show ?thesis by (simp add: openz_def) next case Upd with assms show ?thesis by (simp add: openz_def) next case (Bvar b') show ?thesis proof (cases b') case (Self k) with assms Bvar show ?thesis by (cases "k = n") (simp_all add: openz_def) next case (Param k) with assms Bvar show ?thesis by (cases "k = n") (simp_all add: openz_def) qed qed lemma sopen_eq_Obj: fixes n s p t f T assumes "{n \ [Fvar s,Fvar p]} t = Obj f T" shows "\f'. {n \ [Fvar s, Fvar p]} Obj f' T = Obj f T \ t = Obj f' T" proof (cases t) case Fvar with assms show ?thesis by simp next case Obj with assms show ?thesis by simp next case Call with assms show ?thesis by simp next case Upd with assms show ?thesis by simp next case (Bvar b) thus ?thesis proof (cases b) case (Self k) with assms Bvar show ?thesis by (cases "k = n") simp_all next case (Param k) with assms Bvar show ?thesis by (cases "k = n") simp_all qed qed lemma sopen_eq_Upd: fixes n s p t t1 l t2 assumes "{n \ [Fvar s,Fvar p]} t = Upd t1 l t2" shows "\t1' t2'. {n \ [Fvar s,Fvar p]} t1' = t1 \ {(Suc n) \ [Fvar s,Fvar p]} t2' = t2 \ t = Upd t1' l t2'" proof (cases t) case Fvar with assms show ?thesis by simp next case Obj with assms show ?thesis by simp next case Call with assms show ?thesis by simp next case Upd with assms show ?thesis by simp next case (Bvar b) thus ?thesis proof (cases b) case (Self k) with assms Bvar show ?thesis by (cases "k = n") simp_all next case (Param k) with assms Bvar show ?thesis by (cases "k = n") simp_all qed qed lemma sopen_eq_Call: fixes n s p t t1 l t2 assumes "{n \ [Fvar s,Fvar p]} t = Call t1 l t2" shows "\t1' t2'. {n \ [Fvar s,Fvar p]} t1' = t1 \ {n \ [Fvar s,Fvar p]} t2' = t2 \ t = Call t1' l t2'" proof (cases t) case Fvar with assms show ?thesis by simp next case Obj with assms show ?thesis by simp next case Call with assms show ?thesis by simp next case Upd with assms show ?thesis by simp next case (Bvar b) thus ?thesis proof (cases b) case (Self k) with assms Bvar show ?thesis by (cases "k = n") simp_all next case (Param k) with assms Bvar show ?thesis by (cases "k = n") simp_all qed qed lemma dom_sopenoption_lem[simp]: "dom (\l. sopen_option k s t (f l)) = dom f" by (auto, case_tac "x \ dom f", auto) lemma sopen_option_lem: "\l\dom f. {n \ [s,p]} the(f l) = the (sopen_option n s p (f l))" by auto lemma pred_sopenoption_lem: "(\l\dom (\l. sopen_option n s p (f l)). (P::sterm \ bool) (the (sopen_option n s p (f l)))) = (\l\dom f. (P::sterm \ bool) ({n \ [s,p]} the (f l)))" by (simp, force) lemma sopen_FV[rule_format]: "\n s p. FV ({n \ [s,p]} t) \ FV t \ FV s \ FV p" proof - fix u have "(\n s p. FV ({n \ [s,p]} t) \ FV t \ FV s \ FV p) &(\n s p. FVoption (sopen_option n s p u) \ FVoption u \ FV s \ FV p)" apply (induct u rule: compat_sterm_sterm_option.induct) apply (auto split: bVariable.split) apply (metis (no_types, lifting) FV_Some UnE domI sopen_Some subsetCE) apply blast apply blast done from conjunct1[OF this] show ?thesis by assumption qed lemma sopen_commute[rule_format]: "\n k s p s' p'. n \ k \ {n \ [Fvar s', Fvar p']} {k \ [Fvar s, Fvar p]} t = {k \ [Fvar s, Fvar p]} {n \ [Fvar s', Fvar p']} t" proof - fix u have "(\n k s p s' p'. n \ k \ {n \ [Fvar s', Fvar p']} {k \ [Fvar s, Fvar p]} t = {k \ [Fvar s, Fvar p]} {n \ [Fvar s', Fvar p']} t) &(\n k s p s' p'. n \ k \ sopen_option n (Fvar s') (Fvar p') (sopen_option k (Fvar s) (Fvar p) u) = sopen_option k (Fvar s) (Fvar p) (sopen_option n (Fvar s') (Fvar p') u))" by (rule compat_sterm_sterm_option.induct, simp_all split: bVariable.split) from conjunct1[OF this] show ?thesis by assumption qed lemma sopen_fresh_inj[rule_format]: "\n s p t'. {n \ [Fvar s, Fvar p]} t = {n \ [Fvar s, Fvar p]} t' \ s \ FV t \ s \ FV t' \ p \ FV t \ p \ FV t' \ s \ p \ t = t'" proof - { fix b s p n t assume "(case b of Self i \ if n = i then Fvar s else Bvar b | Param i \ if n = i then Fvar p else Bvar b) = t" hence "Fvar s = t \ Fvar p = t \ Bvar b = t" by (cases b, auto, (rename_tac nat, case_tac "n = nat", auto)+) } note cT = this (* induction *) fix u have "(\n s p t'. {n \ [Fvar s, Fvar p]} t = {n \ [Fvar s, Fvar p]} t' \ s \ FV t \ s \ FV t' \ p \ FV t \ p \ FV t' \ s \ p \ t = t') &(\n s p u'. sopen_option n (Fvar s) (Fvar p) u = sopen_option n (Fvar s) (Fvar p) u' \ s \ FVoption u \ s \ FVoption u' \ p \ FVoption u \ p \ FVoption u' \ s \ p \ u = u')" proof (induct _ t _ u rule: compat_sterm_sterm_option.induct) case (Bvar b) thus ?case proof (auto) fix s p n t assume a: "(case b of Self i \ if n = i then Fvar s else Bvar b | Param i \ if n = i then Fvar p else Bvar b) = {n \ [Fvar s,Fvar p]} t" note cT[OF this] moreover assume "s \ FV t" and "p \ FV t" and "s \ p" ultimately show "Bvar b = t" proof (auto) { fix b' assume "(case b' of Self i \ if n = i then Fvar s else Bvar b' | Param i \ if n = i then Fvar p else Bvar b') = Fvar s" with \s \ p\ have "b' = Self n" by (cases b', auto, (rename_tac nat, case_tac "n = nat", auto)+) }note fvS = this assume eq_s: "Fvar s = {n \ [Fvar s,Fvar p]} t" with sym[OF this] \s \ FV t\ \s \ p\ fvS have "t = Bvar (Self n)" by (cases t, auto) moreover from a sym[OF eq_s] \s \ p\ fvS[of b] have "Self n = b" by simp ultimately show "Bvar b = t" by simp next { fix b' assume "(case b' of Self i \ if n = i then Fvar s else Bvar b' | Param i \ if n = i then Fvar p else Bvar b') = Fvar p" with \s \ p\ have "b' = Param n" by (cases b', auto, (rename_tac nat, case_tac "n = nat", auto)+) }note fvP = this assume eq_p: "Fvar p = {n \ [Fvar s,Fvar p]} t" with sym[OF this] \p \ FV t\ \s \ p\ fvP have "t = Bvar (Param n)" by (cases t, auto) moreover from a sym[OF eq_p] \s \ p\ fvP[of b] have "Param n = b" by simp ultimately show "Bvar b = t" by simp next assume "Bvar b = {n \ [Fvar s, Fvar p]} t" from sym[OF this] show "{n \ [Fvar s,Fvar p]} t = t" proof (cases t, auto) fix b' assume "(case b' of Self i \ if n = i then Fvar s else Bvar b' | Param i \ if n = i then Fvar p else Bvar b') = Bvar b" from cT[OF this] have "Bvar b = Bvar b'" by simp thus "b = b'" by simp qed qed qed next case (Fvar x) thus ?case proof (auto) fix n s p t assume a: "Fvar x = {n \ [Fvar s,Fvar p]} t" and "s \ FV ({n \ [Fvar s,Fvar p]} t)" hence "s \ x" by force moreover assume "p \ FV ({n \ [Fvar s,Fvar p]} t)" with a have "p \ x" by force ultimately show "{n \ [Fvar s,Fvar p]} t = t" using a proof (cases t, auto) fix b assume "Fvar x = (case b of Self i \ if n = i then Fvar s else Bvar b | Param i \ if n = i then Fvar p else Bvar b)" from cT[OF sym[OF this]] \s \ x\ \p \ x\ have False by simp then show "(case b of Self i \ if n = i then Fvar s else Bvar b | Param i \ if n = i then Fvar p else Bvar b) = Bvar b" .. qed qed next case (Obj f T) thus ?case proof (auto) fix n s p t assume "Obj (\l. sopen_option (Suc n) (Fvar s) (Fvar p) (f l)) T = {n \ [Fvar s,Fvar p]} t" and "\l\dom f. s \ FVoption (f l)" and "\l\dom f. p \ FVoption (f l)" and "s \ FV t" and "p \ FV t" and "s \ p" thus "Obj f T = t" using Obj proof (cases t, auto) (* case Bvar *) fix b assume "Obj (\l. sopen_option (Suc n) (Fvar s) (Fvar p) (f l)) T = (case b of Self i \ if n = i then Fvar s else Bvar b | Param i \ if n = i then Fvar p else Bvar b)" from cT[OF sym[OF this]] show False by auto next (* case Obj *) fix f' assume nin_s: "\l\dom f'. s \ FVoption (f' l)" and nin_p: "\l\dom f'. p \ FVoption (f' l)" and ff': "(\l. sopen_option (Suc n) (Fvar s) (Fvar p) (f l)) = (\l. sopen_option (Suc n) (Fvar s) (Fvar p) (f' l))" have "\l. f l = f' l" proof - fix l from ff' have "sopen_option (Suc n) (Fvar s) (Fvar p) (f l) = sopen_option (Suc n) (Fvar s) (Fvar p) (f' l)" by (rule cong, simp) moreover from \\l\dom f. s \ FVoption (f l)\ \\l\dom f. p \ FVoption (f l)\ have "s \ FVoption (f l)" and "p \ FVoption (f l)" by (case_tac "l \ dom f", auto)+ moreover from nin_s nin_p have "s \ FVoption (f' l)" and "p \ FVoption (f' l)" by (case_tac "l \ dom f'", auto)+ ultimately show "f l = f' l" using Obj[of l] \s \ p\ by simp qed thus "f = f'" by (rule ext) qed qed next case (Call t1 l t2) thus ?case proof (auto) fix n s p t assume "s \ FV t1" and "s \ FV t2" and "p \ FV t1" and "p \ FV t2" and "s \ p" and "s \ FV t" and "p \ FV t" and "Call ({n \ [Fvar s,Fvar p]} t1) l ({n \ [Fvar s,Fvar p]} t2) = {n \ [Fvar s,Fvar p]} t" thus "Call t1 l t2 = t" using Call proof (cases t, auto) (* case Bvar *) fix b assume "Call ({n \ [Fvar s,Fvar p]} t1) l ({n \ [Fvar s,Fvar p]} t2) = (case b of Self i \ if n = i then Fvar s else Bvar b | Param i \ if n = i then Fvar p else Bvar b)" from cT[OF sym[OF this]] show False by auto qed qed next case (Upd t1 l t2) thus ?case proof (auto) fix n s p t assume "s \ FV t1" and "s \ FV t2" and "p \ FV t1" and "p \ FV t2" and "s \ p" and "s \ FV t" and "p \ FV t" and "Upd ({n \ [Fvar s,Fvar p]} t1) l ({Suc n \ [Fvar s,Fvar p]} t2) = {n \ [Fvar s,Fvar p]} t" thus "Upd t1 l t2 = t" using Upd proof (cases t, auto) (* case Bvar *) fix b assume "Upd ({n \ [Fvar s,Fvar p]} t1) l ({Suc n \ [Fvar s,Fvar p]} t2) = (case b of Self i \ if n = i then Fvar s else Bvar b | Param i \ if n = i then Fvar p else Bvar b)" from cT[OF sym[OF this]] show False by auto qed qed next case None_sterm thus ?case proof (auto) fix u s p n assume "None = sopen_option n (Fvar s) (Fvar p) u" thus "None = u" by (cases u, auto) qed next case (Some_sterm t) thus ?case proof (auto) fix u s p n assume "Some ({n \ [Fvar s,Fvar p]} t) = sopen_option n (Fvar s) (Fvar p) u" and "s \ FV t" and "p \ FV t" and "s \ p" and "s \ FVoption u" and "p \ FVoption u" with Some_sterm show "Some t = u" by (cases u) auto qed qed from conjunct1[OF this] show ?thesis by assumption qed subsubsection \Variable closing\ primrec sclose :: "[nat, fVariable, fVariable, sterm] \ sterm" ("{_ \ [_,_]} _" [0, 0, 0, 300] 300) and sclose_option :: "[nat, fVariable, fVariable, sterm option] \ sterm option" where sclose_Bvar: "{k \ [s,p]}(Bvar b) = Bvar b" | sclose_Fvar: "{k \ [s,p]}(Fvar x) = (if x = s then (Bvar (Self k)) else (if x = p then (Bvar (Param k)) else (Fvar x)))" | sclose_Call: "{k \ [s,p]}(Call t l a) = Call ({k \ [s,p]}t) l ({k \ [s,p]}a)" | sclose_Upd : "{k \ [s,p]}(Upd t l u) = Upd ({k \ [s,p]}t) l ({(Suc k) \ [s,p]}u)" | sclose_Obj : "{k \ [s,p]}(Obj f T) = Obj (\l. sclose_option (Suc k) s p (f l)) T" | sclose_None: "sclose_option k s p None = None" | sclose_Some: "sclose_option k s p (Some t) = Some ({k \ [s,p]}t)" definition closez :: "[fVariable, fVariable, sterm] \ sterm" ("\[_,_] _" [0, 0, 300]) where "\[s,p] t = {0 \ [s,p]}t" lemma dom_scloseoption_lem[simp]: "dom (\l. sclose_option k s t (f l)) = dom f" by (auto, case_tac "x \ dom f", auto) lemma sclose_option_lem: "\l\dom f. {n \ [s,p]} the(f l) = the (sclose_option n s p (f l))" by auto lemma pred_scloseoption_lem: "(\l\dom (\l. sclose_option n s p (f l)). (P::sterm \ bool) (the (sclose_option n s p (f l)))) = (\l\dom f. (P::sterm \ bool) ({n \ [s,p]} the (f l)))" by (simp, force) lemma sclose_fresh[simp, rule_format]: "\n s p. s \ FV t \ p \ FV t \ {n \ [s,p]} t = t" proof - have "(\n s p. s \ FV t \ p \ FV t \ {n \ [s,p]} t = t) &(\n s p. s \ FVoption u \ p \ FVoption u \ sclose_option n s p u = u)" proof (induct _ t _ u rule: compat_sterm_sterm_option.induct, auto simp: bVariable.split) (* Obj *) fix f n s p assume nin_s: "\l\dom f. s \ FVoption (f l)" and nin_p: "\l\dom f. p \ FVoption (f l)" { fix l from nin_s have "s \ FVoption (f l)" by (case_tac "l \ dom f", auto) } moreover { fix l from nin_p have "p \ FVoption (f l)" by (case_tac "l \ dom f", auto) } moreover assume "\x. \n s. s \ FVoption (f x) \ (\p. p \ FVoption (f x) \ sclose_option n s p (f x) = f x)" ultimately have "\l. sclose_option (Suc n) s p (f l) = f l" by auto with ext show "(\l. sclose_option (Suc n) s p (f l)) = f" by auto qed from conjunct1[OF this] show ?thesis by assumption qed lemma sclose_FV[rule_format]: "\n s p. FV ({n \ [s,p]} t) = FV t - {s} - {p}" proof - have "(\n s p. FV ({n \ [s,p]} t) = FV t - {s} - {p}) &(\n s p. FVoption (sclose_option n s p u) = FVoption u - {s} - {p})" by (rule compat_sterm_sterm_option.induct, simp_all split: bVariable.split, blast+) from conjunct1[OF this] show ?thesis by assumption qed lemma sclose_subset_FV[rule_format]: "FV ({n \ [s,p]} t) \ FV t" by (simp add: sclose_FV, blast) lemma Self_not_in_closed[simp]: "sa \ FV ({n \ [sa,pa]} t)" by (simp add: sclose_FV) lemma Param_not_in_closed[simp]: "pa \ FV ({n \ [sa,pa]} t)" by (simp add: sclose_FV) subsubsection \Substitution\ primrec ssubst :: "[fVariable, sterm, sterm] \ sterm" ("[_ \ _] _" [0, 0, 300] 300) and ssubst_option :: "[fVariable, sterm, sterm option] \ sterm option" where ssubst_Bvar: "[z \ u](Bvar v) = Bvar v" | ssubst_Fvar: "[z \ u](Fvar x) = (if (z = x) then u else (Fvar x))" | ssubst_Call: "[z \ u](Call t l s) = Call ([z \ u]t) l ([z \ u]s)" | ssubst_Upd : "[z \ u](Upd t l s) = Upd ([z \ u]t) l ([z \ u]s)" | ssubst_Obj : "[z \ u](Obj f T) = Obj (\l. ssubst_option z u (f l)) T" | ssubst_None: "ssubst_option z u None = None" | ssubst_Some: "ssubst_option z u (Some t) = Some ([z \ u]t)" lemma dom_ssubstoption_lem[simp]: "dom (\l. ssubst_option z u (f l)) = dom f" by (auto, case_tac "x \ dom f", auto) lemma ssubst_option_lem: "\l\dom f. [z \ u] the(f l) = the (ssubst_option z u (f l))" by auto lemma pred_ssubstoption_lem: "(\l\dom (\l. ssubst_option x t (f l)). (P::sterm \ bool) (the (ssubst_option x t (f l)))) = (\l\dom f. (P::sterm \ bool) ([x \ t] the (f l)))" by (simp, force) lemma ssubst_fresh[simp, rule_format]: "\s sa. sa \ FV t \ [sa \ s] t = t" proof - have "(\s sa. sa \ FV t \ [sa \ s] t = t) &(\s sa. sa \ FVoption u \ ssubst_option sa s u = u)" proof (induct _ t _ u rule: compat_sterm_sterm_option.induct, auto) fix s sa f assume sa: "\l\dom f. sa \ FVoption (f l)" and ssubst: "\x. \s sa. sa \ FVoption (f x) \ ssubst_option sa s (f x) = f x" { fix l from sa have "sa \ FVoption (f l)" by (case_tac "l \ dom f", auto) with ssubst have "ssubst_option sa s (f l) = f l" by auto } with ext show "(\l. ssubst_option sa s (f l)) = f" by auto qed from conjunct1[OF this] show ?thesis by assumption qed lemma ssubst_commute[rule_format]: "\s p sa pa. s \ p \ s \ FV pa \ p \ FV sa \ [s \ sa] [p \ pa] t = [p \ pa] [s \ sa] t" proof - have "(\s p sa pa. s \ p \ s \ FV pa \ p \ FV sa \ [s \ sa] [p \ pa] t = [p \ pa] [s \ sa] t) &(\s p sa pa. s \ p \ s \ FV pa \ p \ FV sa \ ssubst_option s sa (ssubst_option p pa u) = ssubst_option p pa (ssubst_option s sa u))" by (rule compat_sterm_sterm_option.induct, simp_all split: bVariable.split) from conjunct1[OF this] show ?thesis by assumption qed lemma ssubst_FV[rule_format]: "\x s. FV ([x \ s] t) \ FV s \ (FV t - {x})" proof - have "(\x s. FV ([x \ s] t) \ FV s \ (FV t - {x})) &(\x s. FVoption (ssubst_option x s u) \ FV s \ (FVoption u - {x}))" by (rule compat_sterm_sterm_option.induct, simp_all split: bVariable.split, blast+) from conjunct1[OF this] show ?thesis by assumption qed lemma ssubstoption_insert: "l \ dom f \ (\(la::Label). ssubst_option x t' (if la = l then Some t else f la)) = (\(la::Label). ssubst_option x t' (f la))(l \ [x \ t'] t)" by (rule Ltake_eq_all, force, simp add: Ltake_eq_def) subsubsection \Local closure\ inductive lc :: "sterm \ bool" where lc_Fvar[simp, intro!]: "lc (Fvar x)" | lc_Call[simp, intro!]: "\ lc t; lc a \ \ lc (Call t l a)" | lc_Upd[simp, intro!] : "\ lc t; finite L; \s p. s \ L \ p \ L \ s \ p \ lc (u\<^bsup>[Fvar s, Fvar p]\<^esup>) \ \ lc (Upd t l u)" | lc_Obj[simp, intro!] : "\ finite L; \l\dom f. \s p. s \ L \ p \ L \ s \ p \ lc (the(f l)\<^bsup>[Fvar s, Fvar p]\<^esup>) \ \ lc (Obj f T)" definition body :: "sterm \ bool" where "body t \ (\L. finite L \ (\s p. s \ L \ p \ L \ s \ p \ lc (t\<^bsup>[Fvar s, Fvar p]\<^esup>)))" lemma lc_bvar: "lc (Bvar b) = False" by (rule iffI, erule lc.cases, simp_all) lemma lc_obj: "lc (Obj f T) = (\l\dom f. body (the(f l)))" proof fix f T assume "lc (Obj f T)" thus "\l\dom f. body (the(f l))" unfolding body_def by (rule lc.cases, auto) next fix f :: "Label -~> sterm" and T :: type assume "\l\dom f. body (the (f l))" hence "\L. finite L \ (\l\dom f. \s p. s \ L \ p \ L \ s \ p \ lc (the (f l)\<^bsup>[Fvar s, Fvar p]\<^esup>))" proof (induct f rule: fmap_induct) case empty thus ?case by blast next case (insert F x y) thus ?case proof - assume "x \ dom F" hence "\l\dom F. the(F l) = the ((F(x \ y)) l)" by auto with \\l\dom (F(x \ y)). body (the((F(x \ y)) l))\ have "\l\dom F. body (the (F l))" by force from insert(2)[OF this] obtain L where "finite L" and "\l\dom F. \s p. s \ L \ p \ L \ s \ p \ lc (the (F l)\<^bsup>[Fvar s, Fvar p]\<^esup>)" by auto moreover from \\l\dom (F(x \ y)). body (the((F(x \ y)) l))\ have "body y" by force then obtain L' where "finite L'" and "\s p. s \ L' \ p \ L' \ s \ p \ lc (y\<^bsup>[Fvar s, Fvar p]\<^esup>)" by (auto simp: body_def) ultimately show "\L. finite L \ (\l\dom (F(x \ y)). \s p. s \ L \ p \ L \ s \ p \ lc (the ((F(x \ y)) l)\<^bsup>[Fvar s,Fvar p]\<^esup>))" by (rule_tac x = "L \ L'" in exI, auto) qed qed thus "lc (Obj f T)" by auto qed lemma lc_upd: "lc (Upd t l s) = (lc t \ body s)" by (unfold body_def, rule iffI, erule lc.cases, auto) lemma lc_call: "lc (Call t l s) = (lc t \ lc s)" by (rule iffI, erule lc.cases, simp_all) lemma lc_induct[consumes 1, case_names Fvar Call Upd Obj Bnd]: fixes P1 :: "sterm \ bool" and P2 :: "sterm \ bool" assumes "lc t" and "\x. P1 (Fvar x)" and "\t l a. \ lc t; P1 t; lc a; P1 a \ \ P1 (Call t l a)" and "\t l u. \ lc t; P1 t; P2 u \ \ P1 (Upd t l u)" and "\f T. \l\dom f. P2 (the(f l)) \ P1 (Obj f T)" and "\L t. \ finite L; \s p. s \ L \ p \ L \ s \ p \ lc (t\<^bsup>[Fvar s, Fvar p]\<^esup>) \ P1 (t\<^bsup>[Fvar s, Fvar p]\<^esup>) \ \ P2 t" shows "P1 t" using assms by (induct rule: lc.induct, auto) subsubsection \Connections between sopen, sclose, ssubst, lc and body and resulting properties\ lemma ssubst_intro[rule_format]: "\n s p sa pa. sa \ FV t \ pa \ FV t \ sa \ pa \ sa \ FV p \ {n \ [s,p]} t = [sa \ s] [pa \ p] {n \ [Fvar sa, Fvar pa]} t" proof - have "(\n s p sa pa. sa \ FV t \ pa \ FV t \ sa \ pa \ sa \ FV p \ {n \ [s,p]} t = [sa \ s] [pa \ p] {n \ [Fvar sa, Fvar pa]} t) &(\n s p sa pa. sa \ FVoption u \ pa \ FVoption u \ sa \ pa \ sa \ FV p \ sopen_option n s p u = ssubst_option sa s (ssubst_option pa p (sopen_option n (Fvar sa) (Fvar pa) u)))" proof (induct _ t _ u rule: compat_sterm_sterm_option.induct) case Bvar thus ?case by (simp split: bVariable.split) next case Fvar thus ?case by simp next case Upd thus ?case by simp next case Call thus ?case by simp next case None_sterm thus ?case by simp next case (Obj f T) thus ?case proof (clarify) fix n s p sa pa assume "sa \ FV (Obj f T)" and "pa \ FV (Obj f T)" { fix l from \sa \ FV (Obj f T)\ have "sa \ FVoption (f l)" by (case_tac "l \ dom f", auto) } moreover { fix l from \pa \ FV (Obj f T)\ have pa: "pa \ FVoption (f l)" by (case_tac "l \ dom f", auto) } moreover assume "sa \ pa" and "sa \ FV p" ultimately have "\l. sopen_option (Suc n) s p (f l) = ssubst_option sa s (ssubst_option pa p (sopen_option (Suc n) (Fvar sa) (Fvar pa) (f l)))" using Obj by auto with ext show "{n \ [s,p]} Obj f T = [sa \ s] [pa \ p] {n \ [Fvar sa,Fvar pa]} Obj f T" by auto qed next case (Some_sterm t) thus ?case proof (clarify) fix n s p sa pa assume "sa \ FVoption (Some t)" hence "sa \ FV t" by simp moreover assume "pa \ FVoption (Some t)" hence "pa \ FV t" by simp moreover assume "sa \ pa" and "sa \ FV p" ultimately have "{n \ [s,p]} t = [sa \ s] [pa \ p] {n \ [Fvar sa,Fvar pa]} t" using Some_sterm by blast thus "sopen_option n s p (Some t) = ssubst_option sa s (ssubst_option pa p (sopen_option n (Fvar sa) (Fvar pa) (Some t)))" by simp qed qed from conjunct1[OF this] show ?thesis by assumption qed lemma sopen_lc_FV[rule_format]: fixes t assumes "lc t" shows "\n s p. {n \ [Fvar s, Fvar p]} t = t" using assms proof (induct taking: "\t. \n s p. {Suc n \ [Fvar s, Fvar p]} t = t" rule: lc_induct) case Fvar thus ?case by simp next case Call thus ?case by simp next case Upd thus ?case by simp next case (Obj f T) note pred = this show ?case proof (intro strip, simp) fix n s p { fix l have "sopen_option (Suc n) (Fvar s) (Fvar p) (f l) = f l" proof (cases "l \ dom f") case False hence "f l = None" by force thus ?thesis by force next case True with pred show ?thesis by force qed } with ext show "(\l. sopen_option (Suc n) (Fvar s) (Fvar p) (f l)) = f" by auto qed next case (Bnd L t) note cof = this(2) show ?case proof (intro strip) fix n s p from \finite L\ exFresh_s_p_cof[of "L \ FV t \ {s} \ {p}"] obtain sa pa where sapa: "sa \ L \ FV t \ {s} \ {p} \ pa \ L \ FV t \ {s} \ {p} \ sa \ pa" by auto with cof have "{Suc n \ [Fvar s,Fvar p]} (t\<^bsup>[Fvar sa, Fvar pa]\<^esup>) = (t\<^bsup>[Fvar sa, Fvar pa]\<^esup>)" by auto with sopen_commute[OF Suc_not_Zero[of n]] have eq: "{0 \ [Fvar sa,Fvar pa]} {Suc n \ [Fvar s,Fvar p]} t = {0 \ [Fvar sa,Fvar pa]} t" by (simp add: openz_def) from sapa contra_subsetD[OF sopen_FV[of _ "Fvar s" "Fvar p" t]] have "sa \ FV ({Suc n \ [Fvar s,Fvar p]} t)" and "sa \ FV t" and "pa \ FV ({Suc n \ [Fvar s,Fvar p]} t)" and "pa \ FV t" and "sa \ pa" by auto from sopen_fresh_inj[OF eq this] show "{Suc n \ [Fvar s,Fvar p]} t = t" by assumption qed qed lemma sopen_lc[simp]: fixes t n s p assumes "lc t" shows "{n \ [s,p]} t = t" proof - from exFresh_s_p_cof[of "FV t \ FV p"] obtain sa pa where "sa \ FV t" and "pa \ FV t" and "sa \ pa" and "sa \ FV p" and "pa \ FV p" by auto from ssubst_intro[OF this(1-4)] have "{n \ [s,p]} t = [sa \ s] [pa \ p] {n \ [Fvar sa,Fvar pa]} t" by simp with assms have "{n \ [s,p]} t = [sa \ s] [pa \ p] t" using sopen_lc_FV by simp with ssubst_fresh[OF \pa \ FV t\] have "{n \ [s,p]} t = [sa \ s] t" by simp with ssubst_fresh[OF \sa \ FV t\] show "{n \ [s,p]} t = t" by simp qed lemma sopen_twice[rule_format]: "\s p s' p' n. lc s \ lc p \ {n \ [s',p']} {n \ [s,p]} t = {n \ [s,p]} t" proof - have "(\s p s' p' n. lc s \ lc p \ {n \ [s',p']} {n \ [s,p]} t = {n \ [s,p]} t) &(\s p s' p' n. lc s \ lc p \ sopen_option n s' p' (sopen_option n s p u) = sopen_option n s p u)" by (rule compat_sterm_sterm_option.induct, auto simp: bVariable.split) from conjunct1[OF this] show ?thesis by assumption qed lemma sopen_sclose_commute[rule_format]: "\n k s p sa pa. n \ k \ sa \ FV s \ sa \ FV p \ pa \ FV s \ pa \ FV p \ {n \ [s, p]} {k \ [sa,pa]} t = {k \ [sa,pa]} {n \ [s, p]} t" proof - have "(\n k s p sa pa. n \ k \ sa \ FV s \ sa \ FV p \ pa \ FV s \ pa \ FV p \ {n \ [s, p]} {k \ [sa,pa]} t = {k \ [sa,pa]} {n \ [s, p]} t) &(\n k s p sa pa. n \ k \ sa \ FV s \ sa \ FV p \ pa \ FV s \ pa \ FV p \ sopen_option n s p (sclose_option k sa pa u) = sclose_option k sa pa (sopen_option n s p u))" by (rule compat_sterm_sterm_option.induct, simp_all split: bVariable.split) from conjunct1[OF this] show ?thesis by assumption qed lemma sclose_sopen_eq_t[rule_format]: "\n s p. s \ FV t \ p \ FV t \ s \ p \ {n \ [s,p]} {n \ [Fvar s, Fvar p]} t = t" proof - have "(\n s p. s \ FV t \ p \ FV t \ s \ p \ {n \ [s,p]} {n \ [Fvar s, Fvar p]} t = t) &(\n s p. s \ FVoption u \ p \ FVoption u \ s \ p \ sclose_option n s p (sopen_option n (Fvar s) (Fvar p) u) = u)" proof (induct _ t _ u rule: compat_sterm_sterm_option.induct, simp_all split: bVariable.split, auto) (* Obj *) fix f n s p assume nin_s: "\l\dom f. s \ FVoption (f l)" and nin_p: "\l\dom f. p \ FVoption (f l)" { fix l from nin_s have "s \ FVoption (f l)" by (case_tac "l \ dom f", auto) } moreover { fix l from nin_p have "p \ FVoption (f l)" by (case_tac "l \ dom f", auto) } moreover assume "s \ p" and "\x. \n s. s \ FVoption (f x) \ (\p. p \ FVoption (f x) \ s \ p \ sclose_option n s p (sopen_option n (Fvar s) (Fvar p) (f x)) = f x)" ultimately have "\l. sclose_option n s p (sopen_option n (Fvar s) (Fvar p) (f l)) = f l" by auto with ext show "(\l. sclose_option n s p (sopen_option n (Fvar s) (Fvar p) (f l))) = f " by auto qed from conjunct1[OF this] show ?thesis by assumption qed lemma sopen_sclose_eq_t[simp, rule_format]: fixes t assumes "lc t" shows "\n s p. {n \ [Fvar s, Fvar p]} {n \ [s,p]} t = t" using assms proof (induct taking: "\t. \n s p. {Suc n \ [Fvar s, Fvar p]} {Suc n \ [s,p]} t = t" rule: lc_induct) case Fvar thus ?case by simp next case Call thus ?case by simp next case Upd thus ?case by simp next case (Obj f T) note pred = this show ?case proof (intro strip, simp) fix n s p { fix l have "sopen_option (Suc n) (Fvar s) (Fvar p) (sclose_option (Suc n) s p (f l)) = f l" proof (cases "l \ dom f") case False hence "f l = None" by force thus ?thesis by simp next case True with pred show ?thesis by force qed } with ext show "(\l. sopen_option (Suc n) (Fvar s) (Fvar p) (sclose_option (Suc n) s p (f l))) = f" by simp qed next case (Bnd L t) note cof = this(2) show ?case proof (intro strip) fix n s p from \finite L\ exFresh_s_p_cof[of "L \ FV t \ {s} \ {p}"] obtain sa pa where sapa: "sa \ L \ FV t \ {s} \ {p} \ pa \ L \ FV t \ {s} \ {p} \ sa \ pa" by auto with cof have eq: "{Suc n \ [Fvar s,Fvar p]} {Suc n \ [s,p]} (t\<^bsup>[Fvar sa, Fvar pa]\<^esup>) = (t\<^bsup>[Fvar sa, Fvar pa]\<^esup>)" by blast { fix x assume "x \ FV t" from contra_subsetD[OF sclose_subset_FV this] have "x \ FV ({Suc n \ [s,p]} t)" by simp moreover assume "x \ p" and "x \ s" ultimately have "x \ FV ({Suc n \ [s,p]} t) \ FV (Fvar s) \ FV (Fvar p)" by simp from contra_subsetD[OF sopen_FV this] have "x \ FV ({Suc n \ [Fvar s,Fvar p]} {Suc n \ [s,p]} t)" by simp } with sapa have "s \ FV (Fvar sa)" and "s \ FV (Fvar pa)" and "p \ FV (Fvar sa)" and "p \ FV (Fvar pa)" and "sa \ FV ({Suc n \ [Fvar s,Fvar p]} {Suc n \ [s,p]} t)" and "sa \ FV t" and "pa \ FV ({Suc n \ [Fvar s,Fvar p]} {Suc n \ [s,p]} t)" and "pa \ FV t" and "sa \ pa" by auto (* proof: {Suc n \ [Fvar s,Fvar p]} {Suc n \ [s,p]} (t\<^bsup>[Fvar sa,Fvar pa]\<^esup>) = {Suc n \ [Fvar s,Fvar p]} {Suc n \ [s,p]} {0 \ [Fvar sa,Fvar pa]} t = {Suc n \ [Fvar s,Fvar p]} {0 \ [Fvar sa,Fvar pa]} {Suc n \ [s,p]} t = {0 \ [Fvar sa,Fvar pa]} {Suc n \ [Fvar s,Fvar p]} {Suc n \ [s,p]} t = {0 \ [Fvar sa,Fvar pa]} t = (t\<^bsup>[Fvar sa,Fvar pa]\<^esup>) *) from eq sym[OF sopen_sclose_commute[OF not_sym[OF Suc_not_Zero[of n]] this(1-4)]] sopen_commute[OF Suc_not_Zero[of n]] sopen_fresh_inj[OF _ this(5-9)] show "{Suc n \ [Fvar s,Fvar p]} {Suc n \ [s,p]} t = t" by (auto simp: openz_def) qed qed lemma ssubst_sopen_distrib[rule_format]: "\n s p t'. lc t' \ [x \ t'] {n \ [s,p]} t = {n \ [[x \ t']s, [x \ t']p]} [x \ t'] t" proof - have "(\n s p t'. lc t' \ [x \ t'] {n \ [s,p]} t = {n \ [[x \ t']s, [x \ t']p]} [x \ t'] t) &(\n s p t'. lc t' \ ssubst_option x t' (sopen_option n s p u) = sopen_option n ([x \ t']s) ([x \ t']p) (ssubst_option x t' u))" by (rule compat_sterm_sterm_option.induct, simp_all split: bVariable.split) from conjunct1[OF this] show ?thesis by assumption qed lemma ssubst_openz_distrib: "lc t' \ [x \ t'] (t\<^bsup>[s,p]\<^esup>) = (([x \ t'] t)\<^bsup>[[x \ t'] s, [x \ t'] p]\<^esup>)" by (simp add: openz_def ssubst_sopen_distrib) lemma ssubst_sopen_commute: "\ lc t'; x \ FV s; x \ FV p \ \ [x \ t'] {n \ [s,p]} t = {n \ [s,p]} [x \ t'] t" by (frule ssubst_sopen_distrib[of t' x n s p t], simp) lemma sopen_commute_gen: fixes s p s' p' n k t assumes "lc s" and "lc p" and "lc s'" and "lc p'" and "n \ k" shows "{n \ [s,p]} {k \ [s',p']} t = {k \ [s',p']} {n \ [s,p]} t" proof - have "finite (FV s \ FV p \ FV s' \ FV p' \ FV t)" by auto from exFresh_s_p_cof[OF this] obtain sa pa where "sa \ FV s \ FV p \ FV s' \ FV p' \ FV t \ pa \ FV s \ FV p \ FV s' \ FV p' \ FV t \ sa \ pa" by auto moreover hence "finite (FV s \ FV p \ FV s' \ FV p' \ FV t \ {sa} \ {pa})" by auto from exFresh_s_p_cof[OF this] obtain sb pb where "sb \ FV s \ FV p \ FV s' \ FV p' \ FV t \ {sa} \ {pa} \ pb \ FV s \ FV p \ FV s' \ FV p' \ FV t \ {sa} \ {pa} \ sb \ pb" by auto ultimately have "sa \ FV t" and "pa \ FV t" and "sb \ FV t" and "pb \ FV t" and "sa \ FV ({n \ [s,p]} t)" and "pa \ FV ({n \ [s,p]} t)" and "sb \ FV ({k \ [s',p']} t)" and "pb \ FV ({k \ [s',p']} t)" and "sa \ pa" and "sb \ pb" and "sb \ sa" and "sb \ pa" and "pb \ sa" and "pb \ pa" and "sa \ FV s" and "sa \ FV p" and "pa \ FV s" and "pa \ FV p" and "sb \ FV s'" and "sb \ FV p'" and "pb \ FV s'" and "pb \ FV p'" and "sa \ FV p'" and "sb \ FV p" and "sa \ FV (Fvar sb)" and "sa \ FV (Fvar pb)" and "pa \ FV (Fvar sb)" and "pa \ FV (Fvar pb)" and "pb \ FV (Fvar sa)" and "pb \ FV (Fvar pa)" and "sb \ FV (Fvar sa)" and "sb \ FV (Fvar pa)" and "lc s" and "lc p" and "lc s'" and "lc p'" using contra_subsetD[OF sopen_FV] assms(1-4) by auto (* proof: {n \ [s,p]} {k \ [s',p']} t = {n \ [s,p]} [sa \ s'] [pa \ p'] {k \ [Fvar sa,Fvar pa]} t = [sa \ s'] [pa \ p'] [pa \ p'] {k \ [Fvar sa,Fvar pa]} t = [sb \ s] [pb \ p] {n \ [Fvar sb,Fvar pb]} [sa \ s'] [pa \ p'] {k \ [Fvar sa,Fvar pa]} t = [sb \ s] [pb \ p] [sa \ s'] {n \ [Fvar sb,Fvar pb]} [pa \ p'] {k \ [Fvar sa,Fvar pa]} t = [sb \ s] [pb \ p] [sa \ s'] [pa \ p'] {n \ [Fvar sb,Fvar pb]} {k \ [Fvar sa,Fvar pa]} t = [sb \ s] [pb \ p] [sa \ s'] [pa \ p'] {k \ [Fvar sa,Fvar pa]} {n \ [Fvar sb,Fvar pb]} t = [sb \ s] [sa \ s'] [pb \ p] [pa \ p'] {k \ [Fvar sa,Fvar pa]} {n \ [Fvar sb,Fvar pb]} t = [sa \ s'] [sb \ s] [pb \ p] [pa \ p'] {k \ [Fvar sa,Fvar pa]} {n \ [Fvar sb,Fvar pb]} t = [sa \ s'] [sb \ s] [pa \ p'] [pb \ p] {k \ [Fvar sa,Fvar pa]} {n \ [Fvar sb,Fvar pb]} t = [sa \ s'] [pa \ p'] [sb \ s] [pb \ p] {k \ [Fvar sa,Fvar pa]} {n \ [Fvar sb,Fvar pb]} t = [sa \ s'] [pa \ p'] [sb \ s] {k \ [Fvar sa,Fvar pa]} [pb \ p] {n \ [Fvar sb,Fvar pb]} t = [sa \ s'] [pa \ p'] {k \ [Fvar sa,Fvar pa]} [sb \ s] [pb \ p] {n \ [Fvar sb,Fvar pb]} t = [sa \ s'] [pa \ p'] {k \ [Fvar sa,Fvar pa]} {n \ [s,p]} t = {k \ [s',p']} {n \ [s,p]} t *) from ssubst_intro[OF \sa \ FV t\ \pa \ FV t\ \sa \ pa\ \sa \ FV p'\] ssubst_intro[OF \sb \ FV ({k \ [s',p']} t)\ \pb \ FV ({k \ [s',p']} t)\ \sb \ pb\ \sb \ FV p\] sym[OF ssubst_sopen_commute[OF \lc s'\ \sa \ FV (Fvar sb)\ \sa \ FV (Fvar pb)\]] sym[OF ssubst_sopen_commute[OF \lc p'\ \pa \ FV (Fvar sb)\ \pa \ FV (Fvar pb)\]] sopen_commute[OF \n \ k\] ssubst_commute[OF \pb \ sa\ \pb \ FV s'\ \sa \ FV p\] ssubst_commute[OF \sb \ sa\ \sb \ FV s'\ \sa \ FV s\] ssubst_commute[OF \pb \ pa\ \pb \ FV p'\ \pa \ FV p\] ssubst_commute[OF \sb \ pa\ \sb \ FV p'\ \pa \ FV s\] ssubst_sopen_commute[OF \lc s\ \sb \ FV (Fvar sa)\ \sb \ FV (Fvar pa)\] ssubst_sopen_commute[OF \lc p\ \pb \ FV (Fvar sa)\ \pb \ FV (Fvar pa)\] sym[OF ssubst_intro[OF \sb \ FV t\ \pb \ FV t\ \sb \ pb\ \sb \ FV p\]] sym[OF ssubst_intro[OF \sa \ FV ({n \ [s,p]} t)\ \pa \ FV ({n \ [s,p]} t)\ \sa \ pa\ \sa \ FV p'\]] show "{n \ [s,p]} {k \ [s',p']} t = {k \ [s',p']} {n \ [s,p]} t" by force qed lemma ssubst_preserves_lc[simp, rule_format]: fixes t assumes "lc t" shows "\x t'. lc t' \ lc ([x \ t'] t)" proof - define pred_cof where "pred_cof L t \ (\s p. s \ L \ p \ L \ s \ p \ lc (t\<^bsup>[Fvar s, Fvar p]\<^esup>))" for L t { fix x v t assume "lc v" and "\x v. lc v \ (\L. finite L \ pred_cof L ([x \ v] t))" hence "\L. finite L \ pred_cof L ([x \ v] t)" by auto }note Lex = this from assms show ?thesis proof (induct taking: "\t. \x t'. lc t' \ (\L. finite L \ pred_cof L ([x \ t'] t))" rule: lc_induct) case Fvar thus ?case by simp next case Call thus ?case by simp next case (Upd t l u) note pred_t = this(2) and pred_bnd = this(3) show ?case proof (intro strip) fix x t' assume "lc t'" note Lex[OF this pred_bnd] from this[of x] obtain L where "finite L" and "pred_cof L ([x \ t'] u)" by auto with \lc t'\ pred_t show "lc ([x \ t'] Upd t l u)" unfolding pred_cof_def by simp qed next case (Obj f T) note pred = this show ?case proof (intro strip) fix x t' assume "lc t'" define pred_fl where "pred_fl s p b l = lc ([x \ t'] the b\<^bsup>[Fvar s, Fvar p]\<^esup>)" for s p b and l::Label from \lc t'\ fmap_ball_all2[OF pred] have "\l\dom f. \L. finite L \ pred_cof L ([x \ t'] the(f l))" unfolding pred_cof_def by simp with fmap_ex_cof[of f pred_fl] obtain L where "finite L" and "\l\dom f. pred_cof L ([x \ t'] the(f l))" unfolding pred_cof_def pred_fl_def by auto with pred_ssubstoption_lem[of x t' f "pred_cof L"] show "lc ([x \ t'] Obj f T)" unfolding pred_cof_def by simp qed next case (Bnd L t) note pred = this(2) show ?case proof (intro strip) fix x t' assume "lc t'" with \finite L\ show "\L. finite L \ pred_cof L ([x \ t'] t)" unfolding pred_cof_def proof ( rule_tac x = "L \ {x}" in exI, intro conjI, simp, intro strip) fix s p assume sp: "s \ L \ {x} \ p \ L \ {x} \ s \ p" hence "x \ FV (Fvar s)" and "x \ FV (Fvar p)" by auto from sp pred \lc t'\ have "lc ([x \ t'] (t\<^bsup>[Fvar s,Fvar p]\<^esup>))" by blast with ssubst_sopen_commute[OF \lc t'\ \x \ FV (Fvar s)\ \x \ FV (Fvar p)\] show "lc ([x \ t'] t\<^bsup>[Fvar s,Fvar p]\<^esup>)" by (auto simp: openz_def) qed qed qed qed lemma sopen_sclose_eq_ssubst: "\ sa \ pa; sa \ FV p; lc t \ \ {n \ [s,p]} {n \ [sa,pa]} t = [sa \ s] [pa \ p] t" by (rule_tac sa1 = sa and pa1 = pa and t1 = "{n \ [sa,pa]} t" in ssubst[OF ssubst_intro], simp+) lemma ssubst_sclose_commute[rule_format]: "\x n s p t'. s \ FV t' \ p \ FV t' \ x \ s \ x \ p \ [x \ t'] {n \ [s,p]} t = {n \ [s,p]} [x \ t'] t" proof - have "(\x n s p t'. s \ FV t' \ p \ FV t' \ x \ s \ x \ p \ [x \ t'] {n \ [s,p]} t = {n \ [s,p]} [x \ t'] t) &(\x n s p t'. s \ FV t' \ p \ FV t' \ x \ s \ x \ p \ ssubst_option x t' (sclose_option n s p u) = sclose_option n s p (ssubst_option x t' u))" by (rule compat_sterm_sterm_option.induct, simp_all split: bVariable.split) from conjunct1[OF this] show ?thesis by assumption qed lemma body_lc_FV: fixes t s p assumes "body t" shows "lc (t\<^bsup>[Fvar s, Fvar p]\<^esup>)" proof - from assms obtain L where "finite L" and pred_sp: "\s p. s \ L \ p \ L \ s \ p \ lc (t\<^bsup>[Fvar s,Fvar p]\<^esup>)" unfolding body_def by auto hence "finite (L \ FV t \ {s} \ {p})" by simp from exFresh_s_p_cof[OF this] obtain sa pa where sapa: "sa \ L \ FV t \ {s} \ {p} \ pa \ L \ FV t \ {s} \ {p} \ sa \ pa" by auto hence "sa \ FV t" and "pa \ FV t" and "sa \ pa" and "sa \ FV (Fvar p)" by auto from pred_sp sapa have "lc (t\<^bsup>[Fvar sa,Fvar pa]\<^esup>)" by blast with ssubst_intro[OF \sa \ FV t\ \pa \ FV t\ \sa \ pa\ \sa \ FV (Fvar p)\] ssubst_preserves_lc show "lc (t\<^bsup>[Fvar s,Fvar p]\<^esup>)" by (auto simp: openz_def) qed lemma body_lc: fixes t s p assumes "body t" and "lc s" and "lc p" shows "lc (t\<^bsup>[s, p]\<^esup>)" proof - have "finite (FV t \ FV p)" by simp from exFresh_s_p_cof[OF this] obtain sa pa where "sa \ FV t \ FV p \ pa \ FV t \ FV p \ sa \ pa" by auto hence "sa \ FV t" and "pa \ FV t" and "sa \ pa" and "sa \ FV p" by auto from body_lc_FV[OF \body t\] have lc: "lc (t\<^bsup>[Fvar sa,Fvar pa]\<^esup>)" by assumption from ssubst_intro[OF \sa \ FV t\ \pa \ FV t\ \sa \ pa\ \sa \ FV p\] ssubst_preserves_lc[OF lc] \lc s\ \lc p\ show "lc (t\<^bsup>[s,p]\<^esup>)" by (auto simp: openz_def) qed lemma lc_body: fixes t s p assumes "lc t" and "s \ p" shows "body (\[s,p] t)" unfolding body_def proof have "\sa pa. sa \ FV t \ {s} \ {p} \ pa \ FV t \ {s} \ {p} \ sa \ pa \ lc (\[s,p] t\<^bsup>[Fvar sa,Fvar pa]\<^esup>)" proof (intro strip) fix sa :: fVariable and pa :: fVariable assume "sa \ FV t \ {s} \ {p} \ pa \ FV t \ {s} \ {p} \ sa \ pa" hence "s \ FV (Fvar pa)" by auto from sopen_sclose_eq_ssubst[OF \s \ p\ this \lc t\] ssubst_preserves_lc[OF \lc t\] show "lc (\[s,p] t\<^bsup>[Fvar sa,Fvar pa]\<^esup>)" by (simp add: openz_def closez_def) qed thus "finite (FV t \ {s} \ {p}) \ (\sa pa. sa \ FV t \ {s} \ {p} \ pa \ FV t \ {s} \ {p} \ sa \ pa \ lc (\[s,p] t\<^bsup>[Fvar sa,Fvar pa]\<^esup>))" by simp qed lemma ssubst_preserves_lcE_lem[rule_format]: fixes t assumes "lc t" shows "\x u t'. t = [x \ u] t' \ lc u \ lc t'" using assms proof (induct taking: "\t. \x u t'. t = [x \ u] t' \ lc u \ body t'" rule: lc_induct) case Fvar thus ?case by (intro strip, case_tac t', simp_all) next case Call thus ?case by (intro strip, case_tac t', simp_all) next case (Upd t l u) note pred_t = this(2) and pred_u = this(3) show ?case proof (intro strip) fix x v t'' assume "Upd t l u = [x \ v] t''" and "lc v" from this(1) have t'': "(\t' u'. t'' = Upd t' l u') \ (t'' = Fvar x)" proof (cases t'', auto) fix y assume "Upd t l u = (if x = y then v else Fvar y)" thus "y = x" by (case_tac "y = x", auto) qed show "lc t''" proof (cases "t'' = Fvar x") case True thus ?thesis by simp next case False with \Upd t l u = [x \ v] t''\ t'' show ?thesis proof (clarify) fix t' u' assume "Upd t l u = [x \ v] Upd t' l u'" hence "t = [x \ v] t'" and "u = [x \ v] u'" by auto with \lc v\ pred_t pred_u lc_upd[of t' l u'] show "lc (Upd t' l u')" by auto qed qed qed next case (Obj f T) note pred = this show ?case proof (intro strip) fix x v t' assume "Obj f T = [x \ v] t'" and "lc v" from this(1) have t': "(\f'. t' = Obj f' T) \ (t' = Fvar x)" proof (cases t', auto) fix y :: fVariable assume "Obj f T = (if x = y then v else Fvar y)" thus "y = x" by (case_tac "y = x", auto) qed show "lc t'" proof (cases "t' = Fvar x") case True thus ?thesis by simp next case False with \Obj f T = [x \ v] t'\ t' show ?thesis proof (clarify) fix f' assume "Obj f T = [x \ v] Obj f' T" hence ssubst: "\l\dom f. the(f l) = [x \ v] the(f' l)" and "dom f = dom f'" by auto with pred \lc v\ lc_obj[of f' T] show "lc (Obj f' T)" by auto qed qed qed next case (Bnd L t) note pred = this(2) show ?case proof (intro strip) fix x v t' assume "t = [x \ v] t'" and "lc v" from \finite L\ exFresh_s_p_cof[of "L \ {x} \ FV t'"] obtain s p where "s \ L" and "p \ L" and "s \ p" and "x \ FV (Fvar s)" and "x \ FV (Fvar p)" and "s \ FV t'" and "p \ FV t'" by auto from \t = [x \ v] t'\ ssubst_sopen_commute[OF \lc v\ \x \ FV (Fvar s)\ \x \ FV (Fvar p)\] have "(t\<^bsup>[Fvar s, Fvar p]\<^esup>) = [x \ v] (t'\<^bsup>[Fvar s, Fvar p]\<^esup>)" by (auto simp: openz_def) with \s \ L\ \p \ L\ \s \ p\ \lc v\ pred have "lc (t'\<^bsup>[Fvar s, Fvar p]\<^esup>)" by blast from lc_body[OF this \s \ p\] sclose_sopen_eq_t[OF \s \ FV t'\ \p \ FV t'\ \s \ p\] show "body t'" by (auto simp: openz_def closez_def) qed qed lemma ssubst_preserves_lcE: "\ lc ([x \ t'] t); lc t' \ \ lc t" by (drule_tac t = "[x \ t'] t" and x = x and u = t' and t' = t in ssubst_preserves_lcE_lem, simp+) lemma obj_openz_lc: "\ lc (Obj f T); lc p; l \ dom f \ \ lc (the(f l)\<^bsup>[Obj f T, p]\<^esup>)" by (rule_tac s = "Obj f T" and p = p in body_lc, (simp add: lc_obj)+) lemma obj_insert_lc: fixes f T t l assumes "lc (Obj f T)" and "body t" shows "lc (Obj (f(l \ t)) T)" proof (rule ssubst[OF lc_obj], rule ballI) fix l' :: Label assume "l' \ dom (f(l \ t))" with assms show "body (the ((f(l \ t)) l'))" by (cases "l' = l", (auto simp: lc_obj)) qed lemma ssubst_preserves_body[simp]: fixes t t' x assumes "body t" and "lc t'" shows "body ([x \ t'] t)" unfolding body_def proof - have "\s p. s \ FV t' \ {x} \ p \ FV t' \ {x} \ s \ p \ lc ([x \ t'] t\<^bsup>[Fvar s,Fvar p]\<^esup>)" proof (intro strip) fix s :: fVariable and p :: fVariable from body_lc_FV[OF \body t\] have "lc ({0 \ [Fvar s,Fvar p]} t)" by (simp add: openz_def) from ssubst_preserves_lc[OF this \lc t'\] have "lc ([x \ t'] (t\<^bsup>[Fvar s,Fvar p]\<^esup>))" by (simp add: openz_def) moreover assume "s \ FV t' \ {x} \ p \ FV t' \ {x} \ s \ p" hence "x \ FV (Fvar s)" and "x \ FV (Fvar p)" by auto note ssubst_sopen_commute[OF \lc t'\ this] ultimately show "lc ([x \ t'] t\<^bsup>[Fvar s,Fvar p]\<^esup>)" by (simp add: openz_def) qed thus "\L. finite L \ (\s p. s \ L \ p \ L \ s \ p \ lc ([x \ t'] t\<^bsup>[Fvar s,Fvar p]\<^esup>))" by (rule_tac x = "FV t' \ {x}" in exI, simp) qed lemma sopen_preserves_body[simp]: fixes t s p assumes "body t" and "lc s" and "lc p" shows "body ({n \ [s,p]} t)" unfolding body_def proof - have "\sa pa. sa \ FV t \ FV s \ pa \ FV p \ sa \ pa \ lc ({n \ [s,p]} t\<^bsup>[Fvar sa,Fvar pa]\<^esup>)" proof (cases "n = 0") case True thus ?thesis using body_lc[OF \body t\ \lc s\ \lc p\] sopen_twice[OF \lc s\ \lc p\] by (simp add: openz_def) next case False thus ?thesis proof (intro strip) fix sa :: fVariable and pa :: fVariable from body_lc_FV[OF \body t\] have "lc (t\<^bsup>[Fvar sa,Fvar pa]\<^esup>)" by assumption moreover from sopen_commute_gen[OF _ _ \lc s\ \lc p\ not_sym[OF \n \ 0\]] have "{n \ [s,p]} t\<^bsup>[Fvar sa,Fvar pa]\<^esup> = {n \ [s,p]} (t\<^bsup>[Fvar sa,Fvar pa]\<^esup>)" by (simp add: openz_def) ultimately show "lc ({n \ [s,p]} t\<^bsup>[Fvar sa,Fvar pa]\<^esup>)" by simp qed qed thus "\L. finite L \ (\sa pa. sa \ L \ pa \ L \ sa \ pa \ lc ({n \ [s,p]} t\<^bsup>[Fvar sa,Fvar pa]\<^esup>))" by (rule_tac x = "FV t \ FV s \ FV p" in exI, simp) qed subsection \Beta-reduction\ inductive beta :: "[sterm, sterm] \ bool" (infixl "\\<^sub>\" 50) where beta[simp, intro!] : "\ l \ dom f; lc (Obj f T); lc a \ \ Call (Obj f T) l a \\<^sub>\ (the (f l)\<^bsup>[(Obj f T), a]\<^esup>)" | beta_Upd[simp, intro!] : "\ l \ dom f; lc (Obj f T); body t \ \ Upd (Obj f T) l t \\<^sub>\ Obj (f(l \ t)) T" | beta_CallL[simp, intro!]: "\ t \\<^sub>\ t'; lc u \ \ Call t l u \\<^sub>\ Call t' l u" | beta_CallR[simp, intro!]: "\ t \\<^sub>\ t'; lc u \ \ Call u l t \\<^sub>\ Call u l t'" | beta_UpdL[simp, intro!] : "\ t \\<^sub>\ t'; body u \ \ Upd t l u \\<^sub>\ Upd t' l u" | beta_UpdR[simp, intro!] : "\ finite L; \s p. s \ L \ p \ L \ s \ p \ (\t''. t\<^bsup>[Fvar s,Fvar p]\<^esup> \\<^sub>\ t''\ t'= \[s,p]t''); lc u \ \ Upd u l t \\<^sub>\ Upd u l t'" | beta_Obj[simp, intro!] : "\ l \ dom f; finite L; \s p. s \ L \ p \ L \ s \ p \ (\t''. t\<^bsup>[Fvar s,Fvar p]\<^esup> \\<^sub>\ t'' \ t'= \[s,p]t''); \l\dom f. body (the (f l)) \ \ Obj (f(l \ t)) T \\<^sub>\ Obj (f(l \ t')) T" inductive_cases beta_cases [elim!]: "Call s l t \\<^sub>\ u" "Upd s l t \\<^sub>\ u" "Obj s T \\<^sub>\ t" abbreviation beta_reds :: "[sterm, sterm] => bool" (infixl "->>" 50) where "s ->> t == beta^** s t" abbreviation beta_ascii :: "[sterm, sterm] => bool" (infixl "->" 50) where "s -> t == beta s t" notation (latex) beta_reds (infixl "\\<^sub>\\<^sup>*" 50) lemma beta_induct[consumes 1, case_names CallL CallR UpdL UpdR Upd Obj beta Bnd]: fixes t :: sterm and t' :: sterm and P1 :: "sterm \ sterm \ bool" and P2 :: "sterm \ sterm \ bool" assumes "t \\<^sub>\ t'" and "\t t' u l. \ t \\<^sub>\ t'; P1 t t'; lc u \ \ P1 (Call t l u) (Call t' l u)" and "\t t' u l. \ t \\<^sub>\ t'; P1 t t'; lc u \ \ P1 (Call u l t) (Call u l t')" and "\t t' u l. \ t \\<^sub>\ t'; P1 t t'; body u\ \ P1 (Upd t l u) (Upd t' l u)" and "\t t' u l. \ P2 t t'; lc u \ \ P1 (Upd u l t) (Upd u l t')" and "\l f T t. \ l \ dom f; lc (Obj f T); body t \ \ P1 (Upd (Obj f T) l t) (Obj (f(l \ t)) T)" and "\l f t t' T. \ l \ dom f; P2 t t'; \l\dom f. body (the (f l)) \ \ P1 (Obj (f(l \ t)) T) (Obj (f(l \ t')) T)" and "\l f T a. \ l \ dom f; lc (Obj f T); lc a \ \ P1 (Call (Obj f T) l a) (the (f l)\<^bsup>[Obj f T,a]\<^esup>)" and "\L t t'. \ finite L; \s p. s \ L \ p \ L \ s \ p \ (\t''. t\<^bsup>[Fvar s,Fvar p]\<^esup> \\<^sub>\ t'' \ P1 (t\<^bsup>[Fvar s,Fvar p]\<^esup>) t'' \ t' = \[s,p] t'') \ \ P2 t t'" shows "P1 t t'" using assms by (induct rule: beta.induct, auto) lemma Fvar_beta: "Fvar x \\<^sub>\ t \ False" by (erule beta.cases, auto) lemma Obj_beta: assumes "Obj f T \\<^sub>\ z" shows "\l f' t t'. dom f = dom f' \ f = (f'(l \ t)) \ l \ dom f' \ (\L. finite L \ (\s p. s \ L \ p \ L \ s \ p \ (\t''. t\<^bsup>[Fvar s,Fvar p]\<^esup> \\<^sub>\ t'' \ t'= \[s,p]t''))) \ z = Obj (f'(l \ t')) T" proof (cases rule: beta_cases(3)[OF assms]) case (1 l fa L t t') thus ?thesis by (rule_tac x = l in exI, rule_tac x = fa in exI, rule_tac x = t in exI, rule_tac x = t' in exI, auto) qed lemma Upd_beta: "Upd t l u \\<^sub>\ z \ (\t'. t \\<^sub>\ t' \ z = Upd t' l u) \(\u' L. finite L \ (\s p. s \ L \ p \ L \ s \ p \ (\t''. (u\<^bsup>[Fvar s, Fvar p]\<^esup>) \\<^sub>\ t'' \ u' = \[s,p]t'')) \ z = Upd t l u') \(\f T. l \ dom f \ Obj f T = t \ z = Obj (f(l \ u)) T)" by (erule beta_cases, auto) lemma Call_beta: "Call t l u \\<^sub>\ z \ (\t'. t \\<^sub>\ t' \ z = Call t' l u) \ (\u'. u \\<^sub>\ u' \ z = Call t l u') \(\f T. Obj f T = t \ l \ dom f \ z = (the (f l)\<^bsup>[Obj f T, u]\<^esup>))" by (erule beta_cases, auto) subsubsection \Properties\ lemma beta_lc[simp]: fixes t t' assumes "t \\<^sub>\ t'" shows "lc t \ lc t'" using assms proof (induct taking: "\t t'. body t \ body t'" rule: beta_induct) case CallL thus ?case by simp next case CallR thus ?case by simp next case UpdR thus ?case by (simp add: lc_upd) next case UpdL thus ?case by (simp add: lc_upd) next case beta thus ?case by (simp add: obj_openz_lc) next case Upd thus ?case by (simp add: lc_obj lc_upd) next case Obj thus ?case by (simp add: lc_obj) next case (Bnd L t t') note cof = this(2) from \finite L\ exFresh_s_p_cof[of "L \ FV t"] obtain s p where "s \ L" and "s \ FV t" and "p \ L" and "p \ FV t" and "s \ p" by auto with cof obtain t'' where "lc (t\<^bsup>[Fvar s, Fvar p]\<^esup>)" and "lc t''" and "t' = \[s,p] t''" by auto from lc_body[OF this(1) \s \ p\] sclose_sopen_eq_t[OF \s \ FV t\ \p \ FV t\ \s \ p\] this(3) lc_body[OF this(2) \s \ p\] show ?case by (simp add: openz_def closez_def) qed lemma beta_ssubst[rule_format]: fixes t t' assumes "t \\<^sub>\ t'" shows "\x v. lc v \ [x \ v] t \\<^sub>\ [x \ v] t'" proof - define pred_cof where "pred_cof L t t' \ (\s p. s \ L \ p \ L \ s \ p \ (\t''. t\<^bsup>[Fvar s, Fvar p]\<^esup> \\<^sub>\ t'' \ t' = \[s,p] t''))" for L t t' { fix x v t t' assume "lc v" and "\x v. lc v \ (\L. finite L \ pred_cof L ([x \ v] t) ([x \ v] t'))" hence "\L. finite L \ pred_cof L ([x \ v] t) ([x \ v] t')" by auto }note Lex = this { fix x v l and f :: "Label \ sterm option" assume "l \ dom f" hence "l \ dom (\l. ssubst_option x v (f l))" by simp }note domssubst = this { fix x v l T and f :: "Label \ sterm option" assume "lc (Obj f T)" and "lc v" from ssubst_preserves_lc[OF this] have obj: "lc (Obj (\l. ssubst_option x v (f l)) T)" by simp }note lcobj = this from assms show ?thesis proof (induct taking: "\t t'. \x v. lc v \ (\L. finite L \ pred_cof L ([x \ v] t) ([x \ v] t'))" rule: beta_induct) case CallL thus ?case by simp next case CallR thus ?case by simp next case UpdL thus ?case by simp next case (UpdR t t' u l) note pred = this(1) show ?case proof (intro strip) fix x v assume "lc v" from Lex[OF this pred] obtain L where "finite L" and "pred_cof L ([x \ v] t) ([x \ v] t')" by auto with ssubst_preserves_lc[OF \lc u\ \lc v\] show "[x \ v] Upd u l t \\<^sub>\ [x \ v] Upd u l t'" unfolding pred_cof_def by auto qed next case (beta l f T t) thus ?case proof (intro strip, simp) fix x v assume "lc v" from ssubst_preserves_lc[OF \lc t\ this] have "lc ([x \ v] t)" by simp note lem = beta.beta[OF domssubst[OF \l \ dom f\] lcobj[OF \lc (Obj f T)\ \lc v\] this] from \l \ dom f\ have "the (ssubst_option x v (f l)) = [x \ v] the (f l)" by auto with lem[of x] ssubst_openz_distrib[OF \lc v\] show "Call (Obj (\l. ssubst_option x v (f l)) T) l ([x \ v] t) \\<^sub>\ [x \ v] (the (f l)\<^bsup>[Obj f T, t]\<^esup>)" by simp qed next case (Upd l f T t) thus ?case proof (intro strip, simp) fix x v assume "lc v" from ssubst_preserves_body[OF \body t\ \lc v\] have "body ([x \ v] t)" by simp from beta.beta_Upd[OF domssubst[OF \l \ dom f\] lcobj[OF \lc (Obj f T)\ \lc v\] this] ssubstoption_insert[OF \l \ dom f\] show "Upd (Obj (\l. ssubst_option x v (f l)) T) l ([x \ v] t) \\<^sub>\ Obj (\la. ssubst_option x v (if la = l then Some t else f la)) T" by simp qed next case (Obj l f t t' T) note pred = this(2) show ?case proof (intro strip, simp) fix x v assume "lc v" note Lex[OF this pred] from this[of x] obtain L where "finite L" and "pred_cof L ([x \ v] t) ([x \ v] t')" by auto have "\l\dom (\l. ssubst_option x v (f l)). body (the (ssubst_option x v (f l)))" proof (intro strip, simp) fix l' :: Label assume "l' \ dom f" with \\l\dom f. body (the(f l))\ have "body (the (f l'))" by blast note ssubst_preserves_body[OF this \lc v\] with \l' \ dom f\ ssubst_option_lem show "body (the (ssubst_option x v (f l')))" by auto qed from beta.beta_Obj[OF domssubst[OF \l \ dom f\] \finite L\ _ this] ssubstoption_insert[OF \l \ dom f\] \pred_cof L ([x \ v] t) ([x \ v] t')\ show "Obj (\la. ssubst_option x v (if la = l then Some t else f la)) T \\<^sub>\ Obj (\la. ssubst_option x v (if la = l then Some t' else f la)) T" unfolding pred_cof_def by simp qed next case (Bnd L t t') note pred = this(2) show ?case proof (intro strip) fix x v assume "lc v" from \finite L\ show "\L. finite L \ pred_cof L ([x \ v] t) ([x \ v] t')" proof (rule_tac x = "L \ {x} \ FV v" in exI, unfold pred_cof_def, auto) fix s p assume "s \ L" and "p \ L" and "s \ p" with pred \lc v\ obtain t'' where "t\<^bsup>[Fvar s,Fvar p]\<^esup> \\<^sub>\ t''" and ssubst_beta: "[x \ v] (t\<^bsup>[Fvar s,Fvar p]\<^esup>) \\<^sub>\ [x \ v] t''" and "t' = \[s,p] t''" by blast assume "s \ x" and "p \ x" hence "x \ FV (Fvar s)" and "x \ FV (Fvar p)" by auto from ssubst_sopen_commute[OF \lc v\ this] ssubst_beta have "[x \ v] t\<^bsup>[Fvar s,Fvar p]\<^esup> \\<^sub>\ [x \ v] t''" by (simp add: openz_def) moreover assume "s \ FV v" and "p \ FV v" from ssubst_sclose_commute[OF this not_sym[OF \s \ x\] not_sym[OF \p \ x\]] \t' = \[s,p] t''\ have "[x \ v] t' = \[s,p] [x \ v] t''" by (simp add: closez_def) ultimately show "\t''. [x \ v] t\<^bsup>[Fvar s,Fvar p]\<^esup> \\<^sub>\ t'' \ [x \ v] t' = \[s,p] t''" by (rule_tac x = "[x \ v] t''" in exI, simp) qed qed qed qed declare if_not_P [simp] not_less_eq [simp] \ \don't add @{text "r_into_rtrancl[intro!]"}\ lemma beta_preserves_FV[simp, rule_format]: fixes t t' x assumes "t \\<^sub>\ t'" shows "x \ FV t \ x \ FV t'" using assms proof (induct taking: "\t t'. x \ FV t \ x \ FV t'" rule: beta_induct) case CallL thus ?case by simp next case CallR thus ?case by simp next case UpdL thus ?case by simp next case UpdR thus ?case by simp next case Upd thus ?case by simp next case Obj thus ?case by simp next case (beta l f T t) thus ?case proof (intro strip) assume "x \ FV (Call (Obj f T) l t)" with \l \ dom f\ have "x \ FV (the (f l)) \ FV (Obj f T) \ FV t" proof (auto) fix y :: sterm assume "x \ FV y" and "f l = Some y" hence "x \ FVoption (f l)" by auto moreover assume "\l\dom f. x \ FVoption (f l)" ultimately show False using \l \ dom f\ by blast qed from contra_subsetD[OF sopen_FV this] show "x \ FV (the (f l)\<^bsup>[Obj f T,t]\<^esup>)" by (simp add: openz_def) qed next case (Bnd L t t') thus ?case proof (intro strip) assume "x \ FV t" from \finite L\ exFresh_s_p_cof[of "L \ {x}"] obtain s p where sp: "s \ L \ {x} \ p \ L \ {x} \ s \ p" by auto with \x \ FV t\ sopen_FV[of 0 "Fvar s" "Fvar p" t] have "x \ FV (t\<^bsup>[Fvar s, Fvar p]\<^esup>)" by (auto simp: openz_def) with sp Bnd(2) obtain t'' where "x \ FV t''" and "t' = \[s,p] t''" by auto with sclose_subset_FV[of 0 s p t''] show "x \ FV t'" by (auto simp: closez_def) qed qed lemma rtrancl_beta_lc[simp, rule_format]: "t \\<^sub>\\<^sup>* t' \ t \ t' \ lc t \ lc t'" by (erule rtranclp.induct, simp, drule beta_lc, blast) lemma rtrancl_beta_lc2[simp]: "\ t \\<^sub>\\<^sup>* t'; lc t \ \ lc t'" by (case_tac "t = t'", simp+) lemma rtrancl_beta_body: fixes L t t' assumes "finite L" and "\s p. s \ L \ p \ L \ s \ p \ (\t''. t\<^bsup>[Fvar s,Fvar p]\<^esup> \\<^sub>\\<^sup>* t'' \ t' = \[s,p] t'')" and "body t" shows "body t'" proof (cases "t = t'") case True with assms(3) show ?thesis by simp next from exFresh_s_p_cof[OF \finite L\] obtain s p where sp: "s \ L \ p \ L \ s \ p" by auto hence "s \ p" by simp from assms(2) sp obtain t'' where "t\<^bsup>[Fvar s,Fvar p]\<^esup> \\<^sub>\\<^sup>* t''" and "t' = \[s,p] t''" by auto with \body t\ have "lc t''" proof (cases "(t\<^bsup>[Fvar s,Fvar p]\<^esup>) = t''") case True with body_lc[OF \body t\] show "lc t''" by auto next case False with rtrancl_beta_lc[OF \t\<^bsup>[Fvar s,Fvar p]\<^esup> \\<^sub>\\<^sup>* t''\] show "lc t''" by auto qed from lc_body[OF this \s \ p\] \t' = \[s,p] t''\ show "body t'" by simp qed lemma rtrancl_beta_preserves_FV[simp, rule_format]: "t \\<^sub>\\<^sup>* t' \ x \ FV t \ x \ FV t'" proof (induct t t' rule: rtranclp.induct, simp) case (rtrancl_into_rtrancl a b c) thus ?case proof (clarify) assume "x \ FV b" and "x \ FV c" from beta_preserves_FV[OF \b \\<^sub>\ c\ this(1)] this(2) show False by simp qed qed subsubsection \Congruence rules\ lemma rtrancl_beta_CallL [intro!, rule_format]: "\ t \\<^sub>\\<^sup>* t'; lc u \ \ Call t l u \\<^sub>\\<^sup>* Call t' l u" proof (induct t t' rule: rtranclp.induct, simp) case (rtrancl_into_rtrancl a b c) thus ?case proof (auto) from \b \\<^sub>\ c\ \lc u\ have "Call b l u \\<^sub>\ Call c l u" by simp with rtrancl_into_rtrancl(2)[OF \lc u\] show "Call a l u \\<^sub>\\<^sup>* Call c l u" by auto qed qed lemma rtrancl_beta_CallR [intro!, rule_format]: "\ t \\<^sub>\\<^sup>* t'; lc u \ \ Call u l t \\<^sub>\\<^sup>* Call u l t'" proof (induct t t' rule: rtranclp.induct, simp) case (rtrancl_into_rtrancl a b c) thus ?case proof (auto) from \b \\<^sub>\ c\ \lc u\ have "Call u l b \\<^sub>\ Call u l c" by simp with rtrancl_into_rtrancl(2)[OF \lc u\] show "Call u l a \\<^sub>\\<^sup>* Call u l c" by auto qed qed lemma rtrancl_beta_Call [intro!, rule_format]: "\ t \\<^sub>\\<^sup>* t'; lc t; u \\<^sub>\\<^sup>* u'; lc u \ \ Call t l u \\<^sub>\\<^sup>* Call t' l u'" proof (induct t t' rule: rtranclp.induct, blast) case (rtrancl_into_rtrancl a b c) thus ?case proof (auto) from \u \\<^sub>\\<^sup>* u'\ \lc u\ have "lc u'" by auto with \b \\<^sub>\ c\ have "Call b l u' \\<^sub>\ Call c l u'" by simp with rtrancl_into_rtrancl(2)[OF \lc a\ \u \\<^sub>\\<^sup>* u'\ \lc u\] show "Call a l u \\<^sub>\\<^sup>* Call c l u'" by auto qed qed lemma rtrancl_beta_UpdL: "\ t \\<^sub>\\<^sup>* t'; body u \ \ Upd t l u \\<^sub>\\<^sup>* Upd t' l u" proof (induct t t' rule: rtranclp.induct, simp) case (rtrancl_into_rtrancl a b c) thus ?case proof (auto) from \b \\<^sub>\ c\ \body u\ have "Upd b l u \\<^sub>\ Upd c l u" by simp with rtrancl_into_rtrancl(2)[OF \body u\] show "Upd a l u \\<^sub>\\<^sup>* Upd c l u" by auto qed qed lemma beta_binder[rule_format]: fixes t t' assumes "t \\<^sub>\ t'" shows "\L s p. finite L \ s \ L \ p \ L \ s \ p \ (\L'. finite L' \ (\sa pa. sa \ L' \ pa \ L' \ sa \ pa \ (\t''. (\[s,p] t)\<^bsup>[Fvar sa,Fvar pa]\<^esup> \\<^sub>\ t'' \ \[s,p] t' = \[sa,pa] t'')))" proof (intro strip) fix L :: "fVariable set" and s :: fVariable and p :: fVariable assume "s \ p" have "\sa pa. sa \ L \ FV t \ {s} \ {p} \ pa \ L \ FV t \ {s} \ {p} \ sa \ pa \ (\t''. (\[s,p] t)\<^bsup>[Fvar sa,Fvar pa]\<^esup> \\<^sub>\ t'' \ \[s,p] t' = \[sa,pa] t'')" proof (intro strip) fix sa :: fVariable and pa :: fVariable from beta_ssubst[OF \t \\<^sub>\ t'\] have "[p \ Fvar pa] t \\<^sub>\ [p \ Fvar pa] t'" by simp from beta_ssubst[OF this] have betasubst: "[s \ Fvar sa] [p \ Fvar pa] t \\<^sub>\ [s \ Fvar sa] [p \ Fvar pa] t'" by simp from beta_lc[OF \t \\<^sub>\ t'\] have "lc t" and "lc t'" by auto assume sapa: "sa \ L \ FV t \ {s} \ {p} \ pa \ L \ FV t \ {s} \ {p} \ sa \ pa" hence "s \ FV (Fvar pa)" by auto from sopen_sclose_eq_ssubst[OF \s \ p\ this \lc t\] sopen_sclose_eq_ssubst[OF \s \ p\ this \lc t'\] betasubst have "\[s,p] t\<^bsup>[Fvar sa, Fvar pa]\<^esup> \\<^sub>\ (\[s,p] t'\<^bsup>[Fvar sa, Fvar pa]\<^esup>)" by (simp add: openz_def closez_def) moreover { from sapa have "sa \ FV t" by simp from contra_subsetD[OF sclose_subset_FV beta_preserves_FV[OF \t \\<^sub>\ t'\ this]] have "sa \ FV (\[s,p] t')" by (simp add: closez_def) moreover from sapa have "pa \ FV t" by simp from contra_subsetD[OF sclose_subset_FV beta_preserves_FV[OF \t \\<^sub>\ t'\ this]] have "pa \ FV (\[s,p] t')" by (simp add: closez_def) ultimately have "sa \ FV (\[s,p] t')" and "pa \ FV (\[s,p] t')" and "sa \ pa" using sapa by auto note sym[OF sclose_sopen_eq_t[OF this]] } ultimately show "\t''. \[s,p] t\<^bsup>[Fvar sa,Fvar pa]\<^esup> \\<^sub>\ t'' \ \[s,p] t' = \[sa,pa] t''" by (auto simp: openz_def closez_def) qed moreover assume "finite L" ultimately show "\L'. finite L' \ (\sa pa. sa \ L' \ pa \ L' \ sa \ pa \ (\t''. \[s,p] t\<^bsup>[Fvar sa,Fvar pa]\<^esup> \\<^sub>\ t'' \ \[s,p] t' = \[sa,pa] t''))" by (rule_tac x = "L \ FV t \ {s} \ {p}" in exI, simp) qed lemma rtrancl_beta_UpdR: fixes L t t' u l assumes "\s p. s \ L \ p \ L \ s \ p \ (\t''. (t\<^bsup>[Fvar s, Fvar p]\<^esup>) \\<^sub>\\<^sup>* t'' \ t' = \[s,p]t'')" and "finite L" and "lc u" shows "Upd u l t \\<^sub>\\<^sup>* Upd u l t'" proof - from \finite L\ have "finite (L \ FV t)" by simp from exFresh_s_p_cof[OF this] obtain s p where sp: "s \ L \ FV t \ p \ L \ FV t \ s \ p" by auto with assms(1) obtain t'' where "t\<^bsup>[Fvar s,Fvar p]\<^esup> \\<^sub>\\<^sup>* t''" and t': "t' = \[s,p] t''" by auto with \lc u\ have "Upd u l t \\<^sub>\\<^sup>* Upd u l \[s,p] t''" proof (erule_tac rtranclp_induct) from sp have "s \ FV t" and "p \ FV t" and "s \ p" by auto from sclose_sopen_eq_t[OF this] show "Upd u l t \\<^sub>\\<^sup>* Upd u l (\[s,p](t\<^bsup>[Fvar s,Fvar p]\<^esup>))" by (simp add: openz_def closez_def) next fix y :: sterm and z :: sterm assume "y \\<^sub>\ z" from sp have "s \ L" and "p \ L" and "s \ p" by auto from beta_binder[OF \y \\<^sub>\ z\ \finite L\ this] obtain L' where "finite L'" and "\sa pa. sa \ L' \ pa \ L' \ sa \ pa \ (\t''. \[s,p] y\<^bsup>[Fvar sa,Fvar pa]\<^esup> \\<^sub>\ t'' \ \[s,p] z = \[sa,pa] t'')" by auto from beta.beta_UpdR[OF this \lc u\] have "Upd u l (\[s,p] y) \\<^sub>\ Upd u l (\[s,p] z)" by assumption moreover assume "Upd u l t \\<^sub>\\<^sup>* Upd u l (\[s,p] y)" ultimately show "Upd u l t \\<^sub>\\<^sup>* Upd u l (\[s,p] z)" by simp qed with t' show "Upd u l t \\<^sub>\\<^sup>* Upd u l t'" by simp qed lemma rtrancl_beta_Upd: "\ u \\<^sub>\\<^sup>* u'; finite L; \s p. s \ L \ p \ L \ s \ p \ (\t''. t\<^bsup>[Fvar s,Fvar p]\<^esup> \\<^sub>\\<^sup>* t'' \ t' = \[s,p]t''); lc u; body t \ \ Upd u l t \\<^sub>\\<^sup>* Upd u' l t'" proof (induct u u' rule: rtranclp.induct) case rtrancl_refl thus ?case by (simp add: rtrancl_beta_UpdR) next case (rtrancl_into_rtrancl a b c) thus ?case proof (auto) from rtrancl_beta_body[OF \finite L\ rtrancl_into_rtrancl(5) \body t\] \b \\<^sub>\ c\ have "Upd b l t' \\<^sub>\ Upd c l t'" by simp with rtrancl_into_rtrancl(2)[OF \finite L\ rtrancl_into_rtrancl(5) \lc a\ \body t\] show "Upd a l t \\<^sub>\\<^sup>* Upd c l t'" by simp qed qed lemma rtrancl_beta_obj: fixes l f L T t t' assumes "l \ dom f" and "finite L" and "\s p. s \ L \ p \ L \ s \ p \ (\t''. t\<^bsup>[Fvar s,Fvar p]\<^esup> \\<^sub>\\<^sup>* t'' \ t' = \[s,p]t'')" and "\l\dom f. body (the(f l))" and "body t" shows "Obj (f (l \ t)) T \\<^sub>\\<^sup>* Obj (f (l \ t')) T" proof - from \finite L\ have "finite (L \ FV t)" by simp from exFresh_s_p_cof[OF this] obtain s p where sp: "s \ L \ FV t \ p \ L \ FV t \ s \ p" by auto with assms(3) obtain t'' where "t\<^bsup>[Fvar s,Fvar p]\<^esup> \\<^sub>\\<^sup>* t''" and "t' = \[s,p] t''" by auto with \l \ dom f\ \\l\dom f. body (the(f l))\ have "Obj (f(l \ t)) T \\<^sub>\\<^sup>* Obj (f(l \ \[s,p] t'')) T" proof (erule_tac rtranclp_induct) from sp have "s \ FV t" and "p \ FV t" and "s \ p" by auto from sclose_sopen_eq_t[OF this] show "Obj (f(l \ t)) T \\<^sub>\\<^sup>* Obj (f(l \ \[s,p] (t\<^bsup>[Fvar s,Fvar p]\<^esup>))) T" by (simp add: openz_def closez_def) next fix y :: sterm and z :: sterm assume "y \\<^sub>\ z" from sp have "s \ L" and "p \ L" and "s \ p" by auto from beta_binder[OF \y \\<^sub>\ z\ \finite L\ this] obtain L' where "finite L'" and "\sa pa. sa \ L' \ pa \ L' \ sa \ pa \ (\t''. \[s,p] y\<^bsup>[Fvar sa,Fvar pa]\<^esup> \\<^sub>\ t'' \ \[s,p] z = \[sa,pa] t'')" by auto from beta.beta_Obj[OF \l \ dom f\ this \\l\dom f. body (the(f l))\] have "Obj (f(l \ \[s,p] y)) T \\<^sub>\ Obj (f(l \ \[s,p] z)) T" by assumption moreover assume "Obj (f(l \ t)) T \\<^sub>\\<^sup>* Obj (f(l \ \[s,p] y)) T" ultimately show "Obj (f(l \ t)) T \\<^sub>\\<^sup>* Obj (f(l \ \[s,p] z)) T" by simp qed with \t' = \[s,p] t''\ show "Obj (f(l \ t)) T \\<^sub>\\<^sup>* Obj (f(l \ t')) T" by simp qed lemma obj_lem: fixes l f T L t' assumes "l \ dom f" and "finite L" and "\s p. s \ L \ p \ L \ s \ p \ (\t''. ((the(f l))\<^bsup>[Fvar s,Fvar p]\<^esup>) \\<^sub>\\<^sup>* t'' \ t' = \[s,p]t'')" and "\l\dom f. body (the(f l))" shows "Obj f T \\<^sub>\\<^sup>* Obj (f(l \ t')) T" proof (rule_tac P = "\y. Obj y T \\<^sub>\\<^sup>* Obj (f(l \ t')) T" and s = "(f(l \ the(f l)))" in subst) from \l \ dom f\ fun_upd_idem show "f(l \ the (f l)) = f" by force next from \l \ dom f\ \\l\dom f. body (the(f l))\ have "body (the (f l))" by blast with rtrancl_beta_obj[OF \l \ dom f\ \finite L\ assms(3) \\l\dom f. body (the(f l))\] show "Obj (f(l \ the (f l))) T \\<^sub>\\<^sup>* Obj (f(l \ t')) T" by simp qed lemma rtrancl_beta_obj_lem00: fixes L f g assumes "finite L" and "\l\dom f. \s p. s \ L \ p \ L \ s \ p \ (\t''. ((the(f l))\<^bsup>[Fvar s, Fvar p]\<^esup>) \\<^sub>\\<^sup>* t'' \ the(g l) = \[s,p]t'')" and "dom f = dom g" and "\l\dom f. body (the (f l))" shows "\k \ (card (dom f)). (\ob. length ob = (k + 1) \ (\obi. obi \ set ob \ dom (fst(obi)) = dom f \ ((snd obi) \ dom f)) \ (fst (ob!0) = f) \ (card (snd (ob!k)) = k) \ (\i < k. snd (ob!i) \ snd (ob!k)) \ (Obj (fst (ob!0)) T \\<^sub>\\<^sup>* Obj (fst (ob!k)) T) \ (card (snd (ob!k)) = k \ (Ltake_eq (snd (ob!k)) (fst (ob!k)) g) \ (Ltake_eq ((dom f) - (snd (ob!k))) (fst (ob!k)) f)))" proof fix k :: nat show "k \ card (dom f) \ (\ob. length ob = k + 1 \ (\obi. obi \ set ob \ dom (fst obi) = dom f \ snd obi \ dom f) \ fst (ob ! 0) = f \ card (snd (ob ! k)) = k \ (\i snd (ob ! k)) \ Obj (fst (ob ! 0)) T \\<^sub>\\<^sup>* Obj (fst (ob ! k)) T \ (card (snd (ob ! k)) = k \ Ltake_eq (snd (ob ! k)) (fst (ob ! k)) g \ Ltake_eq (dom f - snd (ob ! k)) (fst (ob ! k)) f))" proof (induct k) case 0 thus ?case by (simp, rule_tac x = "[(f,{})]" in exI, simp add: Ltake_eq_def) next case (Suc k) thus ?case proof (clarify) assume "Suc k \ card (dom f)" hence "k < card (dom f)" by arith with Suc.hyps obtain ob where "length ob = k + 1" and mem_ob: "\obi. obi \ set ob \ dom (fst obi) = dom f \ snd obi \ dom f" and "fst (ob ! 0) = f" and "card (snd (ob ! k)) = k" and "\i snd (ob ! k)" and "Obj (fst (ob ! 0)) T \\<^sub>\\<^sup>* Obj (fst (ob ! k)) T" and card_k: "card (snd (ob ! k)) = k \ Ltake_eq (snd (ob ! k)) (fst (ob ! k)) g \ Ltake_eq (dom f - snd (ob ! k)) (fst (ob ! k)) f" by auto from \length ob = k + 1\ have obkmem: "(ob!k) \ set ob" by auto with mem_ob have obksnd: "snd(ob!k) \ dom f" by blast from card_psubset[OF finite_dom_fmap this] \card (snd(ob!k)) = k\ \k < card (dom f)\ have "snd (ob!k) \ dom f" by simp then obtain l' where "l' \ dom f" and "l' \ snd (ob!k)" by auto from obkmem mem_ob have obkfst: "dom (fst(ob!k)) = dom f" by blast (* get witness *) - define ob' where "ob' = ob @ [(fst(ob!k)(l' \ the (g l')), insert l' (snd(ob!k)))]" + define ob' where "ob' = ob @ [((fst(ob!k))(l' \ the (g l')), insert l' (snd(ob!k)))]" from nth_fst[OF \length ob = k + 1\] have first: "ob'!0 = ob!0" by (simp add: ob'_def) from \length ob = k + 1\ nth_last[of ob "Suc k"] - have last: "ob'!Suc k = (fst(ob!k)(l' \ the (g l')), insert l' (snd(ob!k)))" + have last: "ob'!Suc k = ((fst(ob!k))(l' \ the (g l')), insert l' (snd(ob!k)))" by (simp add: ob'_def) from \length ob = k + 1\ nth_append[of ob _ k] have kth: "ob'!k = ob!k" by (auto simp: ob'_def) from \card (snd(ob!k)) = k\ card_k have ass: "\l\(snd(ob!k)). fst(ob!k) l = g l" "\l\(dom f - snd(ob!k)). fst(ob!k) l = f l" by (auto simp: Ltake_eq_def) (* prop#1 *) from \length ob = k + 1\ have "length ob' = Suc k + 1" by (auto simp: ob'_def) (* prop#2 *) moreover have "\obi. obi \ set ob' \ dom (fst obi) = dom f \ snd obi \ dom f" unfolding ob'_def proof (intro strip) fix obi :: "(Label -~> sterm) \ (Label set)" - assume "obi \ set (ob @ [(fst(ob!k)(l' \ the (g l')), insert l' (snd (ob!k)))])" + assume "obi \ set (ob @ [((fst(ob!k))(l' \ the (g l')), insert l' (snd (ob!k)))])" note mem_append_lem'[OF this] thus "dom (fst obi) = dom f \ snd obi \ dom f" proof (rule disjE, simp_all) assume "obi \ set ob" with mem_ob show "dom (fst obi) = dom f \ snd obi \ dom f" by blast next from obkfst obksnd \l' \ dom f\ show "insert l' (dom (fst (ob!k))) = dom f \ l' \ dom f \ snd(ob!k) \ dom f" by blast qed qed (* prop#3 *) moreover from first \fst(ob!0) = f\ have "fst(ob'!0) = f" by simp (* prop#4 *) moreover from obksnd finite_dom_fmap finite_subset have "finite (snd (ob!k))" by auto from card.insert_remove[OF this] have "card (insert l' (snd (ob!k))) = Suc (card (snd(ob!k) - {l'}))" by simp with \l' \ snd (ob!k)\ \card (snd(ob!k)) = k\ last have "card(snd(ob'!Suc k)) = Suc k" by auto (* prop#5 *) moreover have "\i snd (ob'!Suc k)" proof (intro strip) fix i :: nat from last have "snd(ob'!Suc k) = insert l' (snd (ob!k))" by simp with \l' \ snd (ob!k)\ have "snd(ob!k) \ snd(ob'!Suc k)" by auto moreover assume "i < Suc k" with \length ob = k + 1\ have "i < length ob" by simp with nth_append[of ob _ i] have "ob'!i = ob!i" by (simp add: ob'_def) ultimately show "snd(ob'!i) \ snd(ob'!Suc k)" proof (cases "i < k") case True with \\i snd(ob!k)\ \ob'!i = ob!i\ \snd(ob!k) \ snd(ob'!Suc k)\ show "snd (ob'!i) \ snd (ob'!Suc k)" by auto next case False with \i < Suc k\ have "i = k" by arith with \ob'!i = ob!i\ \snd(ob!k) \ snd(ob'!Suc k)\ show "snd (ob'!i) \ snd (ob'!Suc k)" by auto qed qed (* prop#6 -- the main statement *) moreover { from \l' \ dom f\ \l' \ snd(ob!k)\ have "l' \ (dom f - snd(ob!k))" by auto with ass have "the(fst(ob!k) l') = the(f l')" by auto with \l' \ dom f\ assms(2) have sp: "\s p. s \ L \ p \ L \ s \ p \ (\t''. the(fst(ob!k) l')\<^bsup>[Fvar s,Fvar p]\<^esup> \\<^sub>\\<^sup>* t'' \ the (g l') = \[s,p] t'')" by simp moreover have "\l\dom (fst(ob!k)). body (the(fst(ob!k) l))" proof (intro strip) fix la :: Label assume "la \ dom (fst(ob!k))" with obkfst have inf: "la \ dom f" by auto with assms(4) have bodyf: "body (the(f la))" by auto show "body (the(fst(ob!k) la))" proof (cases "la \ snd(ob!k)") case False with inf have "la \ (dom f - snd(ob!k))" by auto with ass have "fst(ob!k) la = f la" by blast with bodyf show "body (the (fst(ob!k) la))" by auto next from exFresh_s_p_cof[OF \finite L\] obtain s p where "s \ L \ p \ L \ s \ p" by auto with assms(2) inf obtain t' where "the (f la)\<^bsup>[Fvar s,Fvar p]\<^esup> \\<^sub>\\<^sup>* t'" and "the (g la) = \[s,p] t'" by blast from body_lc[OF bodyf] have lcf: "lc (the (f la)\<^bsup>[Fvar s,Fvar p]\<^esup>)" by auto hence bodyg: "body (the(g la))" proof (cases "(the (f la)\<^bsup>[Fvar s,Fvar p]\<^esup>) = t'") case True with lcf lc_body \s \ L \ p \ L \ s \ p\ \the(g la) = \[s,p] t'\ show "body (the(g la))" by auto next case False with rtrancl_beta_lc[OF \the (f la)\<^bsup>[Fvar s,Fvar p]\<^esup> \\<^sub>\\<^sup>* t'\] lc_body \s \ L \ p \ L \ s \ p\ \the(g la) = \[s,p] t'\ show "body (the(g la))" by auto qed case True with ass bodyg show "body (the(fst(ob!k) la))" by simp qed qed moreover from \l' \ dom f\ obkfst have "l' \ dom(fst(ob!k))" by auto note obj_lem[OF this \finite L\] ultimately - have "Obj (fst(ob!k)) T \\<^sub>\\<^sup>* Obj (fst(ob!k)(l' \ the (g l'))) T" + have "Obj (fst(ob!k)) T \\<^sub>\\<^sup>* Obj ((fst(ob!k))(l' \ the (g l'))) T" by blast moreover - from last have "fst(ob'!Suc k) = fst(ob!k)(l' \ the (g l'))" + from last have "fst(ob'!Suc k) = (fst(ob!k))(l' \ the (g l'))" by auto ultimately have "Obj (fst(ob'!0)) T \\<^sub>\\<^sup>* Obj (fst(ob'!Suc k)) T" using rtranclp_trans[OF \Obj (fst (ob!0)) T \\<^sub>\\<^sup>* Obj (fst (ob!k)) T\] first kth by auto } (* prop#7 *) moreover from \l' \ dom f\ \dom f = dom g\ have "card (snd(ob'!Suc k)) = Suc k \ Ltake_eq (snd (ob'!Suc k)) (fst (ob'!Suc k)) g \ Ltake_eq (dom f - snd(ob'!Suc k)) (fst(ob'!Suc k)) f" by (auto simp: Ltake_eq_def last ass) ultimately show "\ob. length ob = Suc k + 1 \ (\obi. obi \ set ob \ dom (fst obi) = dom f \ snd obi \ dom f) \ fst (ob ! 0) = f \ card (snd (ob ! Suc k)) = Suc k \ (\i snd (ob ! Suc k)) \ Obj (fst (ob ! 0)) T \\<^sub>\\<^sup>* Obj (fst (ob ! Suc k)) T \ (card (snd (ob ! Suc k)) = Suc k \ Ltake_eq (snd (ob ! Suc k)) (fst (ob ! Suc k)) g \ Ltake_eq (dom f - snd (ob ! Suc k)) (fst (ob ! Suc k)) f)" by (rule_tac x = ob' in exI, simp) qed qed qed lemma rtrancl_beta_obj_n: fixes f g L T assumes "finite L" and "\l\dom f. \s p. s \ L \ p \ L \ s \ p \ (\t''. ((the(f l))\<^bsup>[Fvar s, Fvar p]\<^esup>) \\<^sub>\\<^sup>* t'' \ the(g l) = \[s,p]t'')" and "dom f = dom g" and "\l\dom f. body (the(f l))" shows "Obj f T \\<^sub>\\<^sup>* Obj g T" proof (cases "f = Map.empty") case True with \dom f = dom g\ have "{} = dom g" by simp from \f = Map.empty\ empty_dom[OF this] show ?thesis by simp next from rtrancl_beta_obj_lem00[OF assms] obtain ob :: "((Label -~> sterm) \ (Label set)) list" where "length ob = card(dom f) + 1" and "\obi. obi \ set ob \ dom (fst obi) = dom f \ snd obi \ dom f" and "fst(ob!0) = f" and "card (snd(ob!card(dom f))) = card(dom f)" and "Obj (fst(ob!0)) T \\<^sub>\\<^sup>* Obj (fst(ob!card(dom f))) T" and "Ltake_eq (snd(ob!card(dom f))) (fst(ob!card(dom f))) g" by blast from \length ob = card (dom f) + 1\ have "(ob!card(dom f)) \ set ob" by auto with \\obi. obi \ set ob \ dom (fst obi) = dom f \ snd obi \ dom f\ have "dom (fst(ob!card(dom f))) = dom f" and "snd(ob!card(dom f)) \ dom f" by blast+ { fix l :: Label from \snd(ob!card(dom f)) \ dom f\ \card (snd(ob!card(dom f))) = card(dom f)\ Ltake_eq_dom have "snd(ob!card(dom f)) = dom f" by blast with \Ltake_eq (snd(ob!card (dom f))) (fst(ob!card (dom f))) g\ have "fst(ob!card(dom f)) l = g l" proof (cases "l \ dom f", simp_all add: Ltake_eq_def) assume "l \ dom f" with \dom f = dom g\ \dom (fst(ob!card(dom f))) = dom f\ show "fst(ob!card(dom f)) l = g l" by auto qed } with ext have "fst(ob!card(dom f)) = g" by auto with \fst(ob!0) = f\ \Obj (fst(ob!0)) T \\<^sub>\\<^sup>* Obj (fst(ob!card (dom f))) T\ show "Obj f T \\<^sub>\\<^sup>* Obj g T" by simp qed subsection \Size of sterms\ (* this section defines the size of sterms compared to size, the size of an object is the sum of the size of its fields +1 *) definition fsize0 :: "(Label -~> sterm) \ (sterm \ nat) \ nat" where "fsize0 f sts = foldl (+) 0 (map sts (Finite_Set.fold (\x z. z@[THE y. Some y = f x]) [] (dom f)))" primrec ssize :: "sterm \ nat" and ssize_option :: "sterm option \ nat" where ssize_Bvar : "ssize (Bvar b) = 0" | ssize_Fvar : "ssize (Fvar x) = 0" | ssize_Call : "ssize (Call a l b) = (ssize a) + (ssize b) + Suc 0" | ssize_Upd : "ssize (Upd a l b) = (ssize a) + (ssize b) + Suc 0" | ssize_Obj : "ssize (Obj f T) = Finite_Set.fold (\x y. y + ssize_option (f x)) (Suc 0) (dom f)" | ssize_None : "ssize_option (None) = 0" | ssize_Some : "ssize_option (Some y) = ssize y + Suc 0" (* We need this locale, as all the handy functions are there now *) interpretation comp_fun_commute "(\x y::nat. y + (f x))" by (unfold comp_fun_commute_def, force) lemma SizeOfObjectPos: "ssize (Obj (f::Label -~> sterm) T) > 0" proof (simp) from finite_dom_fmap have "finite (dom f)" by auto thus "0 < Finite_Set.fold (\x y. y + ssize_option (f x)) (Suc 0) (dom f)" proof (induct) case empty thus ?case by simp next case (insert A a) thus ?case by auto qed qed end diff --git a/thys/Monad_Memo_DP/heap_monad/Pair_Memory.thy b/thys/Monad_Memo_DP/heap_monad/Pair_Memory.thy --- a/thys/Monad_Memo_DP/heap_monad/Pair_Memory.thy +++ b/thys/Monad_Memo_DP/heap_monad/Pair_Memory.thy @@ -1,670 +1,670 @@ subsection \Pair Memory\ theory Pair_Memory imports "../state_monad/Memory" begin (* XXX Move *) lemma map_add_mono: "(m1 ++ m2) \\<^sub>m (m1' ++ m2')" if "m1 \\<^sub>m m1'" "m2 \\<^sub>m m2'" "dom m1 \ dom m2' = {}" using that unfolding map_le_def map_add_def dom_def by (auto split: option.splits) lemma map_add_upd2: "f(x \ y) ++ g = (f ++ g)(x \ y)" if "dom f \ dom g = {}" "x \ dom g" apply (subst map_add_comm) defer apply simp apply (subst map_add_comm) using that by auto locale pair_mem_defs = fixes lookup1 lookup2 :: "'a \ ('mem, 'v option) state" and update1 update2 :: "'a \ 'v \ ('mem, unit) state" and move12 :: "'k1 \ ('mem, unit) state" and get_k1 get_k2 :: "('mem, 'k1) state" and P :: "'mem \ bool" fixes key1 :: "'k \ 'k1" and key2 :: "'k \ 'a" begin text \We assume that look-ups happen on the older row, so it is biased towards the second entry.\ definition "lookup_pair k = do { let k' = key1 k; k2 \ get_k2; if k' = k2 then lookup2 (key2 k) else do { k1 \ get_k1; if k' = k1 then lookup1 (key2 k) else State_Monad.return None } } " text \We assume that updates happen on the newer row, so it is biased towards the first entry.\ definition "update_pair k v = do { let k' = key1 k; k1 \ get_k1; if k' = k1 then update1 (key2 k) v else do { k2 \ get_k2; if k' = k2 then update2 (key2 k) v else (move12 k' \ update1 (key2 k) v) } } " sublocale pair: state_mem_defs lookup_pair update_pair . sublocale mem1: state_mem_defs lookup1 update1 . sublocale mem2: state_mem_defs lookup2 update2 . definition "inv_pair heap \ let k1 = fst (State_Monad.run_state get_k1 heap); k2 = fst (State_Monad.run_state get_k2 heap) in (\ k \ dom (mem1.map_of heap). \ k'. key1 k' = k1 \ key2 k' = k) \ (\ k \ dom (mem2.map_of heap). \ k'. key1 k' = k2 \ key2 k' = k) \ k1 \ k2 \ P heap " definition "map_of1 m k = (if key1 k = fst (State_Monad.run_state get_k1 m) then mem1.map_of m (key2 k) else None)" definition "map_of2 m k = (if key1 k = fst (State_Monad.run_state get_k2 m) then mem2.map_of m (key2 k) else None)" end (* Pair Mem Defs *) locale pair_mem = pair_mem_defs + assumes get_state: "State_Monad.run_state get_k1 m = (k, m') \ m' = m" "State_Monad.run_state get_k2 m = (k, m') \ m' = m" assumes move12_correct: "P m \ State_Monad.run_state (move12 k1) m = (x, m') \ mem1.map_of m' \\<^sub>m Map.empty" "P m \ State_Monad.run_state (move12 k1) m = (x, m') \ mem2.map_of m' \\<^sub>m mem1.map_of m" assumes move12_keys: "State_Monad.run_state (move12 k1) m = (x, m') \ fst (State_Monad.run_state get_k1 m') = k1" "State_Monad.run_state (move12 k1) m = (x, m') \ fst (State_Monad.run_state get_k2 m') = fst (State_Monad.run_state get_k1 m)" assumes move12_inv: "lift_p P (move12 k1)" assumes lookup_inv: "lift_p P (lookup1 k')" "lift_p P (lookup2 k')" assumes update_inv: "lift_p P (update1 k' v)" "lift_p P (update2 k' v)" assumes lookup_keys: "P m \ State_Monad.run_state (lookup1 k') m = (v', m') \ fst (State_Monad.run_state get_k1 m') = fst (State_Monad.run_state get_k1 m)" "P m \ State_Monad.run_state (lookup1 k') m = (v', m') \ fst (State_Monad.run_state get_k2 m') = fst (State_Monad.run_state get_k2 m)" "P m \ State_Monad.run_state (lookup2 k') m = (v', m') \ fst (State_Monad.run_state get_k1 m') = fst (State_Monad.run_state get_k1 m)" "P m \ State_Monad.run_state (lookup2 k') m = (v', m') \ fst (State_Monad.run_state get_k2 m') = fst (State_Monad.run_state get_k2 m)" assumes update_keys: "P m \ State_Monad.run_state (update1 k' v) m = (x, m') \ fst (State_Monad.run_state get_k1 m') = fst (State_Monad.run_state get_k1 m)" "P m \ State_Monad.run_state (update1 k' v) m = (x, m') \ fst (State_Monad.run_state get_k2 m') = fst (State_Monad.run_state get_k2 m)" "P m \ State_Monad.run_state (update2 k' v) m = (x, m') \ fst (State_Monad.run_state get_k1 m') = fst (State_Monad.run_state get_k1 m)" "P m \ State_Monad.run_state (update2 k' v) m = (x, m') \ fst (State_Monad.run_state get_k2 m') = fst (State_Monad.run_state get_k2 m)" assumes lookup_correct: "P m \ mem1.map_of (snd (State_Monad.run_state (lookup1 k') m)) \\<^sub>m (mem1.map_of m)" "P m \ mem2.map_of (snd (State_Monad.run_state (lookup1 k') m)) \\<^sub>m (mem2.map_of m)" "P m \ mem1.map_of (snd (State_Monad.run_state (lookup2 k') m)) \\<^sub>m (mem1.map_of m)" "P m \ mem2.map_of (snd (State_Monad.run_state (lookup2 k') m)) \\<^sub>m (mem2.map_of m)" assumes update_correct: "P m \ mem1.map_of (snd (State_Monad.run_state (update1 k' v) m)) \\<^sub>m (mem1.map_of m)(k' \ v)" "P m \ mem2.map_of (snd (State_Monad.run_state (update2 k' v) m)) \\<^sub>m (mem2.map_of m)(k' \ v)" "P m \ mem2.map_of (snd (State_Monad.run_state (update1 k' v) m)) \\<^sub>m mem2.map_of m" "P m \ mem1.map_of (snd (State_Monad.run_state (update2 k' v) m)) \\<^sub>m mem1.map_of m" begin lemma map_of_le_pair: "pair.map_of m \\<^sub>m map_of1 m ++ map_of2 m" if "inv_pair m" using that unfolding pair.map_of_def map_of1_def map_of2_def unfolding lookup_pair_def inv_pair_def map_of_def map_le_def dom_def map_add_def unfolding State_Monad.bind_def by (auto 4 4 simp: mem2.map_of_def mem1.map_of_def Let_def dest: get_state split: prod.split_asm if_split_asm ) lemma pair_le_map_of: "map_of1 m ++ map_of2 m \\<^sub>m pair.map_of m" if "inv_pair m" using that unfolding pair.map_of_def map_of1_def map_of2_def unfolding lookup_pair_def inv_pair_def map_of_def map_le_def dom_def map_add_def unfolding State_Monad.bind_def by (auto simp: mem2.map_of_def mem1.map_of_def State_Monad.run_state_return Let_def dest: get_state split: prod.splits if_split_asm option.split ) lemma map_of_eq_pair: "map_of1 m ++ map_of2 m = pair.map_of m" if "inv_pair m" using that unfolding pair.map_of_def map_of1_def map_of2_def unfolding lookup_pair_def inv_pair_def map_of_def map_le_def dom_def map_add_def unfolding State_Monad.bind_def by (auto 4 4 simp: mem2.map_of_def mem1.map_of_def State_Monad.run_state_return Let_def dest: get_state split: prod.splits option.split ) lemma inv_pair_neq[simp]: False if "inv_pair m" "fst (State_Monad.run_state get_k1 m) = fst (State_Monad.run_state get_k2 m)" using that unfolding inv_pair_def by auto lemma inv_pair_P_D: "P m" if "inv_pair m" using that unfolding inv_pair_def by (auto simp: Let_def) lemma inv_pair_domD[intro]: "dom (map_of1 m) \ dom (map_of2 m) = {}" if "inv_pair m" using that unfolding inv_pair_def map_of1_def map_of2_def by (auto split: if_split_asm) lemma move12_correct1: "map_of1 heap' \\<^sub>m Map.empty" if "State_Monad.run_state (move12 k1) heap = (x, heap')" "P heap" using move12_correct[OF that(2,1)] unfolding map_of1_def by (auto simp: move12_keys map_le_def) lemma move12_correct2: "map_of2 heap' \\<^sub>m map_of1 heap" if "State_Monad.run_state (move12 k1) heap = (x, heap')" "P heap" using move12_correct(2)[OF that(2,1)] that unfolding map_of1_def map_of2_def by (auto simp: move12_keys map_le_def) lemma dom_empty[simp]: "dom (map_of1 heap') = {}" if "State_Monad.run_state (move12 k1) heap = (x, heap')" "P heap" using move12_correct1[OF that] by (auto dest: map_le_implies_dom_le) lemma inv_pair_lookup1: "inv_pair m'" if "State_Monad.run_state (lookup1 k) m = (v, m')" "inv_pair m" using that lookup_inv[of k] inv_pair_P_D[OF \inv_pair m\] unfolding inv_pair_def by (auto 4 4 simp: Let_def lookup_keys dest: lift_p_P lookup_correct[of _ k, THEN map_le_implies_dom_le] ) lemma inv_pair_lookup2: "inv_pair m'" if "State_Monad.run_state (lookup2 k) m = (v, m')" "inv_pair m" using that lookup_inv[of k] inv_pair_P_D[OF \inv_pair m\] unfolding inv_pair_def by (auto 4 4 simp: Let_def lookup_keys dest: lift_p_P lookup_correct[of _ k, THEN map_le_implies_dom_le] ) lemma inv_pair_update1: "inv_pair m'" if "State_Monad.run_state (update1 (key2 k) v) m = (v', m')" "inv_pair m" "fst (State_Monad.run_state get_k1 m) = key1 k" using that update_inv[of "key2 k" v] inv_pair_P_D[OF \inv_pair m\] unfolding inv_pair_def apply (auto simp: Let_def update_keys dest: lift_p_P update_correct[of _ "key2 k" v, THEN map_le_implies_dom_le] ) apply (frule update_correct[of _ "key2 k" v, THEN map_le_implies_dom_le]; auto 13 2; fail) apply (frule update_correct[of _ "key2 k" v, THEN map_le_implies_dom_le]; auto 13 2; fail) done lemma inv_pair_update2: "inv_pair m'" if "State_Monad.run_state (update2 (key2 k) v) m = (v', m')" "inv_pair m" "fst (State_Monad.run_state get_k2 m) = key1 k" using that update_inv[of "key2 k" v] inv_pair_P_D[OF \inv_pair m\] unfolding inv_pair_def apply (auto simp: Let_def update_keys dest: lift_p_P update_correct[of _ "key2 k" v, THEN map_le_implies_dom_le] ) apply (frule update_correct[of _ "key2 k" v, THEN map_le_implies_dom_le]; auto 13 2; fail) apply (frule update_correct[of _ "key2 k" v, THEN map_le_implies_dom_le]; auto 13 2; fail) done lemma inv_pair_move12: "inv_pair m'" if "State_Monad.run_state (move12 k) m = (v', m')" "inv_pair m" "fst (State_Monad.run_state get_k1 m) \ k" using that move12_inv[of "k"] inv_pair_P_D[OF \inv_pair m\] unfolding inv_pair_def apply (auto simp: Let_def move12_keys dest: lift_p_P move12_correct[of _ "k", THEN map_le_implies_dom_le] ) apply (blast dest: move12_correct[of _ "k", THEN map_le_implies_dom_le]) done lemma mem_correct_pair: "mem_correct lookup_pair update_pair inv_pair" if injective: "\ k k'. key1 k = key1 k' \ key2 k = key2 k' \ k = k'" proof (standard, goal_cases) case (1 k) \ \Lookup invariant\ show ?case unfolding lookup_pair_def Let_def by (auto 4 4 intro!: lift_pI dest: get_state inv_pair_lookup1 inv_pair_lookup2 simp: State_Monad.bind_def State_Monad.run_state_return split: if_split_asm prod.split_asm ) next case (2 k v) \ \Update invariant\ show ?case unfolding update_pair_def Let_def apply (auto 4 4 intro!: lift_pI intro: inv_pair_update1 inv_pair_update2 dest: get_state simp: State_Monad.bind_def get_state State_Monad.run_state_return split: if_split_asm prod.split_asm )+ apply (elim inv_pair_update1 inv_pair_move12) apply (((subst get_state, assumption)+)?, auto intro: move12_keys dest: get_state; fail)+ done next case (3 m k) { let ?m = "snd (State_Monad.run_state (lookup2 (key2 k)) m)" have "map_of1 ?m \\<^sub>m map_of1 m" by (smt 3 domIff inv_pair_P_D local.lookup_keys lookup_correct map_le_def map_of1_def surjective_pairing) moreover have "map_of2 ?m \\<^sub>m map_of2 m" by (smt 3 domIff inv_pair_P_D local.lookup_keys lookup_correct map_le_def map_of2_def surjective_pairing) moreover have "dom (map_of1 ?m) \ dom (map_of2 m) = {}" using 3 \map_of1 ?m \\<^sub>m map_of1 m\ inv_pair_domD map_le_implies_dom_le by fastforce moreover have "inv_pair ?m" using 3 inv_pair_lookup2 surjective_pairing by metis ultimately have "pair.map_of ?m \\<^sub>m pair.map_of m" apply (subst map_of_eq_pair[symmetric]) defer apply (subst map_of_eq_pair[symmetric]) by (auto intro: 3 map_add_mono) } moreover { let ?m = "snd (State_Monad.run_state (lookup1 (key2 k)) m)" have "map_of1 ?m \\<^sub>m map_of1 m" by (smt 3 domIff inv_pair_P_D local.lookup_keys lookup_correct map_le_def map_of1_def surjective_pairing) moreover have "map_of2 ?m \\<^sub>m map_of2 m" by (smt 3 domIff inv_pair_P_D local.lookup_keys lookup_correct map_le_def map_of2_def surjective_pairing) moreover have "dom (map_of1 ?m) \ dom (map_of2 m) = {}" using 3 \map_of1 ?m \\<^sub>m map_of1 m\ inv_pair_domD map_le_implies_dom_le by fastforce moreover have "inv_pair ?m" using 3 inv_pair_lookup1 surjective_pairing by metis ultimately have "pair.map_of ?m \\<^sub>m pair.map_of m" apply (subst map_of_eq_pair[symmetric]) defer apply (subst map_of_eq_pair[symmetric]) by (auto intro: 3 map_add_mono) } ultimately show ?case by (auto split:if_split prod.split simp: Let_def lookup_pair_def State_Monad.bind_def State_Monad.run_state_return dest: get_state intro: map_le_refl ) next case prems: (4 m k v) let ?m1 = "snd (State_Monad.run_state (update1 (key2 k) v) m)" let ?m2 = "snd (State_Monad.run_state (update2 (key2 k) v) m)" from prems have disjoint: "dom (map_of1 m) \ dom (map_of2 m) = {}" by (simp add: inv_pair_domD) show ?case apply (auto intro: map_le_refl dest: get_state split: prod.split simp: Let_def update_pair_def State_Monad.bind_def State_Monad.run_state_return ) proof goal_cases case (1 m') then have "m' = m" by (rule get_state) - from 1 prems have "map_of1 ?m1 \\<^sub>m map_of1 m(k \ v)" + from 1 prems have "map_of1 ?m1 \\<^sub>m (map_of1 m)(k \ v)" by (smt inv_pair_P_D map_le_def map_of1_def surjective_pairing domIff fst_conv fun_upd_apply injective update_correct update_keys ) moreover from prems have "map_of2 ?m1 \\<^sub>m map_of2 m" by (smt domIff inv_pair_P_D update_correct update_keys map_le_def map_of2_def surjective_pairing) moreover from prems have "dom (map_of1 ?m1) \ dom (map_of2 m) = {}" by (smt inv_pair_P_D[OF \inv_pair m\] domIff Int_emptyI eq_snd_iff inv_pair_neq map_of1_def map_of2_def update_keys(1) ) moreover from 1 prems have "k \ dom (map_of2 m)" using inv_pair_neq map_of2_def by fastforce moreover from 1 prems have "inv_pair ?m1" using inv_pair_update1 fst_conv surjective_pairing by metis - ultimately show "pair.map_of (snd (State_Monad.run_state (update1 (key2 k) v) m')) \\<^sub>m pair.map_of m(k \ v)" + ultimately show "pair.map_of (snd (State_Monad.run_state (update1 (key2 k) v) m')) \\<^sub>m (pair.map_of m)(k \ v)" unfolding \m' = m\ using disjoint apply (subst map_of_eq_pair[symmetric]) defer apply (subst map_of_eq_pair[symmetric], rule prems) apply (subst map_add_upd2[symmetric]) by (auto intro: map_add_mono) next case (2 k1 m' m'') then have "m' = m" "m'' = m" by (auto dest: get_state) - from 2 prems have "map_of2 ?m2 \\<^sub>m map_of2 m(k \ v)" + from 2 prems have "map_of2 ?m2 \\<^sub>m (map_of2 m)(k \ v)" unfolding \m' = m\ \m'' = m\ by (smt inv_pair_P_D map_le_def map_of2_def surjective_pairing domIff fst_conv fun_upd_apply injective update_correct update_keys ) moreover from prems have "map_of1 ?m2 \\<^sub>m map_of1 m" by (smt domIff inv_pair_P_D update_correct update_keys map_le_def map_of1_def surjective_pairing) - moreover from 2 have "dom (map_of1 ?m2) \ dom (map_of2 m(k \ v)) = {}" + moreover from 2 have "dom (map_of1 ?m2) \ dom ((map_of2 m)(k \ v)) = {}" unfolding \m' = m\ by (smt domIff \map_of1 ?m2 \\<^sub>m map_of1 m\ disjoint_iff_not_equal fst_conv fun_upd_apply map_le_def map_of1_def map_of2_def ) moreover from 2 prems have "inv_pair ?m2" unfolding \m' = m\ using inv_pair_update2 fst_conv surjective_pairing by metis - ultimately show "pair.map_of (snd (State_Monad.run_state (update2 (key2 k) v) m'')) \\<^sub>m pair.map_of m(k \ v)" + ultimately show "pair.map_of (snd (State_Monad.run_state (update2 (key2 k) v) m'')) \\<^sub>m (pair.map_of m)(k \ v)" unfolding \m' = m\ \m'' = m\ apply (subst map_of_eq_pair[symmetric]) defer apply (subst map_of_eq_pair[symmetric], rule prems) apply (subst map_add_upd[symmetric]) by (rule map_add_mono) next case (3 k1 m1 k2 m2 m3) then have "m1 = m" "m2 = m" by (auto dest: get_state) let ?m3 = "snd (State_Monad.run_state (update1 (key2 k) v) m3)" - from 3 prems have "map_of1 ?m3 \\<^sub>m map_of2 m(k \ v)" + from 3 prems have "map_of1 ?m3 \\<^sub>m (map_of2 m)(k \ v)" unfolding \m2 = m\ by (smt inv_pair_P_D map_le_def map_of1_def surjective_pairing domIff fst_conv fun_upd_apply injective inv_pair_move12 move12_correct move12_keys update_correct update_keys ) moreover have "map_of2 ?m3 \\<^sub>m map_of1 m" proof - from prems 3 have "P m" "P m3" unfolding \m1 = m\ \m2 = m\ using inv_pair_P_D[OF prems] by (auto elim: lift_p_P[OF move12_inv]) from 3(3)[unfolded \m2 = m\] have "mem2.map_of ?m3 \\<^sub>m mem1.map_of m" by - (erule map_le_trans[OF update_correct(3)[OF \P m3\] move12_correct(2)[OF \P m\]]) with 3 prems show ?thesis unfolding \m1 = m\ \m2 = m\ map_le_def map_of2_def apply auto apply (frule move12_keys(2), simp) by (metis domI inv_pair_def map_of1_def surjective_pairing inv_pair_move12 move12_keys(2) update_keys(2) ) qed moreover from prems 3 have "dom (map_of1 ?m3) \ dom (map_of1 m) = {}" unfolding \m1 = m\ \m2 = m\ by (smt inv_pair_P_D disjoint_iff_not_equal map_of1_def surjective_pairing domIff fst_conv inv_pair_move12 move12_keys update_keys ) moreover from 3 have "k \ dom (map_of1 m)" by (simp add: domIff map_of1_def) moreover from 3 prems have "inv_pair ?m3" unfolding \m2 = m\ by (metis inv_pair_move12 inv_pair_update1 move12_keys(1) fst_conv surjective_pairing) ultimately show ?case unfolding \m1 = m\ \m2 = m\ using disjoint apply (subst map_of_eq_pair[symmetric]) defer apply (subst map_of_eq_pair[symmetric]) apply (rule prems) apply (subst (2) map_add_comm) defer apply (subst map_add_upd2[symmetric]) apply (auto intro: map_add_mono) done qed qed lemma emptyI: assumes "inv_pair m" "mem1.map_of m \\<^sub>m Map.empty" "mem2.map_of m \\<^sub>m Map.empty" shows "pair.map_of m \\<^sub>m Map.empty" using assms by (auto simp: map_of1_def map_of2_def map_le_def map_of_eq_pair[symmetric]) end (* Pair Mem *) datatype ('k, 'v) pair_storage = Pair_Storage 'k 'k 'v 'v context mem_correct_empty begin context fixes key :: "'a \ 'k" begin text \We assume that look-ups happen on the older row, so it is biased towards the second entry.\ definition "lookup_pair k = State (\ mem. ( case mem of Pair_Storage k1 k2 m1 m2 \ let k' = key k in if k' = k2 then case State_Monad.run_state (lookup k) m2 of (v, m) \ (v, Pair_Storage k1 k2 m1 m) else if k' = k1 then case State_Monad.run_state (lookup k) m1 of (v, m) \ (v, Pair_Storage k1 k2 m m2) else (None, mem) ) ) " text \We assume that updates happen on the newer row, so it is biased towards the first entry.\ definition "update_pair k v = State (\ mem. ( case mem of Pair_Storage k1 k2 m1 m2 \ let k' = key k in if k' = k1 then case State_Monad.run_state (update k v) m1 of (_, m) \ ((), Pair_Storage k1 k2 m m2) else if k' = k2 then case State_Monad.run_state (update k v) m2 of (_, m) \ ((),Pair_Storage k1 k2 m1 m) else case State_Monad.run_state (update k v) empty of (_, m) \ ((), Pair_Storage k' k1 m m1) ) ) " interpretation pair: state_mem_defs lookup_pair update_pair . definition "inv_pair p = (case p of Pair_Storage k1 k2 m1 m2 \ key ` dom (map_of m1) \ {k1} \ key ` dom (map_of m2) \ {k2} \ k1 \ k2 \ P m1 \ P m2 )" lemma map_of_le_pair: "pair.map_of (Pair_Storage k1 k2 m1 m2) \\<^sub>m (map_of m1 ++ map_of m2)" if "inv_pair (Pair_Storage k1 k2 m1 m2)" using that unfolding pair.map_of_def unfolding lookup_pair_def inv_pair_def map_of_def map_le_def dom_def map_add_def apply auto apply (auto 4 6 split: prod.split_asm if_split_asm option.split simp: Let_def) done lemma pair_le_map_of: "map_of m1 ++ map_of m2 \\<^sub>m pair.map_of (Pair_Storage k1 k2 m1 m2)" if "inv_pair (Pair_Storage k1 k2 m1 m2)" using that unfolding pair.map_of_def unfolding lookup_pair_def inv_pair_def map_of_def map_le_def dom_def map_add_def by (auto 4 5 split: prod.split_asm if_split_asm option.split simp: Let_def) lemma map_of_eq_pair: "map_of m1 ++ map_of m2 = pair.map_of (Pair_Storage k1 k2 m1 m2)" if "inv_pair (Pair_Storage k1 k2 m1 m2)" using that unfolding pair.map_of_def unfolding lookup_pair_def inv_pair_def map_of_def map_le_def dom_def map_add_def by (auto 4 7 split: prod.split_asm if_split_asm option.split simp: Let_def) lemma inv_pair_neq[simp, dest]: False if "inv_pair (Pair_Storage k k x y)" using that unfolding inv_pair_def by auto lemma inv_pair_P_D1: "P m1" if "inv_pair (Pair_Storage k1 k2 m1 m2)" using that unfolding inv_pair_def by auto lemma inv_pair_P_D2: "P m2" if "inv_pair (Pair_Storage k1 k2 m1 m2)" using that unfolding inv_pair_def by auto lemma inv_pair_domD[intro]: "dom (map_of m1) \ dom (map_of m2) = {}" if "inv_pair (Pair_Storage k1 k2 m1 m2)" using that unfolding inv_pair_def by fastforce lemma mem_correct_pair: "mem_correct lookup_pair update_pair inv_pair" proof (standard, goal_cases) case (1 k) \ \Lookup invariant\ with lookup_inv[of k] show ?case unfolding lookup_pair_def Let_def by (auto intro!: lift_pI split: pair_storage.split_asm if_split_asm prod.split_asm) (auto dest: lift_p_P simp: inv_pair_def, (force dest!: lookup_correct[of _ k] map_le_implies_dom_le)+ ) next case (2 k v) \ \Update invariant\ with update_inv[of k v] update_correct[OF P_empty, of k v] P_empty show ?case unfolding update_pair_def Let_def by (auto intro!: lift_pI split: pair_storage.split_asm if_split_asm prod.split_asm) (auto dest: lift_p_P simp: inv_pair_def, (force dest: lift_p_P dest!: update_correct[of _ k v] map_le_implies_dom_le)+ ) next case (3 m k) { fix m1 v1 m1' m2 v2 m2' k1 k2 assume assms: "State_Monad.run_state (lookup k) m1 = (v1, m1')" "State_Monad.run_state (lookup k) m2 = (v2, m2')" "inv_pair (Pair_Storage k1 k2 m1 m2)" from assms have "P m1" "P m2" by (auto intro: inv_pair_P_D1 inv_pair_P_D2) have [intro]: "map_of m1' \\<^sub>m map_of m1" "map_of m2' \\<^sub>m map_of m2" using lookup_correct[OF \P m1\, of k] lookup_correct[OF \P m2\, of k] assms by auto from inv_pair_domD[OF assms(3)] have 1: "dom (map_of m1') \ dom (map_of m2) = {}" by (metis (no_types) \map_of m1' \\<^sub>m map_of m1\ disjoint_iff_not_equal domIff map_le_def) have inv1: "inv_pair (Pair_Storage (key k) k2 m1' m2)" if "k2 \ key k" "k1 = key k" using that \P m1\ \P m2\ unfolding inv_pair_def apply clarsimp apply safe subgoal for x' y using assms(1,3) lookup_correct[OF \P m1\, of k, THEN map_le_implies_dom_le] unfolding inv_pair_def by auto subgoal for x' y using assms(3) unfolding inv_pair_def by fastforce using lookup_inv[of k] assms unfolding lift_p_def by force have inv2: "inv_pair (Pair_Storage k1 (key k) m1 m2')" if "k2 = key k" "k1 \ key k" using that \P m1\ \P m2\ unfolding inv_pair_def apply clarsimp apply safe subgoal for x' y using assms(3) unfolding inv_pair_def by fastforce subgoal for x x' y using assms(2,3) lookup_correct[OF \P m2\, of k, THEN map_le_implies_dom_le] unfolding inv_pair_def by fastforce using lookup_inv[of k] assms unfolding lift_p_def by force have A: "pair.map_of (Pair_Storage (key k) k2 m1' m2) \\<^sub>m pair.map_of (Pair_Storage (key k) k2 m1 m2)" if "k2 \ key k" "k1 = key k" using inv1 assms(3) 1 by (auto intro: map_add_mono map_le_refl simp: that map_of_eq_pair[symmetric]) have B: "pair.map_of (Pair_Storage k1 (key k) m1 m2') \\<^sub>m pair.map_of (Pair_Storage k1 (key k) m1 m2)" if "k2 = key k" "k1 \ key k" using inv2 assms(3) that by (auto intro: map_add_mono map_le_refl simp: map_of_eq_pair[symmetric] dest: inv_pair_domD) note A B } with \inv_pair m\ show ?case by (auto split: pair_storage.split if_split prod.split simp: Let_def lookup_pair_def) next case (4 m k v) { fix m1 v1 m1' m2 v2 m2' m3 k1 k2 assume assms: "State_Monad.run_state (update k v) m1 = ((), m1')" "State_Monad.run_state (update k v) m2 = ((), m2')" "State_Monad.run_state (update k v) empty = ((), m3)" "inv_pair (Pair_Storage k1 k2 m1 m2)" from assms have "P m1" "P m2" by (auto intro: inv_pair_P_D1 inv_pair_P_D2) from assms(3) P_empty update_inv[of k v] have "P m3" unfolding lift_p_def by auto - have [intro]: "map_of m1' \\<^sub>m map_of m1(k \ v)" "map_of m2' \\<^sub>m map_of m2(k \ v)" + have [intro]: "map_of m1' \\<^sub>m (map_of m1)(k \ v)" "map_of m2' \\<^sub>m (map_of m2)(k \ v)" using update_correct[OF \P m1\, of k v] update_correct[OF \P m2\, of k v] assms by auto - have "map_of m3 \\<^sub>m map_of empty(k \ v)" + have "map_of m3 \\<^sub>m (map_of empty)(k \ v)" using assms(3) update_correct[OF P_empty, of k v] by auto - also have "\ \\<^sub>m map_of m2(k \ v)" + also have "\ \\<^sub>m (map_of m2)(k \ v)" using empty_correct by (auto elim: map_le_trans intro!: map_le_upd) - finally have "map_of m3 \\<^sub>m map_of m2(k \ v)" . - have 1: "dom (map_of m1) \ dom (map_of m2(k \ v)) = {}" if "k1 \ key k" + finally have "map_of m3 \\<^sub>m (map_of m2)(k \ v)" . + have 1: "dom (map_of m1) \ dom ((map_of m2)(k \ v)) = {}" if "k1 \ key k" using assms(4) that by (force simp: inv_pair_def) have 2: "dom (map_of m3) \ dom (map_of m1) = {}" if "k1 \ key k" - using \local.map_of m3 \\<^sub>m local.map_of empty(k \ v)\ assms(4) that + using \local.map_of m3 \\<^sub>m (map_of empty)(k \ v)\ assms(4) that by (fastforce dest!: map_le_implies_dom_le simp: inv_pair_def) have inv: "inv_pair (Pair_Storage (key k) k1 m3 m1)" if "k2 \ key k" "k1 \ key k" using that \P m1\ \P m2\ \P m3\ unfolding inv_pair_def apply clarsimp apply safe subgoal for x x' y using assms(3) update_correct[OF P_empty, of k v, THEN map_le_implies_dom_le] empty_correct by (auto dest: map_le_implies_dom_le) subgoal for x x' y using assms(4) unfolding inv_pair_def by fastforce done have A: - "pair.map_of (Pair_Storage (key k) k1 m3 m1) \\<^sub>m pair.map_of (Pair_Storage k1 k2 m1 m2)(k \ v)" + "pair.map_of (Pair_Storage (key k) k1 m3 m1) \\<^sub>m (pair.map_of (Pair_Storage k1 k2 m1 m2))(k \ v)" if "k2 \ key k" "k1 \ key k" - using inv assms(4) \map_of m3 \\<^sub>m map_of m2(k \ v)\ 1 + using inv assms(4) \map_of m3 \\<^sub>m (map_of m2)(k \ v)\ 1 apply (simp add: that map_of_eq_pair[symmetric]) apply (subst map_add_upd[symmetric], subst Map.map_add_comm, rule 2, rule that) by (rule map_add_mono; auto) have inv1: "inv_pair (Pair_Storage (key k) k2 m1' m2)" if "k2 \ key k" "k1 = key k" using that \P m1\ \P m2\ unfolding inv_pair_def apply clarsimp apply safe subgoal for x' y using assms(1,4) update_correct[OF \P m1\, of k v, THEN map_le_implies_dom_le] unfolding inv_pair_def by auto subgoal for x' y using assms(4) unfolding inv_pair_def by fastforce using update_inv[of k v] assms unfolding lift_p_def by force have inv2: "inv_pair (Pair_Storage k1 (key k) m1 m2')" if "k2 = key k" "k1 \ key k" using that \P m1\ \P m2\ unfolding inv_pair_def apply clarsimp apply safe subgoal for x' y using assms(4) unfolding inv_pair_def by fastforce subgoal for x x' y using assms(2,4) update_correct[OF \P m2\, of k v, THEN map_le_implies_dom_le] unfolding inv_pair_def by fastforce using update_inv[of k v] assms unfolding lift_p_def by force have C: "pair.map_of (Pair_Storage (key k) k2 m1' m2) \\<^sub>m - pair.map_of (Pair_Storage (key k) k2 m1 m2)(k \ v)" + (pair.map_of (Pair_Storage (key k) k2 m1 m2))(k \ v)" if "k2 \ key k" "k1 = key k" using inv1[OF that] assms(4) \inv_pair m\ by (simp add: that map_of_eq_pair[symmetric]) (subst map_add_upd2[symmetric]; force simp: inv_pair_def intro: map_add_mono map_le_refl) have B: "pair.map_of (Pair_Storage k1 (key k) m1 m2') \\<^sub>m - pair.map_of (Pair_Storage k1 (key k) m1 m2)(k \ v)" + (pair.map_of (Pair_Storage k1 (key k) m1 m2))(k \ v)" if "k2 = key k" "k1 \ key k" using inv2[OF that] assms(4) by (simp add: that map_of_eq_pair[symmetric]) (subst map_add_upd[symmetric]; rule map_add_mono; force simp: inv_pair_def) note A B C } with \inv_pair m\ show ?case by (auto split: pair_storage.split if_split prod.split simp: Let_def update_pair_def) qed end (* Key function *) end (* Lookup & Update w/ Empty *) end (* Theory *) diff --git a/thys/Monad_Memo_DP/heap_monad/State_Heap.thy b/thys/Monad_Memo_DP/heap_monad/State_Heap.thy --- a/thys/Monad_Memo_DP/heap_monad/State_Heap.thy +++ b/thys/Monad_Memo_DP/heap_monad/State_Heap.thy @@ -1,278 +1,278 @@ subsection \Relation Between the State and the Heap Monad\ theory State_Heap imports "../state_monad/DP_CRelVS" "HOL-Imperative_HOL.Imperative_HOL" State_Heap_Misc Heap_Monad_Ext begin definition lift_p :: "(heap \ bool) \ 'a Heap \ bool" where "lift_p P f = (\ heap. P heap \ (case execute f heap of None \ False | Some (_, heap) \ P heap))" context fixes P f heap assumes lift: "lift_p P f" and P: "P heap" begin lemma execute_cases: "case execute f heap of None \ False | Some (_, heap) \ P heap" using lift P unfolding lift_p_def by auto lemma execute_cases': "case execute f heap of Some (_, heap) \ P heap" using execute_cases by (auto split: option.split) lemma lift_p_None[simp, dest]: False if "execute f heap = None" using that execute_cases by auto lemma lift_p_P: "case the (execute f heap) of (_, heap) \ P heap" using execute_cases by (auto split: option.split_asm) lemma lift_p_P': "P heap'" if "the (execute f heap) = (v, heap')" using that lift_p_P by auto lemma lift_p_P'': "P heap'" if "execute f heap = Some (v, heap')" using that lift_p_P by auto lemma lift_p_the_Some[simp]: "execute f heap = Some (v, heap')" if "the (execute f heap) = (v, heap')" using that execute_cases by (auto split: option.split_asm) lemma lift_p_E: obtains v heap' where "execute f heap = Some (v, heap')" "P heap'" using execute_cases by (cases "execute f heap") auto end definition "state_of s \ State (\ heap. the (execute s heap))" locale heap_mem_defs = fixes P :: "heap \ bool" and lookup :: "'k \ 'v option Heap" and update :: "'k \ 'v \ unit Heap" begin definition rel_state :: "('a \ 'b \ bool) \ (heap, 'a) state \ 'b Heap \ bool" where "rel_state R f g \ \ heap. P heap \ (case State_Monad.run_state f heap of (v1, heap1) \ case execute g heap of Some (v2, heap2) \ R v1 v2 \ heap1 = heap2 \ P heap2 | None \ False)" definition "lookup' k \ State (\ heap. the (execute (lookup k) heap))" definition "update' k v \ State (\ heap. the (execute (update k v) heap))" definition "heap_get = Heap_Monad.Heap (\ heap. Some (heap, heap))" definition checkmem :: "'k \ 'v Heap \ 'v Heap" where "checkmem param calc \ Heap_Monad.bind (lookup param) (\ x. case x of Some x \ return x | None \ Heap_Monad.bind calc (\ x. Heap_Monad.bind (update param x) (\ _. return x ) ) ) " definition checkmem' :: "'k \ (unit \ 'v Heap) \ 'v Heap" where "checkmem' param calc \ Heap_Monad.bind (lookup param) (\ x. case x of Some x \ return x | None \ Heap_Monad.bind (calc ()) (\ x. Heap_Monad.bind (update param x) (\ _. return x ) ) ) " lemma checkmem_checkmem': "checkmem' param (\_. calc) = checkmem param calc" unfolding checkmem'_def checkmem_def .. definition map_of_heap where "map_of_heap heap k = fst (the (execute (lookup k) heap))" lemma rel_state_elim: assumes "rel_state R f g" "P heap" obtains heap' v v' where "State_Monad.run_state f heap = (v, heap')" "execute g heap = Some (v', heap')" "R v v'" "P heap'" apply atomize_elim using assms unfolding rel_state_def apply auto apply (cases "State_Monad.run_state f heap") apply auto apply (auto split: option.split_asm) done lemma rel_state_intro: assumes "\ heap v heap'. P heap \ State_Monad.run_state f heap = (v, heap') \ \ v'. R v v' \ execute g heap = Some (v', heap')" "\ heap v heap'. P heap \ State_Monad.run_state f heap = (v, heap') \ P heap'" shows "rel_state R f g" unfolding rel_state_def apply auto apply (frule assms(1)[rotated]) apply (auto intro: assms(2)) done context includes lifting_syntax state_monad_syntax begin lemma transfer_bind[transfer_rule]: "(rel_state R ===> (R ===> rel_state Q) ===> rel_state Q) State_Monad.bind Heap_Monad.bind" unfolding rel_fun_def State_Monad.bind_def Heap_Monad.bind_def by (force elim!: rel_state_elim intro!: rel_state_intro) lemma transfer_return[transfer_rule]: "(R ===> rel_state R) State_Monad.return Heap_Monad.return" unfolding rel_fun_def State_Monad.return_def Heap_Monad.return_def by (fastforce intro: rel_state_intro elim: rel_state_elim simp: execute_heap) lemma fun_app_lifted_transfer: "(rel_state (R ===> rel_state Q) ===> rel_state R ===> rel_state Q) State_Monad_Ext.fun_app_lifted Heap_Monad_Ext.fun_app_lifted" unfolding State_Monad_Ext.fun_app_lifted_def Heap_Monad_Ext.fun_app_lifted_def by transfer_prover lemma transfer_get[transfer_rule]: "rel_state (=) State_Monad.get heap_get" unfolding State_Monad.get_def heap_get_def by (auto intro: rel_state_intro) end (* Lifting Syntax *) end (* Heap Mem Defs *) locale heap_inv = heap_mem_defs _ lookup for lookup :: "'k \ 'v option Heap" + assumes lookup_inv: "lift_p P (lookup k)" assumes update_inv: "lift_p P (update k v)" begin lemma rel_state_lookup: "rel_state (=) (lookup' k) (lookup k)" unfolding rel_state_def lookup'_def using lookup_inv[of k] by (auto intro: lift_p_P') lemma rel_state_update: "rel_state (=) (update' k v) (update k v)" unfolding rel_state_def update'_def using update_inv[of k v] by (auto intro: lift_p_P') context includes lifting_syntax begin lemma transfer_lookup: "((=) ===> rel_state (=)) lookup' lookup" unfolding rel_fun_def by (auto intro: rel_state_lookup) lemma transfer_update: "((=) ===> (=) ===> rel_state (=)) update' update" unfolding rel_fun_def by (auto intro: rel_state_update) lemma transfer_checkmem: "((=) ===> rel_state (=) ===> rel_state (=)) (state_mem_defs.checkmem lookup' update') checkmem" supply [transfer_rule] = transfer_lookup transfer_update unfolding state_mem_defs.checkmem_def checkmem_def by transfer_prover end (* Lifting Syntax *) end (* Heap Invariant *) locale heap_correct = heap_inv + assumes lookup_correct: "P m \ map_of_heap (snd (the (execute (lookup k) m))) \\<^sub>m (map_of_heap m)" and update_correct: "P m \ map_of_heap (snd (the (execute (update k v) m))) \\<^sub>m (map_of_heap m)(k \ v)" begin lemma lookup'_correct: "state_mem_defs.map_of lookup' (snd (State_Monad.run_state (lookup' k) m)) \\<^sub>m (state_mem_defs.map_of lookup' m)" if "P m" using \P m\ unfolding state_mem_defs.map_of_def map_le_def lookup'_def by simp (metis (mono_tags, lifting) domIff lookup_correct map_le_def map_of_heap_def) lemma update'_correct: - "state_mem_defs.map_of lookup' (snd (State_Monad.run_state (update' k v) m)) \\<^sub>m state_mem_defs.map_of lookup' m(k \ v)" + "state_mem_defs.map_of lookup' (snd (State_Monad.run_state (update' k v) m)) \\<^sub>m (state_mem_defs.map_of lookup' m)(k \ v)" if "P m" unfolding state_mem_defs.map_of_def map_le_def lookup'_def update'_def using update_correct[of m k v] that by (auto split: if_split_asm simp: map_le_def map_of_heap_def) lemma lookup'_inv: "DP_CRelVS.lift_p P (lookup' k)" unfolding DP_CRelVS.lift_p_def lookup'_def by (auto elim: lift_p_P'[OF lookup_inv]) lemma update'_inv: "DP_CRelVS.lift_p P (update' k v)" unfolding DP_CRelVS.lift_p_def update'_def by (auto elim: lift_p_P'[OF update_inv]) lemma mem_correct_heap: "mem_correct lookup' update' P" by (intro mem_correct.intro lookup'_correct update'_correct lookup'_inv update'_inv) end (* Heap correct *) context heap_mem_defs begin context includes lifting_syntax begin lemma mem_correct_heap_correct: assumes correct: "mem_correct lookup\<^sub>s update\<^sub>s P" and lookup: "((=) ===> rel_state (=)) lookup\<^sub>s lookup" and update: "((=) ===> (=) ===> rel_state (=)) update\<^sub>s update" shows "heap_correct P update lookup" proof - interpret mem: mem_correct lookup\<^sub>s update\<^sub>s P by (rule correct) have [simp]: "the (execute (lookup k) m) = run_state (lookup\<^sub>s k) m" if "P m" for k m using lookup[THEN rel_funD, OF HOL.refl, of k] \P m\ by (auto elim: rel_state_elim) have [simp]: "the (execute (update k v) m) = run_state (update\<^sub>s k v) m" if "P m" for k v m using update[THEN rel_funD, THEN rel_funD, OF HOL.refl HOL.refl, of k v] \P m\ by (auto elim: rel_state_elim) have [simp]: "map_of_heap m = mem.map_of m" if "P m" for m unfolding map_of_heap_def mem.map_of_def using \P m\ by simp show ?thesis apply standard subgoal for k using mem.lookup_inv[of k] lookup[THEN rel_funD, OF HOL.refl, of k] unfolding lift_p_def DP_CRelVS.lift_p_def by (auto split: option.splits elim: rel_state_elim) subgoal for k v using mem.update_inv[of k] update[THEN rel_funD, THEN rel_funD, OF HOL.refl HOL.refl, of k v] unfolding lift_p_def DP_CRelVS.lift_p_def by (auto split: option.splits elim: rel_state_elim) subgoal premises prems for m k proof - have "P (snd (run_state (lookup\<^sub>s k) m))" by (meson DP_CRelVS.lift_p_P mem.lookup_inv prems prod.exhaust_sel) with mem.lookup_correct[OF \P m\, of k] \P m\ show ?thesis by (simp add: prems) qed subgoal premises prems for m k v proof - have "P (snd (run_state (update\<^sub>s k v) m))" by (meson DP_CRelVS.lift_p_P mem.update_inv prems prod.exhaust_sel) with mem.update_correct[OF \P m\, of k] \P m\ show ?thesis by (simp add: prems) qed done qed end end end (* Theory *) diff --git a/thys/Name_Carrying_Type_Inference/PreSimplyTyped.thy b/thys/Name_Carrying_Type_Inference/PreSimplyTyped.thy --- a/thys/Name_Carrying_Type_Inference/PreSimplyTyped.thy +++ b/thys/Name_Carrying_Type_Inference/PreSimplyTyped.thy @@ -1,769 +1,769 @@ theory PreSimplyTyped imports Fresh Permutation begin type_synonym tvar = nat datatype type = TUnit | TVar tvar | TArr type type | TPair type type datatype 'a ptrm = PUnit | PVar 'a | PApp "'a ptrm" "'a ptrm" | PFn 'a type "'a ptrm" | PPair "'a ptrm" "'a ptrm" | PFst "'a ptrm" | PSnd "'a ptrm" fun ptrm_fvs :: "'a ptrm \ 'a set" where "ptrm_fvs PUnit = {}" | "ptrm_fvs (PVar x) = {x}" | "ptrm_fvs (PApp A B) = ptrm_fvs A \ ptrm_fvs B" | "ptrm_fvs (PFn x _ A) = ptrm_fvs A - {x}" | "ptrm_fvs (PPair A B) = ptrm_fvs A \ ptrm_fvs B" | "ptrm_fvs (PFst P) = ptrm_fvs P" | "ptrm_fvs (PSnd P) = ptrm_fvs P" fun ptrm_apply_prm :: "'a prm \ 'a ptrm \ 'a ptrm" (infixr "\" 150) where "ptrm_apply_prm \ PUnit = PUnit" | "ptrm_apply_prm \ (PVar x) = PVar (\ $ x)" | "ptrm_apply_prm \ (PApp A B) = PApp (ptrm_apply_prm \ A) (ptrm_apply_prm \ B)" | "ptrm_apply_prm \ (PFn x T A) = PFn (\ $ x) T (ptrm_apply_prm \ A)" | "ptrm_apply_prm \ (PPair A B) = PPair (ptrm_apply_prm \ A) (ptrm_apply_prm \ B)" | "ptrm_apply_prm \ (PFst P) = PFst (ptrm_apply_prm \ P)" | "ptrm_apply_prm \ (PSnd P) = PSnd (ptrm_apply_prm \ P)" inductive ptrm_alpha_equiv :: "'a ptrm \ 'a ptrm \ bool" (infix "\" 100) where unit: "PUnit \ PUnit" | var: "(PVar x) \ (PVar x)" | app: "\A \ B; C \ D\ \ (PApp A C) \ (PApp B D)" | fn1: "A \ B \ (PFn x T A) \ (PFn x T B)" | fn2: "\a \ b; A \ [a \ b] \ B; a \ ptrm_fvs B\ \ (PFn a T A) \ (PFn b T B)" | pair: "\A \ B; C \ D\ \ (PPair A C) \ (PPair B D)" | fst: "A \ B \ PFst A \ PFst B" | snd: "A \ B \ PSnd A \ PSnd B" inductive_cases unitE: "PUnit \ Y" inductive_cases varE: "PVar x \ Y" inductive_cases appE: "PApp A B \ Y" inductive_cases fnE: "PFn x T A \ Y" inductive_cases pairE: "PPair A B \ Y" inductive_cases fstE: "PFst P \ Y" inductive_cases sndE: "PSnd P \ Y" lemma ptrm_prm_apply_id: shows "\ \ X = X" by(induction X, simp_all add: prm_apply_id) lemma ptrm_prm_apply_compose: shows "\ \ \ \ X = (\ \ \) \ X" by(induction X, simp_all add: prm_apply_composition) lemma ptrm_size_prm: shows "size X = size (\ \ X)" by(induction X, auto) lemma ptrm_size_alpha_equiv: assumes "X \ Y" shows "size X = size Y" using assms proof(induction rule: ptrm_alpha_equiv.induct) case (fn2 a b A B T) hence "size A = size B" using ptrm_size_prm by metis thus ?case by auto next qed auto lemma ptrm_fvs_finite: shows "finite (ptrm_fvs X)" by(induction X, auto) lemma ptrm_prm_fvs: shows "ptrm_fvs (\ \ X) = \ {$} ptrm_fvs X" proof(induction X) case (PUnit) thus ?case unfolding prm_set_def by simp next case (PVar x) have "ptrm_fvs (\ \ PVar x) = ptrm_fvs (PVar (\ $ x))" by simp moreover have "... = {\ $ x}" by simp moreover have "... = \ {$} {x}" using prm_set_singleton by metis moreover have "... = \ {$} ptrm_fvs (PVar x)" by simp ultimately show ?case by metis next case (PApp A B) have "ptrm_fvs (\ \ PApp A B) = ptrm_fvs (PApp (\ \ A) (\ \ B))" by simp moreover have "... = ptrm_fvs (\ \ A) \ ptrm_fvs (\ \ B)" by simp moreover have "... = \ {$} ptrm_fvs A \ \ {$} ptrm_fvs B" using PApp.IH by metis moreover have "... = \ {$} (ptrm_fvs A \ ptrm_fvs B)" using prm_set_distributes_union by metis moreover have "... = \ {$} ptrm_fvs (PApp A B)" by simp ultimately show ?case by metis next case (PFn x T A) have "ptrm_fvs (\ \ PFn x T A) = ptrm_fvs (PFn (\ $ x) T (\ \ A))" by simp moreover have "... = ptrm_fvs (\ \ A) - {\ $ x}" by simp moreover have "... = \ {$} ptrm_fvs A - {\ $ x}" using PFn.IH by metis moreover have "... = \ {$} ptrm_fvs A - \ {$} {x}" using prm_set_singleton by metis moreover have "... = \ {$} (ptrm_fvs A - {x})" using prm_set_distributes_difference by metis moreover have "... = \ {$} ptrm_fvs (PFn x T A)" by simp ultimately show ?case by metis next case (PPair A B) thus ?case using prm_set_distributes_union ptrm_apply_prm.simps(5) ptrm_fvs.simps(5) by fastforce next case (PFst P) thus ?case by auto next case (PSnd P) thus ?case by auto next qed lemma ptrm_alpha_equiv_fvs: assumes "X \ Y" shows "ptrm_fvs X = ptrm_fvs Y" using assms proof(induction rule: ptrm_alpha_equiv.induct) case (fn2 a b A B T) have "ptrm_fvs (PFn a T A) = ptrm_fvs A - {a}" by simp moreover have "... = ptrm_fvs ([a \ b] \ B) - {a}" using fn2.IH by metis moreover have "... = ([a \ b] {$} ptrm_fvs B) - {a}" using ptrm_prm_fvs by metis moreover have "... = ptrm_fvs B - {b}" proof - consider "b \ ptrm_fvs B" | "b \ ptrm_fvs B" by auto thus ?thesis proof(cases) case 1 have "[a \ b] {$} ptrm_fvs B - {a} = [b \ a] {$} ptrm_fvs B - {a}" using prm_unit_commutes by metis moreover have "... = ((ptrm_fvs B - {b}) \ {a}) - {a}" using prm_set_unit_action \b \ ptrm_fvs B\ \a \ ptrm_fvs B\ by metis moreover have "... = ptrm_fvs B - {b}" using \a \ ptrm_fvs B\ \a \ b\ using Diff_insert0 Diff_insert2 Un_insert_right insert_Diff1 insert_is_Un singletonI sup_bot.right_neutral by blast (* why?! *) ultimately show ?thesis by metis next case 2 hence "[a \ b] {$} ptrm_fvs B - {a} = ptrm_fvs B - {a}" using prm_set_unit_inaction \a \ ptrm_fvs B\ by metis moreover have "... = ptrm_fvs B" using \a \ ptrm_fvs B\ by simp moreover have "... = ptrm_fvs B - {b}" using \b \ ptrm_fvs B\ by simp ultimately show ?thesis by metis next qed qed moreover have "... = ptrm_fvs (PFn b T B)" by simp ultimately show ?case by metis next qed auto lemma ptrm_alpha_equiv_prm: assumes "X \ Y" shows "\ \ X \ \ \ Y" using assms proof(induction rule: ptrm_alpha_equiv.induct) case (unit) thus ?case using ptrm_alpha_equiv.unit by simp next case (var x) thus ?case using ptrm_alpha_equiv.var by simp next case (app A B C D) thus ?case using ptrm_alpha_equiv.app by simp next case (fn1 A B x T) thus ?case using ptrm_alpha_equiv.fn1 by simp next case (fn2 a b A B T) have "\ $ a \ \ $ b" using \a \ b\ using prm_apply_unequal by metis moreover have "\ $ a \ ptrm_fvs (\ \ B)" using \a \ ptrm_fvs B\ using imageE prm_apply_unequal prm_set_def ptrm_prm_fvs by (metis (no_types, lifting)) moreover have "\ \ A \ [\ $ a \ \ $ b] \ \ \ B" using fn2.IH prm_compose_push ptrm_prm_apply_compose by metis ultimately show ?case using ptrm_alpha_equiv.fn2 by simp next case (pair A B C D) thus ?case using ptrm_alpha_equiv.pair by simp next case (fst A B) thus ?case using ptrm_alpha_equiv.fst by simp next case (snd A B) thus ?case using ptrm_alpha_equiv.snd by simp next qed lemma ptrm_swp_transfer: shows "[a \ b] \ X \ Y \ X \ [a \ b] \ Y" proof - have 1: "[a \ b] \ X \ Y \ X \ [a \ b] \ Y" proof - assume "[a \ b] \ X \ Y" hence "\ \ X \ [a \ b] \ Y" using ptrm_alpha_equiv_prm ptrm_prm_apply_compose prm_unit_involution by metis thus ?thesis using ptrm_prm_apply_id by metis qed have 2: "X \ [a \ b] \ Y \ [a \ b] \ X \ Y" proof - assume "X \ [a \ b] \ Y" hence "[a \ b] \ X \ \ \ Y" using ptrm_alpha_equiv_prm ptrm_prm_apply_compose prm_unit_involution by metis thus ?thesis using ptrm_prm_apply_id by metis qed from 1 and 2 show "[a \ b] \ X \ Y \ X \ [a \ b] \ Y" by blast qed lemma ptrm_alpha_equiv_fvs_transfer: assumes "A \ [a \ b] \ B" and "a \ ptrm_fvs B" shows "b \ ptrm_fvs A" proof - from \A \ [a \ b] \ B\ have "[a \ b] \ A \ B" using ptrm_swp_transfer by metis hence "ptrm_fvs B = [a \ b] {$} ptrm_fvs A" using ptrm_alpha_equiv_fvs ptrm_prm_fvs by metis hence "a \ [a \ b] {$} ptrm_fvs A" using \a \ ptrm_fvs B\ by metis hence "b \ [a \ b] {$} ([a \ b] {$} ptrm_fvs A)" using prm_set_notmembership prm_unit_action by metis thus ?thesis using prm_set_apply_compose prm_unit_involution prm_set_id by metis qed lemma ptrm_prm_agreement_equiv: assumes "\a. a \ ds \ \ \ a \ ptrm_fvs M" shows "\ \ M \ \ \ M" using assms proof(induction M arbitrary: \ \) case (PUnit) thus ?case using ptrm_alpha_equiv.unit by simp next case (PVar x) consider "x \ ds \ \" | "x \ ds \ \" by auto thus ?case proof(cases) case 1 hence "x \ ptrm_fvs (PVar x)" using PVar.prems by blast thus ?thesis by simp next case 2 hence "\ $ x = \ $ x" using prm_disagreement_ext by metis thus ?thesis using ptrm_alpha_equiv.var by simp next qed next case (PApp A B) hence "\ \ A \ \ \ A" and "\ \ B \ \ \ B" by auto thus ?case using ptrm_alpha_equiv.app by auto next case (PFn x T A) consider "x \ ds \ \" | "x \ ds \ \" by auto thus ?case proof(cases) case 1 hence *: "\ $ x = \ $ x" using prm_disagreement_ext by metis have "\a. a \ ds \ \ \ a \ ptrm_fvs A" proof - fix a assume "a \ ds \ \" hence "a \ ptrm_fvs (PFn x T A)" using PFn.prems by metis hence "a = x \ a \ ptrm_fvs A" by auto thus "a \ ptrm_fvs A" using \a \ ds \ \\ \x \ ds \ \\ by auto qed thus ?thesis using PFn.IH * ptrm_alpha_equiv.fn1 ptrm_apply_prm.simps(3) by fastforce next case 2 hence "\ $ x \ \ $ x" using prm_disagreement_def CollectD by fastforce moreover have "\ $ x \ ptrm_fvs (\ \ A)" proof - have "y \ (ptrm_fvs A) \ \ $ x \ \ $ y" for y using PFn \\ $ x \ \ $ x\ prm_apply_unequal prm_disagreement_ext ptrm_fvs.simps(4) by (metis Diff_iff empty_iff insert_iff) hence "\ $ x \ \ {$} ptrm_fvs A" unfolding prm_set_def by auto thus ?thesis using ptrm_prm_fvs by metis qed moreover have "\ \ A \ [\ $ x \ \ $ x] \ \ \ A" proof - have "\a. a \ ds \ ([\ $ x \ \ $ x] \ \) \ a \ ptrm_fvs A" proof - fix a assume *: "a \ ds \ ([\ $ x \ \ $ x] \ \)" hence "a \ x" using \\ $ x \ \ $ x\ using prm_apply_composition prm_disagreement_ext prm_unit_action prm_unit_commutes by metis hence "a \ ds \ \" using * prm_apply_composition prm_apply_unequal prm_disagreement_ext prm_unit_inaction by metis thus "a \ ptrm_fvs A" using \a \ x\ PFn.prems by auto qed thus ?thesis using PFn by (simp add: ptrm_prm_apply_compose) qed ultimately show ?thesis using ptrm_alpha_equiv.fn2 by simp next qed next case (PPair A B) hence "\ \ A \ \ \ A" and "\ \ B \ \ \ B" by auto thus ?case using ptrm_alpha_equiv.pair by auto next case (PFst P) hence "\ \ P \ \ \ P" by auto thus ?case using ptrm_alpha_equiv.fst by auto next case (PSnd P) hence "\ \ P \ \ \ P" by auto thus ?case using ptrm_alpha_equiv.snd by auto next qed lemma ptrm_prm_unit_inaction: assumes "a \ ptrm_fvs X" "b \ ptrm_fvs X" shows "[a \ b] \ X \ X" proof - have "(\x. x \ ds [a \ b] \ \ x \ ptrm_fvs X)" proof - fix x assume "x \ ds [a \ b] \" hence "[a \ b] $ x \ \ $ x" unfolding prm_disagreement_def by auto hence "x = a \ x = b" using prm_apply_id prm_unit_inaction by metis thus "x \ ptrm_fvs X" using assms by auto qed hence "[a \ b] \ X \ \ \ X" using ptrm_prm_agreement_equiv by metis thus ?thesis using ptrm_prm_apply_id by metis qed lemma ptrm_alpha_equiv_reflexive: shows "M \ M" by(induction M, auto simp add: ptrm_alpha_equiv.intros) corollary ptrm_alpha_equiv_reflp: shows "reflp ptrm_alpha_equiv" unfolding reflp_def using ptrm_alpha_equiv_reflexive by auto lemma ptrm_alpha_equiv_symmetric: assumes "X \ Y" shows "Y \ X" using assms proof(induction rule: ptrm_alpha_equiv.induct, auto simp add: ptrm_alpha_equiv.intros) case (fn2 a b A B T) have "b \ a" using \a \ b\ by auto have "B \ [b \ a] \ A" using \[a \ b] \ B \ A\ using ptrm_swp_transfer prm_unit_commutes by metis have "b \ ptrm_fvs A" using \a \ ptrm_fvs B\ \A \ [a \ b] \ B\ \a \ b\ using ptrm_alpha_equiv_fvs_transfer by metis show ?case using \b \ a\ \B \ [b \ a] \ A\ \b \ ptrm_fvs A\ using ptrm_alpha_equiv.fn2 by metis next qed corollary ptrm_alpha_equiv_symp: shows "symp ptrm_alpha_equiv" unfolding symp_def using ptrm_alpha_equiv_symmetric by auto lemma ptrm_alpha_equiv_transitive: assumes "X \ Y" and "Y \ Z" shows "X \ Z" using assms proof(induction "size X" arbitrary: X Y Z rule: less_induct) fix X Y Z :: "'a ptrm" assume IH: "\K Y Z :: 'a ptrm. size K < size X \ K \ Y \ Y \ Z \ K \ Z" assume "X \ Y" and "Y \ Z" show "X \ Z" proof(cases X) case (PUnit) hence "Y = PUnit" using \X \ Y\ unitE by metis hence "Z = PUnit" using \Y \ Z\ unitE by metis thus ?thesis using ptrm_alpha_equiv.unit \X = PUnit\ by metis next case (PVar x) hence "PVar x \ Y" using \X \ Y\ by auto hence "Y = PVar x" using varE by metis hence "PVar x \ Z" using \Y \ Z\ by auto hence "Z = PVar x" using varE by metis thus ?thesis using ptrm_alpha_equiv.var \X = PVar x\ by metis next case (PApp A B) obtain C D where "Y = PApp C D" and "A \ C" and "B \ D" using appE \X = PApp A B\ \X \ Y\ by metis hence "PApp C D \ Z" using \Y \ Z\ by simp from this obtain E F where "Z = PApp E F" and "C \ E" and "D \ F" using appE by metis have "size A < size X" and "size B < size X" using \X = PApp A B\ by auto hence "A \ E" and "B \ F" using IH \A \ C\ \C \ E\ \B \ D\ \D \ F\ by auto thus ?thesis using \X = PApp A B\ \Z = PApp E F\ ptrm_alpha_equiv.app by metis next case (PFn x T A) from this have X: "X = PFn x T A". hence *: "size A < size X" by auto obtain y B where "Y = PFn y T B" and Y_cases: "(x = y \ A \ B) \ (x \ y \ A \ [x \ y] \ B \ x \ ptrm_fvs B)" using fnE \X \ Y\ \X = PFn x T A\ by metis obtain z C where Z: "Z = PFn z T C" and Z_cases: "(y = z \ B \ C) \ (y \ z \ B \ [y \ z] \ C \ y \ ptrm_fvs C)" using fnE \Y \ Z\ \Y = PFn y T B\ by metis consider "x = y" "A \ B" and "y = z" "B \ C" | "x = y" "A \ B" and "y \ z" "B \ [y \ z] \ C" "y \ ptrm_fvs C" | "x \ y" "A \ [x \ y] \ B" "x \ ptrm_fvs B" and "y = z" "B \ C" | "x \ y" "A \ [x \ y] \ B" "x \ ptrm_fvs B" and "y \ z" "B \ [y \ z] \ C" "y \ ptrm_fvs C" and "x = z" | "x \ y" "A \ [x \ y] \ B" "x \ ptrm_fvs B" and "y \ z" "B \ [y \ z] \ C" "y \ ptrm_fvs C" and "x \ z" using Y_cases Z_cases by auto thus ?thesis proof(cases) case 1 have "A \ C" using \A \ B\ \B \ C\ IH * by metis have "x = z" using \x = y\ \y = z\ by metis show ?thesis using \A \ C\ \x = z\ X Z using ptrm_alpha_equiv.fn1 by metis next case 2 have "x \ z" using \x = y\ \y \ z\ by metis have "A \ [x \ z] \ C" using \A \ B\ \B \ [y \ z] \ C\ \x = y\ IH * by metis have "x \ ptrm_fvs C" using \x = y\ \y \ ptrm_fvs C\ by metis thus ?thesis using \x \ z\ \A \ [x \ z] \ C\ \x \ ptrm_fvs C\ X Z using ptrm_alpha_equiv.fn2 by metis next case 3 have "x \ z" using \x \ y\ \y = z\ by metis have "[x \ y] \ B \ [x \ y] \ C" using \B \ C\ ptrm_alpha_equiv_prm by metis have "A \ [x \ z] \ C" using \A \ [x \ y] \ B\ \[x \ y] \ B \ [x \ y] \ C\ \y = z\ IH * by metis have "x \ ptrm_fvs C" using \B \ C\ \x \ ptrm_fvs B\ ptrm_alpha_equiv_fvs by metis thus ?thesis using \x \ z\ \A \ [x \ z] \ C\ \x \ ptrm_fvs C\ X Z using ptrm_alpha_equiv.fn2 by metis next case 4 have "[x \ y] \ B \ [x \ y] \ [y \ z] \ C" using \B \ [y \ z] \ C\ ptrm_alpha_equiv_prm by metis hence "A \ [x \ y] \ [y \ z] \ C" using \A \ [x \ y] \ B\ IH * by metis hence "A \ ([x \ y] \ [y \ z]) \ C" using ptrm_prm_apply_compose by metis hence "A \ ([x \ y] \ [y \ x]) \ C" using \x = z\ by metis hence "A \ ([x \ y] \ [x \ y]) \ C" using prm_unit_commutes by metis hence "A \ \ \ C" using \x = z\ prm_unit_involution by metis hence "A \ C" using ptrm_prm_apply_id by metis thus ?thesis using \x = z\ \A \ C\ X Z using ptrm_alpha_equiv.fn1 by metis next case 5 have "x \ ptrm_fvs C" proof - have "ptrm_fvs B = ptrm_fvs ([y \ z] \ C)" using ptrm_alpha_equiv_fvs \B \ [y \ z] \ C\ by metis hence "x \ ptrm_fvs ([y \ z] \ C)" using \x \ ptrm_fvs B\ by auto hence "x \ [y \ z] {$} ptrm_fvs C" using ptrm_prm_fvs by metis consider "z \ ptrm_fvs C" | "z \ ptrm_fvs C" by auto thus ?thesis proof(cases) case 1 hence "[y \ z] {$} ptrm_fvs C = ptrm_fvs C - {z} \ {y}" using prm_set_unit_action prm_unit_commutes and \z \ ptrm_fvs C\ \y \ ptrm_fvs C\ by metis hence "x \ ptrm_fvs C - {z} \ {y}" using \x \ [y \ z] {$} ptrm_fvs C\ by auto hence "x \ ptrm_fvs C - {z}" using \x \ y\ by auto thus ?thesis using \x \ z\ by auto next case 2 hence "[y \ z] {$} ptrm_fvs C = ptrm_fvs C" using prm_set_unit_inaction \y \ ptrm_fvs C\ by metis thus ?thesis using \x \ [y \ z] {$} ptrm_fvs C\ by auto next qed qed have "A \ [x \ z] \ C" proof - have "A \ ([x \ y] \ [y \ z]) \ C" using IH * \A \ [x \ y] \ B\ \B \ [y \ z] \ C\ and ptrm_alpha_equiv_prm ptrm_prm_apply_compose by metis have "([x \ y] \ [y \ z]) \ C \ [x \ z] \ C" proof - have "ds ([x \ y] \ [y \ z]) [x \ z] = {x, y}" using \x \ y\ \y \ z\ \x \ z\ prm_disagreement_composition by metis hence "\a. a \ ds ([x \ y] \ [y \ z]) [x \ z] \ a \ ptrm_fvs C" using \x \ ptrm_fvs C\ \y \ ptrm_fvs C\ using Diff_iff Diff_insert_absorb insert_iff by auto thus ?thesis using ptrm_prm_agreement_equiv by metis qed thus ?thesis using IH * using \A \ ([x \ y] \ [y \ z]) \ C\ \([x \ y] \ [y \ z]) \ C \ [x \ z] \ C\ by metis qed show ?thesis using \x \ z\ \A \ [x \ z] \ C\ \x \ ptrm_fvs C\ X Z using ptrm_alpha_equiv.fn2 by metis next qed next case (PPair A B) obtain C D where "Y = PPair C D" and "A \ C" and "B \ D" using pairE \X = PPair A B\ \X \ Y\ by metis hence "PPair C D \ Z" using \Y \ Z\ by simp from this obtain E F where "Z = PPair E F" and "C \ E" and "D \ F" using pairE by metis have "size A < size X" and "size B < size X" using \X = PPair A B\ by auto hence "A \ E" and "B \ F" using IH \A \ C\ \C \ E\ \B \ D\ \D \ F\ by auto thus ?thesis using \X = PPair A B\ \Z = PPair E F\ ptrm_alpha_equiv.pair by metis next case (PFst P) obtain Q where "Y = PFst Q" "P \ Q" using fstE \X = PFst P\ \X \ Y\ by metis obtain R where "Z = PFst R" "Q \ R" using fstE \Y = PFst Q\ \Y \ Z\ by metis have "size P < size X" using \X = PFst P\ by auto hence "P \ R" using IH \P \ Q\ \Q \ R\ by metis thus ?thesis using \X = PFst P\ \Z = PFst R\ ptrm_alpha_equiv.fst by metis next case (PSnd P) obtain Q where "Y = PSnd Q" "P \ Q" using sndE \X = PSnd P\ \X \ Y\ by metis obtain R where "Z = PSnd R" "Q \ R" using sndE \Y = PSnd Q\ \Y \ Z\ by metis have "size P < size X" using \X = PSnd P\ by auto hence "P \ R" using IH \P \ Q\ \Q \ R\ by metis thus ?thesis using \X = PSnd P\ \Z = PSnd R\ ptrm_alpha_equiv.snd by metis next qed qed corollary ptrm_alpha_equiv_transp: shows "transp ptrm_alpha_equiv" unfolding transp_def using ptrm_alpha_equiv_transitive by auto type_synonym 'a typing_ctx = "'a \ type" fun ptrm_infer_type :: "'a typing_ctx \ 'a ptrm \ type option" where "ptrm_infer_type \ PUnit = Some TUnit" | "ptrm_infer_type \ (PVar x) = \ x" | "ptrm_infer_type \ (PApp A B) = (case (ptrm_infer_type \ A, ptrm_infer_type \ B) of (Some (TArr \ \), Some \') \ (if \ = \' then Some \ else None) | _ \ None )" | "ptrm_infer_type \ (PFn x \ A) = (case ptrm_infer_type (\(x \ \)) A of Some \ \ Some (TArr \ \) | None \ None )" | "ptrm_infer_type \ (PPair A B) = (case (ptrm_infer_type \ A, ptrm_infer_type \ B) of (Some \, Some \) \ Some (TPair \ \) | _ \ None )" | "ptrm_infer_type \ (PFst P) = (case ptrm_infer_type \ P of (Some (TPair \ \)) \ Some \ | _ \ None )" | "ptrm_infer_type \ (PSnd P) = (case ptrm_infer_type \ P of (Some (TPair \ \)) \ Some \ | _ \ None )" lemma ptrm_infer_type_swp_types: assumes "a \ b" shows "ptrm_infer_type (\(a \ T, b \ S)) X = ptrm_infer_type (\(a \ S, b \ T)) ([a \ b] \ X)" using assms proof(induction X arbitrary: T S \) case (PUnit) thus ?case by simp next case (PVar x) consider "x = a" | "x = b" | "x \ a \ x \ b" by auto thus ?case proof(cases) assume "x = a" thus ?thesis using \a \ b\ by (simp add: prm_unit_action) next assume "x = b" thus ?thesis using \a \ b\ using prm_unit_action prm_unit_commutes fun_upd_same fun_upd_twist by (metis ptrm_apply_prm.simps(2) ptrm_infer_type.simps(2)) next assume "x \ a \ x \ b" thus ?thesis by (simp add: prm_unit_inaction) next qed next case (PApp A B) thus ?case by simp next case (PFn x \ A) hence *: "ptrm_infer_type (\(a \ T, b \ S)) A = ptrm_infer_type (\(a \ S, b \ T)) ([a \ b] \ A)" for T S \ by metis consider "x = a" | "x = b" | "x \ a \ x \ b" by auto thus ?case proof(cases) case 1 hence "ptrm_infer_type (\(a \ S, b \ T)) ([a \ b] \ PFn x \ A) = ptrm_infer_type (\(a \ S, b \ T)) (PFn b \ ([a \ b] \ A))" using prm_unit_action ptrm_apply_prm.simps(4) by metis moreover have "... = (case ptrm_infer_type (\(a \ S, b \ \)) ([a \ b] \ A) of None \ None | Some \ \ Some (TArr \ \))" by simp moreover have "... = (case ptrm_infer_type (\(a \ \, b \ S)) A of None \ None | Some \ \ Some (TArr \ \))" using * by metis moreover have "... = (case ptrm_infer_type (\(b \ S, a \ T, a \ \)) A of None \ None | Some \ \ Some (TArr \ \))" using \a \ b\ fun_upd_twist fun_upd_upd by metis moreover have "... = ptrm_infer_type (\(b \ S, a \ T)) (PFn x \ A)" using \x = a\ by simp moreover have "... = ptrm_infer_type (\(a \ T, b \ S)) (PFn x \ A)" using \a \ b\ fun_upd_twist by metis ultimately show ?thesis by metis next case 2 hence "ptrm_infer_type (\(a \ S, b \ T)) ([a \ b] \ PFn x \ A) = ptrm_infer_type (\(a \ S, b \ T)) (PFn a \ ([a \ b] \ A))" using prm_unit_action prm_unit_commutes ptrm_apply_prm.simps(4) by metis - moreover have "... = (case ptrm_infer_type (\(a \ S, b \ T)(a \ \)) ([a \ b] \ A) of None \ None | Some \ \ Some (TArr \ \))" + moreover have "... = (case ptrm_infer_type (\(a \ S, b \ T, a \ \)) ([a \ b] \ A) of None \ None | Some \ \ Some (TArr \ \))" by simp moreover have "... = (case ptrm_infer_type (\(a \ \, b \ T)) ([a \ b] \ A) of None \ None | Some \ \ Some (TArr \ \))" using fun_upd_upd fun_upd_twist \a \ b\ by metis moreover have "... = (case ptrm_infer_type (\(a \ T, b \ \)) A of None \ None | Some \ \ Some (TArr \ \))" using * by metis - moreover have "... = (case ptrm_infer_type (\(a \ T, b \ S)(b \ \)) A of None \ None | Some \ \ Some (TArr \ \))" + moreover have "... = (case ptrm_infer_type (\(a \ T, b \ S, b \ \)) A of None \ None | Some \ \ Some (TArr \ \))" using \a \ b\ fun_upd_upd by metis moreover have "... = ptrm_infer_type (\(b \ S, a \ T)) (PFn x \ A)" using \x = b\ by simp moreover have "... = ptrm_infer_type (\(a \ T, b \ S)) (PFn x \ A)" using \a \ b\ fun_upd_twist by metis ultimately show ?thesis by metis next case 3 hence "x \ a" "x \ b" by auto hence "ptrm_infer_type (\(a \ S, b \ T)) ([a \ b] \ PFn x \ A) = ptrm_infer_type (\(a \ S, b \ T)) (PFn x \ ([a \ b] \ A))" by (simp add: prm_unit_inaction) - moreover have "... = (case ptrm_infer_type (\(a \ S, b \ T)(x \ \)) ([a \ b] \ A) of None \ None | Some \ \ Some (TArr \ \))" + moreover have "... = (case ptrm_infer_type (\(a \ S, b \ T, x \ \)) ([a \ b] \ A) of None \ None | Some \ \ Some (TArr \ \))" by simp moreover have "... = (case ptrm_infer_type (\(x \ \, a \ S, b \ T)) ([a \ b] \ A) of None \ None | Some \ \ Some (TArr \ \))" using \x \ a\ \x \ b\ fun_upd_twist by metis moreover have "... = (case ptrm_infer_type (\(x \ \, a \ T, b \ S)) A of None \ None | Some \ \ Some (TArr \ \))" using * by metis moreover have "... = (case ptrm_infer_type (\(a \ T, b \ S, x \ \)) A of None \ None | Some \ \ Some (TArr \ \))" using \x \ a\ \x \ b\ fun_upd_twist by metis moreover have "... = ptrm_infer_type (\(a \ T, b \ S)) (PFn x \ A)" by simp ultimately show ?thesis by metis next qed next case (PPair A B) thus ?case by simp next case (PFst P) thus ?case by simp next case (PSnd P) thus ?case by simp next qed lemma ptrm_infer_type_swp: assumes "a \ b" "b \ ptrm_fvs X" shows "ptrm_infer_type (\(a \ \)) X = ptrm_infer_type (\(b \ \)) ([a \ b] \ X)" using assms proof(induction X arbitrary: \ \) case (PUnit) thus ?case by simp next case (PVar x) hence "x \ b" by simp consider "x = a" | "x \ a" by auto thus ?case proof(cases) case 1 hence "[a \ b] \ (PVar x) = PVar b" and "ptrm_infer_type (\(a \ \)) (PVar x) = Some \" using prm_unit_action by auto thus ?thesis by auto next case 2 hence *: "[a \ b] \ PVar x = PVar x" using \x \ b\ prm_unit_inaction by simp consider "\\. \ x = Some \" | "\ x = None" by auto thus ?thesis proof(cases) assume "\\. \ x = Some \" from this obtain \ where "\ x = Some \" by auto thus ?thesis using * \x \ a\ \x \ b\ by auto next assume "\ x = None" thus ?thesis using * \x \ a\ \x \ b\ by auto qed next qed next case (PApp A B) have "b \ ptrm_fvs A" and "b \ ptrm_fvs B" using PApp.prems by auto hence "ptrm_infer_type (\(a \ \)) A = ptrm_infer_type (\(b \ \)) ([a \ b] \ A)" and "ptrm_infer_type (\(a \ \)) B = ptrm_infer_type (\(b \ \)) ([a \ b] \ B)" using PApp.IH assms by metis+ thus ?case by (metis ptrm_apply_prm.simps(3) ptrm_infer_type.simps(3)) next case (PFn x \ A) consider "b \ x \ b \ ptrm_fvs A" | "b = x" using PFn.prems by auto thus ?case proof(cases) case 1 hence "b \ x" "b \ ptrm_fvs A" by auto hence *: "\\ \. ptrm_infer_type (\(a \ \)) A = ptrm_infer_type (\(b \ \)) ([a \ b] \ A)" using PFn.IH assms by metis consider "a = x" | "a \ x" by auto thus ?thesis proof(cases) case 1 hence "ptrm_infer_type (\(a \ \)) (PFn x \ A) = ptrm_infer_type (\(a \ \)) (PFn a \ A)" and " ptrm_infer_type (\(b \ \)) ([a \ b] \ PFn x \ A) = ptrm_infer_type (\(b \ \)) (PFn b \ ([a \ b] \ A))" by (auto simp add: prm_unit_action) thus ?thesis using * ptrm_infer_type.simps(4) fun_upd_upd by metis next case 2 hence "ptrm_infer_type (\(b \ \)) ([a \ b] \ PFn x \ A) = ptrm_infer_type (\(b \ \)) (PFn x \ ([a \ b] \ A))" using \b \ x\ by (simp add: prm_unit_inaction) moreover have "... = (case ptrm_infer_type (\(b \ \, x \ \)) ([a \ b] \ A) of None \ None | Some \' \ Some (TArr \ \'))" by simp moreover have "... = (case ptrm_infer_type (\(x \ \, b \ \)) ([a \ b] \ A) of None \ None | Some \' \ Some (TArr \ \'))" using \b \ x\ fun_upd_twist by metis moreover have "... = (case ptrm_infer_type (\(x \ \, a \ \)) A of None \ None | Some \' \ Some (TArr \ \'))" using * by metis moreover have "... = (case ptrm_infer_type (\(a \ \, x \ \)) A of None \ None | Some \' \ Some (TArr \ \'))" using \a \ x\ fun_upd_twist by metis moreover have "... = ptrm_infer_type (\(a \ \)) (PFn x \ A)" by simp ultimately show ?thesis by metis next qed next case 2 hence "a \ x" using assms by auto have " - ptrm_infer_type (\(a \ \)(b \ \)) A = - ptrm_infer_type (\(b \ \)(a \ \)) ([a \ b] \ A) + ptrm_infer_type (\(a \ \, b \ \)) A = + ptrm_infer_type (\(b \ \, a \ \)) ([a \ b] \ A) " using ptrm_infer_type_swp_types using \a \ b\ fun_upd_twist by metis thus ?thesis using \b = x\ prm_unit_action prm_unit_commutes using ptrm_infer_type.simps(4) ptrm_apply_prm.simps(4) by metis next qed next case (PPair A B) thus ?case by simp next case (PFst P) thus ?case by simp next case (PSnd P) thus ?case by simp next qed lemma ptrm_infer_type_alpha_equiv: assumes "X \ Y" shows "ptrm_infer_type \ X = ptrm_infer_type \ Y" using assms proof(induction arbitrary: \) case (fn2 a b A B T \) hence "ptrm_infer_type (\(a \ T)) A = ptrm_infer_type (\(b \ T)) B" using ptrm_infer_type_swp prm_unit_commutes by metis thus ?case by simp next qed auto end diff --git a/thys/Name_Carrying_Type_Inference/SimplyTyped.thy b/thys/Name_Carrying_Type_Inference/SimplyTyped.thy --- a/thys/Name_Carrying_Type_Inference/SimplyTyped.thy +++ b/thys/Name_Carrying_Type_Inference/SimplyTyped.thy @@ -1,3264 +1,3264 @@ theory SimplyTyped imports PreSimplyTyped begin quotient_type 'a trm = "'a ptrm" / ptrm_alpha_equiv proof(rule equivpI) show "reflp ptrm_alpha_equiv" using ptrm_alpha_equiv_reflp. show "symp ptrm_alpha_equiv" using ptrm_alpha_equiv_symp. show "transp ptrm_alpha_equiv" using ptrm_alpha_equiv_transp. qed lift_definition Unit :: "'a trm" is PUnit. lift_definition Var :: "'a \ 'a trm" is PVar. lift_definition App :: "'a trm \ 'a trm \ 'a trm" is PApp using ptrm_alpha_equiv.app. lift_definition Fn :: "'a \ type \ 'a trm \ 'a trm" is PFn using ptrm_alpha_equiv.fn1. lift_definition Pair :: "'a trm \ 'a trm \ 'a trm" is PPair using ptrm_alpha_equiv.pair. lift_definition Fst :: "'a trm \ 'a trm" is PFst using ptrm_alpha_equiv.fst. lift_definition Snd :: "'a trm \ 'a trm" is PSnd using ptrm_alpha_equiv.snd. lift_definition fvs :: "'a trm \ 'a set" is ptrm_fvs using ptrm_alpha_equiv_fvs. lift_definition prm :: "'a prm \ 'a trm \ 'a trm" (infixr "\" 150) is ptrm_apply_prm using ptrm_alpha_equiv_prm. lift_definition depth :: "'a trm \ nat" is size using ptrm_size_alpha_equiv. lemma depth_prm: shows "depth (\ \ A) = depth A" by(transfer, metis ptrm_size_prm) lemma depth_app: shows "depth A < depth (App A B)" "depth B < depth (App A B)" by(transfer, auto)+ lemma depth_fn: shows "depth A < depth (Fn x T A)" by(transfer, auto) lemma depth_pair: shows "depth A < depth (Pair A B)" "depth B < depth (Pair A B)" by(transfer, auto)+ lemma depth_fst: shows "depth P < depth (Fst P)" by(transfer, auto) lemma depth_snd: shows "depth P < depth (Snd P)" by(transfer, auto) lemma unit_not_var: shows "Unit \ Var x" proof(transfer) fix x :: 'a show "\ PUnit \ PVar x" proof(rule classical) assume "\\ PUnit \ PVar x" hence False using unitE by fastforce thus ?thesis by blast qed qed lemma unit_not_app: shows "Unit \ App A B" proof(transfer) fix A B :: "'a ptrm" show "\ PUnit \ PApp A B" proof(rule classical) assume "\\ PUnit \ PApp A B" hence False using unitE by fastforce thus ?thesis by blast qed qed lemma unit_not_fn: shows "Unit \ Fn x T A" proof(transfer) fix x :: 'a and T A show "\ PUnit \ PFn x T A" proof(rule classical) assume "\\ PUnit \ PFn x T A" hence False using unitE by fastforce thus ?thesis by blast qed qed lemma unit_not_pair: shows "Unit \ Pair A B" proof(transfer) fix A B :: "'a ptrm" show "\ PUnit \ PPair A B" proof(rule classical) assume "\\ PUnit \ PPair A B" hence False using unitE by fastforce thus ?thesis by blast qed qed lemma unit_not_fst: shows "Unit \ Fst P" proof(transfer) fix P :: "'a ptrm" show "\ PUnit \ PFst P" proof(rule classical) assume "\\ PUnit \ PFst P" hence False using unitE by fastforce thus ?thesis by blast qed qed lemma unit_not_snd: shows "Unit \ Snd P" proof(transfer) fix P :: "'a ptrm" show "\ PUnit \ PSnd P" proof(rule classical) assume "\\ PUnit \ PSnd P" hence False using unitE by fastforce thus ?thesis by blast qed qed lemma var_not_app: shows "Var x \ App A B" proof(transfer) fix x :: 'a and A B show "\PVar x \ PApp A B" proof(rule classical) assume "\\PVar x \ PApp A B" hence False using varE by fastforce thus ?thesis by blast qed qed lemma var_not_fn: shows "Var x \ Fn y T A" proof(transfer) fix x y :: 'a and T A show "\PVar x \ PFn y T A" proof(rule classical) assume "\\PVar x \ PFn y T A" hence False using varE by fastforce thus ?thesis by blast qed qed lemma var_not_pair: shows "Var x \ Pair A B" proof(transfer) fix x :: 'a and A B show "\PVar x \ PPair A B" proof(rule classical) assume "\\PVar x \ PPair A B" hence False using varE by fastforce thus ?thesis by blast qed qed lemma var_not_fst: shows "Var x \ Fst P" proof(transfer) fix x :: 'a and P show "\PVar x \ PFst P" proof(rule classical) assume "\\PVar x \ PFst P" hence False using varE by fastforce thus ?thesis by blast qed qed lemma var_not_snd: shows "Var x \ Snd P" proof(transfer) fix x :: 'a and P show "\PVar x \ PSnd P" proof(rule classical) assume "\\PVar x \ PSnd P" hence False using varE by fastforce thus ?thesis by blast qed qed lemma app_not_fn: shows "App A B \ Fn y T X" proof(transfer) fix y :: 'a and A B T X show "\PApp A B \ PFn y T X" proof(rule classical) assume "\\PApp A B \ PFn y T X" hence False using appE by auto thus ?thesis by blast qed qed lemma app_not_pair: shows "App A B \ Pair C D" proof(transfer) fix A B C D :: "'a ptrm" show "\PApp A B \ PPair C D" proof(rule classical) assume "\\PApp A B \ PPair C D" hence False using appE by auto thus ?thesis by blast qed qed lemma app_not_fst: shows "App A B \ Fst P" proof(transfer) fix A B P :: "'a ptrm" show "\PApp A B \ PFst P" proof(rule classical) assume "\\PApp A B \ PFst P" hence False using appE by auto thus ?thesis by blast qed qed lemma app_not_snd: shows "App A B \ Snd P" proof(transfer) fix A B P :: "'a ptrm" show "\PApp A B \ PSnd P" proof(rule classical) assume "\\PApp A B \ PSnd P" hence False using appE by auto thus ?thesis by blast qed qed lemma fn_not_pair: shows "Fn x T A \ Pair C D" proof(transfer) fix x :: 'a and T A C D show "\PFn x T A \ PPair C D" proof(rule classical) assume "\\PFn x T A \ PPair C D" hence False using fnE by fastforce thus ?thesis by blast qed qed lemma fn_not_fst: shows "Fn x T A \ Fst P" proof(transfer) fix x :: 'a and T A P show "\PFn x T A \ PFst P" proof(rule classical) assume "\\PFn x T A \ PFst P" hence False using fnE by fastforce thus ?thesis by blast qed qed lemma fn_not_snd: shows "Fn x T A \ Snd P" proof(transfer) fix x :: 'a and T A P show "\PFn x T A \ PSnd P" proof(rule classical) assume "\\PFn x T A \ PSnd P" hence False using fnE by fastforce thus ?thesis by blast qed qed lemma pair_not_fst: shows "Pair A B \ Fst P" proof(transfer) fix A B P :: "'a ptrm" show "\PPair A B \ PFst P" proof(rule classical) assume "\\PPair A B \ PFst P" hence False using pairE by auto thus ?thesis by blast qed qed lemma pair_not_snd: shows "Pair A B \ Snd P" proof(transfer) fix A B P :: "'a ptrm" show "\PPair A B \ PSnd P" proof(rule classical) assume "\\PPair A B \ PSnd P" hence False using pairE by auto thus ?thesis by blast qed qed lemma fst_not_snd: shows "Fst P \ Snd Q" proof(transfer) fix P Q :: "'a ptrm" show "\PFst P \ PSnd Q" proof(rule classical) assume "\\PFst P \ PSnd Q" hence False using fstE by auto thus ?thesis by blast qed qed lemma trm_simp: shows "Var x = Var y \ x = y" "App A B = App C D \ A = C" "App A B = App C D \ B = D" "Fn x T A = Fn y S B \ (x = y \ T = S \ A = B) \ (x \ y \ T = S \ x \ fvs B \ A = [x \ y] \ B)" "Pair A B = Pair C D \ A = C" "Pair A B = Pair C D \ B = D" "Fst P = Fst Q \ P = Q" "Snd P = Snd Q \ P = Q" proof - show "Var x = Var y \ x = y" by (transfer, insert ptrm.inject varE, fastforce) show "App A B = App C D \ A = C" by (transfer, insert ptrm.inject appE, auto) show "App A B = App C D \ B = D" by (transfer, insert ptrm.inject appE, auto) show "Pair A B = Pair C D \ A = C" by (transfer, insert ptrm.inject pairE, auto) show "Pair A B = Pair C D \ B = D" by (transfer, insert ptrm.inject pairE, auto) show "Fst P = Fst Q \ P = Q" by (transfer, insert ptrm.inject fstE, auto) show "Snd P = Snd Q \ P = Q" by (transfer, insert ptrm.inject sndE, auto) show "Fn x T A = Fn y S B \ (x = y \ T = S \ A = B) \ (x \ y \ T = S \ x \ fvs B \ A = [x \ y] \ B)" proof(transfer') fix x y :: 'a and T S :: type and A B :: "'a ptrm" assume *: "PFn x T A \ PFn y S B" thus "x = y \ T = S \ A \ B \ x \ y \ T = S \ x \ ptrm_fvs B \ A \ [x \ y] \ B" proof(induction rule: fnE[where x=x and T=T and A=A and Y="PFn y S B"], metis *) case (2 C) thus ?case by simp next case (3 z C) thus ?case by simp next qed qed qed lemma fn_eq: assumes "x \ y" "x \ fvs B" "A = [x \ y] \ B" shows "Fn x T A = Fn y T B" using assms by(transfer', metis ptrm_alpha_equiv.fn2) lemma trm_prm_simp: shows "\ \ Unit = Unit" "\ \ Var x = Var (\ $ x)" "\ \ App A B = App (\ \ A) (\ \ B)" "\ \ Fn x T A = Fn (\ $ x) T (\ \ A)" "\ \ Pair A B = Pair (\ \ A) (\ \ B)" "\ \ Fst P = Fst (\ \ P)" "\ \ Snd P = Snd (\ \ P)" apply (transfer, auto simp add: ptrm_alpha_equiv_reflexive) apply (transfer', auto simp add: ptrm_alpha_equiv_reflexive) apply ((transfer, auto simp add: ptrm_alpha_equiv_reflexive)+) done lemma trm_prm_apply_compose: shows "\ \ \ \ A = (\ \ \) \ A" by(transfer', metis ptrm_prm_apply_compose ptrm_alpha_equiv_reflexive) lemma fvs_finite: shows "finite (fvs M)" by(transfer, metis ptrm_fvs_finite) lemma fvs_simp: shows "fvs Unit = {}" and "fvs (Var x) = {x}" "fvs (App A B) = fvs A \ fvs B" "fvs (Fn x T A) = fvs A - {x}" "fvs (Pair A B) = fvs A \ fvs B" "fvs (Fst P) = fvs P" "fvs (Snd P) = fvs P" by((transfer, simp)+) lemma var_prm_action: shows "[a \ b] \ Var a = Var b" by(transfer', simp add: prm_unit_action ptrm_alpha_equiv.intros) lemma var_prm_inaction: assumes "a \ x" "b \ x" shows "[a \ b] \ Var x = Var x" using assms by(transfer', simp add: prm_unit_inaction ptrm_alpha_equiv.intros) lemma trm_prm_apply_id: shows "\ \ M = M" by(transfer', auto simp add: ptrm_prm_apply_id) lemma trm_prm_unit_inaction: assumes "a \ fvs X" "b \ fvs X" shows "[a \ b] \ X = X" using assms by(transfer', metis ptrm_prm_unit_inaction) lemma trm_prm_agreement_equiv: assumes "\a. a \ ds \ \ \ a \ fvs M" shows "\ \ M = \ \ M" using assms by(transfer, metis ptrm_prm_agreement_equiv) lemma trm_induct: fixes P :: "'a trm \ bool" assumes "P Unit" "\x. P (Var x)" "\A B. \P A; P B\ \ P (App A B)" "\x T A. P A \ P (Fn x T A)" "\A B. \P A; P B\ \ P (Pair A B)" "\A. P A \ P (Fst A)" "\A. P A \ P (Snd A)" shows "P M" proof - have "\X. P (abs_trm X)" proof(rule ptrm.induct) show "P (abs_trm PUnit)" using assms(1) Unit.abs_eq by metis show "P (abs_trm (PVar x))" for x using assms(2) Var.abs_eq by metis show "\P (abs_trm A); P (abs_trm B)\ \ P (abs_trm (PApp A B))" for A B using assms(3) App.abs_eq by metis show "P (abs_trm A) \ P (abs_trm (PFn x T A))" for x T A using assms(4) Fn.abs_eq by metis show "\P (abs_trm A); P (abs_trm B)\ \ P (abs_trm (PPair A B))" for A B using assms(5) Pair.abs_eq by metis show "P (abs_trm A) \ P (abs_trm (PFst A))" for A using assms(6) Fst.abs_eq by metis show "P (abs_trm A) \ P (abs_trm (PSnd A))" for A using assms(7) Snd.abs_eq by metis qed thus ?thesis using trm.abs_induct by auto qed lemma trm_cases: assumes "M = Unit \ P M" "\x. M = Var x \ P M" "\A B. M = App A B \ P M" "\x T A. M = Fn x T A \ P M" "\A B. M = Pair A B \ P M" "\A. M = Fst A \ P M" "\A. M = Snd A \ P M" shows "P M" using assms by(induction rule: trm_induct, auto) lemma trm_depth_induct: assumes "P Unit" "\x. P (Var x)" "\A B. \\K. depth K < depth (App A B) \ P K\ \ P (App A B)" "\M x T A. (\K. depth K < depth (Fn x T A) \ P K) \ P (Fn x T A)" "\A B. \\K. depth K < depth (Pair A B) \ P K\ \ P (Pair A B)" "\A. \\K. depth K < depth (Fst A) \ P K\ \ P (Fst A)" "\A. \\K. depth K < depth (Snd A) \ P K\ \ P (Snd A)" shows "P M" proof(induction "depth M" arbitrary: M rule: less_induct) fix M :: "'a trm" assume IH: "depth K < depth M \ P K" for K hence " M = Unit \ P M" "\x. M = Var x \ P M" "\A B. M = App A B \ P M" "\x T A. M = Fn x T A \ P M" "\A B. M = Pair A B \ P M" "\A. M = Fst A \ P M" "\A. M = Snd A \ P M" using assms by blast+ thus "P M" using trm_cases[where M=M] by blast qed context fresh begin lemma fresh_fn: fixes x :: 'a and S :: "'a set" assumes "finite S" shows "\y B. y \ S \ B = [y \ x] \ A \ (Fn x T A = Fn y T B)" proof - have *: "finite ({x} \ fvs A \ S)" using fvs_finite assms by auto obtain y where "y = fresh_in ({x} \ fvs A \ S)" by auto hence "y \ ({x} \ fvs A \ S)" using fresh_axioms * unfolding class.fresh_def by metis hence "y \ x" "y \ fvs A" "y \ S" by auto obtain B where B: "B = [y \ x] \ A" by auto hence "Fn x T A = Fn y T B" using fn_eq \y \ x\ \y \ fvs A\ by metis thus ?thesis using \y \ x\ \y \ S\ B by metis qed lemma trm_strong_induct: fixes P :: "'a set \ 'a trm \ bool" assumes "P S Unit" "\x. P S (Var x)" "\A B. \P S A; P S B\ \ P S (App A B)" "\x T. x \ S \ (\A. P S A \ P S (Fn x T A))" "\A B. \P S A; P S B\ \ P S (Pair A B)" "\A. P S A \ P S (Fst A)" "\A. P S A \ P S (Snd A)" "finite S" shows "P S M" proof - have "\\. P S (\ \ M)" proof(induction M rule: trm_induct) case 1 thus ?case using assms(1) trm_prm_simp(1) by metis next case (2 x) thus ?case using assms(2) trm_prm_simp(2) by metis next case (3 A B) thus ?case using assms(3) trm_prm_simp(3) by metis next case (4 x T A) have "finite S" "finite (fvs (\ \ A))" "finite {\ $ x}" using \finite S\ fvs_finite by auto hence "finite (S \ fvs (\ \ A) \ {\ $ x})" by auto obtain y where "y = fresh_in (S \ fvs (\ \ A) \ {\ $ x})" by auto hence "y \ S \ fvs (\ \ A) \ {\ $ x}" using fresh_axioms unfolding class.fresh_def using \finite (S \ fvs (\ \ A) \ {\ $ x})\ by metis hence "y \ \ $ x" "y \ fvs (\ \ A)" "y \ S" by auto hence *: "\A. P S A \ P S (Fn y T A)" using assms(4) by metis have "P S (([y \ \ $ x] \ \) \ A)" using 4 by metis hence "P S (Fn y T (([y \ \ $ x] \ \) \ A))" using * by metis moreover have "(Fn y T (([y \ \ $ x] \ \) \ A)) = Fn (\ $ x) T (\ \ A)" using trm_prm_apply_compose fn_eq \y \ \ $ x\ \y \ fvs (\ \ A)\ by metis ultimately show ?case using trm_prm_simp(4) by metis next case (5 A B) thus ?case using assms(5) trm_prm_simp(5) by metis next case (6 A) thus ?case using assms(6) trm_prm_simp(6) by metis next case (7 A) thus ?case using assms(7) trm_prm_simp(7) by metis next qed hence "P S (\ \ M)" by metis thus "P S M" using trm_prm_apply_id by metis qed lemma trm_strong_cases: fixes P :: "'a set \ 'a trm \ bool" assumes " M = Unit \ P S M" "\x. M = Var x \ P S M" "\A B. M = App A B \ P S M" "\x T A. M = Fn x T A \ x \ S \ P S M" "\A B. M = Pair A B \ P S M" "\A. M = Fst A \ P S M" "\A. M = Snd A \ P S M" "finite S" shows "P S M" using assms by(induction S M rule: trm_strong_induct, metis+) lemma trm_strong_depth_induct: fixes P :: "'a set \ 'a trm \ bool" assumes "P S Unit" "\x. P S (Var x)" "\A B. \\K. depth K < depth (App A B) \ P S K\ \ P S (App A B)" "\x T. x \ S \ (\A. (\K. depth K < depth (Fn x T A) \ P S K) \ P S (Fn x T A))" "\A B. \\K. depth K < depth (Pair A B) \ P S K\ \ P S (Pair A B)" "\A. \\K. depth K < depth (Fst A) \ P S K\ \ P S (Fst A)" "\A. \\K. depth K < depth (Snd A) \ P S K\ \ P S (Snd A)" "finite S" shows "P S M" proof(induction "depth M" arbitrary: M rule: less_induct) fix M :: "'a trm" assume IH: "depth K < depth M \ P S K" for K hence " M = Unit \ P S M" "\x. M = Var x \ P S M" "\A B. M = App A B \ P S M" "\x T A. M = Fn x T A \ x \ S \ P S M" "\A B. M = Pair A B \ P S M" "\A. M = Fst A \ P S M" "\A. M = Snd A \ P S M" "finite S" using assms IH by metis+ thus "P S M" using trm_strong_cases[where M=M] by blast qed lemma trm_prm_fvs: shows "fvs (\ \ M) = \ {$} fvs M" by(transfer, metis ptrm_prm_fvs) inductive typing :: "'a typing_ctx \ 'a trm \ type \ bool" ("_ \ _ : _") where tunit: "\ \ Unit : TUnit" | tvar: "\ x = Some \ \ \ \ Var x : \" | tapp: "\\ \ f : (TArr \ \); \ \ x : \\ \ \ \ App f x : \" | tfn: "\(x \ \) \ A : \ \ \ \ Fn x \ A : (TArr \ \)" | tpair: "\\ \ A : \; \ \ B : \\ \ \ \ Pair A B : (TPair \ \)" | tfst: "\ \ P : (TPair \ \) \ \ \ Fst P : \" | tsnd: "\ \ P : (TPair \ \) \ \ \ Snd P : \" lemma typing_prm: assumes "\ \ M : \" "\y. y \ fvs M \ \ y = \ (\ $ y)" shows "\ \ \ \ M : \" using assms proof(induction arbitrary: \ rule: typing.induct) case (tunit \) thus ?case using typing.tunit trm_prm_simp(1) by metis next case (tvar \ x \) thus ?case using typing.tvar trm_prm_simp(2) fvs_simp(2) singletonI by metis next case (tapp \ A \ \ B) thus ?case using typing.tapp trm_prm_simp(3) fvs_simp(3) UnCI by metis next case (tfn \ x \ A \) have "y \ fvs A \ (\(x \ \)) y = (\(\ $ x \ \)) (\ $ y)" for y proof(cases "y = x") case True thus ?thesis using fun_upd_apply by simp next case False assume "y \ fvs A" hence "y \ fvs (Fn x \ A)" using fvs_simp(4) \y \ x\ DiffI singletonD by fastforce hence "\ y = \ (\ $ y)" using tfn.prems by metis thus ?thesis by (simp add: prm_apply_unequal) next qed hence "\(\ $ x \ \) \ \ \ A : \" using tfn.IH by metis thus ?case using trm_prm_simp(4) typing.tfn by metis next case (tpair \ A B) thus ?case using typing.tpair trm_prm_simp(5) fvs_simp(5) UnCI by metis next case (tfst \ P \ \) thus ?case using typing.tfst trm_prm_simp(6) fvs_simp(6) by metis next case (tsnd \ P \ \) thus ?case using typing.tsnd trm_prm_simp(7) fvs_simp(7) by metis next qed lemma typing_swp: assumes "\(a \ \) \ M : \" "b \ fvs M" shows "\(b \ \) \ [a \ b] \ M : \" proof - have "y \ fvs M \ (\(a \ \)) y = (\(b \ \)) ([a \ b] $ y)" for y proof - assume "y \ fvs M" hence "y \ b" using assms(2) by auto consider "y = a" | "y \ a" by auto thus "(\(a \ \)) y = (\(b \ \)) ([a \ b] $ y)" by(cases, simp add: prm_unit_action, simp add: prm_unit_inaction \y \ b\) qed thus ?thesis using typing_prm assms(1) by metis qed lemma typing_unitE: assumes "\ \ Unit : \" shows "\ = TUnit" using assms apply cases apply blast apply (auto simp add: unit_not_var unit_not_app unit_not_fn unit_not_pair unit_not_fst unit_not_snd) done lemma typing_varE: assumes "\ \ Var x : \" shows "\ x = Some \" using assms apply cases prefer 2 apply (metis trm_simp(1)) apply (metis unit_not_var) apply (auto simp add: var_not_app var_not_fn var_not_pair var_not_fst var_not_snd) done lemma typing_appE: assumes "\ \ App A B : \" shows "\\. (\ \ A : (TArr \ \)) \ (\ \ B : \)" using assms apply cases prefer 3 apply (metis trm_simp(2, 3)) apply (metis unit_not_app) apply (metis var_not_app) apply (auto simp add: app_not_fn app_not_pair app_not_fst app_not_snd) done lemma typing_fnE: assumes "\ \ Fn x T A : \" shows "\\. \ = (TArr T \) \ (\(x \ T) \ A : \)" using assms proof(cases) case (tfn y S B \) from this consider "x = y \ T = S \ A = B" | "x \ y \ T = S \ x \ fvs B \ A = [x \ y] \ B" using trm_simp(4) by metis thus ?thesis proof(cases) case 1 thus ?thesis using tfn by metis next case 2 thus ?thesis using tfn typing_swp prm_unit_commutes by metis next qed next qed ( metis unit_not_fn, metis var_not_fn, metis app_not_fn, metis fn_not_pair, metis fn_not_fst, metis fn_not_snd ) lemma typing_pairE: assumes "\ \ Pair A B : \" shows "\\ \. \ = (TPair \ \) \ (\ \ A : \) \ (\ \ B : \)" using assms proof(cases) case (tpair A \ B \) thus ?thesis using trm_simp(5) trm_simp(6) by blast next qed ( metis unit_not_pair, metis var_not_pair, metis app_not_pair, metis fn_not_pair, metis pair_not_fst, metis pair_not_snd ) lemma typing_fstE: assumes "\ \ Fst P : \" shows "\\. (\ \ P : (TPair \ \))" using assms proof(cases) case (tfst P \) thus ?thesis using trm_simp(7) by blast next qed ( metis unit_not_fst, metis var_not_fst, metis app_not_fst, metis fn_not_fst, metis pair_not_fst, metis fst_not_snd ) lemma typing_sndE: assumes "\ \ Snd P : \" shows "\\. (\ \ P : (TPair \ \))" using assms proof(cases) case (tsnd P \) thus ?thesis using trm_simp(8) by blast next qed ( metis unit_not_snd, metis var_not_snd, metis app_not_snd, metis fn_not_snd, metis pair_not_snd, metis fst_not_snd ) theorem typing_weaken: assumes "\ \ M : \" "y \ fvs M" shows "\(y \ \) \ M : \" using assms proof(induction rule: typing.induct) case (tunit \) thus ?case using typing.tunit by metis next case (tvar \ x \) hence "y \ x" using fvs_simp(2) singletonI by force hence "(\(y \ \)) x = Some \" using tvar.hyps fun_upd_apply by simp thus ?case using typing.tvar by metis next case (tapp \ f \ \' x) from \y \ fvs (App f x)\ have "y \ fvs f" "y \ fvs x" using fvs_simp(3) Un_iff by force+ hence "\(y \ \) \ f : (TArr \ \')" "\(y \ \) \ x : \" using tapp.IH by metis+ thus ?case using typing.tapp by metis next case (tfn \ x \ A \') from \y \ fvs (Fn x \ A)\ consider "y = x" | "y \ x \ y \ fvs A" using fvs_simp(4) DiffI empty_iff insert_iff by fastforce thus ?case proof(cases) case 1 - hence "(\(y \ \)(x \ \)) \ A : \'" using tfn.hyps fun_upd_upd by simp + hence "(\(y \ \, x \ \)) \ A : \'" using tfn.hyps fun_upd_upd by simp thus ?thesis using typing.tfn by metis next case 2 hence "y \ x" "y \ fvs A" by auto hence "\(x \ \, y \ \) \ A : \'" using tfn.IH by metis hence "\(y \ \, x \ \) \ A : \'" using \y \ x\ fun_upd_twist by metis thus ?thesis using typing.tfn by metis next qed next case (tpair \ A \ B \) thus ?case using typing.tpair fvs_simp(5) UnCI by metis next case (tfst \ P \ \) thus ?case using typing.tfst fvs_simp(6) by metis next case (tsnd \ P \ \) thus ?case using typing.tsnd fvs_simp(7) by metis next qed lift_definition infer :: "'a typing_ctx \ 'a trm \ type option" is ptrm_infer_type using ptrm_infer_type_alpha_equiv. export_code infer fresh_nat_inst.fresh_in_nat in Haskell lemma infer_simp: shows "infer \ Unit = Some TUnit" "infer \ (Var x) = \ x" "infer \ (App A B) = (case (infer \ A, infer \ B) of (Some (TArr \ \), Some \') \ (if \ = \' then Some \ else None) | _ \ None )" "infer \ (Fn x \ A) = (case infer (\(x \ \)) A of Some \ \ Some (TArr \ \) | None \ None )" "infer \ (Pair A B) = (case (infer \ A, infer \ B) of (Some \, Some \) \ Some (TPair \ \) | _ \ None )" "infer \ (Fst P) = (case infer \ P of (Some (TPair \ \)) \ Some \ | _ \ None )" "infer \ (Snd P) = (case infer \ P of (Some (TPair \ \)) \ Some \ | _ \ None )" by((transfer, simp)+) lemma infer_unitE: assumes "infer \ Unit = Some \" shows "\ = TUnit" using assms by(transfer, simp) lemma infer_varE: assumes "infer \ (Var x) = Some \" shows "\ x = Some \" using assms by(transfer, simp) lemma infer_appE: assumes "infer \ (App A B) = Some \" shows "\\. infer \ A = Some (TArr \ \) \ infer \ B = Some \" using assms proof(transfer) fix \ :: "'a typing_ctx" and A B \ assume H: "ptrm_infer_type \ (PApp A B) = Some \" have "ptrm_infer_type \ A \ None" proof(rule classical, auto) assume "ptrm_infer_type \ A = None" hence "ptrm_infer_type \ (PApp A B) = None" by auto thus False using H by auto qed from this obtain T where *: "ptrm_infer_type \ A = Some T" by auto have "T \ TVar x" for x proof(rule classical, auto) fix x assume "T = TVar x" hence "ptrm_infer_type \ A = Some (TVar x)" using * by metis hence "ptrm_infer_type \ (PApp A B) = None" by simp thus False using H by auto qed moreover have "T \ TUnit" proof(rule classical, auto) fix x assume "T = TUnit" hence "ptrm_infer_type \ A = Some TUnit" using * by metis hence "ptrm_infer_type \ (PApp A B) = None" by simp thus False using H by auto qed moreover have "T \ TPair \ \" for \ \ proof(rule classical, auto) fix \ \ assume "T = TPair \ \" hence "ptrm_infer_type \ A = Some (TPair \ \)" using * by metis hence "ptrm_infer_type \ (PApp A B) = None" by simp thus False using H by auto qed ultimately obtain \ \' where "T = TArr \ \'" by(cases T, blast, auto) hence *: "ptrm_infer_type \ A = Some (TArr \ \')" using * by metis have "ptrm_infer_type \ B \ None" proof(rule classical, auto) assume "ptrm_infer_type \ B = None" hence "ptrm_infer_type \ (PApp A B) = None" using * by auto thus False using H by auto qed from this obtain \' where **: "ptrm_infer_type \ B = Some \'" by auto have "\ = \'" proof(rule classical) assume "\ \ \'" hence "ptrm_infer_type \ (PApp A B) = None" using * ** by simp hence False using H by auto thus "\ = \'" by blast qed hence **: "ptrm_infer_type \ B = Some \" using ** by auto have "ptrm_infer_type \ (PApp A B) = Some \'" using * ** by auto hence "\ = \'" using H by auto hence *: "ptrm_infer_type \ A = Some (TArr \ \)" using * by auto show "\\. ptrm_infer_type \ A = Some (TArr \ \) \ ptrm_infer_type \ B = Some \" using * ** by auto qed lemma infer_fnE: assumes "infer \ (Fn x T A) = Some \" shows "\\. \ = TArr T \ \ infer (\(x \ T)) A = Some \" using assms proof(transfer) fix x :: 'a and \ T A \ assume H: "ptrm_infer_type \ (PFn x T A) = Some \" have "ptrm_infer_type (\(x \ T)) A \ None" proof(rule classical, auto) assume "ptrm_infer_type (\(x \ T)) A = None" hence "ptrm_infer_type \ (PFn x T A) = None" by auto thus False using H by auto qed from this obtain \ where *: "ptrm_infer_type (\(x \ T)) A = Some \" by auto have "ptrm_infer_type \ (PFn x T A) = Some (TArr T \)" using * by auto hence "\ = TArr T \" using H by auto thus "\\. \ = TArr T \ \ ptrm_infer_type (\(x \ T)) A = Some \" using * by auto qed lemma infer_pairE: assumes "infer \ (Pair A B) = Some \" shows "\T S. \ = TPair T S \ infer \ A = Some T \ infer \ B = Some S" using assms proof(transfer) fix A B :: "'a ptrm" and \ \ assume H: "ptrm_infer_type \ (PPair A B) = Some \" have "ptrm_infer_type \ A \ None" proof(rule classical, auto) assume "ptrm_infer_type \ A = None" hence "ptrm_infer_type \ (PPair A B) = None" by auto thus False using H by auto qed moreover have "ptrm_infer_type \ B \ None" proof(rule classical, auto) assume "ptrm_infer_type \ B = None" hence "ptrm_infer_type \ (PPair A B) = None" by (simp add: option.case_eq_if) thus False using H by auto qed ultimately obtain T S where "\ = TPair T S" "ptrm_infer_type \ A = Some T" "ptrm_infer_type \ B = Some S" using H by auto thus "\T S. \ = TPair T S \ ptrm_infer_type \ A = Some T \ ptrm_infer_type \ B = Some S" by auto qed lemma infer_fstE: assumes "infer \ (Fst P) = Some \" shows "\T S. infer \ P = Some (TPair T S) \ \ = T" using assms proof(transfer) fix P :: "'a ptrm" and \ \ assume H: "ptrm_infer_type \ (PFst P) = Some \" have "ptrm_infer_type \ P \ None" proof(rule classical, auto) assume "ptrm_infer_type \ P = None" thus False using H by simp qed moreover have "ptrm_infer_type \ P \ Some TUnit" proof(rule classical, auto) assume "ptrm_infer_type \ P = Some TUnit" thus False using H by simp qed moreover have "ptrm_infer_type \ P \ Some (TVar x)" for x proof(rule classical, auto) assume "ptrm_infer_type \ P = Some (TVar x)" thus False using H by simp qed moreover have "ptrm_infer_type \ P \ Some (TArr T S)" for T S proof(rule classical, auto) assume "ptrm_infer_type \ P = Some (TArr T S)" thus False using H by simp qed ultimately obtain T S where "ptrm_infer_type \ P = Some (TPair T S)" using type.distinct type.exhaust option.exhaust by metis moreover hence "ptrm_infer_type \ (PFst P) = Some T" by simp ultimately show "\T S. ptrm_infer_type \ P = Some (TPair T S) \ \ = T" using H by auto qed lemma infer_sndE: assumes "infer \ (Snd P) = Some \" shows "\T S. infer \ P = Some (TPair T S) \ \ = S" using assms proof(transfer) fix P :: "'a ptrm" and \ \ assume H: "ptrm_infer_type \ (PSnd P) = Some \" have "ptrm_infer_type \ P \ None" proof(rule classical, auto) assume "ptrm_infer_type \ P = None" thus False using H by simp qed moreover have "ptrm_infer_type \ P \ Some TUnit" proof(rule classical, auto) assume "ptrm_infer_type \ P = Some TUnit" thus False using H by simp qed moreover have "ptrm_infer_type \ P \ Some (TVar x)" for x proof(rule classical, auto) assume "ptrm_infer_type \ P = Some (TVar x)" thus False using H by simp qed moreover have "ptrm_infer_type \ P \ Some (TArr T S)" for T S proof(rule classical, auto) assume "ptrm_infer_type \ P = Some (TArr T S)" thus False using H by simp qed ultimately obtain T S where "ptrm_infer_type \ P = Some (TPair T S)" using type.distinct type.exhaust option.exhaust by metis moreover hence "ptrm_infer_type \ (PSnd P) = Some S" by simp ultimately show "\T S. ptrm_infer_type \ P = Some (TPair T S) \ \ = S" using H by auto qed lemma infer_sound: assumes "infer \ M = Some \" shows "\ \ M : \" using assms proof(induction M arbitrary: \ \ rule: trm_induct) case 1 thus ?case using infer_unitE typing.tunit by metis next case (2 x) hence "\ x = Some \" using infer_varE by metis thus ?case using typing.tvar by metis next case (3 A B) from \infer \ (App A B) = Some \\ obtain \ where "infer \ A = Some (TArr \ \)" and "infer \ B = Some \" using infer_appE by metis thus ?case using "3.IH" typing.tapp by metis next case (4 x T A \ \) from \infer \ (Fn x T A) = Some \\ obtain \ where "\ = TArr T \" and "infer (\(x \ T)) A = Some \" using infer_fnE by metis thus ?case using "4.IH" typing.tfn by metis next case (5 A B \ \) thus ?case using typing.tpair infer_pairE by metis next case (6 P \ \) thus ?case using typing.tfst infer_fstE by metis next case (7 P \ \) thus ?case using typing.tsnd infer_sndE by metis next qed lemma infer_complete: assumes "\ \ M : \" shows "infer \ M = Some \" using assms proof(induction) case (tfn \ x \ A \) thus ?case by (simp add: infer_simp(4) tfn.IH) next qed (auto simp add: infer_simp) theorem infer_valid: shows "(\ \ M : \) = (infer \ M = Some \)" using infer_sound infer_complete by blast inductive substitutes :: "'a trm \ 'a \ 'a trm \ 'a trm \ bool" where unit: "substitutes Unit y M Unit" | var1: "x = y \ substitutes (Var x) y M M" | var2: "x \ y \ substitutes (Var x) y M (Var x)" | app: "\substitutes A x M A'; substitutes B x M B'\ \ substitutes (App A B) x M (App A' B')" | fn: "\x \ y; y \ fvs M; substitutes A x M A'\ \ substitutes (Fn y T A) x M (Fn y T A')" | pair: "\substitutes A x M A'; substitutes B x M B'\ \ substitutes (Pair A B) x M (Pair A' B')" | fst: "substitutes P x M P' \ substitutes (Fst P) x M (Fst P')" | snd: "substitutes P x M P' \ substitutes (Snd P) x M (Snd P')" lemma substitutes_prm: assumes "substitutes A x M A'" shows "substitutes (\ \ A) (\ $ x) (\ \ M) (\ \ A')" using assms proof(induction) case (unit y M) thus ?case using substitutes.unit trm_prm_simp(1) by metis case (var1 x y M) thus ?case using substitutes.var1 trm_prm_simp(2) by metis next case (var2 x y M) thus ?case using substitutes.var2 trm_prm_simp(2) prm_apply_unequal by metis next case (app A x M A' B B') thus ?case using substitutes.app trm_prm_simp(3) by metis next case (fn x y M A A' T) thus ?case using substitutes.fn trm_prm_simp(4) prm_apply_unequal prm_set_notmembership trm_prm_fvs by metis next case (pair A x M A' B B') thus ?case using substitutes.pair trm_prm_simp(5) by metis next case (fst P x M P') thus ?case using substitutes.fst trm_prm_simp(6) by metis next case (snd P x M P') thus ?case using substitutes.snd trm_prm_simp(7) by metis next qed lemma substitutes_fvs: assumes "substitutes A x M A'" shows "fvs A' \ fvs A - {x} \ fvs M" using assms proof(induction) case (unit y M) thus ?case using fvs_simp(1) by auto case (var1 x y M) thus ?case by auto next case (var2 x y M) thus ?case using fvs_simp(2) Un_subset_iff Un_upper2 insert_Diff_if insert_is_Un singletonD sup_commute by metis next case (app A x M A' B B') hence "fvs A' \ fvs B' \ (fvs A - {x} \ fvs M) \ (fvs B - {x} \ fvs M)" by auto hence "fvs A' \ fvs B' \ (fvs A \ fvs B) - {x} \ fvs M" by auto thus ?case using fvs_simp(3) by metis next case (fn x y M A A' T) hence "fvs A' - {y} \ fvs A - {y} - {x} \ fvs M" by auto thus ?case using fvs_simp(4) by metis next case (pair A x M A' B B') hence "fvs A' \ fvs B' \ (fvs A - {x} \ fvs M) \ (fvs B - {x} \ fvs M)" by auto hence "fvs A' \ fvs B' \ (fvs A \ fvs B) - {x} \ fvs M" by auto thus ?case using fvs_simp(5) by metis next case (fst P x M P') thus ?case using fvs_simp(6) by fastforce next case (snd P x M P') thus ?case using fvs_simp(7) by fastforce next qed inductive_cases substitutes_unitE': "substitutes Unit y M X" lemma substitutes_unitE: assumes "substitutes Unit y M X" shows "X = Unit" by( rule substitutes_unitE'[where y=y and M=M and X=X], metis assms, auto simp add: unit_not_var unit_not_app unit_not_fn unit_not_pair unit_not_fst unit_not_snd ) inductive_cases substitutes_varE': "substitutes (Var x) y M X" lemma substitutes_varE: assumes "substitutes (Var x) y M X" shows "(x = y \ M = X) \ (x \ y \ X = Var x)" by( rule substitutes_varE'[where x=x and y=y and M=M and X=X], metis assms, metis unit_not_var, metis trm_simp(1), metis trm_simp(1), auto simp add: var_not_app var_not_fn var_not_pair var_not_fst var_not_snd ) inductive_cases substitutes_appE': "substitutes (App A B) x M X" lemma substitutes_appE: assumes "substitutes (App A B) x M X" shows "\A' B'. substitutes A x M A' \ substitutes B x M B' \ X = App A' B'" by( cases rule: substitutes_appE'[where A=A and B=B and x=x and M=M and X=X], metis assms, metis unit_not_app, metis var_not_app, metis var_not_app, metis trm_simp(2,3), auto simp add: app_not_fn app_not_pair app_not_fst app_not_snd ) inductive_cases substitutes_fnE': "substitutes (Fn y T A) x M X" lemma substitutes_fnE: assumes "substitutes (Fn y T A) x M X" "y \ x" "y \ fvs M" shows "\A'. substitutes A x M A' \ X = Fn y T A'" using assms proof(induction rule: substitutes_fnE'[where y=y and T=T and A=A and x=x and M=M and X=X]) case (6 z B B' S) consider "y = z \ T = S \ A = B" | "y \ z \ T = S \ y \ fvs B \ A = [y \ z] \ B" using \Fn y T A = Fn z S B\ trm_simp(4) by metis thus ?case proof(cases) case 1 thus ?thesis using 6 by metis next case 2 hence "y \ z" "T = S" "y \ fvs B" "A = [y \ z] \ B" by auto have "substitutes ([y \ z] \ B) ([y \ z] $ x) ([y \ z] \ M) ([y \ z] \ B')" using substitutes_prm \substitutes B x M B'\ by metis hence "substitutes A ([y \ z] $ x) ([y \ z] \ M) ([y \ z] \ B')" using \A = [y \ z] \ B\ by metis hence "substitutes A x ([y \ z] \ M) ([y \ z] \ B')" using \y \ x\ \x \ z\ prm_unit_inaction by metis hence *: "substitutes A x M ([y \ z] \ B')" using \y \ fvs M\ \z \ fvs M\ trm_prm_unit_inaction by metis have "y \ fvs B'" using substitutes_fvs \substitutes B x M B'\ \y \ fvs B\ \y \ fvs M\ Diff_subset UnE rev_subsetD by blast hence "X = Fn y T ([y \ z] \ B')" using \X = Fn z S B'\ \y \ z\ \T = S\ fn_eq by metis thus ?thesis using * by auto next qed next qed ( metis assms(1), metis unit_not_fn, metis var_not_fn, metis var_not_fn, metis app_not_fn, metis fn_not_pair, metis fn_not_fst, metis fn_not_snd ) inductive_cases substitutes_pairE': "substitutes (Pair A B) x M X" lemma substitutes_pairE: assumes "substitutes (Pair A B) x M X" shows "\A' B'. substitutes A x M A' \ substitutes B x M B' \ X = Pair A' B'" proof(cases rule: substitutes_pairE'[where A=A and B=B and x=x and M=M and X=X]) case (7 A A' B B') thus ?thesis using trm_simp(5) trm_simp(6) by blast next qed ( metis assms, metis unit_not_pair, metis var_not_pair, metis var_not_pair, metis app_not_pair, metis fn_not_pair, metis pair_not_fst, metis pair_not_snd ) inductive_cases substitutes_fstE': "substitutes (Fst P) x M X" lemma substitutes_fstE: assumes "substitutes (Fst P) x M X" shows "\P'. substitutes P x M P' \ X = Fst P'" proof(cases rule: substitutes_fstE'[where P=P and x=x and M=M and X=X]) case (8 P P') thus ?thesis using trm_simp(7) by blast next qed ( metis assms, metis unit_not_fst, metis var_not_fst, metis var_not_fst, metis app_not_fst, metis fn_not_fst, metis pair_not_fst, metis fst_not_snd ) inductive_cases substitutes_sndE': "substitutes (Snd P) x M X" lemma substitutes_sndE: assumes "substitutes (Snd P) x M X" shows "\P'. substitutes P x M P' \ X = Snd P'" proof(cases rule: substitutes_sndE'[where P=P and x=x and M=M and X=X]) case (9 P P') thus ?thesis using trm_simp(8) by blast next qed ( metis assms, metis unit_not_snd, metis var_not_snd, metis var_not_snd, metis app_not_snd, metis fn_not_snd, metis pair_not_snd, metis fst_not_snd ) lemma substitutes_total: shows "\X. substitutes A x M X" proof(induction A rule: trm_strong_induct[where S="{x} \ fvs M"]) show "finite ({x} \ fvs M)" using fvs_finite by auto next case 1 obtain X :: "'a trm" where "X = Unit" by auto thus ?case using substitutes.unit by metis next case (2 y) consider "x = y" | "x \ y" by auto thus ?case proof(cases) case 1 obtain X where "X = M" by auto hence "substitutes (Var y) x M X" using \x = y\ substitutes.var1 by metis thus ?thesis by auto next case 2 obtain X where "X = (Var y)" by auto hence "substitutes (Var y) x M X" using \x \ y\ substitutes.var2 by metis thus ?thesis by auto next qed next case (3 A B) from this obtain A' B' where A': "substitutes A x M A'" and B': "substitutes B x M B'" by auto obtain X where "X = App A' B'" by auto hence "substitutes (App A B) x M X" using A' B' substitutes.app by metis thus ?case by auto next case (4 y T A) from this obtain A' where A': "substitutes A x M A'" by auto from \y \ ({x} \ fvs M)\ have "y \ x" "y \ fvs M" by auto obtain X where "X = Fn y T A'" by auto hence "substitutes (Fn y T A) x M X" using substitutes.fn \y \ x\ \y \ fvs M\ A' by metis thus ?case by auto next case (5 A B) from this obtain A' B' where "substitutes A x M A'" "substitutes B x M B'" by auto from this obtain X where "X = Pair A' B'" by auto hence "substitutes (Pair A B) x M X" using substitutes.pair \substitutes A x M A'\ \substitutes B x M B'\ by metis thus ?case by auto next case (6 P) from this obtain P' where "substitutes P x M P'" by auto from this obtain X where "X = Fst P'" by auto hence "substitutes (Fst P) x M X" using substitutes.fst \substitutes P x M P'\ by metis thus ?case by auto next case (7 P) from this obtain P' where "substitutes P x M P'" by auto from this obtain X where "X = Snd P'" by auto hence "substitutes (Snd P) x M X" using substitutes.snd \substitutes P x M P'\ by metis thus ?case by auto next qed lemma substitutes_unique: assumes "substitutes A x M B" "substitutes A x M C" shows "B = C" using assms proof(induction A arbitrary: B C rule: trm_strong_induct[where S="{x} \ fvs M"]) show "finite ({x} \ fvs M)" using fvs_finite by auto next case 1 thus ?case using substitutes_unitE by metis next case (2 y) thus ?case using substitutes_varE by metis next case (3 X Y) thus ?case using substitutes_appE by metis next case (4 y T A) hence "y \ x" and "y \ fvs M" by auto thus ?case using 4 substitutes_fnE by metis next case (5 A B C D) thus ?case using substitutes_pairE by metis next case (6 P Q R) thus ?case using substitutes_fstE by metis next case (7 P Q R) thus ?case using substitutes_sndE by metis next qed lemma substitutes_function: shows "\! X. substitutes A x M X" using substitutes_total substitutes_unique by metis definition subst :: "'a trm \ 'a \ 'a trm \ 'a trm" ("_[_ ::= _]") where "subst A x M \ (THE X. substitutes A x M X)" lemma subst_simp_unit: shows "Unit[x ::= M] = Unit" unfolding subst_def by(rule, metis substitutes.unit, metis substitutes_function substitutes.unit) lemma subst_simp_var1: shows "(Var x)[x ::= M] = M" unfolding subst_def by(rule, metis substitutes.var1, metis substitutes_function substitutes.var1) lemma subst_simp_var2: assumes "x \ y" shows "(Var x)[y ::= M] = Var x" unfolding subst_def by( rule, metis substitutes.var2 assms, metis substitutes_function substitutes.var2 assms ) lemma subst_simp_app: shows "(App A B)[x ::= M] = App (A[x ::= M]) (B[x ::= M])" unfolding subst_def proof obtain A' B' where A': "A' = (A[x ::= M])" and B': "B' = (B[x ::= M])" by auto hence "substitutes A x M A'" "substitutes B x M B'" unfolding subst_def using substitutes_function theI by metis+ hence "substitutes (App A B) x M (App A' B')" using substitutes.app by metis thus *: "substitutes (App A B) x M (App (THE X. substitutes A x M X) (THE X. substitutes B x M X))" using A' B' unfolding subst_def by metis fix X assume "substitutes (App A B) x M X" thus "X = App (THE X. substitutes A x M X) (THE X. substitutes B x M X)" using substitutes_function * by metis qed lemma subst_simp_fn: assumes "x \ y" "y \ fvs M" shows "(Fn y T A)[x ::= M] = Fn y T (A[x ::= M])" unfolding subst_def proof obtain A' where A': "A' = (A[x ::= M])" by auto hence "substitutes A x M A'" unfolding subst_def using substitutes_function theI by metis hence "substitutes (Fn y T A) x M (Fn y T A')" using substitutes.fn assms by metis thus *: "substitutes (Fn y T A) x M (Fn y T (THE X. substitutes A x M X))" using A' unfolding subst_def by metis fix X assume "substitutes (Fn y T A) x M X" thus "X = Fn y T (THE X. substitutes A x M X)" using substitutes_function * by metis qed lemma subst_simp_pair: shows "(Pair A B)[x ::= M] = Pair (A[x ::= M]) (B[x ::= M])" unfolding subst_def proof obtain A' B' where A': "A' = (A[x ::= M])" and B': "B' = (B[x ::= M])" by auto hence "substitutes A x M A'" "substitutes B x M B'" unfolding subst_def using substitutes_function theI by metis+ hence "substitutes (Pair A B) x M (Pair A' B')" using substitutes.pair by metis thus *: "substitutes (Pair A B) x M (Pair (THE X. substitutes A x M X) (THE X. substitutes B x M X))" using A' B' unfolding subst_def by metis fix X assume "substitutes (Pair A B) x M X" thus "X = Pair (THE X. substitutes A x M X) (THE X. substitutes B x M X)" using substitutes_function * by metis qed lemma subst_simp_fst: shows "(Fst P)[x ::= M] = Fst (P[x ::= M])" unfolding subst_def proof obtain P' where P': "P' = (P[x ::= M])" by auto hence "substitutes P x M P'" unfolding subst_def using substitutes_function theI by metis hence "substitutes (Fst P) x M (Fst P')" using substitutes.fst by metis thus *: "substitutes (Fst P) x M (Fst (THE X. substitutes P x M X))" using P' unfolding subst_def by metis fix X assume "substitutes (Fst P) x M X" thus "X = Fst (THE X. substitutes P x M X)" using substitutes_function * by metis qed lemma subst_simp_snd: shows "(Snd P)[x ::= M] = Snd (P[x ::= M])" unfolding subst_def proof obtain P' where P': "P' = (P[x ::= M])" by auto hence "substitutes P x M P'" unfolding subst_def using substitutes_function theI by metis hence "substitutes (Snd P) x M (Snd P')" using substitutes.snd by metis thus *: "substitutes (Snd P) x M (Snd (THE X. substitutes P x M X))" using P' unfolding subst_def by metis fix X assume "substitutes (Snd P) x M X" thus "X = Snd (THE X. substitutes P x M X)" using substitutes_function * by metis qed lemma subst_prm: shows "(\ \ (M[z ::= N])) = ((\ \ M)[\ $ z ::= \ \ N])" unfolding subst_def using substitutes_prm substitutes_function the1_equality by (metis (full_types)) lemma subst_fvs: shows "fvs (M[z ::= N]) \ (fvs M - {z}) \ fvs N" unfolding subst_def using substitutes_fvs substitutes_function theI2 by metis lemma subst_free: assumes "y \ fvs M" shows "M[y ::= N] = M" using assms proof(induction M rule: trm_strong_induct[where S="{y} \ fvs N"]) show "finite ({y} \ fvs N)" using fvs_finite by auto case 1 thus ?case using subst_simp_unit by metis next case (2 x) thus ?case using subst_simp_var2 by (simp add: fvs_simp) next case (3 A B) thus ?case using subst_simp_app by (simp add: fvs_simp) next case (4 x T A) hence "y \ x" "x \ fvs N" by auto have "y \ fvs A - {x}" using \y \ x\ \y \ fvs (Fn x T A)\ fvs_simp(4) by metis hence "y \ fvs A" using \y \ x\ by auto hence "A[y ::= N] = A" using "4.IH" by blast thus ?case using \y \ x\ \y \ fvs A\ \x \ fvs N\ subst_simp_fn by metis next case (5 A B) thus ?case using subst_simp_pair by (simp add: fvs_simp) next case (6 P) thus ?case using subst_simp_fst by (simp add: fvs_simp) next case (7 P) thus ?case using subst_simp_snd by (simp add: fvs_simp) next qed lemma subst_swp: assumes "y \ fvs A" shows "([y \ z] \ A)[y ::= M] = (A[z ::= M])" using assms proof(induction A rule: trm_strong_induct[where S="{y, z} \ fvs M"]) show "finite ({y, z} \ fvs M)" using fvs_finite by auto next case 1 thus ?case using trm_prm_simp(1) subst_simp_unit by metis next case (2 x) hence "y \ x" using fvs_simp(2) by force consider "x = z" | "x \ z" by auto thus ?case proof(cases) case 1 thus ?thesis using subst_simp_var1 trm_prm_simp(2) prm_unit_action prm_unit_commutes by metis next case 2 thus ?thesis using subst_simp_var2 trm_prm_simp(2) prm_unit_inaction \y \ x\ by metis next qed next case (3 A B) from \y \ fvs (App A B)\ have "y \ fvs A" "y \ fvs B" by (auto simp add: fvs_simp(3)) thus ?case using "3.IH" subst_simp_app trm_prm_simp(3) by metis next case (4 x T A) hence "x \ y" "x \ z" "x \ fvs M" by auto hence "y \ fvs A" using \y \ fvs (Fn x T A)\ fvs_simp(4) by force hence *: "([y \ z] \ A)[y ::= M] = (A[z ::= M])" using "4.IH" by metis have "([y \ z] \ Fn x T A)[y ::= M] = ((Fn ([y \ z] $ x) T ([y \ z] \ A))[y ::= M])" using trm_prm_simp(4) by metis also have "... = ((Fn x T ([y \ z] \ A))[y ::= M])" using prm_unit_inaction \x \ y\ \x \ z\ by metis also have "... = Fn x T (([y \ z] \ A)[y ::= M])" using subst_simp_fn \x \ y\ \x \ fvs M\ by metis also have "... = Fn x T (A[z ::= M])" using * by metis also have "... = ((Fn x T A)[z ::= M])" using subst_simp_fn \x \ z\ \x \ fvs M\ by metis finally show ?case. next case (5 A B) from \y \ fvs (Pair A B)\ have "y \ fvs A" "y \ fvs B" by (auto simp add: fvs_simp(5)) hence "([y \ z] \ A)[y ::= M] = (A[z ::= M])" "([y \ z] \ B)[y ::= M] = (B[z ::= M])" using "5.IH" by metis+ thus ?case using trm_prm_simp(5) subst_simp_pair by metis next case (6 P) from \y \ fvs (Fst P)\ have "y \ fvs P" by (simp add: fvs_simp(6)) hence "([y \ z] \ P)[y ::= M] = (P[z ::= M])" using "6.IH" by metis thus ?case using trm_prm_simp(6) subst_simp_fst by metis next case (7 P) from \y \ fvs (Snd P)\ have "y \ fvs P" by (simp add: fvs_simp(7)) hence "([y \ z] \ P)[y ::= M] = (P[z ::= M])" using "7.IH" by metis thus ?case using trm_prm_simp(7) subst_simp_snd by metis next qed lemma barendregt: assumes "y \ z" "y \ fvs L" shows "M[y ::= N][z ::= L] = (M[z ::= L][y ::= N[z ::= L]])" using assms proof(induction M rule: trm_strong_induct[where S="{y, z} \ fvs N \ fvs L"]) show "finite ({y, z} \ fvs N \ fvs L)" using fvs_finite by auto next case 1 thus ?case using subst_simp_unit by metis next case (2 x) consider "x = y" | "x = z" | "x \ y \ x \ z" by auto thus ?case proof(cases) case 1 hence "x = y" "x \ z" using assms(1) by auto thus ?thesis using subst_simp_var1 subst_simp_var2 by metis next case 2 hence "x \ y" "x = z" using assms(1) by auto thus ?thesis using subst_free \y \ fvs L\ subst_simp_var1 subst_simp_var2 by metis next case 3 then show ?thesis using subst_simp_var2 by metis next qed next case (3 A B) thus ?case using subst_simp_app by metis next case (4 x T A) hence *: "A[y ::= N][z ::= L] = (A[z ::= L][y ::= N[z ::= L]])" by blast from \x \ {y, z} \ fvs N \ fvs L\ have "x \ y" "x \ z" "x \ fvs N" "x \ fvs L" by auto hence "x \ fvs (N[z ::= L])" using subst_fvs by auto have "(Fn x T A)[y ::= N][z ::= L] = Fn x T (A[y ::= N][z ::= L])" using subst_simp_fn \x \ y\ \x \ z\ \x \ fvs N\ \x \ fvs L\ by metis also have "... = Fn x T (A[z ::= L][y ::= N[z ::= L]])" using * by metis also have "... = ((Fn x T A)[z ::= L][y ::= N[z ::= L]])" using subst_simp_fn \x \ y\ \x \ z\ \x \ fvs (N[z ::= L])\ \x \ fvs L\ by metis finally show ?case. next case (5 A B) thus ?case using subst_simp_pair by metis next case (6 P) thus ?case using subst_simp_fst by metis next case (7 P) thus ?case using subst_simp_snd by metis next qed lemma typing_subst: assumes "\(z \ \) \ M : \" "\ \ N : \" shows "\ \ M[z ::= N] : \" using assms proof(induction M arbitrary: \ \ rule: trm_strong_depth_induct[where S="{z} \ fvs N"]) show "finite ({z} \ fvs N)" using fvs_finite by auto next case 1 thus ?case using subst_simp_unit typing.tunit typing_unitE by metis next case (2 x) hence *: "(\(z \ \)) x = Some \" using typing_varE by metis consider "x = z" | "x \ z" by auto thus ?case proof(cases) case 1 hence "(\(x \ \)) x = Some \" using * by metis hence "\ = \" by auto thus ?thesis using \\ \ N : \\ subst_simp_var1 \x = z\ by metis next case 2 hence "\ x = Some \" using * by auto hence "\ \ Var x : \" using typing.tvar by metis thus ?thesis using \x \ z\ subst_simp_var2 by metis next qed next case (3 A B) have *: "depth A < depth (App A B) \ depth B < depth (App A B)" using depth_app by auto from \\(z \ \) \ App A B : \\ obtain \' where "\(z \ \) \ A : (TArr \' \)" "\(z \ \) \ B : \'" using typing_appE by metis hence "\ \ (A[z ::= N]) : (TArr \' \)" "\ \ (B[z ::= N]) : \'" using 3 * by metis+ hence "\ \ App (A[z ::= N]) (B[z ::= N]) : \" using typing.tapp by metis thus ?case using subst_simp_app by metis next case (4 x T A) hence "x \ z" "x \ fvs N" by auto hence *: "\(x \ T) \ N : \" using typing_weaken 4 by metis have **: "depth A < depth (Fn x T A)" using depth_fn. from \\(z \ \) \ Fn x T A : \\ obtain \' where "\ = TArr T \'" - "\(z \ \)(x \ T) \ A : \'" + "\(z \ \, x \ T) \ A : \'" using typing_fnE by metis - hence "\(x \ T)(z \ \) \ A : \'" using \x \ z\ fun_upd_twist by metis + hence "\(x \ T, z \ \) \ A : \'" using \x \ z\ fun_upd_twist by metis hence "\(x \ T) \ A[z ::= N] : \'" using 4 * ** by metis hence "\ \ Fn x T (A[z ::= N]) : \" using typing.tfn \\ = TArr T \'\ by metis thus ?case using \x \ z\ \x \ fvs N\ subst_simp_fn by metis next case (5 A B) from this obtain T S where "\ = TPair T S" "\(z \ \) \ A : T" "\(z \ \) \ B : S" using typing_pairE by metis moreover have "depth A < depth (Pair A B)" and "depth B < depth (Pair A B)" using depth_pair by auto ultimately have "\ \ A[z ::= N] : T" "\ \ B[z ::= N] : S" using 5 by metis+ hence "\ \ Pair (A[z ::= N]) (B[z ::= N]) : \" using \\ = TPair T S\ typing.tpair by metis thus ?case using subst_simp_pair by metis next case (6 P) from this obtain \' where "\(z \ \) \ P : (TPair \ \')" using typing_fstE by metis moreover have "depth P < depth (Fst P)" using depth_fst by metis ultimately have "\ \ P[z ::= N] : (TPair \ \')" using 6 by metis hence "\ \ Fst (P[z ::= N]) : \" using typing.tfst by metis thus ?case using subst_simp_fst by metis next case (7 P) from this obtain \' where "\(z \ \) \ P : (TPair \' \)" using typing_sndE by metis moreover have "depth P < depth (Snd P)" using depth_snd by metis ultimately have "\ \ P[z ::= N] : (TPair \' \)" using 7 by metis hence "\ \ Snd (P[z ::= N]) : \" using typing.tsnd by metis thus ?case using subst_simp_snd by metis next qed inductive beta_reduction :: "'a trm \ 'a trm \ bool" ("_ \\ _") where beta: "(App (Fn x T A) M) \\ (A[x ::= M])" | app1: "A \\ A' \ (App A B) \\ (App A' B)" | app2: "B \\ B' \ (App A B) \\ (App A B')" | fn: "A \\ A' \ (Fn x T A) \\ (Fn x T A')" | pair1: "A \\ A' \ (Pair A B) \\ (Pair A' B)" | pair2: "B \\ B' \ (Pair A B) \\ (Pair A B')" | fst1: "P \\ P' \ (Fst P) \\ (Fst P')" | fst2: "(Fst (Pair A B)) \\ A" | snd1: "P \\ P' \ (Snd P) \\ (Snd P')" | snd2: "(Snd (Pair A B)) \\ B" lemma beta_reduction_fvs: assumes "M \\ M'" shows "fvs M' \ fvs M" using assms proof(induction) case (beta x T A M) thus ?case using fvs_simp(3) fvs_simp(4) subst_fvs by metis next case (app1 A A' B) hence "fvs A' \ fvs B \ fvs A \ fvs B" by auto thus ?case using fvs_simp(3) by metis next case (app2 B B' A) hence "fvs A \ fvs B' \ fvs A \ fvs B" by auto thus ?case using fvs_simp(3) by metis next case (fn A A' x T) hence "fvs A' - {x} \ fvs A - {x}" by auto thus ?case using fvs_simp(4) by metis next case (pair1 A A' B) hence "fvs A' \ fvs B \ fvs A \ fvs B" by auto thus ?case using fvs_simp(5) by metis next case (pair2 B B' A) hence "fvs A \ fvs B' \ fvs A \ fvs B" by auto thus ?case using fvs_simp(5) by metis next case (fst1 P P') thus ?case using fvs_simp(6) by metis next case (fst2 A B) thus ?case using fvs_simp(5, 6) by force next case (snd1 P P') thus ?case using fvs_simp(7) by metis next case (snd2 A B) thus ?case using fvs_simp(5, 7) by force next qed lemma beta_reduction_prm: assumes "M \\ M'" shows "(\ \ M) \\ (\ \ M')" using assms by(induction, auto simp add: beta_reduction.intros trm_prm_simp subst_prm) lemma beta_reduction_left_unitE: assumes "Unit \\ M'" shows "False" using assms by(cases, auto simp add: unit_not_app unit_not_fn unit_not_pair unit_not_fst unit_not_snd) lemma beta_reduction_left_varE: assumes "(Var x) \\ M'" shows "False" using assms by(cases, auto simp add: var_not_app var_not_fn var_not_pair var_not_fst var_not_snd) lemma beta_reduction_left_appE: assumes "(App A B) \\ M'" shows " (\x T X. A = (Fn x T X) \ M' = (X[x ::= B])) \ (\A'. (A \\ A') \ M' = App A' B) \ (\B'. (B \\ B') \ M' = App A B') " using assms by( cases, metis trm_simp(2, 3), metis trm_simp(2, 3), metis trm_simp(2, 3), auto simp add: app_not_fn app_not_pair app_not_fst app_not_snd ) lemma beta_reduction_left_fnE: assumes "(Fn x T A) \\ M'" shows "\A'. (A \\ A') \ M' = (Fn x T A')" using assms proof(cases) case (fn B B' y S) consider "x = y \ T = S \ A = B" | "x \ y \ T = S \ x \ fvs B \ A = [x \ y] \ B" using trm_simp(4) \Fn x T A = Fn y S B\ by metis thus ?thesis proof(cases) case 1 thus ?thesis using fn by auto next case 2 thus ?thesis using fn beta_reduction_fvs beta_reduction_prm fn_eq by fastforce next qed next qed ( metis app_not_fn, metis app_not_fn, metis app_not_fn, auto simp add: fn_not_pair fn_not_fst fn_not_snd ) lemma beta_reduction_left_pairE: assumes "(Pair A B) \\ M'" shows "(\A'. (A \\ A') \ M' = (Pair A' B)) \ (\B'. (B \\ B') \ M' = (Pair A B'))" using assms apply cases prefer 5 apply (metis trm_simp(5, 6)) prefer 5 apply (metis trm_simp(5, 6)) apply (metis app_not_pair, metis app_not_pair, metis app_not_pair, metis fn_not_pair, metis pair_not_fst, metis pair_not_fst, metis pair_not_snd, metis pair_not_snd) done lemma beta_reduction_left_fstE: assumes "(Fst P) \\ M'" shows "(\P'. (P \\ P') \ M' = (Fst P')) \ (\A B. P = (Pair A B) \ M' = A)" using assms proof(cases) case (fst1 P P') thus ?thesis using trm_simp(7) by blast next case (fst2 B) thus ?thesis using trm_simp(7) by blast next qed ( metis app_not_fst, metis app_not_fst, metis app_not_fst, metis fn_not_fst, metis pair_not_fst, metis pair_not_fst, metis fst_not_snd, metis fst_not_snd ) lemma beta_reduction_left_sndE: assumes "(Snd P) \\ M'" shows "(\P'. (P \\ P') \ M' = (Snd P')) \ (\A B. P = Pair A B \ M' = B)" using assms proof(cases) case (snd1 P P') thus ?thesis using trm_simp(8) by blast next case (snd2 A) thus ?thesis using trm_simp(8) by blast next qed ( metis app_not_snd, metis app_not_snd, metis app_not_snd, metis fn_not_snd, metis pair_not_snd, metis pair_not_snd, metis fst_not_snd, metis fst_not_snd ) lemma preservation': assumes "\ \ M : \" and "M \\ M'" shows "\ \ M' : \" using assms proof(induction arbitrary: M' rule: typing.induct) case (tunit \) thus ?case using beta_reduction_left_unitE by metis next case (tvar \ x \) thus ?case using beta_reduction_left_varE by metis next case (tapp \ A \ \ B M') from \(App A B) \\ M'\ consider "(\x T X. A = (Fn x T X) \ M' = (X[x ::= B]))" | "(\A'. (A \\ A') \ M' = App A' B)" | "(\B'. (B \\ B') \ M' = App A B')" using beta_reduction_left_appE by metis thus ?case proof(cases) case 1 from this obtain x T X where "A = Fn x T X" and *: "M' = (X[x ::= B])" by auto have "\(x \ \) \ X : \" using typing_fnE \\ \ A : (TArr \ \)\ \A = Fn x T X\ type.inject by blast hence "\ \ (X[x ::= B]) : \" using typing_subst \\ \ B : \\ by metis thus ?thesis using * by auto next case 2 from this obtain A' where "A \\ A'" and *: "M' = (App A' B)" by auto hence "\ \ A' : (TArr \ \)" using tapp.IH(1) by metis hence "\ \ (App A' B) : \" using \\ \ B : \\ typing.tapp by metis thus ?thesis using * by auto next case 3 from this obtain B' where "B \\ B'" and *: "M' = (App A B')" by auto hence "\ \ B' : \" using tapp.IH(2) by metis hence "\ \ (App A B') : \" using \\ \ A : (TArr \ \)\ typing.tapp by metis thus ?thesis using * by auto next qed next case (tfn \ x T A \) from this obtain A' where "A \\ A'" and *: "M' = (Fn x T A')" using beta_reduction_left_fnE by metis hence "\(x \ T) \ A' : \" using tfn.IH by metis hence "\ \ (Fn x T A') : (TArr T \)" using typing.tfn by metis thus ?case using * by auto next case (tpair \ A \ B \) from this consider "\A'. (A \\ A') \ M' = (Pair A' B)" | "\B'. (B \\ B') \ M' = (Pair A B')" using beta_reduction_left_pairE by metis thus ?case proof(cases) case 1 from this obtain A' where "A \\ A'" and "M' = Pair A' B" by auto thus ?thesis using tpair typing.tpair by metis next case 2 from this obtain B' where "B \\ B'" and "M' = Pair A B'" by auto thus ?thesis using tpair typing.tpair by metis next qed next case (tfst \ P \ \) from this consider "\P'. (P \\ P') \ M' = Fst P'" | "\A B. P = Pair A B \ M' = A" using beta_reduction_left_fstE by metis thus ?case proof(cases) case 1 from this obtain P' where "P \\ P'" and "M' = Fst P'" by auto thus ?thesis using tfst typing.tfst by metis next case 2 from this obtain A B where "P = Pair A B" and "M' = A" by auto thus ?thesis using \\ \ P : (TPair \ \)\ typing_pairE by blast next qed next case (tsnd \ P \ \) from this consider "\P'. (P \\ P') \ M' = Snd P'" | "\A B. P = Pair A B \ M' = B" using beta_reduction_left_sndE by metis thus ?case proof(cases) case 1 from this obtain P' where "P \\ P'" and "M' = Snd P'" by auto thus ?thesis using tsnd typing.tsnd by metis next case 2 from this obtain A B where "P = Pair A B" and "M' = B" by auto thus ?thesis using \\ \ P : (TPair \ \)\ typing_pairE by blast next qed next qed inductive NF :: "'a trm \ bool" where unit: "NF Unit" | var: "NF (Var x)" | app: "\A \ Fn x T A'; NF A; NF B\ \ NF (App A B)" | fn: "NF A \ NF (Fn x T A)" | pair: "\NF A; NF B\ \ NF (Pair A B)" | fst: "\P \ Pair A B; NF P\ \ NF (Fst P)" | snd: "\P \ Pair A B; NF P\ \ NF (Snd P)" theorem progress: assumes "\ \ M : \" shows "NF M \ (\M'. (M \\ M'))" using assms proof(induction M arbitrary: \ \ rule: trm_induct) case 1 thus ?case using NF.unit by metis next case (2 x) thus ?case using NF.var by metis next case (3 A B) from \\ \ App A B : \\ obtain \ where "\ \ A : (TArr \ \)" and "\ \ B : \" using typing_appE by metis hence A: "NF A \ (\A'. (A \\ A'))" and B: "NF B \ (\B'. (B \\ B'))" using "3.IH" by auto consider "NF B" | "\B'. (B \\ B')" using B by auto thus ?case proof(cases) case 1 consider "NF A" | "\A'. (A \\ A')" using A by auto thus ?thesis proof(cases) case 1 consider "\x T A'. A = Fn x T A'" | "\x T A'. A = Fn x T A'" by auto thus ?thesis proof(cases) case 1 from this obtain x T A' where "A = Fn x T A'" by auto hence "(App A B) \\ (A'[x ::= B])" using beta_reduction.beta by metis thus ?thesis by blast next case 2 thus ?thesis using \NF A\ \NF B\ NF.app by metis next qed next case 2 thus ?thesis using beta_reduction.app1 by metis next qed next case 2 thus ?thesis using beta_reduction.app2 by metis next qed next case (4 x T A) from \\ \ Fn x T A : \\ obtain \ where "\ = TArr T \" and "\(x \ T) \ A : \" using typing_fnE by metis from \\(x \ T) \ A : \\ consider "NF A" | "\A'. (A \\ A')" using "4.IH" by metis thus ?case proof(cases) case 1 thus ?thesis using NF.fn by metis next case 2 from this obtain A' where "A \\ A'" by auto obtain M' where "M' = Fn x T A'" by auto hence "(Fn x T A) \\ M'" using \A \\ A'\ beta_reduction.fn by metis thus ?thesis by auto next qed next case (5 A B) thus ?case using typing_pairE beta_reduction.pair1 beta_reduction.pair2 NF.pair by meson next case (6 P) from this consider "NF P" | "\P'. (P \\ P')" using typing_fstE by metis thus ?case proof(cases) case 1 consider "\A B. P = Pair A B" | "\A B. P = Pair A B" by auto thus ?thesis proof(cases) case 1 from this obtain A B where "P = Pair A B" by auto hence "(Fst P) \\ A" using beta_reduction.fst2 by metis thus ?thesis by auto next case 2 thus ?thesis using \NF P\ NF.fst by metis next qed next case 2 thus ?thesis using beta_reduction.fst1 by metis next qed next case (7 P) from this consider "NF P" | "\P'. (P \\ P')" using typing_sndE by metis thus ?case proof(cases) case 1 consider "\A B. P = Pair A B" | "\A B. P = Pair A B" by auto thus ?thesis proof(cases) case 1 from this obtain A B where "P = Pair A B" by auto hence "(Snd P) \\ B" using beta_reduction.snd2 by metis thus ?thesis by auto next case 2 thus ?thesis using \NF P\ NF.snd by metis next qed next case 2 thus ?thesis using beta_reduction.snd1 by metis next qed next qed inductive beta_reduces :: "'a trm \ 'a trm \ bool" ("_ \\\<^sup>* _") where reflexive: "M \\\<^sup>* M" | transitive: "\M \\\<^sup>* M'; M' \\ M''\ \ M \\\<^sup>* M''" lemma beta_reduces_composition: assumes "A' \\\<^sup>* A''" and "A \\\<^sup>* A'" shows "A \\\<^sup>* A''" using assms proof(induction) case (reflexive B) thus ?case. next case (transitive B B' B'') thus ?case using beta_reduces.transitive by metis next qed lemma beta_reduces_fvs: assumes "A \\\<^sup>* A'" shows "fvs A' \ fvs A" using assms proof(induction) case (reflexive M) thus ?case by auto next case (transitive M M' M'') hence "fvs M'' \ fvs M'" using beta_reduction_fvs by metis thus ?case using \fvs M' \ fvs M\ by auto next qed lemma beta_reduces_app_left: assumes "A \\\<^sup>* A'" shows "(App A B) \\\<^sup>* (App A' B)" using assms proof(induction) case (reflexive A) thus ?case using beta_reduces.reflexive. next case (transitive A A' A'') thus ?case using beta_reduces.transitive beta_reduction.app1 by metis next qed lemma beta_reduces_app_right: assumes "B \\\<^sup>* B'" shows "(App A B) \\\<^sup>* (App A B')" using assms proof(induction) case (reflexive B) thus ?case using beta_reduces.reflexive. next case (transitive B B' B'') thus ?case using beta_reduces.transitive beta_reduction.app2 by metis next qed lemma beta_reduces_fn: assumes "A \\\<^sup>* A'" shows "(Fn x T A) \\\<^sup>* (Fn x T A')" using assms proof(induction) case (reflexive A) thus ?case using beta_reduces.reflexive. next case (transitive A A' A'') thus ?case using beta_reduces.transitive beta_reduction.fn by metis next qed lemma beta_reduces_pair_left: assumes "A \\\<^sup>* A'" shows "(Pair A B) \\\<^sup>* (Pair A' B)" using assms proof(induction) case (reflexive M) thus ?case using beta_reduces.reflexive. next case (transitive M M' M'') thus ?case using beta_reduces.transitive beta_reduction.pair1 by metis next qed lemma beta_reduces_pair_right: assumes "B \\\<^sup>* B'" shows "(Pair A B) \\\<^sup>* (Pair A B')" using assms proof(induction) case (reflexive M) thus ?case using beta_reduces.reflexive. next case (transitive M M' M'') thus ?case using beta_reduces.transitive beta_reduction.pair2 by metis next qed lemma beta_reduces_fst: assumes "P \\\<^sup>* P'" shows "(Fst P) \\\<^sup>* (Fst P')" using assms proof(induction) case (reflexive M) thus ?case using beta_reduces.reflexive. next case (transitive M M' M'') thus ?case using beta_reduces.transitive beta_reduction.fst1 by metis next qed lemma beta_reduces_snd: assumes "P \\\<^sup>* P'" shows "(Snd P) \\\<^sup>* (Snd P')" using assms proof(induction) case (reflexive M) thus ?case using beta_reduces.reflexive. next case (transitive M M' M'') thus ?case using beta_reduces.transitive beta_reduction.snd1 by metis next qed theorem preservation: assumes "M \\\<^sup>* M'" "\ \ M : \" shows "\ \ M' : \" using assms proof(induction) case (reflexive M) thus ?case. next case (transitive M M' M'') thus ?case using preservation' by metis next qed theorem safety: assumes "M \\\<^sup>* M'" "\ \ M : \" shows "NF M' \ (\M''. (M' \\ M''))" using assms proof(induction) case (reflexive M) thus ?case using progress by metis next case (transitive M M' M'') hence "\ \ M' : \" using preservation by metis hence "\ \ M'' : \" using preservation' \M' \\ M''\ by metis thus ?case using progress by metis next qed inductive parallel_reduction :: "'a trm \ 'a trm \ bool" ("_ >> _") where refl: "A >> A" | beta: "\A >> A'; B >> B'\ \ (App (Fn x T A) B) >> (A'[x ::= B'])" | eta: "A >> A' \ (Fn x T A) >> (Fn x T A')" | app: "\A >> A'; B >> B'\ \ (App A B) >> (App A' B')" | pair: "\A >> A'; B >> B'\ \ (Pair A B) >> (Pair A' B')" | fst1: "P >> P' \ (Fst P) >> (Fst P')" | fst2: "A >> A' \ (Fst (Pair A B)) >> A'" | snd1: "P >> P' \ (Snd P) >> (Snd P')" | snd2: "B >> B' \ (Snd (Pair A B)) >> B'" lemma parallel_reduction_prm: assumes "A >> A'" shows "(\ \ A) >> (\ \ A')" using assms apply induction apply (rule parallel_reduction.refl) apply (metis parallel_reduction.beta subst_prm trm_prm_simp(3, 4)) apply (metis parallel_reduction.eta trm_prm_simp(4)) apply (metis parallel_reduction.app trm_prm_simp(3)) apply (metis parallel_reduction.pair trm_prm_simp(5)) apply (metis parallel_reduction.fst1 trm_prm_simp(6)) apply (metis parallel_reduction.fst2 trm_prm_simp(5, 6)) apply (metis parallel_reduction.snd1 trm_prm_simp(7)) apply (metis parallel_reduction.snd2 trm_prm_simp(5, 7)) done lemma parallel_reduction_fvs: assumes "A >> A'" shows "fvs A' \ fvs A" using assms proof(induction) case (refl A) thus ?case by auto next case (beta A A' B B' x T) have *:"fvs (App (Fn x T A) B) = fvs A - {x} \ fvs B" using fvs_simp(3, 4) by metis have "fvs (A'[x ::= B']) \ fvs A' - {x} \ fvs B'" using subst_fvs. also have "... \ fvs A - {x} \ fvs B" using beta.IH by auto finally show ?case using fvs_simp(3, 4) by metis next case (eta A A' x T) thus ?case using fvs_simp(4) Un_Diff subset_Un_eq by metis next case (app A A' B B') thus ?case using fvs_simp(3) Un_mono by metis next case (pair A A' B B') thus ?case using fvs_simp(5) Un_mono by metis next case (fst1 P P') thus ?case using fvs_simp(6) by force next case (fst2 A A' B) thus ?case using fvs_simp(5, 6) by force next case (snd1 P P') thus ?case using fvs_simp(7) by force next case (snd2 B B' A) thus ?case using fvs_simp(5, 7) by force next qed inductive_cases parallel_reduction_unitE': "Unit >> A" lemma parallel_reduction_unitE: assumes "Unit >> A" shows "A = Unit" using assms apply (rule parallel_reduction_unitE'[where A=A]) apply blast apply (auto simp add: unit_not_app unit_not_fn unit_not_pair unit_not_fst unit_not_snd) done inductive_cases parallel_reduction_varE': "(Var x) >> A" lemma parallel_reduction_varE: assumes "(Var x) >> A" shows "A = Var x" using assms apply (rule parallel_reduction_varE'[where x=x and A=A]) apply blast apply (auto simp add: var_not_app var_not_fn var_not_pair var_not_fst var_not_snd) done inductive_cases parallel_reduction_fnE': "(Fn x T A) >> X" lemma parallel_reduction_fnE: assumes "(Fn x T A) >> X" shows "X = Fn x T A \ (\A'. (A >> A') \ X = Fn x T A')" using assms proof(induction rule: parallel_reduction_fnE'[where x=x and T=T and A=A and X=X]) case (4 B B' y S) from this consider "x = y \ T = S \ A = B" | "x \ y \ T = S \ x \ fvs B \ A = [x \ y] \ B" using trm_simp(4) by metis thus ?case proof(cases) case 1 hence "x = y" "T = S" "A = B" by auto thus ?thesis using 4 by metis next case 2 hence "x \ y" "T = S" "x \ fvs B" "A = [x \ y] \ B" by auto hence "x \ fvs B'" "A >> ([x \ y] \ B')" using parallel_reduction_fvs parallel_reduction_prm \B >> B'\ by auto thus ?thesis using fn_eq \X = Fn y S B'\ \x \ y\ \T = S\ by metis next qed next qed ( metis assms, blast, metis app_not_fn, metis app_not_fn, metis fn_not_pair, metis fn_not_fst, metis fn_not_fst, metis fn_not_snd, metis fn_not_snd ) inductive_cases parallel_reduction_redexE': "(App (Fn x T A) B) >> X" lemma parallel_reduction_redexE: assumes "(App (Fn x T A) B) >> X" shows " (X = App (Fn x T A) B) \ (\A' B'. (A >> A') \ (B >> B') \ X = (A'[x ::= B'])) \ (\A' B'. ((Fn x T A) >> (Fn x T A')) \ (B >> B') \ X = (App (Fn x T A') B')) " using assms proof(induction rule: parallel_reduction_redexE'[where x=x and T=T and A=A and B=B and X=X]) case (5 C C' D D') from \App (Fn x T A) B = App C D\ have C: "C = Fn x T A" and D: "D = B" by (metis trm_simp(2), metis trm_simp(3)) from C and \C >> C'\ obtain A' where C': "C' = Fn x T A'" using parallel_reduction_fnE by metis thus ?thesis using C C' D \C >> C'\ \D >> D'\ \X = App C' D'\ by metis next case (3 C C' D D' y S) from \App (Fn x T A) B = App (Fn y S C) D\ have "Fn x T A = Fn y S C" and B: "B = D" by (metis trm_simp(2), metis trm_simp(3)) from this consider "x = y \ T = S \ A = C" | "x \ y \ T = S \ A = [x \ y] \ C \ x \ fvs C" using trm_simp(4) by metis thus ?case proof(cases) case 1 thus ?thesis using \C >> C'\ \X = (C'[y ::= D'])\ \D >> D'\ B by metis next case 2 hence "x \ y" "T = S" and A: "A = [x \ y] \ C" "x \ fvs C" by auto have "x \ fvs C'" using parallel_reduction_fvs \x \ fvs C\ \C >> C'\ by force have "A >> ([x \ y] \ C')" using parallel_reduction_prm \C >> C'\ A by metis moreover have "X = (([x \ y] \ C')[x ::= D'])" using \X = (C'[y ::= D'])\ subst_swp \x \ fvs C'\ by metis ultimately show ?thesis using \D >> D'\ B by metis next qed next qed ( metis assms, blast, metis app_not_fn, metis app_not_pair, metis app_not_fst, metis app_not_fst, metis app_not_snd, metis app_not_snd ) inductive_cases parallel_reduction_nonredexE': "(App A B) >> X" lemma parallel_reduction_nonredexE: assumes "(App A B) >> X" and "\x T A'. A \ Fn x T A'" shows "\A' B'. (A >> A') \ (B >> B') \ X = (App A' B')" using assms proof(induction rule: parallel_reduction_nonredexE'[where A=A and B=B and X=X]) case (5 C C' D D') hence "A = C" "B = D" using trm_simp(2, 3) by auto thus ?case using \C >> C'\ \D >> D'\ \X = App C' D'\ by metis next qed ( metis assms(1), metis parallel_reduction.refl, metis trm_simp(2, 3) assms(2), metis app_not_fn, metis app_not_pair, metis app_not_fst, metis app_not_fst, metis app_not_snd, metis app_not_snd ) inductive_cases parallel_reduction_pairE': "(Pair A B) >> X" lemma parallel_reduction_pairE: assumes "(Pair A B) >> X" shows "\A' B'. (A >> A') \ (B >> B') \ (X = Pair A' B')" using assms proof(induction rule: parallel_reduction_pairE'[where A=A and B=B and X=X]) case 2 thus ?case using parallel_reduction.refl by blast next case (6 A A' B B') thus ?case using parallel_reduction.pair trm_simp(5, 6) by fastforce next qed ( metis assms, metis app_not_pair, metis fn_not_pair, metis app_not_pair, metis pair_not_fst, metis pair_not_fst, metis pair_not_snd, metis pair_not_snd ) inductive_cases parallel_reduction_fstE': "(Fst P) >> X" lemma parallel_reduction_fstE: assumes "(Fst P) >> X" shows "(\P'. (P >> P') \ X = (Fst P')) \ (\A A' B. (P = Pair A B) \ (A >> A') \ X = A')" using assms proof(induction rule: parallel_reduction_fstE'[where P=P and X=X]) case (7 P P') thus ?case using parallel_reduction.fst1 trm_simp(7) by metis next case (8 A B) thus ?case using parallel_reduction.fst2 trm_simp(7) by metis next qed ( metis assms, insert parallel_reduction.refl, metis, metis app_not_fst, metis fn_not_fst, metis app_not_fst, metis pair_not_fst, metis fst_not_snd, metis fst_not_snd ) inductive_cases parallel_reduction_sndE': "(Snd P) >> X" lemma parallel_reduction_sndE: assumes "(Snd P) >> X" shows "(\P'. (P >> P') \ X = (Snd P')) \ (\A B B'. (P = Pair A B) \ (B >> B') \ X = B')" using assms proof(induction rule: parallel_reduction_sndE'[where P=P and X=X]) case (9 P P') thus ?case using parallel_reduction.snd1 trm_simp(8) by metis next case (10 A B) thus ?case using parallel_reduction.snd2 trm_simp(8) by metis next qed ( metis assms, insert parallel_reduction.refl, metis, metis app_not_snd, metis fn_not_snd, metis app_not_snd, metis pair_not_snd, metis fst_not_snd, metis fst_not_snd ) lemma parallel_reduction_subst_inner: assumes "M >> M'" shows "X[z ::= M] >> (X[z ::= M'])" using assms proof(induction X rule: trm_strong_induct[where S="{z} \ fvs M \ fvs M'"]) show "finite ({z} \ fvs M \ fvs M')" using fvs_finite by auto case 1 thus ?case using subst_simp_unit parallel_reduction.refl by metis next case (2 x) thus ?case by(cases "x = z", metis subst_simp_var1, metis subst_simp_var2 parallel_reduction.refl) next case (3 A B) thus ?case using subst_simp_app parallel_reduction.app by metis next case (4 x T A) hence "x \ z" "x \ fvs M" "x \ fvs M'" by auto thus ?case using 4 subst_simp_fn parallel_reduction.eta by metis next case (5 A B) thus ?case using subst_simp_pair parallel_reduction.pair by metis next case (6 P) thus ?case using subst_simp_fst parallel_reduction.fst1 by metis next case (7 P) thus ?case using subst_simp_snd parallel_reduction.snd1 by metis next qed lemma parallel_reduction_subst: assumes "X >> X'" "M >> M'" shows "X[z ::= M] >> (X'[z ::= M'])" using assms proof(induction X arbitrary: X' rule: trm_strong_depth_induct[where S="{z} \ fvs M \ fvs M'"]) show "finite ({z} \ fvs M \ fvs M')" using fvs_finite by auto next case 1 hence "X' = Unit" using parallel_reduction_unitE by metis thus ?case using parallel_reduction.refl subst_simp_unit by metis next case (2 x) hence "X' = Var x" using parallel_reduction_varE by metis thus ?case using parallel_reduction_subst_inner \M >> M'\ by metis next case (3 C D) consider "\x T A. C = Fn x T A" | "\x T A. C = Fn x T A" by metis thus ?case proof(cases) case 1 from this obtain x T A where C: "C = Fn x T A" by auto have "depth C < depth (App C D)" "depth D < depth (App C D)" using depth_app by auto consider "X' = App (Fn x T A) D" | "\A' D'. ((Fn x T A) >> (Fn x T A')) \ (D >> D') \ X' = (App (Fn x T A') D')" | "\A' D'. (A >> A') \ (D >> D') \ X' = (A'[x ::= D'])" using parallel_reduction_redexE \(App C D) >> X'\ C by metis thus ?thesis proof(cases) case 1 thus ?thesis using parallel_reduction_subst_inner \M >> M'\ C by metis next case 2 from this obtain A' D' where "(Fn x T A) >> (Fn x T A')" "D >> D'" and X': "X' = App (Fn x T A') D'" by auto have *: "((Fn x T A)[z ::= M]) >> ((Fn x T A')[z ::= M'])" using "3.IH" \depth C < depth (App C D)\ C \(Fn x T A) >> (Fn x T A')\ \M >> M'\ by metis have **: "(D[z ::= M]) >> (D'[z ::= M'])" using "3.IH" \depth D < depth (App C D)\ \D >> D'\ \M >> M'\ by metis have "(App C D)[z ::= M] = App ((Fn x T A)[z ::= M]) (D[z ::= M])" using subst_simp_app C by metis moreover have "... >> (App ((Fn x T A')[z ::= M']) (D'[z ::= M']))" using * ** parallel_reduction.app by metis moreover have "... = ((App (Fn x T A') D')[z ::= M'])" using subst_simp_app by metis moreover have "... = (X'[z ::= M'])" using X' by metis ultimately show ?thesis by metis next case 3 from this obtain A' D' where "A >> A'" "D >> D'" and X': "X' = (A'[x ::= D'])" by auto have "depth A < depth (App C D)" using C depth_app depth_fn dual_order.strict_trans by fastforce have "finite ({z} \ fvs M \ fvs M' \ fvs A')" using fvs_finite by auto from this obtain y where "y \ {z} \ fvs M \ fvs M' \ fvs A'" and C: "C = Fn y T ([y \ x] \ A)" using fresh_fn C by metis hence "y \ z" "y \ fvs M" "y \ fvs M'" "y \ fvs A'" by auto have "([y \ x] \ A) >> ([y \ x] \ A')" using parallel_reduction_prm \A >> A'\ by metis hence *: "([y \ x] \ A)[z ::= M] >> (([y \ x] \ A')[z ::= M'])" using "3.IH" \depth A < depth (App C D)\ depth_prm using \([y \ x] \ A) >> ([y \ x] \A')\ \M >> M'\ by metis have **: "(D[z ::= M]) >> (D'[z ::= M'])" using "3.IH" \depth D < depth (App C D)\ \D >> D'\ \M >> M'\ by metis have "(App C D)[z ::= M] = (App ((Fn y T ([y \ x] \ A))[z ::= M]) (D[z ::= M]))" using C subst_simp_app by metis moreover have "... = (App (Fn y T (([y \ x] \ A)[z ::= M])) (D[z ::= M]))" using \y \ z\ \y \ fvs M\ subst_simp_fn by metis moreover have "... >> (([y \ x] \ A')[z ::= M'][y ::= D'[z ::= M']])" using parallel_reduction.beta * ** by metis moreover have "... = (([y \ x] \ A')[y ::= D'][z ::= M'])" using barendregt \y \ z\ \y \ fvs M'\ by metis moreover have "... = (A'[x ::= D'][z ::= M'])" using subst_swp \y \ fvs A'\ by metis moreover have "... = (X'[z ::= M'])" using X' by metis ultimately show ?thesis by metis next qed next case 2 from this obtain C' D' where "C >> C'" "D >> D'" and X': "X' = App C' D'" using parallel_reduction_nonredexE \(App C D) >> X'\ by metis have "depth C < depth (App C D)" "depth D < depth (App C D)" using depth_app by auto hence *: "(C[z ::= M]) >> (C'[z ::= M'])" and **: "(D[z ::= M]) >> (D'[z ::= M'])" using "3.IH" \C >> C'\ \D >> D'\ \M >> M'\ by metis+ have "(App C D)[z ::= M] = App (C[z ::= M]) (D[z ::= M])" using subst_simp_app by metis moreover have "... >> (App (C'[z ::= M']) (D'[z ::= M']))" using parallel_reduction.app * ** by metis moreover have "... = ((App C' D')[z ::= M'])" using subst_simp_app by metis moreover have "... = (X'[z ::= M'])" using X' by metis ultimately show ?thesis by metis next qed next case (4 x T A) hence "x \ z" "x \ fvs M" "x \ fvs M'" by auto from \(Fn x T A) >> X'\ consider "X' = Fn x T A" | "\A'. (A >> A') \ X' = Fn x T A'" using parallel_reduction_fnE by metis thus ?case proof(cases) case 1 thus ?thesis using parallel_reduction_subst_inner \M >> M'\ by metis next case 2 from this obtain A' where "A >> A'" and X': "X' = Fn x T A'" by auto hence *: "(A[z ::= M]) >> (A'[z ::= M'])" using "4.IH" depth_fn \A >> A'\ \M >> M'\ by metis have "((Fn x T A)[z ::= M]) = (Fn x T (A[z ::= M]))" using subst_simp_fn \x \ z\ \x \ fvs M\ by metis moreover have "... >> (Fn x T (A'[z ::= M']))" using parallel_reduction.eta * by metis moreover have "... = ((Fn x T A')[z ::= M'])" using subst_simp_fn \x \ z\ \x \ fvs M'\ by metis moreover have "... = (X'[z ::= M'])" using X' by metis ultimately show ?thesis by metis next qed next case (5 A B) from \(Pair A B) >> X'\ consider "X' = Pair A B" | "\A' B'. (A >> A') \ (B >> B') \ X' = Pair A' B'" using parallel_reduction_pairE by metis thus ?case proof(cases) case 1 thus ?thesis using parallel_reduction_subst_inner \M >> M'\ by metis next case 2 from this obtain A' B' where "A >> A'" "B >> B'" and X': "X' = Pair A' B'" by auto have *: "(A[z ::= M]) >> (A'[z ::= M'])" and **: "(B[z ::= M]) >> (B'[z ::= M'])" using "5.IH" \A >> A'\ \B >> B'\ \M >> M'\ by (metis depth_pair(1), metis depth_pair(2)) have "(Pair A B)[z ::= M] = (Pair (A[z ::= M]) (B[z ::= M]))" using subst_simp_pair by metis moreover have "... >> (Pair (A'[z ::= M']) (B'[z ::= M']))" using parallel_reduction.pair * ** by metis moreover have "... = ((Pair A' B')[z ::= M'])" using subst_simp_pair by metis moreover have "... = (X'[z ::= M'])" using X' by metis ultimately show ?thesis by metis next qed next case (6 P) from \(Fst P) >> X'\ consider "\P'. (P >> P') \ X' = Fst P'" | "\A A' B. P = Pair A B \ (A >> A') \ X' = A'" using parallel_reduction_fstE by metis thus ?case proof(cases) case 1 from this obtain P' where "P >> P'" and X': "X' = Fst P'" by auto have *: "(P[z ::= M]) >> (P'[z ::= M'])" using "6.IH" depth_fst \P >> P'\ \M >> M'\ by metis have "(Fst P)[z ::= M] = Fst (P[z ::= M])" using subst_simp_fst by metis moreover have "... >> (Fst (P'[z ::= M']))" using parallel_reduction.fst1 * by metis moreover have "... = ((Fst P')[z ::= M'])" using subst_simp_fst by metis moreover have "... = (X'[z ::= M'])" using X' by metis ultimately show ?thesis by metis next case 2 from this obtain A A' B where P: "P = Pair A B" "A >> A'" and X': "X' = A'" by auto have "depth A < depth (Fst P)" using P depth_fst depth_pair dual_order.strict_trans by fastforce hence *: "(A[z ::= M]) >> (A'[z ::= M'])" using "6.IH" \A >> A'\ \M >> M'\ by metis have "(Fst P)[z ::= M] = (Fst (Pair (A[z ::= M]) (B[z ::= M])))" using P subst_simp_fst subst_simp_pair by metis moreover have "... >> (A'[z ::= M'])" using parallel_reduction.fst2 * by metis moreover have "... = (X'[z ::= M'])" using X' by metis ultimately show ?thesis by metis next qed next case (7 P) from \(Snd P) >> X'\ consider "\P'. (P >> P') \ X' = Snd P'" | "\A B B'. P = Pair A B \ (B >> B') \ X' = B'" using parallel_reduction_sndE by metis thus ?case proof(cases) case 1 from this obtain P' where "P >> P'" and X': "X' = Snd P'" by auto have *: "(P[z ::= M]) >> (P'[z ::= M'])" using "7.IH" depth_snd \P >> P'\ \M >> M'\ by metis have "(Snd P)[z ::= M] = Snd (P[z ::= M])" using subst_simp_snd by metis moreover have "... >> (Snd (P'[z ::= M']))" using parallel_reduction.snd1 * by metis moreover have "... = ((Snd P')[z ::= M'])" using subst_simp_snd by metis moreover have "... = (X'[z ::= M'])" using X' by metis ultimately show ?thesis by metis next case 2 from this obtain A B B' where P: "P = Pair A B" "B >> B'" and X': "X' = B'" by auto have "depth B < depth (Snd P)" using P depth_snd depth_pair dual_order.strict_trans by fastforce hence *: "(B[z ::= M]) >> (B'[z ::= M'])" using "7.IH" \B >> B'\ \M >> M'\ by metis have "(Snd P)[z ::= M] = (Snd (Pair (A[z ::= M]) (B[z ::= M])))" using P subst_simp_snd subst_simp_pair by metis moreover have "... >> (B'[z ::= M'])" using parallel_reduction.snd2 * by metis moreover have "... = (X'[z ::= M'])" using X' by metis ultimately show ?thesis by metis next qed next qed inductive complete_development :: "'a trm \ 'a trm \ bool" ("_ >>> _") where unit: "Unit >>> Unit" | var: "(Var x) >>> (Var x)" | beta: "\A >>> A'; B >>> B'\ \ (App (Fn x T A) B) >>> (A'[x ::= B'])" | eta: "A >>> A' \ (Fn x T A) >>> (Fn x T A')" | app: "\A >>> A'; B >>> B'; (\x T M. A \ Fn x T M)\ \ (App A B) >>> (App A' B')" | pair: "\A >>> A'; B >>> B'\ \ (Pair A B) >>> (Pair A' B')" | fst1: "\P >>> P'; (\A B. P \ Pair A B)\ \ (Fst P) >>> (Fst P')" | fst2: "A >>> A' \ (Fst (Pair A B)) >>> A'" | snd1: "\P >>> P'; (\A B. P \ Pair A B)\ \ (Snd P) >>> (Snd P')" | snd2: "B >>> B' \ (Snd (Pair A B)) >>> B'" lemma not_fn_prm: assumes "\x T M. A \ Fn x T M" shows "\ \ A \ Fn x T M" proof(rule classical) fix x T M obtain \' where *: "\' = prm_inv \" by auto assume "\(\ \ A \ Fn x T M)" hence "\ \ A = Fn x T M" by blast hence "\' \ (\ \ A) = \' \ Fn x T M" by fastforce hence "(\' \ \) \ A = \' \ Fn x T M" using trm_prm_apply_compose by metis hence "A = \' \ Fn x T M" using * prm_inv_compose trm_prm_apply_id by metis hence "A = Fn (\' $ x) T (\' \ M)" using trm_prm_simp(4) by metis hence False using assms by blast thus ?thesis by blast qed lemma not_pair_prm: assumes "\A B. P \ Pair A B" shows "\ \ P \ Pair A B" proof(rule classical) fix A B obtain \' where *: "\' = prm_inv \" by auto assume "\(\ \ P \ Pair A B)" hence "\ \ P = Pair A B" by blast hence "\' \ \ \ P = \' \ (Pair A B)" by fastforce hence "(\' \ \) \ P = \' \ (Pair A B)" using trm_prm_apply_compose by metis hence "P = \' \ (Pair A B)" using * prm_inv_compose trm_prm_apply_id by metis hence "P = Pair (\' \ A) (\' \ B)" using trm_prm_simp(5) by metis hence False using assms by blast thus ?thesis by blast qed lemma complete_development_prm: assumes "A >>> A'" shows "(\ \ A) >>> (\ \ A')" using assms proof(induction) case unit thus ?case using complete_development.unit trm_prm_simp(1) by metis next case (var x) thus ?case using complete_development.var trm_prm_simp(2) by metis next case (beta A A' B B' x T) thus ?case using complete_development.beta subst_prm trm_prm_simp(3, 4) by metis next case (eta A A' x T) thus ?case using complete_development.eta trm_prm_simp(4) by metis next case (app A A' B B') thus ?case using complete_development.app by (simp add: trm_prm_simp(3) not_fn_prm) next case (pair A A' B B') thus ?case using complete_development.pair trm_prm_simp(5) by metis next case (fst1 P P') hence "\ \ P \ Pair A B" for A B using not_pair_prm by metis thus ?case using trm_prm_simp(6) fst1.IH complete_development.fst1 by metis next case (fst2 A A' B) thus ?case using trm_prm_simp(5, 6) complete_development.fst2 by metis next case (snd1 P P') hence "\ \ P \ Pair A B" for A B using not_pair_prm by metis thus ?case using trm_prm_simp(7) snd1.IH complete_development.snd1 by metis next case (snd2 B B' A) thus ?case using trm_prm_simp(5, 7) complete_development.snd2 by metis next qed lemma complete_development_fvs: assumes "A >>> A'" shows "fvs A' \ fvs A" using assms proof(induction) case unit thus ?case using fvs_simp by auto next case (var x) thus ?case using fvs_simp by auto next case (beta A A' B B' x T) have "fvs (A'[x ::= B']) \ fvs A' - {x} \ fvs B'" using subst_fvs. also have "... \ fvs A - {x} \ fvs B" using beta.IH by auto also have "... \ fvs (Fn x T A) \ fvs B" using fvs_simp(4) subset_refl by force also have "... \ fvs (App (Fn x T A) B)" using fvs_simp(3) subset_refl by force finally show ?case. next case (eta A A' x T) thus ?case using fvs_simp(4) using Un_Diff subset_Un_eq by (metis (no_types, lifting)) next case (app A A' B B') thus ?case using fvs_simp(3) Un_mono by metis next case (pair A A' B B') thus ?case using fvs_simp(5) Un_mono by metis next case (fst1 P P') thus ?case using fvs_simp(6) by force next case (fst2 A A' B) thus ?case by (simp add: fvs_simp(5, 6) sup.coboundedI1) next case (snd1 P P') thus ?case using fvs_simp(7) by fastforce next case (snd2 B B' A) thus ?case using fvs_simp(5, 7) by fastforce next qed lemma complete_development_fnE: assumes "(Fn x T A) >>> X" shows "\A'. (A >>> A') \ X = Fn x T A'" using assms proof(cases) case (eta B B' y S) hence "T = S" using trm_simp(4) by metis from eta consider "x = y \ A = B" | "x \ y \ x \ fvs B \ A = [x \ y] \ B" using trm_simp(4) by metis thus ?thesis proof(cases) case 1 hence "x = y" and "A = B" by auto obtain A' where "A' = B'" by auto hence "A >>> A'" and "X = Fn x T A'" using eta \A = B\ \x = y\ \T = S\ by auto thus ?thesis by auto next case 2 hence "x \ y" "x \ fvs B" and A: "A = [x \ y] \ B" by auto hence "x \ fvs B'" using \B >>> B'\ complete_development_fvs by auto obtain A' where A': "A' = [x \ y] \ B'" by auto hence "A >>> A'" using A \B >>> B'\ complete_development_prm by auto have "X = Fn x T A'" using fn_eq \x \ y\ \x \ fvs B'\ A' \X = Fn y S B'\ \T = S\ by metis thus ?thesis using \A >>> A'\ by auto next qed next qed ( metis unit_not_fn, metis var_not_fn, metis app_not_fn, metis app_not_fn, metis fn_not_pair, metis fn_not_fst, metis fn_not_fst, metis fn_not_snd, metis fn_not_snd ) lemma complete_development_pairE: assumes "(Pair A B) >>> X" shows "\A' B'. (A >>> A') \ (B >>> B') \ X = Pair A' B'" using assms apply cases apply (metis unit_not_pair, metis var_not_pair, metis app_not_pair, metis fn_not_pair, metis app_not_pair) apply (metis trm_simp(5, 6)) apply (metis pair_not_fst, metis pair_not_fst, metis pair_not_snd, metis pair_not_snd) done lemma complete_development_exists: shows "\X. (A >>> X)" proof(induction A rule: trm_induct) case 1 obtain X :: "'a trm" where "X = Unit" by auto hence "Unit >>> X" using complete_development.unit by metis thus ?case by auto next case (2 x) obtain X where "X = Var x" by auto hence "(Var x) >>> X" using complete_development.var by metis thus ?case by auto next case (3 A B) from this obtain A' B' where A': "A >>> A'" and B': "B >>> B'" by auto consider "(\x T M. A = Fn x T M)" | "\(\x T M. A = Fn x T M)" by auto thus ?case proof(cases) case 1 from this obtain x T M where A: "A = Fn x T M" by auto from this obtain M' where "A' = Fn x T M'" and "M >>> M'" using complete_development_fnE A' by metis obtain X where "X = (M'[x ::= B'])" by auto hence "(App A B) >>> X" using complete_development.beta \M >>> M'\ B' A by metis thus ?thesis by auto next case 2 thus ?thesis using complete_development.app A' B' by metis next qed next case (4 x T A) from this obtain A' where A': "A >>> A'" by auto obtain X where "X = Fn x T A'" by auto hence "(Fn x T A) >>> X" using complete_development.eta A' by metis thus ?case by auto next case (5 A B) thus ?case using complete_development.pair by blast next case (6 P) consider "\A B. P = Pair A B" | "\A B. P = Pair A B" by auto thus ?case proof(cases) case 1 from this obtain A B X where "P = Pair A B" "P >>> X" using 6 by auto from this obtain A' B' where "A >>> A'" "B >>> B'" "X = Pair A' B'" using complete_development_pairE by metis thus ?thesis using complete_development.fst2 \P = Pair A B\ by metis next case 2 thus ?thesis using complete_development.fst1 6 by blast next qed next case (7 P) consider "\A B. P = Pair A B" | "\A B. P = Pair A B" by auto thus ?case proof(cases) case 1 from this obtain A B X where "P = Pair A B" "P >>> X" using 7 by auto from this obtain A' B' where "A >>> A'" "B >>> B'" "X = Pair A' B'" using complete_development_pairE by metis thus ?thesis using complete_development.snd2 \P = Pair A B\ by metis next case 2 thus ?thesis using complete_development.snd1 7 by blast next qed next qed lemma complete_development_triangle: assumes "A >>> D" "A >> B" shows "B >> D" using assms proof(induction arbitrary: B rule: complete_development.induct) case unit thus ?case using parallel_reduction_unitE parallel_reduction.refl by metis next case (var x) thus ?case using parallel_reduction_varE parallel_reduction.refl by metis next case (beta A A' C C' x T) hence "A >> A'" "C >> C'" using parallel_reduction.refl by metis+ from \(App (Fn x T A) C) >> B\ consider "B = App (Fn x T A) C" | "\A'' C''. (A >> A'') \ (C >> C'') \ B = (A''[x ::= C''])" | "\A'' C''. ((Fn x T A) >> (Fn x T A'')) \ (C >> C'') \ B = (App (Fn x T A'') C'')" using parallel_reduction_redexE by metis thus ?case proof(cases) case 1 thus ?thesis using parallel_reduction.beta \A >> A'\ \C >> C'\ by metis next case 2 from this obtain A'' C'' where "A >> A''" "C >> C''" and B: "B = (A''[x ::= C''])" by auto hence "A'' >> A'" "C'' >> C'" using beta.IH by metis+ thus ?thesis using B parallel_reduction_subst by metis next case 3 from this obtain A'' C'' where "(Fn x T A) >> (Fn x T A'')" "C >> C''" and B: "B = App (Fn x T A'') C''" by auto hence "C'' >> C'" using beta.IH by metis have "A'' >> A'" proof - thm parallel_reduction_fnE from \(Fn x T A) >> (Fn x T A'')\ consider "Fn x T A = Fn x T A''" | "\A'''. (A >> A''') \ Fn x T A'' = Fn x T A'''" using parallel_reduction_fnE by metis hence "A >> A''" proof(cases) case 1 hence "A = A''" using trm_simp(4) by metis thus ?thesis using parallel_reduction.refl by metis next case 2 from this obtain A''' where "A >> A'''" "Fn x T A'' = Fn x T A'''" by auto thus ?thesis using trm_simp(4) by metis next qed thus ?thesis using beta.IH by metis qed thus ?thesis using B \C'' >> C'\ parallel_reduction.beta by metis next qed next case (eta A A' x T B) from this consider "B = Fn x T A" | "\A''. (A >> A'') \ B = Fn x T A''" using parallel_reduction_fnE by metis thus ?case proof(cases) case 1 thus ?thesis using eta.IH parallel_reduction.refl parallel_reduction.eta by metis next case 2 from this obtain A'' where "A >> A''" and "B = Fn x T A''" by auto thus ?thesis using eta.IH parallel_reduction.eta by metis next qed next case (app A A' C C') from this obtain A'' C'' where "A >> A''" "C >> C''" and B: "B = App A'' C''" using parallel_reduction_nonredexE by metis hence "A'' >> A'" "C'' >> C'" using app.IH by metis+ thus ?case using B parallel_reduction.app by metis next case (pair A A' C C') from \(Pair A C) >> B\ and parallel_reduction_pairE obtain A'' C'' where "A >> A''" "C >> C''" "B = Pair A'' C''" by metis thus ?case using pair.IH parallel_reduction.pair by metis next case (fst1 P P') from this obtain P'' where "P >> P''" "B = Fst P''" using parallel_reduction_fstE by blast thus ?case using fst1.IH parallel_reduction.fst1 by metis next case (fst2 A A' C B) from this consider "\P''. ((Pair A C) >> P'') \ B = Fst P''" | "\A''. (A >> A'') \ (B = A'')" using parallel_reduction_fstE[where P="(Pair A C)" and X=B] using trm_simp(5) by metis thus ?case proof(cases) case 1 from this obtain P'' where "(Pair A C) >> P''" and "B = Fst P''" by auto from this obtain A'' C'' where "P'' = Pair A'' C''" "A >> A''" "C >> C''" using parallel_reduction_pairE by metis thus ?thesis using fst2 parallel_reduction.fst2 \B = Fst P''\ by metis next case 2 from this obtain A'' where "A >> A''" "B = A''" by metis thus ?thesis using fst2 by metis next qed next case (snd1 P P') from this obtain P'' where "P >> P''" "B = Snd P''" using parallel_reduction_sndE by blast thus ?case using snd1.IH parallel_reduction.snd1 by metis next case (snd2 C A' A B) from this consider "\P''. ((Pair A C) >> P'') \ B = Snd P''" | "\C''. (C >> C'') \ (B = C'')" using parallel_reduction_sndE[where P="(Pair A C)" and X=B] using trm_simp(5, 6) by metis thus ?case proof(cases) case 1 from this obtain P'' where "(Pair A C) >> P''" and "B = Snd P''" by auto from this obtain A'' C'' where "P'' = Pair A'' C''" "A >> A''" "C >> C''" using parallel_reduction_pairE by metis thus ?thesis using snd2 parallel_reduction.snd2 \B = Snd P''\ by metis next case 2 from this obtain C'' where "C >> C''" "B = C''" by metis thus ?thesis using snd2 by metis next qed next qed lemma parallel_reduction_diamond: assumes "A >> B" "A >> C" shows "\D. (B >> D) \ (C >> D)" proof - obtain D where "A >>> D" using complete_development_exists by metis hence "(B >> D) \ (C >> D)" using complete_development_triangle assms by auto thus ?thesis by blast qed inductive parallel_reduces :: "'a trm \ 'a trm \ bool" ("_ >>\<^sup>* _") where reflexive: "A >>\<^sup>* A" | transitive: "\A >>\<^sup>* A'; A' >> A''\ \ A >>\<^sup>* A''" lemma parallel_reduces_diamond': assumes "A >>\<^sup>* C" "A >> B" shows "\D. (B >>\<^sup>* D) \ (C >> D)" using assms proof(induction) case (reflexive A) thus ?case using parallel_reduces.reflexive by metis next case (transitive A A' A'') from this obtain C where "B >>\<^sup>* C" "A' >> C" by metis from \A' >> C\ \A' >> A''\ obtain D where "C >> D" "A'' >> D" using parallel_reduction_diamond by metis thus ?case using parallel_reduces.transitive \B >>\<^sup>* C\ by metis next qed lemma parallel_reduces_diamond: assumes "A >>\<^sup>* B" "A >>\<^sup>* C" shows "\D. (B >>\<^sup>* D) \ (C >>\<^sup>* D)" using assms proof(induction) case (reflexive A) thus ?case using parallel_reduces.reflexive by metis next case (transitive A A' A'') from this obtain C' where "A' >>\<^sup>* C'" "C >>\<^sup>* C'" by metis from this obtain D where "A'' >>\<^sup>* D" "C' >> D" using \A' >> A''\ \A' >>\<^sup>* C'\ parallel_reduces_diamond' by metis thus ?case using parallel_reduces.transitive \C >>\<^sup>* C'\ by metis next qed lemma beta_reduction_is_parallel_reduction: assumes "A \\ B" shows "A >> B" using assms apply induction apply (metis parallel_reduction.beta parallel_reduction.refl) apply (metis parallel_reduction.app parallel_reduction.refl) apply (metis parallel_reduction.app parallel_reduction.refl) apply (metis parallel_reduction.eta) apply (metis parallel_reduction.pair parallel_reduction.refl) apply (metis parallel_reduction.pair parallel_reduction.refl) apply (metis parallel_reduction.fst1) apply (metis parallel_reduction.fst2 parallel_reduction.refl) apply (metis parallel_reduction.snd1) apply (metis parallel_reduction.snd2 parallel_reduction.refl) done lemma parallel_reduction_is_beta_reduction: assumes "A >> B" shows "A \\\<^sup>* B" using assms proof(induction) case (refl A) thus ?case using beta_reduces.reflexive. next case (beta A A' B B' x T) hence "(App (Fn x T A) B) \\\<^sup>* (App (Fn x T A') B')" using \A \\\<^sup>* A'\ beta_reduces_fn beta_reduces_app_left beta_reduces_app_right beta_reduces_composition by metis moreover have "(App (Fn x T A') B') \\ (A'[x ::= B'])" using beta_reduction.beta. ultimately show ?case using beta_reduces.transitive by metis next case (eta A A' x T) thus ?case using beta_reduces_fn by metis next case (app A A' B B') thus ?case using beta_reduces_app_left beta_reduces_app_right beta_reduces_composition by metis next case (pair A A' B B') thus ?case using beta_reduces_pair_left beta_reduces_pair_right beta_reduces_composition by metis next case (fst1 P P') thus ?case using beta_reduces_fst by metis next case (fst2 A A' B) thus ?case using beta_reduces_pair_left beta_reduction.fst2 beta_reduces.intros beta_reduces_composition by blast next case (snd1 P P') thus ?case using beta_reduces_snd by metis next case (snd2 B B' A) thus ?case using beta_reduces_pair_left beta_reduction.snd2 beta_reduces.intros beta_reduces_composition by blast next qed lemma parallel_beta_reduces_equivalent: shows "(A >>\<^sup>* B) = (A \\\<^sup>* B)" proof - have \: "(A >>\<^sup>* B) \ (A \\\<^sup>* B)" proof(induction rule: parallel_reduces.induct) case (reflexive A) thus ?case using beta_reduces.reflexive. next case (transitive A A' A'') thus ?case using beta_reduces_composition parallel_reduction_is_beta_reduction by metis next qed have \: "(A \\\<^sup>* B) \ (A >>\<^sup>* B)" proof(induction rule: beta_reduces.induct) case (reflexive A) thus ?case using parallel_reduces.reflexive. next case (transitive A A' A'') thus ?case using parallel_reduces.transitive beta_reduction_is_parallel_reduction by metis next qed from \\ show "(A >>\<^sup>* B) = (A \\\<^sup>* B)" by blast qed theorem confluence: assumes "A \\\<^sup>* B" "A \\\<^sup>* C" shows "\D. (B \\\<^sup>* D) \ (C \\\<^sup>* D)" proof - from assms have "A >>\<^sup>* B" "A >>\<^sup>* C" using parallel_beta_reduces_equivalent by metis+ hence "\D. (B >>\<^sup>* D) \ (C >>\<^sup>* D)" using parallel_reduces_diamond by metis thus "\D. (B \\\<^sup>* D) \ (C \\\<^sup>* D)" using parallel_beta_reduces_equivalent by metis qed end end diff --git a/thys/SATSolverVerification/ConflictAnalysis.thy b/thys/SATSolverVerification/ConflictAnalysis.thy --- a/thys/SATSolverVerification/ConflictAnalysis.thy +++ b/thys/SATSolverVerification/ConflictAnalysis.thy @@ -1,4221 +1,4221 @@ theory ConflictAnalysis imports AssertLiteral begin (******************************************************************************) (* A P P L Y C O N F L I C T *) (******************************************************************************) lemma clauseFalseInPrefixToLastAssertedLiteral: assumes "isLastAssertedLiteral l (oppositeLiteralList c) (elements M)" and "clauseFalse c (elements M)" and "uniq (elements M)" shows "clauseFalse c (elements (prefixToLevel (elementLevel l M) M))" proof- { fix l'::Literal assume "l' el c" hence "literalFalse l' (elements M)" using \clauseFalse c (elements M)\ by (simp add: clauseFalseIffAllLiteralsAreFalse) hence "literalTrue (opposite l') (elements M)" by simp have "opposite l' el oppositeLiteralList c" using \l' el c\ using literalElListIffOppositeLiteralElOppositeLiteralList[of "l'" "c"] by simp have "elementLevel (opposite l') M \ elementLevel l M" using lastAssertedLiteralHasHighestElementLevel[of "l" "oppositeLiteralList c" "M"] using \isLastAssertedLiteral l (oppositeLiteralList c) (elements M)\ using \uniq (elements M)\ using \opposite l' el oppositeLiteralList c\ using \literalTrue (opposite l') (elements M)\ by auto hence "opposite l' el (elements (prefixToLevel (elementLevel l M) M))" using elementLevelLtLevelImpliesMemberPrefixToLevel[of "opposite l'" "M" "elementLevel l M"] using \literalTrue (opposite l') (elements M)\ by simp } thus ?thesis by (simp add: clauseFalseIffAllLiteralsAreFalse) qed lemma InvariantNoDecisionsWhenConflictEnsuresCurrentLevelCl: assumes "InvariantNoDecisionsWhenConflict F M (currentLevel M)" "clause el F" "clauseFalse clause (elements M)" "uniq (elements M)" "currentLevel M > 0" shows "clause \ [] \ (let Cl = getLastAssertedLiteral (oppositeLiteralList clause) (elements M) in InvariantClCurrentLevel Cl M)" proof- have "clause \ []" proof- { assume "\ ?thesis" hence "clauseFalse clause (elements (prefixToLevel ((currentLevel M) - 1) M))" by simp hence False using \InvariantNoDecisionsWhenConflict F M (currentLevel M)\ using \currentLevel M > 0\ using \clause el F\ unfolding InvariantNoDecisionsWhenConflict_def by (simp add: formulaFalseIffContainsFalseClause) } thus ?thesis by auto qed moreover let ?Cl = "getLastAssertedLiteral (oppositeLiteralList clause) (elements M)" have "elementLevel ?Cl M = currentLevel M" proof- have "elementLevel ?Cl M \ currentLevel M" using elementLevelLeqCurrentLevel[of "?Cl" "M"] by simp moreover have "elementLevel ?Cl M \ currentLevel M" proof- { assume "elementLevel ?Cl M < currentLevel M" have "isLastAssertedLiteral ?Cl (oppositeLiteralList clause) (elements M)" using getLastAssertedLiteralCharacterization[of "clause" "elements M"] using \uniq (elements M)\ using \clauseFalse clause (elements M)\ using \clause \ []\ by simp hence "clauseFalse clause (elements (prefixToLevel (elementLevel ?Cl M) M))" using clauseFalseInPrefixToLastAssertedLiteral[of "?Cl" "clause" "M"] using \clauseFalse clause (elements M)\ using \uniq (elements M)\ by simp hence "False" using \clause el F\ using \InvariantNoDecisionsWhenConflict F M (currentLevel M)\ using \currentLevel M > 0\ unfolding InvariantNoDecisionsWhenConflict_def using \elementLevel ?Cl M < currentLevel M\ by (simp add: formulaFalseIffContainsFalseClause) } thus ?thesis by force qed ultimately show ?thesis by simp qed ultimately show ?thesis unfolding InvariantClCurrentLevel_def by (simp add: Let_def) qed lemma InvariantsClAfterApplyConflict: assumes "getConflictFlag state" "InvariantUniq (getM state)" "InvariantNoDecisionsWhenConflict (getF state) (getM state) (currentLevel (getM state))" "InvariantEquivalentZL (getF state) (getM state) F0" "InvariantConflictClauseCharacterization (getConflictFlag state) (getConflictClause state) (getF state) (getM state)" "currentLevel (getM state) > 0" shows "let state' = applyConflict state in InvariantCFalse (getConflictFlag state') (getM state') (getC state') \ InvariantCEntailed (getConflictFlag state') F0 (getC state') \ InvariantClCharacterization (getCl state') (getC state') (getM state') \ InvariantClCurrentLevel (getCl state') (getM state') \ InvariantCnCharacterization (getCn state') (getC state') (getM state') \ InvariantUniqC (getC state')" proof- let ?M0 = "elements (prefixToLevel 0 (getM state))" let ?oppM0 = "oppositeLiteralList ?M0" let ?clause' = "nth (getF state) (getConflictClause state)" let ?clause'' = "list_diff ?clause' ?oppM0" let ?clause = "remdups ?clause''" let ?l = "getLastAssertedLiteral (oppositeLiteralList ?clause') (elements (getM state))" have "clauseFalse ?clause' (elements (getM state))" "?clause' el (getF state)" using \getConflictFlag state\ using \InvariantConflictClauseCharacterization (getConflictFlag state) (getConflictClause state) (getF state) (getM state)\ unfolding InvariantConflictClauseCharacterization_def by (auto simp add: Let_def) have "?clause' \ []" "elementLevel ?l (getM state) = currentLevel (getM state)" using InvariantNoDecisionsWhenConflictEnsuresCurrentLevelCl[of "getF state" "getM state" "?clause'"] using \?clause' el (getF state)\ using \clauseFalse ?clause' (elements (getM state))\ using \InvariantNoDecisionsWhenConflict (getF state) (getM state) (currentLevel (getM state))\ using \currentLevel (getM state) > 0\ using \InvariantUniq (getM state)\ unfolding InvariantUniq_def unfolding InvariantClCurrentLevel_def by (auto simp add: Let_def) have "isLastAssertedLiteral ?l (oppositeLiteralList ?clause') (elements (getM state))" using \?clause' \ []\ using \clauseFalse ?clause' (elements (getM state))\ using \InvariantUniq (getM state)\ unfolding InvariantUniq_def using getLastAssertedLiteralCharacterization[of "?clause'" "elements (getM state)"] by simp hence "?l el (oppositeLiteralList ?clause')" unfolding isLastAssertedLiteral_def by simp hence "opposite ?l el ?clause'" using literalElListIffOppositeLiteralElOppositeLiteralList[of "opposite ?l" "?clause'"] by auto have "\ ?l el ?M0" proof- { assume "\ ?thesis" hence "elementLevel ?l (getM state) = 0" using prefixToLevelElementsElementLevel[of "?l" "0" "getM state"] by simp hence False using \elementLevel ?l (getM state) = currentLevel (getM state)\ using \currentLevel (getM state) > 0\ by simp } thus ?thesis by auto qed hence "\ opposite ?l el ?oppM0" using literalElListIffOppositeLiteralElOppositeLiteralList[of "?l" "elements (prefixToLevel 0 (getM state))"] by simp have "opposite ?l el ?clause''" using \opposite ?l el ?clause'\ using \\ opposite ?l el ?oppM0\ using listDiffIff[of "opposite ?l" "?clause'" "?oppM0"] by simp hence "?l el (oppositeLiteralList ?clause'')" using literalElListIffOppositeLiteralElOppositeLiteralList[of "opposite ?l" "?clause''"] by simp have "set (oppositeLiteralList ?clause'') \ set (oppositeLiteralList ?clause')" proof fix x assume "x \ set (oppositeLiteralList ?clause'')" thus "x \ set (oppositeLiteralList ?clause')" using literalElListIffOppositeLiteralElOppositeLiteralList[of "opposite x" "?clause''"] using literalElListIffOppositeLiteralElOppositeLiteralList[of "opposite x" "?clause'"] using listDiffIff[of "opposite x" "?clause'" "oppositeLiteralList (elements (prefixToLevel 0 (getM state)))"] by auto qed have "isLastAssertedLiteral ?l (oppositeLiteralList ?clause'') (elements (getM state))" using \?l el (oppositeLiteralList ?clause'')\ using \set (oppositeLiteralList ?clause'') \ set (oppositeLiteralList ?clause')\ using \isLastAssertedLiteral ?l (oppositeLiteralList ?clause') (elements (getM state))\ using isLastAssertedLiteralSubset[of "?l" "oppositeLiteralList ?clause'" "elements (getM state)" "oppositeLiteralList ?clause''"] by auto moreover have "set (oppositeLiteralList ?clause) = set (oppositeLiteralList ?clause'')" unfolding oppositeLiteralList_def by simp ultimately have "isLastAssertedLiteral ?l (oppositeLiteralList ?clause) (elements (getM state))" unfolding isLastAssertedLiteral_def by auto hence "?l el (oppositeLiteralList ?clause)" unfolding isLastAssertedLiteral_def by simp hence "opposite ?l el ?clause" using literalElListIffOppositeLiteralElOppositeLiteralList[of "opposite ?l" "?clause"] by simp hence "?clause \ []" by auto have "clauseFalse ?clause'' (elements (getM state))" proof- { fix l::Literal assume "l el ?clause''" hence "l el ?clause'" using listDiffIff[of "l" "?clause'" "?oppM0"] by simp hence "literalFalse l (elements (getM state))" using \clauseFalse ?clause' (elements (getM state))\ by (simp add: clauseFalseIffAllLiteralsAreFalse) } thus ?thesis by (simp add: clauseFalseIffAllLiteralsAreFalse) qed hence "clauseFalse ?clause (elements (getM state))" by (simp add: clauseFalseIffAllLiteralsAreFalse) let ?l' = "getLastAssertedLiteral (oppositeLiteralList ?clause) (elements (getM state))" have "isLastAssertedLiteral ?l' (oppositeLiteralList ?clause) (elements (getM state))" using \?clause \ []\ using \clauseFalse ?clause (elements (getM state))\ using \InvariantUniq (getM state)\ unfolding InvariantUniq_def using getLastAssertedLiteralCharacterization[of "?clause" "elements (getM state)"] by simp with \isLastAssertedLiteral ?l (oppositeLiteralList ?clause) (elements (getM state))\ have "?l = ?l'" using lastAssertedLiteralIsUniq by simp have "formulaEntailsClause (getF state) ?clause'" using \?clause' el (getF state)\ by (simp add: formulaEntailsItsClauses) let ?F0 = "(getF state) @ val2form ?M0" have "formulaEntailsClause ?F0 ?clause'" using \formulaEntailsClause (getF state) ?clause'\ by (simp add: formulaEntailsClauseAppend) hence "formulaEntailsClause ?F0 ?clause''" using \formulaEntailsClause (getF state) ?clause'\ using formulaEntailsClauseRemoveEntailedLiteralOpposites[of "?F0" "?clause'" "?M0"] using val2formIsEntailed[of "getF state" "?M0" "[]"] by simp hence "formulaEntailsClause ?F0 ?clause" unfolding formulaEntailsClause_def by (simp add: clauseTrueIffContainsTrueLiteral) hence "formulaEntailsClause F0 ?clause" using \InvariantEquivalentZL (getF state) (getM state) F0\ unfolding InvariantEquivalentZL_def unfolding formulaEntailsClause_def unfolding equivalentFormulae_def by auto show ?thesis using \isLastAssertedLiteral ?l' (oppositeLiteralList ?clause) (elements (getM state))\ using \?l = ?l'\ using \elementLevel ?l (getM state) = currentLevel (getM state)\ using \clauseFalse ?clause (elements (getM state))\ using \formulaEntailsClause F0 ?clause\ unfolding applyConflict_def unfolding setConflictAnalysisClause_def unfolding InvariantClCharacterization_def unfolding InvariantClCurrentLevel_def unfolding InvariantCFalse_def unfolding InvariantCEntailed_def unfolding InvariantCnCharacterization_def unfolding InvariantUniqC_def by (auto simp add: findLastAssertedLiteral_def countCurrentLevelLiterals_def Let_def uniqDistinct distinct_remdups_id) qed (******************************************************************************) (* A P P L Y E X P L A I N *) (******************************************************************************) lemma CnEqual1IffUIP: assumes "InvariantClCharacterization (getCl state) (getC state) (getM state)" "InvariantClCurrentLevel (getCl state) (getM state)" "InvariantCnCharacterization (getCn state) (getC state) (getM state)" shows "(getCn state = 1) = isUIP (opposite (getCl state)) (getC state) (getM state)" proof- let ?clls = "filter (\ l. elementLevel (opposite l) (getM state) = currentLevel (getM state)) (remdups (getC state))" let ?Cl = "getCl state" have "isLastAssertedLiteral ?Cl (oppositeLiteralList (getC state)) (elements (getM state))" using \InvariantClCharacterization (getCl state) (getC state) (getM state)\ unfolding InvariantClCharacterization_def . hence "literalTrue ?Cl (elements (getM state))" "?Cl el (oppositeLiteralList (getC state))" unfolding isLastAssertedLiteral_def by auto hence "opposite ?Cl el getC state" using literalElListIffOppositeLiteralElOppositeLiteralList[of "opposite ?Cl" "getC state"] by simp hence "opposite ?Cl el ?clls" using \InvariantClCurrentLevel (getCl state) (getM state)\ unfolding InvariantClCurrentLevel_def by auto hence "?clls \ []" by force hence "length ?clls > 0" by simp have "uniq ?clls" by (simp add: uniqDistinct) { assume "getCn state \ 1" hence "length ?clls > 1" using assms using \length ?clls > 0\ unfolding InvariantCnCharacterization_def by (simp (no_asm)) then obtain literal1::Literal and literal2::Literal where "literal1 el ?clls" "literal2 el ?clls" "literal1 \ literal2" using \uniq ?clls\ using \?clls \ []\ using lengthGtOneTwoDistinctElements[of "?clls"] by auto then obtain literal::Literal where "literal el ?clls" "literal \ opposite ?Cl" using \opposite ?Cl el ?clls\ by auto hence "\ isUIP (opposite ?Cl) (getC state) (getM state)" using \opposite ?Cl el ?clls\ unfolding isUIP_def by auto } moreover { assume "getCn state = 1" hence "length ?clls = 1" using \InvariantCnCharacterization (getCn state) (getC state) (getM state)\ unfolding InvariantCnCharacterization_def by auto { fix literal::Literal assume "literal el (getC state)" "literal \ opposite ?Cl" have "elementLevel (opposite literal) (getM state) < currentLevel (getM state)" proof- have "elementLevel (opposite literal) (getM state) \ currentLevel (getM state)" using elementLevelLeqCurrentLevel[of "opposite literal" "getM state"] by simp moreover have "elementLevel (opposite literal) (getM state) \ currentLevel (getM state)" proof- { assume "\ ?thesis" with \literal el (getC state)\ have "literal el ?clls" by simp hence "False" using \length ?clls = 1\ using \opposite ?Cl el ?clls\ using \literal \ opposite ?Cl\ using lengthOneImpliesOnlyElement[of "?clls" "opposite ?Cl"] by auto } thus ?thesis by auto qed ultimately show ?thesis by simp qed } hence "isUIP (opposite ?Cl) (getC state) (getM state)" using \isLastAssertedLiteral ?Cl (oppositeLiteralList (getC state)) (elements (getM state))\ using \opposite ?Cl el ?clls\ unfolding isUIP_def by auto } ultimately show ?thesis by auto qed lemma InvariantsClAfterApplyExplain: assumes "InvariantUniq (getM state)" "InvariantCFalse (getConflictFlag state) (getM state) (getC state)" "InvariantClCharacterization (getCl state) (getC state) (getM state)" "InvariantClCurrentLevel (getCl state) (getM state)" "InvariantCEntailed (getConflictFlag state) F0 (getC state)" "InvariantCnCharacterization (getCn state) (getC state) (getM state)" "InvariantEquivalentZL (getF state) (getM state) F0" "InvariantGetReasonIsReason (getReason state) (getF state) (getM state) (set (getQ state))" "getCn state \ 1" "getConflictFlag state" "currentLevel (getM state) > 0" shows "let state' = applyExplain (getCl state) state in InvariantCFalse (getConflictFlag state') (getM state') (getC state') \ InvariantCEntailed (getConflictFlag state') F0 (getC state') \ InvariantClCharacterization (getCl state') (getC state') (getM state') \ InvariantClCurrentLevel (getCl state') (getM state') \ InvariantCnCharacterization (getCn state') (getC state') (getM state') \ InvariantUniqC (getC state')" proof- let ?Cl = "getCl state" let ?oppM0 = "oppositeLiteralList (elements (prefixToLevel 0 (getM state)))" have "isLastAssertedLiteral ?Cl (oppositeLiteralList (getC state)) (elements (getM state))" using \InvariantClCharacterization (getCl state) (getC state) (getM state)\ unfolding InvariantClCharacterization_def . hence "literalTrue ?Cl (elements (getM state))" "?Cl el (oppositeLiteralList (getC state))" unfolding isLastAssertedLiteral_def by auto hence "opposite ?Cl el getC state" using literalElListIffOppositeLiteralElOppositeLiteralList[of "opposite ?Cl" "getC state"] by simp have "clauseFalse (getC state) (elements (getM state))" using \getConflictFlag state\ using \InvariantCFalse (getConflictFlag state) (getM state) (getC state)\ unfolding InvariantCFalse_def by simp have "\ isUIP (opposite ?Cl) (getC state) (getM state)" using CnEqual1IffUIP[of "state"] using assms by simp have "\ ?Cl el (decisions (getM state))" proof- { assume "\ ?thesis" hence "isUIP (opposite ?Cl) (getC state) (getM state)" using \InvariantUniq (getM state)\ using \isLastAssertedLiteral ?Cl (oppositeLiteralList (getC state)) (elements (getM state))\ using \clauseFalse (getC state) (elements (getM state))\ using lastDecisionThenUIP[of "getM state" "opposite ?Cl" "getC state"] unfolding InvariantUniq_def by simp with \\ isUIP (opposite ?Cl) (getC state) (getM state)\ have "False" by simp } thus ?thesis by auto qed have "elementLevel ?Cl (getM state) = currentLevel (getM state)" using \InvariantClCurrentLevel (getCl state) (getM state)\ unfolding InvariantClCurrentLevel_def by simp hence "elementLevel ?Cl (getM state) > 0" using \currentLevel (getM state) > 0\ by simp obtain reason where "isReason (nth (getF state) reason) ?Cl (elements (getM state))" "getReason state ?Cl = Some reason" "0 \ reason \ reason < length (getF state)" using \InvariantGetReasonIsReason (getReason state) (getF state) (getM state) (set (getQ state))\ unfolding InvariantGetReasonIsReason_def using \literalTrue ?Cl (elements (getM state))\ using \\ ?Cl el (decisions (getM state))\ using \elementLevel ?Cl (getM state) > 0\ by auto let ?res = "resolve (getC state) (getF state ! reason) (opposite ?Cl)" obtain ol::Literal where "ol el (getC state)" "ol \ opposite ?Cl" "elementLevel (opposite ol) (getM state) \ elementLevel ?Cl (getM state)" using \isLastAssertedLiteral ?Cl (oppositeLiteralList (getC state)) (elements (getM state))\ using \\ isUIP (opposite ?Cl) (getC state) (getM state)\ unfolding isUIP_def by auto hence "ol el ?res" unfolding resolve_def by simp hence "?res \ []" by auto have "opposite ol el (oppositeLiteralList ?res)" using \ol el ?res\ using literalElListIffOppositeLiteralElOppositeLiteralList[of "ol" "?res"] by simp have "opposite ol el (oppositeLiteralList (getC state))" using \ol el (getC state)\ using literalElListIffOppositeLiteralElOppositeLiteralList[of "ol" "getC state"] by simp have "literalFalse ol (elements (getM state))" using \clauseFalse (getC state) (elements (getM state))\ using \ol el getC state\ by (simp add: clauseFalseIffAllLiteralsAreFalse) have "elementLevel (opposite ol) (getM state) = elementLevel ?Cl (getM state)" using \elementLevel (opposite ol) (getM state) \ elementLevel ?Cl (getM state)\ using \isLastAssertedLiteral ?Cl (oppositeLiteralList (getC state)) (elements (getM state))\ using lastAssertedLiteralHasHighestElementLevel[of "?Cl" "oppositeLiteralList (getC state)" "getM state"] using \InvariantUniq (getM state)\ unfolding InvariantUniq_def using \opposite ol el (oppositeLiteralList (getC state))\ using \literalFalse ol (elements (getM state))\ by auto hence "elementLevel (opposite ol) (getM state) = currentLevel (getM state)" using \elementLevel ?Cl (getM state) = currentLevel (getM state)\ by simp have "InvariantCFalse (getConflictFlag state) (getM state) ?res" using \InvariantCFalse (getConflictFlag state) (getM state) (getC state)\ using InvariantCFalseAfterExplain[of "getConflictFlag state" "getM state" "getC state" "?Cl" "nth (getF state) reason" "?res"] using \isReason (nth (getF state) reason) ?Cl (elements (getM state))\ using \opposite ?Cl el (getC state)\ by simp hence "clauseFalse ?res (elements (getM state))" using \getConflictFlag state\ unfolding InvariantCFalse_def by simp let ?rc = "nth (getF state) reason" let ?M0 = "elements (prefixToLevel 0 (getM state))" let ?F0 = "(getF state) @ (val2form ?M0)" let ?C' = "list_diff ?res ?oppM0" let ?C = "remdups ?C'" have "formulaEntailsClause (getF state) ?rc" using \0 \ reason \ reason < length (getF state)\ using nth_mem[of "reason" "getF state"] by (simp add: formulaEntailsItsClauses) hence "formulaEntailsClause ?F0 ?rc" by (simp add: formulaEntailsClauseAppend) hence "formulaEntailsClause F0 ?rc" using \InvariantEquivalentZL (getF state) (getM state) F0\ unfolding InvariantEquivalentZL_def unfolding formulaEntailsClause_def unfolding equivalentFormulae_def by simp hence "formulaEntailsClause F0 ?res" using \getConflictFlag state\ using \InvariantCEntailed (getConflictFlag state) F0 (getC state)\ using InvariantCEntailedAfterExplain[of "getConflictFlag state" "F0" "getC state" "nth (getF state) reason" "?res" "getCl state"] unfolding InvariantCEntailed_def by auto hence "formulaEntailsClause ?F0 ?res" using \InvariantEquivalentZL (getF state) (getM state) F0\ unfolding InvariantEquivalentZL_def unfolding formulaEntailsClause_def unfolding equivalentFormulae_def by simp hence "formulaEntailsClause ?F0 ?C" using formulaEntailsClauseRemoveEntailedLiteralOpposites[of "?F0" "?res" "?M0"] using val2formIsEntailed[of "getF state" "?M0" "[]"] unfolding formulaEntailsClause_def by (auto simp add: clauseTrueIffContainsTrueLiteral) hence "formulaEntailsClause F0 ?C" using \InvariantEquivalentZL (getF state) (getM state) F0\ unfolding InvariantEquivalentZL_def unfolding formulaEntailsClause_def unfolding equivalentFormulae_def by simp let ?ll = "getLastAssertedLiteral (oppositeLiteralList ?res) (elements (getM state))" have "isLastAssertedLiteral ?ll (oppositeLiteralList ?res) (elements (getM state))" using \?res \ []\ using \clauseFalse ?res (elements (getM state))\ using \InvariantUniq (getM state)\ unfolding InvariantUniq_def using getLastAssertedLiteralCharacterization[of "?res" "elements (getM state)"] by simp hence "elementLevel (opposite ol) (getM state) \ elementLevel ?ll (getM state)" using \opposite ol el (oppositeLiteralList (getC state))\ using lastAssertedLiteralHasHighestElementLevel[of "?ll" "oppositeLiteralList ?res" "getM state"] using \InvariantUniq (getM state)\ using \opposite ol el (oppositeLiteralList ?res)\ using \literalFalse ol (elements (getM state))\ unfolding InvariantUniq_def by simp hence "elementLevel ?ll (getM state) = currentLevel (getM state)" using \elementLevel (opposite ol) (getM state) = currentLevel (getM state)\ using elementLevelLeqCurrentLevel[of "?ll" "getM state"] by simp have "?ll el (oppositeLiteralList ?res)" using \isLastAssertedLiteral ?ll (oppositeLiteralList ?res) (elements (getM state))\ unfolding isLastAssertedLiteral_def by simp hence "opposite ?ll el ?res" using literalElListIffOppositeLiteralElOppositeLiteralList[of "opposite ?ll" "?res"] by simp have "\ ?ll el (elements (prefixToLevel 0 (getM state)))" proof- { assume "\ ?thesis" hence "elementLevel ?ll (getM state) = 0" using prefixToLevelElementsElementLevel[of "?ll" "0" "getM state"] by simp hence False using \elementLevel ?ll (getM state) = currentLevel (getM state)\ using \currentLevel (getM state) > 0\ by simp } thus ?thesis by auto qed hence "\ opposite ?ll el ?oppM0" using literalElListIffOppositeLiteralElOppositeLiteralList[of "?ll" "elements (prefixToLevel 0 (getM state))"] by simp have "opposite ?ll el ?C'" using \opposite ?ll el ?res\ using \\ opposite ?ll el ?oppM0\ using listDiffIff[of "opposite ?ll" "?res" "?oppM0"] by simp hence "?ll el (oppositeLiteralList ?C')" using literalElListIffOppositeLiteralElOppositeLiteralList[of "opposite ?ll" "?C'"] by simp have "set (oppositeLiteralList ?C') \ set (oppositeLiteralList ?res)" proof fix x assume "x \ set (oppositeLiteralList ?C')" thus "x \ set (oppositeLiteralList ?res)" using literalElListIffOppositeLiteralElOppositeLiteralList[of "opposite x" "?C'"] using literalElListIffOppositeLiteralElOppositeLiteralList[of "opposite x" "?res"] using listDiffIff[of "opposite x" "?res" "?oppM0"] by auto qed have "isLastAssertedLiteral ?ll (oppositeLiteralList ?C') (elements (getM state))" using \?ll el (oppositeLiteralList ?C')\ using \set (oppositeLiteralList ?C') \ set (oppositeLiteralList ?res)\ using \isLastAssertedLiteral ?ll (oppositeLiteralList ?res) (elements (getM state))\ using isLastAssertedLiteralSubset[of "?ll" "oppositeLiteralList ?res" "elements (getM state)" "oppositeLiteralList ?C'"] by auto moreover have "set (oppositeLiteralList ?C) = set (oppositeLiteralList ?C')" unfolding oppositeLiteralList_def by simp ultimately have "isLastAssertedLiteral ?ll (oppositeLiteralList ?C) (elements (getM state))" unfolding isLastAssertedLiteral_def by auto hence "?ll el (oppositeLiteralList ?C)" unfolding isLastAssertedLiteral_def by simp hence "opposite ?ll el ?C" using literalElListIffOppositeLiteralElOppositeLiteralList[of "opposite ?ll" "?C"] by simp hence "?C \ []" by auto have "clauseFalse ?C' (elements (getM state))" proof- { fix l::Literal assume "l el ?C'" hence "l el ?res" using listDiffIff[of "l" "?res" "?oppM0"] by simp hence "literalFalse l (elements (getM state))" using \clauseFalse ?res (elements (getM state))\ by (simp add: clauseFalseIffAllLiteralsAreFalse) } thus ?thesis by (simp add: clauseFalseIffAllLiteralsAreFalse) qed hence "clauseFalse ?C (elements (getM state))" by (simp add: clauseFalseIffAllLiteralsAreFalse) let ?l' = "getLastAssertedLiteral (oppositeLiteralList ?C) (elements (getM state))" have "isLastAssertedLiteral ?l' (oppositeLiteralList ?C) (elements (getM state))" using \?C \ []\ using \clauseFalse ?C (elements (getM state))\ using \InvariantUniq (getM state)\ unfolding InvariantUniq_def using getLastAssertedLiteralCharacterization[of "?C" "elements (getM state)"] by simp with \isLastAssertedLiteral ?ll (oppositeLiteralList ?C) (elements (getM state))\ have "?ll = ?l'" using lastAssertedLiteralIsUniq by simp show ?thesis using \isLastAssertedLiteral ?l' (oppositeLiteralList ?C) (elements (getM state))\ using \?ll = ?l'\ using \elementLevel ?ll (getM state) = currentLevel (getM state)\ using \getReason state ?Cl = Some reason\ using \clauseFalse ?C (elements (getM state))\ using \formulaEntailsClause F0 ?C\ unfolding applyExplain_def unfolding InvariantCFalse_def unfolding InvariantCEntailed_def unfolding InvariantClCharacterization_def unfolding InvariantClCurrentLevel_def unfolding InvariantCnCharacterization_def unfolding InvariantUniqC_def unfolding setConflictAnalysisClause_def by (simp add: findLastAssertedLiteral_def countCurrentLevelLiterals_def Let_def uniqDistinct distinct_remdups_id) qed (******************************************************************************) (* A P P L Y E X P L A I N U I P *) (******************************************************************************) definition "multLessState = {(state1, state2). (getM state1 = getM state2) \ (getC state1, getC state2) \ multLess (getM state1)}" lemma ApplyExplainUIPTermination: assumes "InvariantUniq (getM state)" "InvariantGetReasonIsReason (getReason state) (getF state) (getM state) (set (getQ state))" "InvariantCFalse (getConflictFlag state) (getM state) (getC state)" "InvariantClCurrentLevel (getCl state) (getM state)" "InvariantClCharacterization (getCl state) (getC state) (getM state)" "InvariantCnCharacterization (getCn state) (getC state) (getM state)" "InvariantCEntailed (getConflictFlag state) F0 (getC state)" "InvariantEquivalentZL (getF state) (getM state) F0" "getConflictFlag state" "currentLevel (getM state) > 0" shows "applyExplainUIP_dom state" using assms proof (induct rule: wf_induct[of "multLessState"]) case 1 thus ?case unfolding wf_eq_minimal proof- show "\Q (state::State). state \ Q \ (\ stateMin \ Q. \state'. (state', stateMin) \ multLessState \ state' \ Q)" proof- { fix Q :: "State set" and state :: State assume "state \ Q" let ?M = "(getM state)" let ?Q1 = "{C::Clause. \ state. state \ Q \ (getM state) = ?M \ (getC state) = C}" from \state \ Q\ have "getC state \ ?Q1" by auto with wfMultLess[of "?M"] obtain Cmin where "Cmin \ ?Q1" "\C'. (C', Cmin) \ multLess ?M \ C' \ ?Q1" unfolding wf_eq_minimal apply (erule_tac x="?Q1" in allE) apply (erule_tac x="getC state" in allE) by auto from \Cmin \ ?Q1\ obtain stateMin where "stateMin \ Q" "(getM stateMin) = ?M" "getC stateMin = Cmin" by auto have "\state'. (state', stateMin) \ multLessState \ state' \ Q" proof fix state' show "(state', stateMin) \ multLessState \ state' \ Q" proof assume "(state', stateMin) \ multLessState" with \getM stateMin = ?M\ have "getM state' = getM stateMin" "(getC state', getC stateMin) \ multLess ?M" unfolding multLessState_def by auto from \\C'. (C', Cmin) \ multLess ?M \ C' \ ?Q1\ \(getC state', getC stateMin) \ multLess ?M\ \getC stateMin = Cmin\ have "getC state' \ ?Q1" by simp with \getM state' = getM stateMin\ \getM stateMin = ?M\ show "state' \ Q" by auto qed qed with \stateMin \ Q\ have "\ stateMin \ Q. (\state'. (state', stateMin) \ multLessState \ state' \ Q)" by auto } thus ?thesis by auto qed qed next case (2 state') note ih = this show ?case proof (cases "getCn state' = 1") case True show ?thesis apply (rule applyExplainUIP_dom.intros) using True by simp next case False let ?state'' = "applyExplain (getCl state') state'" have "InvariantGetReasonIsReason (getReason ?state'') (getF ?state'') (getM ?state'') (set (getQ ?state''))" "InvariantUniq (getM ?state'')" "InvariantEquivalentZL (getF ?state'') (getM ?state'') F0" "getConflictFlag ?state''" "currentLevel (getM ?state'') > 0" using ih unfolding applyExplain_def unfolding setConflictAnalysisClause_def by (auto split: option.split simp add: findLastAssertedLiteral_def countCurrentLevelLiterals_def Let_def) moreover have "InvariantCFalse (getConflictFlag ?state'') (getM ?state'') (getC ?state'')" "InvariantClCharacterization (getCl ?state'') (getC ?state'') (getM ?state'')" "InvariantCnCharacterization (getCn ?state'') (getC ?state'') (getM ?state'')" "InvariantClCurrentLevel (getCl ?state'') (getM ?state'')" "InvariantCEntailed (getConflictFlag ?state'') F0 (getC ?state'')" using InvariantsClAfterApplyExplain[of "state'" "F0"] using ih using False by (auto simp add:Let_def) moreover have "(?state'', state') \ multLessState" proof- have "getM ?state'' = getM state'" unfolding applyExplain_def unfolding setConflictAnalysisClause_def by (auto split: option.split simp add: findLastAssertedLiteral_def countCurrentLevelLiterals_def Let_def) let ?Cl = "getCl state'" let ?oppM0 = "oppositeLiteralList (elements (prefixToLevel 0 (getM state')))" have "isLastAssertedLiteral ?Cl (oppositeLiteralList (getC state')) (elements (getM state'))" using ih unfolding InvariantClCharacterization_def by simp hence "literalTrue ?Cl (elements (getM state'))" "?Cl el (oppositeLiteralList (getC state'))" unfolding isLastAssertedLiteral_def by auto hence "opposite ?Cl el getC state'" using literalElListIffOppositeLiteralElOppositeLiteralList[of "opposite ?Cl" "getC state'"] by simp have "clauseFalse (getC state') (elements (getM state'))" using ih unfolding InvariantCFalse_def by simp have "\ ?Cl el (decisions (getM state'))" proof- { assume "\ ?thesis" hence "isUIP (opposite ?Cl) (getC state') (getM state')" using ih using \isLastAssertedLiteral ?Cl (oppositeLiteralList (getC state')) (elements (getM state'))\ using \clauseFalse (getC state') (elements (getM state'))\ using lastDecisionThenUIP[of "getM state'" "opposite ?Cl" "getC state'"] unfolding InvariantUniq_def unfolding isUIP_def by simp with \getCn state' \ 1\ have "False" using CnEqual1IffUIP[of "state'"] using ih by simp } thus ?thesis by auto qed have "elementLevel ?Cl (getM state') = currentLevel (getM state')" using ih unfolding InvariantClCurrentLevel_def by simp hence "elementLevel ?Cl (getM state') > 0" using ih by simp obtain reason where "isReason (nth (getF state') reason) ?Cl (elements (getM state'))" "getReason state' ?Cl = Some reason" "0 \ reason \ reason < length (getF state')" using ih unfolding InvariantGetReasonIsReason_def using \literalTrue ?Cl (elements (getM state'))\ using \\ ?Cl el (decisions (getM state'))\ using \elementLevel ?Cl (getM state') > 0\ by auto let ?res = "resolve (getC state') (getF state' ! reason) (opposite ?Cl)" have "getC ?state'' = (remdups (list_diff ?res ?oppM0))" unfolding applyExplain_def unfolding setConflictAnalysisClause_def using \getReason state' ?Cl = Some reason\ by (simp add: Let_def findLastAssertedLiteral_def countCurrentLevelLiterals_def) have "(?res, getC state') \ multLess (getM state')" using multLessResolve[of "?Cl" "getC state'" "nth (getF state') reason" "getM state'"] using \opposite ?Cl el (getC state')\ using \isReason (nth (getF state') reason) ?Cl (elements (getM state'))\ by simp hence "(list_diff ?res ?oppM0, getC state') \ multLess (getM state')" by (simp add: multLessListDiff) have "(remdups (list_diff ?res ?oppM0), getC state') \ multLess (getM state')" using \(list_diff ?res ?oppM0, getC state') \ multLess (getM state')\ by (simp add: multLessRemdups) thus ?thesis using \getC ?state'' = (remdups (list_diff ?res ?oppM0))\ using \getM ?state'' = getM state'\ unfolding multLessState_def by simp qed ultimately have "applyExplainUIP_dom ?state''" using ih by auto thus ?thesis using applyExplainUIP_dom.intros[of "state'"] using False by simp qed qed lemma ApplyExplainUIPPreservedVariables: assumes "applyExplainUIP_dom state" shows "let state' = applyExplainUIP state in (getM state' = getM state) \ (getF state' = getF state) \ (getQ state' = getQ state) \ (getWatch1 state' = getWatch1 state) \ (getWatch2 state' = getWatch2 state) \ (getWatchList state' = getWatchList state) \ (getConflictFlag state' = getConflictFlag state) \ (getConflictClause state' = getConflictClause state) \ (getSATFlag state' = getSATFlag state) \ (getReason state' = getReason state)" (is "let state' = applyExplainUIP state in ?p state state'") using assms proof(induct state rule: applyExplainUIP_dom.induct) case (step state') note ih = this show ?case proof (cases "getCn state' = 1") case True with applyExplainUIP.simps[of "state'"] have "applyExplainUIP state' = state'" by simp thus ?thesis by (auto simp only: Let_def) next case False let ?state' = "applyExplainUIP (applyExplain (getCl state') state')" from applyExplainUIP.simps[of "state'"] False have "applyExplainUIP state' = ?state'" by (simp add: Let_def) have "?p state' (applyExplain (getCl state') state')" unfolding applyExplain_def unfolding setConflictAnalysisClause_def by (auto split: option.split simp add: findLastAssertedLiteral_def countCurrentLevelLiterals_def Let_def) thus ?thesis using ih using False using \applyExplainUIP state' = ?state'\ by (simp add: Let_def) qed qed lemma isUIPApplyExplainUIP: assumes "applyExplainUIP_dom state" "InvariantUniq (getM state)" "InvariantCFalse (getConflictFlag state) (getM state) (getC state)" "InvariantCEntailed (getConflictFlag state) F0 (getC state)" "InvariantClCharacterization (getCl state) (getC state) (getM state)" "InvariantCnCharacterization (getCn state) (getC state) (getM state)" "InvariantClCurrentLevel (getCl state) (getM state)" "InvariantGetReasonIsReason (getReason state) (getF state) (getM state) (set (getQ state))" "InvariantEquivalentZL (getF state) (getM state) F0" "getConflictFlag state" "currentLevel (getM state) > 0" shows "let state' = (applyExplainUIP state) in isUIP (opposite (getCl state')) (getC state') (getM state')" using assms proof(induct state rule: applyExplainUIP_dom.induct) case (step state') note ih = this show ?case proof (cases "getCn state' = 1") case True with applyExplainUIP.simps[of "state'"] have "applyExplainUIP state' = state'" by simp thus ?thesis using ih using CnEqual1IffUIP[of "state'"] using True by (simp add: Let_def) next case False let ?state'' = "applyExplain (getCl state') state'" let ?state' = "applyExplainUIP ?state''" from applyExplainUIP.simps[of "state'"] False have "applyExplainUIP state' = ?state'" by (simp add: Let_def) moreover have "InvariantUniq (getM ?state'')" "InvariantGetReasonIsReason (getReason ?state'') (getF ?state'') (getM ?state'') (set (getQ ?state''))" "InvariantEquivalentZL (getF ?state'') (getM ?state'') F0" "getConflictFlag ?state''" "currentLevel (getM ?state'') > 0" using ih unfolding applyExplain_def unfolding setConflictAnalysisClause_def by (auto split: option.split simp add: findLastAssertedLiteral_def countCurrentLevelLiterals_def Let_def) moreover have "InvariantCFalse (getConflictFlag ?state'') (getM ?state'') (getC ?state'')" "InvariantCEntailed (getConflictFlag ?state'') F0 (getC ?state'')" "InvariantClCharacterization (getCl ?state'') (getC ?state'') (getM ?state'')" "InvariantCnCharacterization (getCn ?state'') (getC ?state'') (getM ?state'')" "InvariantClCurrentLevel (getCl ?state'') (getM ?state'')" using False using ih using InvariantsClAfterApplyExplain[of "state'" "F0"] by (auto simp add: Let_def) ultimately show ?thesis using ih(2) using False by (simp add: Let_def) qed qed lemma InvariantsClAfterExplainUIP: assumes "applyExplainUIP_dom state" "InvariantUniq (getM state)" "InvariantCFalse (getConflictFlag state) (getM state) (getC state)" "InvariantCEntailed (getConflictFlag state) F0 (getC state)" "InvariantClCharacterization (getCl state) (getC state) (getM state)" "InvariantCnCharacterization (getCn state) (getC state) (getM state)" "InvariantClCurrentLevel (getCl state) (getM state)" "InvariantUniqC (getC state)" "InvariantGetReasonIsReason (getReason state) (getF state) (getM state) (set (getQ state))" "InvariantEquivalentZL (getF state) (getM state) F0" "getConflictFlag state" "currentLevel (getM state) > 0" shows "let state' = applyExplainUIP state in InvariantCFalse (getConflictFlag state') (getM state') (getC state') \ InvariantCEntailed (getConflictFlag state') F0 (getC state') \ InvariantClCharacterization (getCl state') (getC state') (getM state') \ InvariantCnCharacterization (getCn state') (getC state') (getM state') \ InvariantClCurrentLevel (getCl state') (getM state') \ InvariantUniqC (getC state')" using assms proof(induct state rule: applyExplainUIP_dom.induct) case (step state') note ih = this show ?case proof (cases "getCn state' = 1") case True with applyExplainUIP.simps[of "state'"] have "applyExplainUIP state' = state'" by simp thus ?thesis using assms using ih by (auto simp only: Let_def) next case False let ?state'' = "applyExplain (getCl state') state'" let ?state' = "applyExplainUIP ?state''" from applyExplainUIP.simps[of "state'"] False have "applyExplainUIP state' = ?state'" by (simp add: Let_def) moreover have "InvariantUniq (getM ?state'')" "InvariantGetReasonIsReason (getReason ?state'') (getF ?state'') (getM ?state'') (set (getQ ?state''))" "InvariantEquivalentZL (getF ?state'') (getM ?state'') F0" "getConflictFlag ?state''" "currentLevel (getM ?state'') > 0" using ih unfolding applyExplain_def unfolding setConflictAnalysisClause_def by (auto split: option.split simp add: findLastAssertedLiteral_def countCurrentLevelLiterals_def Let_def) moreover have "InvariantCFalse (getConflictFlag ?state'') (getM ?state'') (getC ?state'')" "InvariantCEntailed (getConflictFlag ?state'') F0 (getC ?state'')" "InvariantClCharacterization (getCl ?state'') (getC ?state'') (getM ?state'')" "InvariantCnCharacterization (getCn ?state'') (getC ?state'') (getM ?state'')" "InvariantClCurrentLevel (getCl ?state'') (getM ?state'')" "InvariantUniqC (getC ?state'')" using False using ih using InvariantsClAfterApplyExplain[of "state'" "F0"] by (auto simp add: Let_def) ultimately show ?thesis using False using ih(2) by simp qed qed (******************************************************************************) (* G E T B A C K J U M P L E V E L *) (******************************************************************************) lemma oneElementSetCharacterization: shows "(set l = {a}) = ((remdups l) = [a])" proof (induct l) case Nil thus ?case by simp next case (Cons a' l') show ?case proof (cases "l' = []") case True thus ?thesis by simp next case False then obtain b where "b \ set l'" by force show ?thesis proof assume "set (a' # l') = {a}" hence "a' = a" "set l' \ {a}" by auto hence "b = a" using \b \ set l'\ by auto hence "{a} \ set l'" using \b \ set l'\ by auto hence "set l' = {a}" using \set l' \ {a}\ by auto thus "remdups (a' # l') = [a]" using \a' = a\ using Cons by simp next assume "remdups (a' # l') = [a]" thus "set (a' # l') = {a}" using set_remdups[of "a' # l'"] by auto qed qed qed lemma uniqOneElementCharacterization: assumes "uniq l" shows "(l = [a]) = (set l = {a})" using assms using uniqDistinct[of "l"] using oneElementSetCharacterization[of "l" "a"] using distinct_remdups_id[of "l"] by auto lemma isMinimalBackjumpLevelGetBackjumpLevel: assumes "InvariantUniq (getM state)" "InvariantCFalse (getConflictFlag state) (getM state) (getC state)" "InvariantClCharacterization (getCl state) (getC state) (getM state)" "InvariantCllCharacterization (getCl state) (getCll state) (getC state) (getM state)" "InvariantClCurrentLevel (getCl state) (getM state)" "InvariantUniqC (getC state)" "getConflictFlag state" "isUIP (opposite (getCl state)) (getC state) (getM state)" "currentLevel (getM state) > 0" shows "isMinimalBackjumpLevel (getBackjumpLevel state) (opposite (getCl state)) (getC state) (getM state)" proof- let ?oppC = "oppositeLiteralList (getC state)" let ?Cl = "getCl state" have "isLastAssertedLiteral ?Cl ?oppC (elements (getM state))" using \InvariantClCharacterization (getCl state) (getC state) (getM state)\ unfolding InvariantClCharacterization_def by simp have "elementLevel ?Cl (getM state) > 0" using \InvariantClCurrentLevel (getCl state) (getM state)\ using \currentLevel (getM state) > 0\ unfolding InvariantClCurrentLevel_def by simp have "clauseFalse (getC state) (elements (getM state))" using \getConflictFlag state\ using \InvariantCFalse (getConflictFlag state) (getM state) (getC state)\ unfolding InvariantCFalse_def by simp show ?thesis proof (cases "getC state = [opposite ?Cl]") case True thus ?thesis using backjumpLevelZero[of "opposite ?Cl" "oppositeLiteralList ?oppC" "getM state"] using \isLastAssertedLiteral ?Cl ?oppC (elements (getM state))\ using True using \elementLevel ?Cl (getM state) > 0\ unfolding getBackjumpLevel_def unfolding isMinimalBackjumpLevel_def by (simp add: Let_def) next let ?Cll = "getCll state" case False with \InvariantCllCharacterization (getCl state) (getCll state) (getC state) (getM state)\ \InvariantUniqC (getC state)\ have "isLastAssertedLiteral ?Cll (removeAll ?Cl ?oppC) (elements (getM state))" unfolding InvariantCllCharacterization_def unfolding InvariantUniqC_def using uniqOneElementCharacterization[of "getC state" "opposite ?Cl"] by simp hence "?Cll el ?oppC" "?Cll \ ?Cl" unfolding isLastAssertedLiteral_def by auto hence "opposite ?Cll el (getC state)" using literalElListIffOppositeLiteralElOppositeLiteralList[of "?Cll" "?oppC"] by auto show ?thesis using backjumpLevelLastLast[of "opposite ?Cl" "getC state" "getM state" "opposite ?Cll"] using \isUIP (opposite (getCl state)) (getC state) (getM state)\ using \clauseFalse (getC state) (elements (getM state))\ using \isLastAssertedLiteral ?Cll (removeAll ?Cl ?oppC) (elements (getM state))\ using \InvariantUniq (getM state)\ using \InvariantUniqC (getC state)\ using uniqOneElementCharacterization[of "getC state" "opposite ?Cl"] unfolding InvariantUniqC_def unfolding InvariantUniq_def using False using \opposite ?Cll el (getC state)\ unfolding getBackjumpLevel_def unfolding isMinimalBackjumpLevel_def by (auto simp add: Let_def) qed qed (******************************************************************************) (* A P P L Y L E A R N *) (******************************************************************************) lemma applyLearnPreservedVariables: "let state' = applyLearn state in getM state' = getM state \ getQ state' = getQ state \ getC state' = getC state \ getCl state' = getCl state \ getConflictFlag state' = getConflictFlag state \ getConflictClause state' = getConflictClause state \ getF state' = (if getC state = [opposite (getCl state)] then getF state else (getF state @ [getC state]) )" proof (cases "getC state = [opposite (getCl state)]") case True thus ?thesis unfolding applyLearn_def unfolding setWatch1_def unfolding setWatch2_def by (simp add:Let_def) next case False thus ?thesis unfolding applyLearn_def unfolding setWatch1_def unfolding setWatch2_def by (simp add:Let_def) qed lemma WatchInvariantsAfterApplyLearn: assumes "InvariantUniq (getM state)" and "InvariantWatchesEl (getF state) (getWatch1 state) (getWatch2 state)" and "InvariantWatchesDiffer (getF state) (getWatch1 state) (getWatch2 state)" and "InvariantWatchCharacterization (getF state) (getWatch1 state) (getWatch2 state) (getM state)" and "InvariantWatchListsContainOnlyClausesFromF (getWatchList state) (getF state)" and "InvariantWatchListsUniq (getWatchList state)" and "InvariantWatchListsCharacterization (getWatchList state) (getWatch1 state) (getWatch2 state)" and "InvariantClCharacterization (getCl state) (getC state) (getM state)" and "getConflictFlag state" "InvariantCFalse (getConflictFlag state) (getM state) (getC state)" "InvariantUniqC (getC state)" shows "let state' = (applyLearn state) in InvariantWatchesEl (getF state') (getWatch1 state') (getWatch2 state') \ InvariantWatchesDiffer (getF state') (getWatch1 state') (getWatch2 state') \ InvariantWatchCharacterization (getF state') (getWatch1 state') (getWatch2 state') (getM state') \ InvariantWatchListsContainOnlyClausesFromF (getWatchList state') (getF state') \ InvariantWatchListsUniq (getWatchList state') \ InvariantWatchListsCharacterization (getWatchList state') (getWatch1 state') (getWatch2 state')" proof (cases "getC state \ [opposite (getCl state)]") case False thus ?thesis using assms unfolding applyLearn_def unfolding InvariantCllCharacterization_def by (simp add: Let_def) next case True let ?oppC = "oppositeLiteralList (getC state)" let ?l = "getCl state" let ?ll = "getLastAssertedLiteral (removeAll ?l ?oppC) (elements (getM state))" have "clauseFalse (getC state) (elements (getM state))" using \getConflictFlag state\ using \InvariantCFalse (getConflictFlag state) (getM state) (getC state)\ unfolding InvariantCFalse_def by simp from True have "set (getC state) \ {opposite ?l}" using \InvariantUniqC (getC state)\ using uniqOneElementCharacterization[of "getC state" "opposite ?l"] unfolding InvariantUniqC_def by (simp add: Let_def) have "isLastAssertedLiteral ?l ?oppC (elements (getM state))" using \InvariantClCharacterization (getCl state) (getC state) (getM state)\ unfolding InvariantClCharacterization_def by simp have "opposite ?l el (getC state)" using \isLastAssertedLiteral ?l ?oppC (elements (getM state))\ unfolding isLastAssertedLiteral_def using literalElListIffOppositeLiteralElOppositeLiteralList[of "?l" "?oppC"] by simp have "removeAll ?l ?oppC \ []" proof- { assume "\ ?thesis" hence "set ?oppC \ {?l}" using set_removeAll[of "?l" "?oppC"] by auto have "set (getC state) \ {opposite ?l}" proof fix x assume "x \ set (getC state)" hence "opposite x \ set ?oppC" using literalElListIffOppositeLiteralElOppositeLiteralList[of "x" "getC state"] by simp hence "opposite x \ {?l}" using \set ?oppC \ {?l}\ by auto thus "x \ {opposite ?l}" using oppositeSymmetry[of "x" "?l"] by force qed hence False using \set (getC state) \ {opposite ?l}\ using \opposite ?l el getC state\ by (auto simp add: Let_def) } thus ?thesis by auto qed have "clauseFalse (oppositeLiteralList (removeAll ?l ?oppC)) (elements (getM state))" using \clauseFalse (getC state) (elements (getM state))\ using oppositeLiteralListRemove[of "?l" "?oppC"] by (simp add: clauseFalseIffAllLiteralsAreFalse) moreover have "oppositeLiteralList (removeAll ?l ?oppC) \ []" using \removeAll ?l ?oppC \ []\ using oppositeLiteralListNonempty by simp ultimately have "isLastAssertedLiteral ?ll (removeAll ?l ?oppC) (elements (getM state))" using \InvariantUniq (getM state)\ unfolding InvariantUniq_def using getLastAssertedLiteralCharacterization[of "oppositeLiteralList (removeAll ?l ?oppC)" "elements (getM state)"] by auto hence "?ll el (removeAll ?l ?oppC)" unfolding isLastAssertedLiteral_def by auto hence "?ll el ?oppC" "?ll \ ?l" by auto hence "opposite ?ll el (getC state)" using literalElListIffOppositeLiteralElOppositeLiteralList[of "?ll" "?oppC"] by auto let ?state' = "applyLearn state" have "InvariantWatchesEl (getF ?state') (getWatch1 ?state') (getWatch2 ?state')" proof- { fix clause::nat assume "0 \ clause \ clause < length (getF ?state')" have "\w1 w2. getWatch1 ?state' clause = Some w1 \ getWatch2 ?state' clause = Some w2 \ w1 el (getF ?state' ! clause) \ w2 el (getF ?state' ! clause)" proof (cases "clause < length (getF state)") case True thus ?thesis using \InvariantWatchesEl (getF state) (getWatch1 state) (getWatch2 state)\ unfolding InvariantWatchesEl_def using \set (getC state) \ {opposite ?l}\ unfolding applyLearn_def unfolding setWatch1_def unfolding setWatch2_def by (auto simp add:Let_def nth_append) next case False with \0 \ clause \ clause < length (getF ?state')\ have "clause = length (getF state)" using \getC state \ [opposite ?l]\ unfolding applyLearn_def unfolding setWatch1_def unfolding setWatch2_def by (auto simp add: Let_def) moreover have "getWatch1 ?state' clause = Some (opposite ?l)" "getWatch2 ?state' clause = Some (opposite ?ll)" using \clause = length (getF state)\ using \set (getC state) \ {opposite ?l}\ unfolding applyLearn_def unfolding setWatch1_def unfolding setWatch2_def by (auto simp add: Let_def) moreover have "getF ?state' ! clause = (getC state)" using \clause = length (getF state)\ using \set (getC state) \ {opposite ?l}\ unfolding applyLearn_def unfolding setWatch1_def unfolding setWatch2_def by (auto simp add: Let_def) ultimately show ?thesis using \opposite ?l el (getC state)\ \opposite ?ll el (getC state)\ by force qed } thus ?thesis unfolding InvariantWatchesEl_def by auto qed moreover have "InvariantWatchesDiffer (getF ?state') (getWatch1 ?state') (getWatch2 ?state')" proof- { fix clause::nat assume "0 \ clause \ clause < length (getF ?state')" have "getWatch1 ?state' clause \ getWatch2 ?state' clause" proof (cases "clause < length (getF state)") case True thus ?thesis using \InvariantWatchesDiffer (getF state) (getWatch1 state) (getWatch2 state)\ unfolding InvariantWatchesDiffer_def using \set (getC state) \ {opposite ?l}\ unfolding applyLearn_def unfolding setWatch1_def unfolding setWatch2_def by (auto simp add:Let_def nth_append) next case False with \0 \ clause \ clause < length (getF ?state')\ have "clause = length (getF state)" using \getC state \ [opposite ?l]\ unfolding applyLearn_def unfolding setWatch1_def unfolding setWatch2_def by (auto simp add: Let_def) moreover have "getWatch1 ?state' clause = Some (opposite ?l)" "getWatch2 ?state' clause = Some (opposite ?ll)" using \clause = length (getF state)\ using \set (getC state) \ {opposite ?l}\ unfolding applyLearn_def unfolding setWatch1_def unfolding setWatch2_def by (auto simp add: Let_def) moreover have "getF ?state' ! clause = (getC state)" using \clause = length (getF state)\ using \set (getC state) \ {opposite ?l}\ unfolding applyLearn_def unfolding setWatch1_def unfolding setWatch2_def by (auto simp add: Let_def) ultimately show ?thesis using \?ll \ ?l\ by force qed } thus ?thesis unfolding InvariantWatchesDiffer_def by auto qed moreover have "InvariantWatchCharacterization (getF ?state') (getWatch1 ?state') (getWatch2 ?state') (getM ?state')" proof- { fix clause::nat and w1::Literal and w2::Literal assume *: "0 \ clause \ clause < length (getF ?state')" assume **: "Some w1 = getWatch1 ?state' clause" "Some w2 = getWatch2 ?state' clause" have "watchCharacterizationCondition w1 w2 (getM ?state') (getF ?state' ! clause) \ watchCharacterizationCondition w2 w1 (getM ?state') (getF ?state' ! clause)" proof (cases "clause < length (getF state)") case True thus ?thesis using \InvariantWatchCharacterization (getF state) (getWatch1 state) (getWatch2 state) (getM state)\ unfolding InvariantWatchCharacterization_def using \set (getC state) \ {opposite ?l}\ using ** unfolding applyLearn_def unfolding setWatch1_def unfolding setWatch2_def by (auto simp add:Let_def nth_append) next case False with \0 \ clause \ clause < length (getF ?state')\ have "clause = length (getF state)" using \getC state \ [opposite ?l]\ unfolding applyLearn_def unfolding setWatch1_def unfolding setWatch2_def by (auto simp add: Let_def) moreover have "getWatch1 ?state' clause = Some (opposite ?l)" "getWatch2 ?state' clause = Some (opposite ?ll)" using \clause = length (getF state)\ using \set (getC state) \ {opposite ?l}\ unfolding applyLearn_def unfolding setWatch1_def unfolding setWatch2_def by (auto simp add: Let_def) moreover have "\ l. l el (getC state) \ l \ opposite ?l \ l \ opposite ?ll \ elementLevel (opposite l) (getM state) \ elementLevel ?l (getM state) \ elementLevel (opposite l) (getM state) \ elementLevel ?ll (getM state)" proof- { fix l assume "l el (getC state)" "l \ opposite ?l" "l \ opposite ?ll" hence "opposite l el ?oppC" using literalElListIffOppositeLiteralElOppositeLiteralList[of "l" "getC state"] by simp moreover from \l \ opposite ?l\ have "opposite l \ ?l" using oppositeSymmetry[of "l" "?l"] by blast ultimately have "opposite l el (removeAll ?l ?oppC)" by simp from \clauseFalse (getC state) (elements (getM state))\ have "literalFalse l (elements (getM state))" using \l el (getC state)\ by (simp add: clauseFalseIffAllLiteralsAreFalse) hence "elementLevel (opposite l) (getM state) \ elementLevel ?l (getM state) \ elementLevel (opposite l) (getM state) \ elementLevel ?ll (getM state)" using \InvariantUniq (getM state)\ unfolding InvariantUniq_def using \isLastAssertedLiteral ?l ?oppC (elements (getM state))\ using lastAssertedLiteralHasHighestElementLevel[of "?l" "?oppC" "getM state"] using \isLastAssertedLiteral ?ll (removeAll ?l ?oppC) (elements (getM state))\ using lastAssertedLiteralHasHighestElementLevel[of "?ll" "(removeAll ?l ?oppC)" "getM state"] using \opposite l el ?oppC\ \opposite l el (removeAll ?l ?oppC)\ by simp } thus ?thesis by simp qed moreover have "getF ?state' ! clause = (getC state)" using \clause = length (getF state)\ using \set (getC state) \ {opposite ?l}\ unfolding applyLearn_def unfolding setWatch1_def unfolding setWatch2_def by (auto simp add: Let_def) moreover have "getM ?state' = getM state" using \set (getC state) \ {opposite ?l}\ unfolding applyLearn_def unfolding setWatch1_def unfolding setWatch2_def by (auto simp add: Let_def) ultimately show ?thesis using \clauseFalse (getC state) (elements (getM state))\ using ** unfolding watchCharacterizationCondition_def by (auto simp add: clauseFalseIffAllLiteralsAreFalse) qed } thus ?thesis unfolding InvariantWatchCharacterization_def by auto qed moreover have "InvariantWatchListsContainOnlyClausesFromF (getWatchList ?state') (getF ?state')" proof- { fix clause::nat and literal::Literal assume "clause \ set (getWatchList ?state' literal)" have "clause < length (getF ?state')" proof(cases "clause \ set (getWatchList state literal)") case True thus ?thesis using \InvariantWatchListsContainOnlyClausesFromF (getWatchList state) (getF state)\ unfolding InvariantWatchListsContainOnlyClausesFromF_def using \set (getC state) \ {opposite ?l}\ unfolding applyLearn_def unfolding setWatch1_def unfolding setWatch2_def by (auto simp add:Let_def nth_append) (force)+ next case False with \clause \ set (getWatchList ?state' literal)\ have "clause = length (getF state)" using \set (getC state) \ {opposite ?l}\ unfolding applyLearn_def unfolding setWatch1_def unfolding setWatch2_def by (auto simp add:Let_def nth_append split: if_split_asm) thus ?thesis using \set (getC state) \ {opposite ?l}\ unfolding applyLearn_def unfolding setWatch1_def unfolding setWatch2_def by (auto simp add:Let_def nth_append) qed } thus ?thesis unfolding InvariantWatchListsContainOnlyClausesFromF_def by simp qed moreover have "InvariantWatchListsUniq (getWatchList ?state')" unfolding InvariantWatchListsUniq_def proof fix l::Literal show "uniq (getWatchList ?state' l)" proof(cases "l = opposite ?l \ l = opposite ?ll") case True hence "getWatchList ?state' l = (length (getF state)) # getWatchList state l" using \set (getC state) \ {opposite ?l}\ unfolding applyLearn_def unfolding setWatch1_def unfolding setWatch2_def using \?ll \ ?l\ by (auto simp add:Let_def nth_append) moreover have "length (getF state) \ set (getWatchList state l)" using \InvariantWatchListsContainOnlyClausesFromF (getWatchList state) (getF state)\ unfolding InvariantWatchListsContainOnlyClausesFromF_def by auto ultimately show ?thesis using \InvariantWatchListsUniq (getWatchList state)\ unfolding InvariantWatchListsUniq_def by (simp add: uniqAppendIff) next case False hence "getWatchList ?state' l = getWatchList state l" using \set (getC state) \ {opposite ?l}\ unfolding applyLearn_def unfolding setWatch1_def unfolding setWatch2_def by (auto simp add:Let_def nth_append) thus ?thesis using \InvariantWatchListsUniq (getWatchList state)\ unfolding InvariantWatchListsUniq_def by simp qed qed moreover have "InvariantWatchListsCharacterization (getWatchList ?state') (getWatch1 ?state') (getWatch2 ?state')" proof- { fix c::nat and l::Literal have "(c \ set (getWatchList ?state' l)) = (Some l = getWatch1 ?state' c \ Some l = getWatch2 ?state' c)" proof (cases "c = length (getF state)") case False thus ?thesis using \InvariantWatchListsCharacterization (getWatchList state) (getWatch1 state) (getWatch2 state)\ unfolding InvariantWatchListsCharacterization_def using \set (getC state) \ {opposite ?l}\ unfolding applyLearn_def unfolding setWatch1_def unfolding setWatch2_def by (auto simp add:Let_def nth_append) next case True have "length (getF state) \ set (getWatchList state l)" using \InvariantWatchListsContainOnlyClausesFromF (getWatchList state) (getF state)\ unfolding InvariantWatchListsContainOnlyClausesFromF_def by auto thus ?thesis using \c = length (getF state)\ using \InvariantWatchListsCharacterization (getWatchList state) (getWatch1 state) (getWatch2 state)\ unfolding InvariantWatchListsCharacterization_def using \set (getC state) \ {opposite ?l}\ unfolding applyLearn_def unfolding setWatch1_def unfolding setWatch2_def by (auto simp add:Let_def nth_append) qed } thus ?thesis unfolding InvariantWatchListsCharacterization_def by simp qed moreover have "InvariantClCharacterization (getCl ?state') (getC ?state') (getM ?state')" using \InvariantClCharacterization (getCl state) (getC state) (getM state)\ using \set (getC state) \ {opposite ?l}\ unfolding applyLearn_def unfolding setWatch1_def unfolding setWatch2_def by (auto simp add:Let_def) moreover have "InvariantCllCharacterization (getCl ?state') (getCll ?state') (getC ?state') (getM ?state')" unfolding InvariantCllCharacterization_def using \isLastAssertedLiteral ?ll (removeAll ?l ?oppC) (elements (getM state))\ using \set (getC state) \ {opposite ?l}\ unfolding applyLearn_def unfolding setWatch1_def unfolding setWatch2_def by (auto simp add:Let_def) ultimately show ?thesis by simp qed lemma InvariantCllCharacterizationAfterApplyLearn: assumes "InvariantUniq (getM state)" "InvariantClCharacterization (getCl state) (getC state) (getM state)" "InvariantCFalse (getConflictFlag state) (getM state) (getC state)" "InvariantUniqC (getC state)" "getConflictFlag state" shows "let state' = applyLearn state in InvariantCllCharacterization (getCl state') (getCll state') (getC state') (getM state')" proof (cases "getC state \ [opposite (getCl state)]") case False thus ?thesis using assms unfolding applyLearn_def unfolding InvariantCllCharacterization_def by (simp add: Let_def) next case True let ?oppC = "oppositeLiteralList (getC state)" let ?l = "getCl state" let ?ll = "getLastAssertedLiteral (removeAll ?l ?oppC) (elements (getM state))" have "clauseFalse (getC state) (elements (getM state))" using \getConflictFlag state\ using \InvariantCFalse (getConflictFlag state) (getM state) (getC state)\ unfolding InvariantCFalse_def by simp from True have "set (getC state) \ {opposite ?l}" using \InvariantUniqC (getC state)\ using uniqOneElementCharacterization[of "getC state" "opposite ?l"] unfolding InvariantUniqC_def by (simp add: Let_def) have "isLastAssertedLiteral ?l ?oppC (elements (getM state))" using \InvariantClCharacterization (getCl state) (getC state) (getM state)\ unfolding InvariantClCharacterization_def by simp have "opposite ?l el (getC state)" using \isLastAssertedLiteral ?l ?oppC (elements (getM state))\ unfolding isLastAssertedLiteral_def using literalElListIffOppositeLiteralElOppositeLiteralList[of "?l" "?oppC"] by simp have "removeAll ?l ?oppC \ []" proof- { assume "\ ?thesis" hence "set ?oppC \ {?l}" using set_removeAll[of "?l" "?oppC"] by auto have "set (getC state) \ {opposite ?l}" proof fix x assume "x \ set (getC state)" hence "opposite x \ set ?oppC" using literalElListIffOppositeLiteralElOppositeLiteralList[of "x" "getC state"] by simp hence "opposite x \ {?l}" using \set ?oppC \ {?l}\ by auto thus "x \ {opposite ?l}" using oppositeSymmetry[of "x" "?l"] by force qed hence False using \set (getC state) \ {opposite ?l}\ using \opposite ?l el getC state\ by (auto simp add: Let_def) } thus ?thesis by auto qed have "clauseFalse (oppositeLiteralList (removeAll ?l ?oppC)) (elements (getM state))" using \clauseFalse (getC state) (elements (getM state))\ using oppositeLiteralListRemove[of "?l" "?oppC"] by (simp add: clauseFalseIffAllLiteralsAreFalse) moreover have "oppositeLiteralList (removeAll ?l ?oppC) \ []" using \removeAll ?l ?oppC \ []\ using oppositeLiteralListNonempty by simp ultimately have "isLastAssertedLiteral ?ll (removeAll ?l ?oppC) (elements (getM state))" using getLastAssertedLiteralCharacterization[of "oppositeLiteralList (removeAll ?l ?oppC)" "elements (getM state)"] using \InvariantUniq (getM state)\ unfolding InvariantUniq_def by auto thus ?thesis using \set (getC state) \ {opposite ?l}\ unfolding applyLearn_def unfolding setWatch1_def unfolding setWatch2_def unfolding InvariantCllCharacterization_def by (auto simp add:Let_def) qed lemma InvariantConflictClauseCharacterizationAfterApplyLearn: assumes "getConflictFlag state" "InvariantConflictClauseCharacterization (getConflictFlag state) (getConflictClause state) (getF state) (getM state)" shows "let state' = applyLearn state in InvariantConflictClauseCharacterization (getConflictFlag state') (getConflictClause state') (getF state') (getM state')" proof- have "getConflictClause state < length (getF state)" using assms unfolding InvariantConflictClauseCharacterization_def by (auto simp add: Let_def) hence "nth ((getF state) @ [getC state]) (getConflictClause state) = nth (getF state) (getConflictClause state)" by (simp add: nth_append) thus ?thesis using \InvariantConflictClauseCharacterization (getConflictFlag state) (getConflictClause state) (getF state) (getM state)\ unfolding InvariantConflictClauseCharacterization_def unfolding applyLearn_def unfolding setWatch1_def unfolding setWatch2_def by (auto simp add: Let_def clauseFalseAppendValuation) qed lemma InvariantGetReasonIsReasonAfterApplyLearn: assumes "InvariantGetReasonIsReason (getReason state) (getF state) (getM state) (set (getQ state))" shows "let state' = applyLearn state in InvariantGetReasonIsReason (getReason state') (getF state') (getM state') (set (getQ state')) " proof (cases "getC state = [opposite (getCl state)]") case True thus ?thesis unfolding applyLearn_def using assms by (simp add: Let_def) next case False have "InvariantGetReasonIsReason (getReason state) ((getF state) @ [getC state]) (getM state) (set (getQ state))" using assms using nth_append[of "getF state" "[getC state]"] unfolding InvariantGetReasonIsReason_def by auto thus ?thesis using False unfolding applyLearn_def unfolding setWatch1_def unfolding setWatch2_def by (simp add: Let_def) qed lemma InvariantQCharacterizationAfterApplyLearn: assumes "getConflictFlag state" "InvariantQCharacterization (getConflictFlag state) (getQ state) (getF state) (getM state)" shows "let state' = applyLearn state in InvariantQCharacterization (getConflictFlag state') (getQ state') (getF state') (getM state')" using assms unfolding InvariantQCharacterization_def unfolding applyLearn_def unfolding setWatch1_def unfolding setWatch2_def by (simp add: Let_def) lemma InvariantUniqQAfterApplyLearn: assumes "InvariantUniqQ (getQ state)" shows "let state' = applyLearn state in InvariantUniqQ (getQ state')" using assms unfolding applyLearn_def unfolding setWatch1_def unfolding setWatch2_def by (simp add: Let_def) lemma InvariantConflictFlagCharacterizationAfterApplyLearn: assumes "getConflictFlag state" "InvariantConflictFlagCharacterization (getConflictFlag state) (getF state) (getM state)" shows "let state' = applyLearn state in InvariantConflictFlagCharacterization (getConflictFlag state') (getF state') (getM state')" using assms unfolding InvariantConflictFlagCharacterization_def unfolding applyLearn_def unfolding setWatch1_def unfolding setWatch2_def by (auto simp add: Let_def formulaFalseIffContainsFalseClause) lemma InvariantNoDecisionsWhenConflictNorUnitAfterApplyLearn: assumes "InvariantUniq (getM state)" "InvariantConsistent (getM state)" "InvariantNoDecisionsWhenConflict (getF state) (getM state) (currentLevel (getM state))" "InvariantNoDecisionsWhenUnit (getF state) (getM state) (currentLevel (getM state))" "InvariantCFalse (getConflictFlag state) (getM state) (getC state)" and "InvariantClCharacterization (getCl state) (getC state) (getM state)" "InvariantClCurrentLevel (getCl state) (getM state)" "InvariantUniqC (getC state)" "getConflictFlag state" "isUIP (opposite (getCl state)) (getC state) (getM state)" "currentLevel (getM state) > 0" shows "let state' = applyLearn state in InvariantNoDecisionsWhenConflict (getF state) (getM state') (currentLevel (getM state')) \ InvariantNoDecisionsWhenUnit (getF state) (getM state') (currentLevel (getM state')) \ InvariantNoDecisionsWhenConflict [getC state] (getM state') (getBackjumpLevel state') \ InvariantNoDecisionsWhenUnit [getC state] (getM state') (getBackjumpLevel state')" proof- let ?state' = "applyLearn state" let ?l = "getCl state" have "clauseFalse (getC state) (elements (getM state))" using \getConflictFlag state\ using \InvariantCFalse (getConflictFlag state) (getM state) (getC state)\ unfolding InvariantCFalse_def by simp have "getM ?state' = getM state" "getC ?state' = getC state" "getCl ?state' = getCl state" "getConflictFlag ?state' = getConflictFlag state" unfolding applyLearn_def unfolding setWatch2_def unfolding setWatch1_def by (auto simp add: Let_def) hence "InvariantNoDecisionsWhenConflict (getF state) (getM ?state') (currentLevel (getM ?state')) \ InvariantNoDecisionsWhenUnit (getF state) (getM ?state') (currentLevel (getM ?state'))" using \InvariantNoDecisionsWhenConflict (getF state) (getM state) (currentLevel (getM state))\ using \InvariantNoDecisionsWhenUnit (getF state) (getM state) (currentLevel (getM state))\ by simp moreover have "InvariantCllCharacterization (getCl ?state') (getCll ?state') (getC ?state') (getM ?state')" using assms using InvariantCllCharacterizationAfterApplyLearn[of "state"] by (simp add: Let_def) hence "isMinimalBackjumpLevel (getBackjumpLevel ?state') (opposite ?l) (getC ?state') (getM ?state')" using assms using \getM ?state' = getM state\ \getC ?state' = getC state\ \getCl ?state' = getCl state\ \getConflictFlag ?state' = getConflictFlag state\ using isMinimalBackjumpLevelGetBackjumpLevel[of "?state'"] unfolding isUIP_def unfolding SatSolverVerification.isUIP_def by (simp add: Let_def) hence "getBackjumpLevel ?state' < elementLevel ?l (getM ?state')" unfolding isMinimalBackjumpLevel_def unfolding isBackjumpLevel_def by simp hence "getBackjumpLevel ?state' < currentLevel (getM ?state')" using elementLevelLeqCurrentLevel[of "?l" "getM ?state'"] by simp have "InvariantNoDecisionsWhenConflict [getC state] (getM ?state') (getBackjumpLevel ?state') \ InvariantNoDecisionsWhenUnit [getC state] (getM ?state') (getBackjumpLevel ?state')" proof- { fix clause::Clause assume "clause el [getC state]" hence "clause = getC state" by simp have "(\ level'. level' < (getBackjumpLevel ?state') \ \ clauseFalse clause (elements (prefixToLevel level' (getM ?state')))) \ (\ level'. level' < (getBackjumpLevel ?state') \ \ (\ l. isUnitClause clause l (elements (prefixToLevel level' (getM ?state')))))" (is "?false \ ?unit") proof(cases "getC state = [opposite ?l]") case True thus ?thesis using \getM ?state' = getM state\ \getC ?state' = getC state\ \getCl ?state' = getCl state\ unfolding getBackjumpLevel_def by (simp add: Let_def) next case False hence "getF ?state' = getF state @ [getC state]" unfolding applyLearn_def unfolding setWatch2_def unfolding setWatch1_def by (auto simp add: Let_def) show ?thesis proof- have "?unit" using \clause = getC state\ using \InvariantUniq (getM state)\ using \InvariantConsistent (getM state)\ using \getM ?state' = getM state\ \getC ?state' = getC state\ using \clauseFalse (getC state) (elements (getM state))\ using \isMinimalBackjumpLevel (getBackjumpLevel ?state') (opposite ?l) (getC ?state') (getM ?state')\ using isMinimalBackjumpLevelEnsuresIsNotUnitBeforePrefix[of "getM ?state'" "getC ?state'" "getBackjumpLevel ?state'" "opposite ?l"] unfolding InvariantUniq_def unfolding InvariantConsistent_def by simp moreover have "isUnitClause (getC state) (opposite ?l) (elements (prefixToLevel (getBackjumpLevel ?state') (getM state)))" using \InvariantUniq (getM state)\ using \InvariantConsistent (getM state)\ using \isMinimalBackjumpLevel (getBackjumpLevel ?state') (opposite ?l) (getC ?state') (getM ?state')\ using \getM ?state' = getM state\ \getC ?state' = getC state\ using \clauseFalse (getC state) (elements (getM state))\ using isBackjumpLevelEnsuresIsUnitInPrefix[of "getM ?state'" "getC ?state'" "getBackjumpLevel ?state'" "opposite ?l"] unfolding isMinimalBackjumpLevel_def unfolding InvariantUniq_def unfolding InvariantConsistent_def by simp hence "\ clauseFalse (getC state) (elements (prefixToLevel (getBackjumpLevel ?state') (getM state)))" unfolding isUnitClause_def by (auto simp add: clauseFalseIffAllLiteralsAreFalse) have "?false" proof fix level' show "level' < getBackjumpLevel ?state' \ \ clauseFalse clause (elements (prefixToLevel level' (getM ?state')))" proof assume "level' < getBackjumpLevel ?state'" show "\ clauseFalse clause (elements (prefixToLevel level' (getM ?state')))" proof- have "isPrefix (prefixToLevel level' (getM state)) (prefixToLevel (getBackjumpLevel ?state') (getM state))" using \level' < getBackjumpLevel ?state'\ using isPrefixPrefixToLevelLowerLevel[of "level'" "getBackjumpLevel ?state'" "getM state"] by simp then obtain s where "prefixToLevel level' (getM state) @ s = prefixToLevel (getBackjumpLevel ?state') (getM state)" unfolding isPrefix_def by auto hence "prefixToLevel (getBackjumpLevel ?state') (getM state) = prefixToLevel level' (getM state) @ s" by (rule sym) thus ?thesis using \getM ?state' = getM state\ using \clause = getC state\ using \\ clauseFalse (getC state) (elements (prefixToLevel (getBackjumpLevel ?state') (getM state)))\ unfolding isPrefix_def by (auto simp add: clauseFalseIffAllLiteralsAreFalse) qed qed qed ultimately show ?thesis by simp qed qed } thus ?thesis unfolding InvariantNoDecisionsWhenConflict_def unfolding InvariantNoDecisionsWhenUnit_def by (auto simp add: formulaFalseIffContainsFalseClause) qed ultimately show ?thesis by (simp add: Let_def) qed lemma InvariantEquivalentZLAfterApplyLearn: assumes "InvariantEquivalentZL (getF state) (getM state) F0" and "InvariantCEntailed (getConflictFlag state) F0 (getC state)" and "getConflictFlag state" shows "let state' = applyLearn state in InvariantEquivalentZL (getF state') (getM state') F0" proof- let ?M0 = "val2form (elements (prefixToLevel 0 (getM state)))" have "equivalentFormulae F0 (getF state @ ?M0)" using \InvariantEquivalentZL (getF state) (getM state) F0\ using equivalentFormulaeSymmetry[of "F0" "getF state @ ?M0"] unfolding InvariantEquivalentZL_def by simp moreover have "formulaEntailsClause (getF state @ ?M0) (getC state)" using assms unfolding InvariantEquivalentZL_def unfolding InvariantCEntailed_def unfolding equivalentFormulae_def unfolding formulaEntailsClause_def by auto ultimately have "equivalentFormulae F0 ((getF state @ ?M0) @ [getC state])" using extendEquivalentFormulaWithEntailedClause[of "F0" "getF state @ ?M0" "getC state"] by simp hence "equivalentFormulae ((getF state @ ?M0) @ [getC state]) F0" by (simp add: equivalentFormulaeSymmetry) have "equivalentFormulae ((getF state) @ [getC state] @ ?M0) F0" proof- { fix valuation::Valuation have "formulaTrue ((getF state @ ?M0) @ [getC state]) valuation = formulaTrue ((getF state) @ [getC state] @ ?M0) valuation" by (simp add: formulaTrueIffAllClausesAreTrue) } thus ?thesis using \equivalentFormulae ((getF state @ ?M0) @ [getC state]) F0\ unfolding equivalentFormulae_def by auto qed thus ?thesis using assms unfolding InvariantEquivalentZL_def unfolding applyLearn_def unfolding setWatch1_def unfolding setWatch2_def by (auto simp add: Let_def) qed lemma InvariantVarsFAfterApplyLearn: assumes "InvariantCFalse (getConflictFlag state) (getM state) (getC state)" "getConflictFlag state" "InvariantVarsF (getF state) F0 Vbl" "InvariantVarsM (getM state) F0 Vbl" shows "let state' = applyLearn state in InvariantVarsF (getF state') F0 Vbl " proof- from assms have "clauseFalse (getC state) (elements (getM state))" unfolding InvariantCFalse_def by simp hence "vars (getC state) \ vars (elements (getM state))" using valuationContainsItsFalseClausesVariables[of "getC state" "elements (getM state)"] by simp thus ?thesis using applyLearnPreservedVariables[of "state"] using assms using varsAppendFormulae[of "getF state" "[getC state]"] unfolding InvariantVarsF_def unfolding InvariantVarsM_def by (auto simp add: Let_def) qed (******************************************************************************) (* A P P L Y B A C K J U M P *) (******************************************************************************) lemma applyBackjumpEffect: assumes "InvariantConsistent (getM state)" "InvariantUniq (getM state)" "InvariantWatchesEl (getF state) (getWatch1 state) (getWatch2 state)" and "InvariantWatchListsContainOnlyClausesFromF (getWatchList state) (getF state)" and "getConflictFlag state" "InvariantCFalse (getConflictFlag state) (getM state) (getC state)" and "InvariantCEntailed (getConflictFlag state) F0 (getC state)" and "InvariantClCharacterization (getCl state) (getC state) (getM state)" and "InvariantCllCharacterization (getCl state) (getCll state) (getC state) (getM state)" and "InvariantClCurrentLevel (getCl state) (getM state)" "InvariantUniqC (getC state)" "isUIP (opposite (getCl state)) (getC state) (getM state)" "currentLevel (getM state) > 0" shows "let l = (getCl state) in let bClause = (getC state) in let bLiteral = opposite l in let level = getBackjumpLevel state in let prefix = prefixToLevel level (getM state) in let state'' = applyBackjump state in (formulaEntailsClause F0 bClause \ isUnitClause bClause bLiteral (elements prefix) \ (getM state'') = prefix @ [(bLiteral, False)]) \ getF state'' = getF state" proof- let ?l = "getCl state" let ?level = "getBackjumpLevel state" let ?prefix = "prefixToLevel ?level (getM state)" let ?state' = "state\ getConflictFlag := False, getQ := [], getM := ?prefix \" let ?state'' = "applyBackjump state" have "clauseFalse (getC state) (elements (getM state))" using \getConflictFlag state\ using \InvariantCFalse (getConflictFlag state) (getM state) (getC state)\ unfolding InvariantCFalse_def by simp have "formulaEntailsClause F0 (getC state)" using \getConflictFlag state\ using \InvariantCEntailed (getConflictFlag state) F0 (getC state)\ unfolding InvariantCEntailed_def by simp have "isBackjumpLevel ?level (opposite ?l) (getC state) (getM state)" using assms using isMinimalBackjumpLevelGetBackjumpLevel[of "state"] unfolding isMinimalBackjumpLevel_def by (simp add: Let_def) then have "isUnitClause (getC state) (opposite ?l) (elements ?prefix)" using assms using \clauseFalse (getC state) (elements (getM state))\ using isBackjumpLevelEnsuresIsUnitInPrefix[of "getM state" "getC state" "?level" "opposite ?l"] unfolding InvariantConsistent_def unfolding InvariantUniq_def by simp moreover have "getM ?state'' = ?prefix @ [(opposite ?l, False)]" "getF ?state'' = getF state" unfolding applyBackjump_def using assms using assertLiteralEffect unfolding setReason_def by (auto simp add: Let_def) ultimately show ?thesis using \formulaEntailsClause F0 (getC state)\ by (simp add: Let_def) qed lemma applyBackjumpPreservedVariables: assumes "InvariantWatchListsContainOnlyClausesFromF (getWatchList state) (getF state)" "InvariantWatchesEl (getF state) (getWatch1 state) (getWatch2 state)" shows "let state' = applyBackjump state in getSATFlag state' = getSATFlag state" using assms unfolding applyBackjump_def unfolding setReason_def by (auto simp add: Let_def assertLiteralEffect) lemma InvariantWatchCharacterizationInBackjumpPrefix: assumes "InvariantWatchCharacterization (getF state) (getWatch1 state) (getWatch2 state) (getM state)" shows "let l = getCl state in let level = getBackjumpLevel state in let prefix = prefixToLevel level (getM state) in let state' = state\ getConflictFlag := False, getQ := [], getM := prefix \ in InvariantWatchCharacterization (getF state') (getWatch1 state') (getWatch2 state') (getM state')" proof- let ?l = "getCl state" let ?level = "getBackjumpLevel state" let ?prefix = "prefixToLevel ?level (getM state)" let ?state' = "state\ getConflictFlag := False, getQ := [], getM := ?prefix \" { fix c w1 w2 assume "c < length (getF state)" "Some w1 = getWatch1 state c" "Some w2 = getWatch2 state c" with \InvariantWatchCharacterization (getF state) (getWatch1 state) (getWatch2 state) (getM state)\ have "watchCharacterizationCondition w1 w2 (getM state) (nth (getF state) c)" "watchCharacterizationCondition w2 w1 (getM state) (nth (getF state) c)" unfolding InvariantWatchCharacterization_def by auto let ?clause = "nth (getF state) c" let "?a state w1 w2" = "\ l. l el ?clause \ literalTrue l (elements (getM state)) \ elementLevel l (getM state) \ elementLevel (opposite w1) (getM state)" let "?b state w1 w2" = "\ l. l el ?clause \ l \ w1 \ l \ w2 \ literalFalse l (elements (getM state)) \ elementLevel (opposite l) (getM state) \ elementLevel (opposite w1) (getM state)" have "watchCharacterizationCondition w1 w2 (getM ?state') ?clause \ watchCharacterizationCondition w2 w1 (getM ?state') ?clause" proof- { assume "literalFalse w1 (elements (getM ?state'))" hence "literalFalse w1 (elements (getM state))" using isPrefixPrefixToLevel[of "?level" "getM state"] using isPrefixElements[of "prefixToLevel ?level (getM state)" "getM state"] using prefixIsSubset[of "elements (prefixToLevel ?level (getM state))" "elements (getM state)"] by auto from \literalFalse w1 (elements (getM ?state'))\ have "elementLevel (opposite w1) (getM state) \ ?level" using prefixToLevelElementsElementLevel[of "opposite w1" "?level" "getM state"] by simp from \literalFalse w1 (elements (getM ?state'))\ have "elementLevel (opposite w1) (getM ?state') = elementLevel (opposite w1) (getM state)" using elementLevelPrefixElement by simp have "?a ?state' w1 w2 \ ?b ?state' w1 w2" proof (cases "?a state w1 w2") case True then obtain l where "l el ?clause" "literalTrue l (elements (getM state))" "elementLevel l (getM state) \ elementLevel (opposite w1) (getM state)" by auto have "literalTrue l (elements (getM ?state'))" using \elementLevel (opposite w1) (getM state) \ ?level\ using elementLevelLtLevelImpliesMemberPrefixToLevel[of "l" "getM state" "?level"] using \elementLevel l (getM state) \ elementLevel (opposite w1) (getM state)\ using \literalTrue l (elements (getM state))\ by simp moreover from \literalTrue l (elements (getM ?state'))\ have "elementLevel l (getM ?state') = elementLevel l (getM state)" using elementLevelPrefixElement by simp ultimately show ?thesis using \elementLevel (opposite w1) (getM ?state') = elementLevel (opposite w1) (getM state)\ using \elementLevel l (getM state) \ elementLevel (opposite w1) (getM state)\ using \l el ?clause\ by auto next case False { fix l assume "l el ?clause" "l \ w1" "l \ w2" hence "literalFalse l (elements (getM state))" "elementLevel (opposite l) (getM state) \ elementLevel (opposite w1) (getM state)" using \literalFalse w1 (elements (getM state))\ using False using \watchCharacterizationCondition w1 w2 (getM state) ?clause\ unfolding watchCharacterizationCondition_def by auto have "literalFalse l (elements (getM ?state')) \ elementLevel (opposite l) (getM ?state') \ elementLevel (opposite w1) (getM ?state')" proof- have "literalFalse l (elements (getM ?state'))" using \elementLevel (opposite w1) (getM state) \ ?level\ using elementLevelLtLevelImpliesMemberPrefixToLevel[of "opposite l" "getM state" "?level"] using \elementLevel (opposite l) (getM state) \ elementLevel (opposite w1) (getM state)\ using \literalFalse l (elements (getM state))\ by simp moreover from \literalFalse l (elements (getM ?state'))\ have "elementLevel (opposite l) (getM ?state') = elementLevel (opposite l) (getM state)" using elementLevelPrefixElement by simp ultimately show ?thesis using \elementLevel (opposite w1) (getM ?state') = elementLevel (opposite w1) (getM state)\ using \elementLevel (opposite l) (getM state) \ elementLevel (opposite w1) (getM state)\ using \l el ?clause\ by auto qed } thus ?thesis by auto qed } moreover { assume "literalFalse w2 (elements (getM ?state'))" hence "literalFalse w2 (elements (getM state))" using isPrefixPrefixToLevel[of "?level" "getM state"] using isPrefixElements[of "prefixToLevel ?level (getM state)" "getM state"] using prefixIsSubset[of "elements (prefixToLevel ?level (getM state))" "elements (getM state)"] by auto from \literalFalse w2 (elements (getM ?state'))\ have "elementLevel (opposite w2) (getM state) \ ?level" using prefixToLevelElementsElementLevel[of "opposite w2" "?level" "getM state"] by simp from \literalFalse w2 (elements (getM ?state'))\ have "elementLevel (opposite w2) (getM ?state') = elementLevel (opposite w2) (getM state)" using elementLevelPrefixElement by simp have "?a ?state' w2 w1 \ ?b ?state' w2 w1" proof (cases "?a state w2 w1") case True then obtain l where "l el ?clause" "literalTrue l (elements (getM state))" "elementLevel l (getM state) \ elementLevel (opposite w2) (getM state)" by auto have "literalTrue l (elements (getM ?state'))" using \elementLevel (opposite w2) (getM state) \ ?level\ using elementLevelLtLevelImpliesMemberPrefixToLevel[of "l" "getM state" "?level"] using \elementLevel l (getM state) \ elementLevel (opposite w2) (getM state)\ using \literalTrue l (elements (getM state))\ by simp moreover from \literalTrue l (elements (getM ?state'))\ have "elementLevel l (getM ?state') = elementLevel l (getM state)" using elementLevelPrefixElement by simp ultimately show ?thesis using \elementLevel (opposite w2) (getM ?state') = elementLevel (opposite w2) (getM state)\ using \elementLevel l (getM state) \ elementLevel (opposite w2) (getM state)\ using \l el ?clause\ by auto next case False { fix l assume "l el ?clause" "l \ w1" "l \ w2" hence "literalFalse l (elements (getM state))" "elementLevel (opposite l) (getM state) \ elementLevel (opposite w2) (getM state)" using \literalFalse w2 (elements (getM state))\ using False using \watchCharacterizationCondition w2 w1 (getM state) ?clause\ unfolding watchCharacterizationCondition_def by auto have "literalFalse l (elements (getM ?state')) \ elementLevel (opposite l) (getM ?state') \ elementLevel (opposite w2) (getM ?state')" proof- have "literalFalse l (elements (getM ?state'))" using \elementLevel (opposite w2) (getM state) \ ?level\ using elementLevelLtLevelImpliesMemberPrefixToLevel[of "opposite l" "getM state" "?level"] using \elementLevel (opposite l) (getM state) \ elementLevel (opposite w2) (getM state)\ using \literalFalse l (elements (getM state))\ by simp moreover from \literalFalse l (elements (getM ?state'))\ have "elementLevel (opposite l) (getM ?state') = elementLevel (opposite l) (getM state)" using elementLevelPrefixElement by simp ultimately show ?thesis using \elementLevel (opposite w2) (getM ?state') = elementLevel (opposite w2) (getM state)\ using \elementLevel (opposite l) (getM state) \ elementLevel (opposite w2) (getM state)\ using \l el ?clause\ by auto qed } thus ?thesis by auto qed } ultimately show ?thesis unfolding watchCharacterizationCondition_def by auto qed } thus ?thesis unfolding InvariantWatchCharacterization_def by auto qed lemma InvariantConsistentAfterApplyBackjump: assumes "InvariantConsistent (getM state)" "InvariantUniq (getM state)" "InvariantWatchesEl (getF state) (getWatch1 state) (getWatch2 state)" and "InvariantWatchListsContainOnlyClausesFromF (getWatchList state) (getF state)" and "getConflictFlag state" "InvariantCFalse (getConflictFlag state) (getM state) (getC state)" and "InvariantUniqC (getC state)" "InvariantCEntailed (getConflictFlag state) F0 (getC state)" and "InvariantClCharacterization (getCl state) (getC state) (getM state)" and "InvariantCllCharacterization (getCl state) (getCll state) (getC state) (getM state)" and "InvariantClCurrentLevel (getCl state) (getM state)" "currentLevel (getM state) > 0" "isUIP (opposite (getCl state)) (getC state) (getM state)" shows "let state' = applyBackjump state in InvariantConsistent (getM state')" proof- let ?l = "getCl state" let ?bClause = "getC state" let ?bLiteral = "opposite ?l" let ?level = "getBackjumpLevel state" let ?prefix = "prefixToLevel ?level (getM state)" let ?state'' = "applyBackjump state" have "formulaEntailsClause F0 ?bClause" and "isUnitClause ?bClause ?bLiteral (elements ?prefix)" and "getM ?state'' = ?prefix @ [(?bLiteral, False)]" using assms using applyBackjumpEffect[of "state"] by (auto simp add: Let_def) thus ?thesis using \InvariantConsistent (getM state)\ using InvariantConsistentAfterBackjump[of "getM state" "?prefix" "?bClause" "?bLiteral" "getM ?state''"] using isPrefixPrefixToLevel by (auto simp add: Let_def) qed lemma InvariantUniqAfterApplyBackjump: assumes "InvariantConsistent (getM state)" "InvariantUniq (getM state)" "InvariantWatchesEl (getF state) (getWatch1 state) (getWatch2 state)" and "InvariantWatchListsContainOnlyClausesFromF (getWatchList state) (getF state)" and "getConflictFlag state" "InvariantCFalse (getConflictFlag state) (getM state) (getC state)" and "InvariantUniqC (getC state)" "InvariantCEntailed (getConflictFlag state) F0 (getC state)" and "InvariantClCharacterization (getCl state) (getC state) (getM state)" and "InvariantCllCharacterization (getCl state) (getCll state) (getC state) (getM state)" and "InvariantClCurrentLevel (getCl state) (getM state)" "currentLevel (getM state) > 0" "isUIP (opposite (getCl state)) (getC state) (getM state)" shows "let state' = applyBackjump state in InvariantUniq (getM state')" proof- let ?l = "getCl state" let ?bClause = "getC state" let ?bLiteral = "opposite ?l" let ?level = "getBackjumpLevel state" let ?prefix = "prefixToLevel ?level (getM state)" let ?state'' = "applyBackjump state" have "clauseFalse (getC state) (elements (getM state))" using \getConflictFlag state\ using \InvariantCFalse (getConflictFlag state) (getM state) (getC state)\ unfolding InvariantCFalse_def by simp have "isUnitClause ?bClause ?bLiteral (elements ?prefix)" and "getM ?state'' = ?prefix @ [(?bLiteral, False)]" using assms using applyBackjumpEffect[of "state"] by (auto simp add: Let_def) thus ?thesis using \InvariantUniq (getM state)\ using InvariantUniqAfterBackjump[of "getM state" "?prefix" "?bClause" "?bLiteral" "getM ?state''"] using isPrefixPrefixToLevel by (auto simp add: Let_def) qed lemma WatchInvariantsAfterApplyBackjump: assumes "InvariantConsistent (getM state)" "InvariantUniq (getM state)" "InvariantWatchesEl (getF state) (getWatch1 state) (getWatch2 state)" and "InvariantWatchesDiffer (getF state) (getWatch1 state) (getWatch2 state)" and "InvariantWatchCharacterization (getF state) (getWatch1 state) (getWatch2 state) (getM state)" and "InvariantWatchListsContainOnlyClausesFromF (getWatchList state) (getF state)" and "InvariantWatchListsUniq (getWatchList state)" and "InvariantWatchListsCharacterization (getWatchList state) (getWatch1 state) (getWatch2 state)" "getConflictFlag state" "InvariantUniqC (getC state)" "InvariantCFalse (getConflictFlag state) (getM state) (getC state)" and "InvariantCEntailed (getConflictFlag state) F0 (getC state)" and "InvariantClCharacterization (getCl state) (getC state) (getM state)" and "InvariantCllCharacterization (getCl state) (getCll state) (getC state) (getM state)" and "InvariantClCurrentLevel (getCl state) (getM state)" "isUIP (opposite (getCl state)) (getC state) (getM state)" "currentLevel (getM state) > 0" shows "let state' = (applyBackjump state) in InvariantWatchesEl (getF state') (getWatch1 state') (getWatch2 state') \ InvariantWatchesDiffer (getF state') (getWatch1 state') (getWatch2 state') \ InvariantWatchCharacterization (getF state') (getWatch1 state') (getWatch2 state') (getM state') \ InvariantWatchListsContainOnlyClausesFromF (getWatchList state') (getF state') \ InvariantWatchListsUniq (getWatchList state') \ InvariantWatchListsCharacterization (getWatchList state') (getWatch1 state') (getWatch2 state')" (is "let state' = (applyBackjump state) in ?inv state'") proof- let ?l = "getCl state" let ?level = "getBackjumpLevel state" let ?prefix = "prefixToLevel ?level (getM state)" let ?state' = "state\ getConflictFlag := False, getQ := [], getM := ?prefix \" let ?state'' = "setReason (opposite (getCl state)) (length (getF state) - 1) ?state'" let ?state0 = "assertLiteral (opposite (getCl state)) False ?state''" have "getF ?state' = getF state" "getWatchList ?state' = getWatchList state" "getWatch1 ?state' = getWatch1 state" "getWatch2 ?state' = getWatch2 state" unfolding setReason_def by (auto simp add: Let_def) moreover have "InvariantWatchCharacterization (getF ?state') (getWatch1 ?state') (getWatch2 ?state') (getM ?state')" using assms using InvariantWatchCharacterizationInBackjumpPrefix[of "state"] unfolding setReason_def by (simp add: Let_def) moreover have "InvariantConsistent (?prefix @ [(opposite ?l, False)])" using assms using InvariantConsistentAfterApplyBackjump[of "state" "F0"] using assertLiteralEffect unfolding applyBackjump_def unfolding setReason_def by (auto simp add: Let_def split: if_split_asm) moreover have "InvariantUniq (?prefix @ [(opposite ?l, False)])" using assms using InvariantUniqAfterApplyBackjump[of "state" "F0"] using assertLiteralEffect unfolding applyBackjump_def unfolding setReason_def by (auto simp add: Let_def split: if_split_asm) ultimately show ?thesis using assms using WatchInvariantsAfterAssertLiteral[of "?state''" "opposite ?l" "False"] using WatchInvariantsAfterAssertLiteral[of "?state'" "opposite ?l" "False"] using InvariantWatchCharacterizationAfterAssertLiteral[of "?state''" "opposite ?l" "False"] using InvariantWatchCharacterizationAfterAssertLiteral[of "?state'" "opposite ?l" "False"] unfolding applyBackjump_def unfolding setReason_def by (auto simp add: Let_def) qed lemma InvariantUniqQAfterApplyBackjump: assumes "InvariantWatchListsContainOnlyClausesFromF (getWatchList state) (getF state)" and "InvariantWatchesEl (getF state) (getWatch1 state) (getWatch2 state)" shows "let state' = applyBackjump state in InvariantUniqQ (getQ state')" proof- let ?l = "getCl state" let ?level = "getBackjumpLevel state" let ?prefix = "prefixToLevel ?level (getM state)" let ?state' = "state\ getConflictFlag := False, getQ := [], getM := ?prefix \" let ?state'' = "setReason (opposite (getCl state)) (length (getF state) - 1) ?state'" show ?thesis using assms unfolding applyBackjump_def using InvariantUniqQAfterAssertLiteral[of "?state'" "opposite ?l" "False"] using InvariantUniqQAfterAssertLiteral[of "?state''" "opposite ?l" "False"] unfolding InvariantUniqQ_def unfolding setReason_def by (auto simp add: Let_def) qed lemma invariantQCharacterizationAfterApplyBackjump_1: assumes "InvariantConsistent (getM state)" "InvariantUniq (getM state)" "InvariantWatchListsContainOnlyClausesFromF (getWatchList state) (getF state)" and "InvariantWatchListsUniq (getWatchList state)" and "InvariantWatchListsCharacterization (getWatchList state) (getWatch1 state) (getWatch2 state)" "InvariantWatchesEl (getF state) (getWatch1 state) (getWatch2 state)" and "InvariantWatchesDiffer (getF state) (getWatch1 state) (getWatch2 state)" and "InvariantWatchCharacterization (getF state) (getWatch1 state) (getWatch2 state) (getM state)" and "InvariantConflictFlagCharacterization (getConflictFlag state) (getF state) (getM state)" and "InvariantQCharacterization (getConflictFlag state) (getQ state) (getF state) (getM state)" and "InvariantUniqC (getC state)" "getC state = [opposite (getCl state)]" "InvariantNoDecisionsWhenUnit (getF state) (getM state) (currentLevel (getM state))" "InvariantNoDecisionsWhenConflict (getF state) (getM state) (currentLevel (getM state))" "getConflictFlag state" "InvariantCFalse (getConflictFlag state) (getM state) (getC state)" "InvariantCEntailed (getConflictFlag state) F0 (getC state)" and "InvariantClCharacterization (getCl state) (getC state) (getM state)" and "InvariantCllCharacterization (getCl state) (getCll state) (getC state) (getM state)" and "InvariantClCurrentLevel (getCl state) (getM state)" "currentLevel (getM state) > 0" "isUIP (opposite (getCl state)) (getC state) (getM state)" shows "let state'' = (applyBackjump state) in InvariantQCharacterization (getConflictFlag state'') (getQ state'') (getF state'') (getM state'')" proof- let ?l = "getCl state" let ?level = "getBackjumpLevel state" let ?prefix = "prefixToLevel ?level (getM state)" let ?state' = "state\ getConflictFlag := False, getQ := [], getM := ?prefix \" let ?state'' = "setReason (opposite (getCl state)) (length (getF state) - 1) ?state'" let ?state'1 = "assertLiteral (opposite ?l) False ?state'" let ?state''1 = "assertLiteral (opposite ?l) False ?state''" have "?level < elementLevel ?l (getM state)" using assms using isMinimalBackjumpLevelGetBackjumpLevel[of "state"] unfolding isMinimalBackjumpLevel_def unfolding isBackjumpLevel_def by (simp add: Let_def) hence "?level < currentLevel (getM state)" using elementLevelLeqCurrentLevel[of "?l" "getM state"] by simp hence "InvariantQCharacterization (getConflictFlag ?state') (getQ ?state') (getF ?state') (getM ?state')" "InvariantConflictFlagCharacterization (getConflictFlag ?state') (getF ?state') (getM ?state')" unfolding InvariantQCharacterization_def unfolding InvariantConflictFlagCharacterization_def using \InvariantNoDecisionsWhenConflict (getF state) (getM state) (currentLevel (getM state))\ using \InvariantNoDecisionsWhenUnit (getF state) (getM state) (currentLevel (getM state))\ unfolding InvariantNoDecisionsWhenConflict_def unfolding InvariantNoDecisionsWhenUnit_def unfolding applyBackjump_def by (auto simp add: Let_def set_conv_nth) moreover have "InvariantConsistent (?prefix @ [(opposite ?l, False)])" using assms using InvariantConsistentAfterApplyBackjump[of "state" "F0"] using assertLiteralEffect unfolding applyBackjump_def unfolding setReason_def by (auto simp add: Let_def split: if_split_asm) moreover have "InvariantWatchCharacterization (getF ?state') (getWatch1 ?state') (getWatch2 ?state') (getM ?state')" using InvariantWatchCharacterizationInBackjumpPrefix[of "state"] using assms by (simp add: Let_def) moreover have "\ opposite ?l el (getQ ?state'1)" "\ opposite ?l el (getQ ?state''1)" using assertedLiteralIsNotUnit[of "?state'" "opposite ?l" "False"] using assertedLiteralIsNotUnit[of "?state''" "opposite ?l" "False"] using \InvariantQCharacterization (getConflictFlag ?state') (getQ ?state') (getF ?state') (getM ?state')\ using \InvariantConsistent (?prefix @ [(opposite ?l, False)])\ using \InvariantWatchCharacterization (getF ?state') (getWatch1 ?state') (getWatch2 ?state') (getM ?state')\ unfolding applyBackjump_def unfolding setReason_def using assms by (auto simp add: Let_def split: if_split_asm) hence "removeAll (opposite ?l) (getQ ?state'1) = getQ ?state'1" "removeAll (opposite ?l) (getQ ?state''1) = getQ ?state''1" using removeAll_id[of "opposite ?l" "getQ ?state'1"] using removeAll_id[of "opposite ?l" "getQ ?state''1"] unfolding setReason_def by auto ultimately show ?thesis using assms using InvariantWatchCharacterizationInBackjumpPrefix[of "state"] using InvariantQCharacterizationAfterAssertLiteral[of "?state'" "opposite ?l" "False"] using InvariantQCharacterizationAfterAssertLiteral[of "?state''" "opposite ?l" "False"] unfolding applyBackjump_def unfolding setReason_def by (auto simp add: Let_def) qed lemma invariantQCharacterizationAfterApplyBackjump_2: fixes state::State assumes "InvariantConsistent (getM state)" "InvariantUniq (getM state)" "InvariantWatchListsContainOnlyClausesFromF (getWatchList state) (getF state)" and "InvariantWatchListsUniq (getWatchList state)" and "InvariantWatchListsCharacterization (getWatchList state) (getWatch1 state) (getWatch2 state)" "InvariantWatchesEl (getF state) (getWatch1 state) (getWatch2 state)" and "InvariantWatchesDiffer (getF state) (getWatch1 state) (getWatch2 state)" and "InvariantWatchCharacterization (getF state) (getWatch1 state) (getWatch2 state) (getM state)" and "InvariantConflictFlagCharacterization (getConflictFlag state) (getF state) (getM state)" and "InvariantQCharacterization (getConflictFlag state) (getQ state) (getF state) (getM state)" and "InvariantUniqC (getC state)" "getC state \ [opposite (getCl state)]" "InvariantNoDecisionsWhenUnit (butlast (getF state)) (getM state) (currentLevel (getM state))" "InvariantNoDecisionsWhenConflict (butlast (getF state)) (getM state) (currentLevel (getM state))" "getF state \ []" "last (getF state) = getC state" "getConflictFlag state" "InvariantCFalse (getConflictFlag state) (getM state) (getC state)" and "InvariantCEntailed (getConflictFlag state) F0 (getC state)" and "InvariantClCharacterization (getCl state) (getC state) (getM state)" and "InvariantCllCharacterization (getCl state) (getCll state) (getC state) (getM state)" and "InvariantClCurrentLevel (getCl state) (getM state)" "currentLevel (getM state) > 0" "isUIP (opposite (getCl state)) (getC state) (getM state)" shows "let state'' = (applyBackjump state) in InvariantQCharacterization (getConflictFlag state'') (getQ state'') (getF state'') (getM state'')" proof- let ?l = "getCl state" let ?level = "getBackjumpLevel state" let ?prefix = "prefixToLevel ?level (getM state)" let ?state' = "state\ getConflictFlag := False, getQ := [], getM := ?prefix \" let ?state'' = "setReason (opposite (getCl state)) (length (getF state) - 1) ?state'" have "?level < elementLevel ?l (getM state)" using assms using isMinimalBackjumpLevelGetBackjumpLevel[of "state"] unfolding isMinimalBackjumpLevel_def unfolding isBackjumpLevel_def by (simp add: Let_def) hence "?level < currentLevel (getM state)" using elementLevelLeqCurrentLevel[of "?l" "getM state"] by simp have "isUnitClause (last (getF state)) (opposite ?l) (elements ?prefix)" using \last (getF state) = getC state\ using isMinimalBackjumpLevelGetBackjumpLevel[of "state"] using \InvariantUniq (getM state)\ using \InvariantConsistent (getM state)\ using \getConflictFlag state\ using \InvariantUniqC (getC state)\ using \InvariantCFalse (getConflictFlag state) (getM state) (getC state)\ using isBackjumpLevelEnsuresIsUnitInPrefix[of "getM state" "getC state" "getBackjumpLevel state" "opposite ?l"] using \InvariantClCharacterization (getCl state) (getC state) (getM state)\ using \InvariantCllCharacterization (getCl state) (getCll state) (getC state) (getM state)\ using \InvariantClCurrentLevel (getCl state) (getM state)\ using \currentLevel (getM state) > 0\ using \isUIP (opposite (getCl state)) (getC state) (getM state)\ unfolding isMinimalBackjumpLevel_def unfolding InvariantUniq_def unfolding InvariantConsistent_def unfolding InvariantCFalse_def by (simp add: Let_def) hence "\ clauseFalse (last (getF state)) (elements ?prefix)" unfolding isUnitClause_def by (auto simp add: clauseFalseIffAllLiteralsAreFalse) have "InvariantConsistent (?prefix @ [(opposite ?l, False)])" using assms using InvariantConsistentAfterApplyBackjump[of "state" "F0"] using assertLiteralEffect unfolding applyBackjump_def unfolding setReason_def by (auto simp add: Let_def split: if_split_asm) have "InvariantUniq (?prefix @ [(opposite ?l, False)])" using assms using InvariantUniqAfterApplyBackjump[of "state" "F0"] using assertLiteralEffect unfolding applyBackjump_def unfolding setReason_def by (auto simp add: Let_def split: if_split_asm) let ?state'1 = "?state' \ getQ := getQ ?state' @ [opposite ?l]\" let ?state'2 = "assertLiteral (opposite ?l) False ?state'1" let ?state''1 = "?state'' \ getQ := getQ ?state'' @ [opposite ?l]\" let ?state''2 = "assertLiteral (opposite ?l) False ?state''1" have "InvariantQCharacterization (getConflictFlag ?state') ((getQ ?state') @ [opposite ?l]) (getF ?state') (getM ?state')" proof- have "\ l c. c el (butlast (getF state)) \ \ isUnitClause c l (elements (getM ?state'))" using \InvariantNoDecisionsWhenUnit (butlast (getF state)) (getM state) (currentLevel (getM state))\ using \?level < currentLevel (getM state)\ unfolding InvariantNoDecisionsWhenUnit_def by simp have "\ l. ((\ c. c el (getF state) \ isUnitClause c l (elements (getM ?state'))) = (l = opposite ?l))" proof fix l show "(\ c. c el (getF state) \ isUnitClause c l (elements (getM ?state'))) = (l = opposite ?l)" (is "?lhs = ?rhs") proof assume "?lhs" then obtain c::Clause where "c el (getF state)" and "isUnitClause c l (elements ?prefix)" by auto show "?rhs" proof (cases "c el (butlast (getF state))") case True thus ?thesis using \\ l c. c el (butlast (getF state)) \ \ isUnitClause c l (elements (getM ?state'))\ using \isUnitClause c l (elements ?prefix)\ by auto next case False from \getF state \ []\ have "butlast (getF state) @ [last (getF state)] = getF state" using append_butlast_last_id[of "getF state"] by simp hence "getF state = butlast (getF state) @ [last (getF state)]" by (rule sym) with \c el getF state\ have "c el butlast (getF state) \ c el [last (getF state)]" using set_append[of "butlast (getF state)" "[last (getF state)]"] by auto hence "c = last (getF state)" using \\ c el (butlast (getF state))\ by simp thus ?thesis using \isUnitClause (last (getF state)) (opposite ?l) (elements ?prefix)\ using \isUnitClause c l (elements ?prefix)\ unfolding isUnitClause_def by auto qed next from \getF state \ []\ have "last (getF state) el (getF state)" by auto assume "?rhs" thus "?lhs" using \isUnitClause (last (getF state)) (opposite ?l) (elements ?prefix)\ using \last (getF state) el (getF state)\ by auto qed qed thus ?thesis unfolding InvariantQCharacterization_def by simp qed hence "InvariantQCharacterization (getConflictFlag ?state'1) (getQ ?state'1) (getF ?state'1) (getM ?state'1)" by simp hence "InvariantQCharacterization (getConflictFlag ?state''1) (getQ ?state''1) (getF ?state''1) (getM ?state''1)" unfolding setReason_def by simp have "InvariantWatchCharacterization (getF ?state'1) (getWatch1 ?state'1) (getWatch2 ?state'1) (getM ?state'1)" using InvariantWatchCharacterizationInBackjumpPrefix[of "state"] using assms by (simp add: Let_def) hence "InvariantWatchCharacterization (getF ?state''1) (getWatch1 ?state''1) (getWatch2 ?state''1) (getM ?state''1)" unfolding setReason_def by simp have "InvariantWatchCharacterization (getF ?state') (getWatch1 ?state') (getWatch2 ?state') (getM ?state')" using InvariantWatchCharacterizationInBackjumpPrefix[of "state"] using assms by (simp add: Let_def) hence "InvariantWatchCharacterization (getF ?state'') (getWatch1 ?state'') (getWatch2 ?state'') (getM ?state'')" unfolding setReason_def by simp have "InvariantConflictFlagCharacterization (getConflictFlag ?state'1) (getF ?state'1) (getM ?state'1)" proof- { fix c::Clause assume "c el (getF state)" have "\ clauseFalse c (elements ?prefix)" proof (cases "c el (butlast (getF state))") case True thus ?thesis using \InvariantNoDecisionsWhenConflict (butlast (getF state)) (getM state) (currentLevel (getM state))\ using \?level < currentLevel (getM state)\ unfolding InvariantNoDecisionsWhenConflict_def by (simp add: formulaFalseIffContainsFalseClause) next case False from \getF state \ []\ have "butlast (getF state) @ [last (getF state)] = getF state" using append_butlast_last_id[of "getF state"] by simp hence "getF state = butlast (getF state) @ [last (getF state)]" by (rule sym) with \c el getF state\ have "c el butlast (getF state) \ c el [last (getF state)]" using set_append[of "butlast (getF state)" "[last (getF state)]"] by auto hence "c = last (getF state)" using \\ c el (butlast (getF state))\ by simp thus ?thesis using \\ clauseFalse (last (getF state)) (elements ?prefix)\ by simp qed } thus ?thesis unfolding InvariantConflictFlagCharacterization_def by (simp add: formulaFalseIffContainsFalseClause) qed hence "InvariantConflictFlagCharacterization (getConflictFlag ?state''1) (getF ?state''1) (getM ?state''1)" unfolding setReason_def by simp have "InvariantQCharacterization (getConflictFlag ?state'2) (removeAll (opposite ?l) (getQ ?state'2)) (getF ?state'2) (getM ?state'2)" using assms using \InvariantConsistent (?prefix @ [(opposite ?l, False)])\ using \InvariantUniq (?prefix @ [(opposite ?l, False)])\ using \InvariantConflictFlagCharacterization (getConflictFlag ?state'1) (getF ?state'1) (getM ?state'1)\ using \InvariantWatchCharacterization (getF ?state'1) (getWatch1 ?state'1) (getWatch2 ?state'1) (getM ?state'1)\ using \InvariantQCharacterization (getConflictFlag ?state'1) (getQ ?state'1) (getF ?state'1) (getM ?state'1)\ using InvariantQCharacterizationAfterAssertLiteral[of "?state'1" "opposite ?l" "False"] by (simp add: Let_def) have "InvariantQCharacterization (getConflictFlag ?state''2) (removeAll (opposite ?l) (getQ ?state''2)) (getF ?state''2) (getM ?state''2)" using assms using \InvariantConsistent (?prefix @ [(opposite ?l, False)])\ using \InvariantUniq (?prefix @ [(opposite ?l, False)])\ using \InvariantConflictFlagCharacterization (getConflictFlag ?state''1) (getF ?state''1) (getM ?state''1)\ using \InvariantWatchCharacterization (getF ?state''1) (getWatch1 ?state''1) (getWatch2 ?state''1) (getM ?state''1)\ using \InvariantQCharacterization (getConflictFlag ?state''1) (getQ ?state''1) (getF ?state''1) (getM ?state''1)\ using InvariantQCharacterizationAfterAssertLiteral[of "?state''1" "opposite ?l" "False"] unfolding setReason_def by (simp add: Let_def) let ?stateB = "applyBackjump state" show ?thesis proof (cases "getBackjumpLevel state > 0") case False let ?state01 = "state\getConflictFlag := False, getM := ?prefix\" have "InvariantWatchesEl (getF ?state01) (getWatch1 ?state01) (getWatch2 ?state01)" using \InvariantWatchesEl (getF state) (getWatch1 state) (getWatch2 state)\ unfolding InvariantWatchesEl_def by auto have "InvariantWatchListsContainOnlyClausesFromF (getWatchList ?state01) (getF ?state01)" using \InvariantWatchListsContainOnlyClausesFromF (getWatchList state) (getF state)\ unfolding InvariantWatchListsContainOnlyClausesFromF_def by auto have "assertLiteral (opposite ?l) False (state \getConflictFlag := False, getQ := [], getM := ?prefix \) = assertLiteral (opposite ?l) False (state \getConflictFlag := False, getM := ?prefix, getQ := [] \)" using arg_cong[of "state \getConflictFlag := False, getQ := [], getM := ?prefix \" "state \getConflictFlag := False, getM := ?prefix, getQ := [] \" "\ x. assertLiteral (opposite ?l) False x"] by simp hence "getConflictFlag ?stateB = getConflictFlag ?state'2" "getF ?stateB = getF ?state'2" "getM ?stateB = getM ?state'2" unfolding applyBackjump_def using AssertLiteralStartQIreleveant[of "?state01" "opposite ?l" "False" "[]" "[opposite ?l]"] using \InvariantWatchesEl (getF ?state01) (getWatch1 ?state01) (getWatch2 ?state01)\ using \InvariantWatchListsContainOnlyClausesFromF (getWatchList ?state01) (getF ?state01)\ using \\ getBackjumpLevel state > 0\ by (auto simp add: Let_def) have "set (getQ ?stateB) = set (removeAll (opposite ?l) (getQ ?state'2))" proof- have "set (getQ ?stateB) = set(getQ ?state'2) - {opposite ?l}" proof- let ?ulSet = "{ ul. (\ uc. uc el (getF ?state'1) \ ?l el uc \ isUnitClause uc ul ((elements (getM ?state'1)) @ [opposite ?l])) }" have "set (getQ ?state'2) = {opposite ?l} \ ?ulSet" using assertLiteralQEffect[of "?state'1" "opposite ?l" "False"] using assms using \InvariantConsistent (?prefix @ [(opposite ?l, False)])\ using \InvariantUniq (?prefix @ [(opposite ?l, False)])\ using \InvariantWatchCharacterization (getF ?state'1) (getWatch1 ?state'1) (getWatch2 ?state'1) (getM ?state'1)\ by (simp add:Let_def) moreover have "set (getQ ?stateB) = ?ulSet" using assertLiteralQEffect[of "?state'" "opposite ?l" "False"] using assms using \InvariantConsistent (?prefix @ [(opposite ?l, False)])\ using \InvariantUniq (?prefix @ [(opposite ?l, False)])\ using \InvariantWatchCharacterization (getF ?state') (getWatch1 ?state') (getWatch2 ?state') (getM ?state')\ using \\ getBackjumpLevel state > 0\ unfolding applyBackjump_def by (simp add:Let_def) moreover have "\ (opposite ?l) \ ?ulSet" using assertedLiteralIsNotUnit[of "?state'" "opposite ?l" "False"] using assms using \InvariantConsistent (?prefix @ [(opposite ?l, False)])\ using \InvariantUniq (?prefix @ [(opposite ?l, False)])\ using \InvariantWatchCharacterization (getF ?state') (getWatch1 ?state') (getWatch2 ?state') (getM ?state')\ using \set (getQ ?stateB) = ?ulSet\ using \\ getBackjumpLevel state > 0\ unfolding applyBackjump_def by (simp add: Let_def) ultimately show ?thesis by simp qed thus ?thesis by simp qed show ?thesis using \InvariantQCharacterization (getConflictFlag ?state'2) (removeAll (opposite ?l) (getQ ?state'2)) (getF ?state'2) (getM ?state'2)\ using \set (getQ ?stateB) = set (removeAll (opposite ?l) (getQ ?state'2))\ using \getConflictFlag ?stateB = getConflictFlag ?state'2\ using \getF ?stateB = getF ?state'2\ using \getM ?stateB = getM ?state'2\ unfolding InvariantQCharacterization_def by (simp add: Let_def) next case True let ?state02 = "setReason (opposite (getCl state)) (length (getF state) - 1) state\getConflictFlag := False, getM := ?prefix\" have "InvariantWatchesEl (getF ?state02) (getWatch1 ?state02) (getWatch2 ?state02)" using \InvariantWatchesEl (getF state) (getWatch1 state) (getWatch2 state)\ unfolding InvariantWatchesEl_def unfolding setReason_def by auto have "InvariantWatchListsContainOnlyClausesFromF (getWatchList ?state02) (getF ?state02)" using \InvariantWatchListsContainOnlyClausesFromF (getWatchList state) (getF state)\ unfolding InvariantWatchListsContainOnlyClausesFromF_def unfolding setReason_def by auto let ?stateTmp' = "assertLiteral (opposite (getCl state)) False (setReason (opposite (getCl state)) (length (getF state) - 1) state \getConflictFlag := False, getM := prefixToLevel (getBackjumpLevel state) (getM state), getQ := []\ )" let ?stateTmp'' = "assertLiteral (opposite (getCl state)) False (setReason (opposite (getCl state)) (length (getF state) - 1) state \getConflictFlag := False, getM := prefixToLevel (getBackjumpLevel state) (getM state), getQ := [opposite (getCl state)]\ )" have "getM ?stateTmp' = getM ?stateTmp''" "getF ?stateTmp' = getF ?stateTmp''" "getSATFlag ?stateTmp' = getSATFlag ?stateTmp''" "getConflictFlag ?stateTmp' = getConflictFlag ?stateTmp''" using AssertLiteralStartQIreleveant[of "?state02" "opposite ?l" "False" "[]" "[opposite ?l]"] using \InvariantWatchesEl (getF ?state02) (getWatch1 ?state02) (getWatch2 ?state02)\ using \InvariantWatchListsContainOnlyClausesFromF (getWatchList ?state02) (getF ?state02)\ by (auto simp add: Let_def) moreover have "?stateB = ?stateTmp'" using \getBackjumpLevel state > 0\ using arg_cong[of "state \ getConflictFlag := False, getQ := [], getM := ?prefix, - getReason := getReason state(opposite ?l \ length (getF state) - 1) + getReason := (getReason state)(opposite ?l \ length (getF state) - 1) \" "state \ - getReason := getReason state(opposite ?l \ length (getF state) - 1), + getReason := (getReason state)(opposite ?l \ length (getF state) - 1), getConflictFlag := False, getM := prefixToLevel (getBackjumpLevel state) (getM state), getQ := [] \" "\ x. assertLiteral (opposite ?l) False x"] unfolding applyBackjump_def unfolding setReason_def by (auto simp add: Let_def) moreover have "?stateTmp'' = ?state''2" unfolding setReason_def - using arg_cong[of "state \getReason := getReason state(opposite ?l \ length (getF state) - 1), + using arg_cong[of "state \getReason := (getReason state)(opposite ?l \ length (getF state) - 1), getConflictFlag := False, getM := ?prefix, getQ := [opposite ?l]\" "state \getConflictFlag := False, getM := prefixToLevel (getBackjumpLevel state) (getM state), - getReason := getReason state(opposite ?l \ length (getF state) - 1), + getReason := (getReason state)(opposite ?l \ length (getF state) - 1), getQ := [opposite ?l]\" "\ x. assertLiteral (opposite ?l) False x"] by simp ultimately have "getConflictFlag ?stateB = getConflictFlag ?state''2" "getF ?stateB = getF ?state''2" "getM ?stateB = getM ?state''2" by auto have "set (getQ ?stateB) = set (removeAll (opposite ?l) (getQ ?state''2))" proof- have "set (getQ ?stateB) = set(getQ ?state''2) - {opposite ?l}" proof- let ?ulSet = "{ ul. (\ uc. uc el (getF ?state''1) \ ?l el uc \ isUnitClause uc ul ((elements (getM ?state''1)) @ [opposite ?l])) }" have "set (getQ ?state''2) = {opposite ?l} \ ?ulSet" using assertLiteralQEffect[of "?state''1" "opposite ?l" "False"] using assms using \InvariantConsistent (?prefix @ [(opposite ?l, False)])\ using \InvariantUniq (?prefix @ [(opposite ?l, False)])\ using \InvariantWatchCharacterization (getF ?state''1) (getWatch1 ?state''1) (getWatch2 ?state''1) (getM ?state''1)\ unfolding setReason_def by (simp add:Let_def) moreover have "set (getQ ?stateB) = ?ulSet" using assertLiteralQEffect[of "?state''" "opposite ?l" "False"] using assms using \InvariantConsistent (?prefix @ [(opposite ?l, False)])\ using \InvariantUniq (?prefix @ [(opposite ?l, False)])\ using \InvariantWatchCharacterization (getF ?state'') (getWatch1 ?state'') (getWatch2 ?state'') (getM ?state'')\ using \getBackjumpLevel state > 0\ unfolding applyBackjump_def unfolding setReason_def by (simp add:Let_def) moreover have "\ (opposite ?l) \ ?ulSet" using assertedLiteralIsNotUnit[of "?state''" "opposite ?l" "False"] using assms using \InvariantConsistent (?prefix @ [(opposite ?l, False)])\ using \InvariantUniq (?prefix @ [(opposite ?l, False)])\ using \InvariantWatchCharacterization (getF ?state'') (getWatch1 ?state'') (getWatch2 ?state'') (getM ?state'')\ using \set (getQ ?stateB) = ?ulSet\ using \getBackjumpLevel state > 0\ unfolding applyBackjump_def unfolding setReason_def by (simp add: Let_def) ultimately show ?thesis by simp qed thus ?thesis by simp qed show ?thesis using \InvariantQCharacterization (getConflictFlag ?state''2) (removeAll (opposite ?l) (getQ ?state''2)) (getF ?state''2) (getM ?state''2)\ using \set (getQ ?stateB) = set (removeAll (opposite ?l) (getQ ?state''2))\ using \getConflictFlag ?stateB = getConflictFlag ?state''2\ using \getF ?stateB = getF ?state''2\ using \getM ?stateB = getM ?state''2\ unfolding InvariantQCharacterization_def by (simp add: Let_def) qed qed lemma InvariantConflictFlagCharacterizationAfterApplyBackjump_1: assumes "InvariantConsistent (getM state)" "InvariantUniq (getM state)" "InvariantWatchListsContainOnlyClausesFromF (getWatchList state) (getF state)" and "InvariantWatchListsUniq (getWatchList state)" and "InvariantWatchListsCharacterization (getWatchList state) (getWatch1 state) (getWatch2 state)" "InvariantWatchesEl (getF state) (getWatch1 state) (getWatch2 state)" and "InvariantWatchesDiffer (getF state) (getWatch1 state) (getWatch2 state)" and "InvariantWatchCharacterization (getF state) (getWatch1 state) (getWatch2 state) (getM state)" and "InvariantUniqC (getC state)" "getC state = [opposite (getCl state)]" "InvariantNoDecisionsWhenConflict (getF state) (getM state) (currentLevel (getM state))" "getConflictFlag state" "InvariantCFalse (getConflictFlag state) (getM state) (getC state)" and "InvariantCEntailed (getConflictFlag state) F0 (getC state)" and "InvariantClCharacterization (getCl state) (getC state) (getM state)" and "InvariantCllCharacterization (getCl state) (getCll state) (getC state) (getM state)" and "InvariantClCurrentLevel (getCl state) (getM state)" "currentLevel (getM state) > 0" "isUIP (opposite (getCl state)) (getC state) (getM state)" shows "let state' = (applyBackjump state) in InvariantConflictFlagCharacterization (getConflictFlag state') (getF state') (getM state')" proof- let ?l = "getCl state" let ?level = "getBackjumpLevel state" let ?prefix = "prefixToLevel ?level (getM state)" let ?state' = "state\ getConflictFlag := False, getQ := [], getM := ?prefix \" let ?state'' = "setReason (opposite ?l) (length (getF state) - 1) ?state'" let ?stateB = "applyBackjump state" have "?level < elementLevel ?l (getM state)" using assms using isMinimalBackjumpLevelGetBackjumpLevel[of "state"] unfolding isMinimalBackjumpLevel_def unfolding isBackjumpLevel_def by (simp add: Let_def) hence "?level < currentLevel (getM state)" using elementLevelLeqCurrentLevel[of "?l" "getM state"] by simp hence "InvariantConflictFlagCharacterization (getConflictFlag ?state') (getF ?state') (getM ?state')" using \InvariantNoDecisionsWhenConflict (getF state) (getM state) (currentLevel (getM state))\ unfolding InvariantNoDecisionsWhenConflict_def unfolding InvariantConflictFlagCharacterization_def by simp moreover have "InvariantConsistent (?prefix @ [(opposite ?l, False)])" using assms using InvariantConsistentAfterApplyBackjump[of "state" "F0"] using assertLiteralEffect unfolding applyBackjump_def unfolding setReason_def by (auto simp add: Let_def split: if_split_asm) ultimately show ?thesis using InvariantConflictFlagCharacterizationAfterAssertLiteral[of "?state'"] using InvariantConflictFlagCharacterizationAfterAssertLiteral[of "?state''"] using InvariantWatchCharacterizationInBackjumpPrefix[of "state"] using assms unfolding applyBackjump_def unfolding setReason_def using assertLiteralEffect by (auto simp add: Let_def) qed lemma InvariantConflictFlagCharacterizationAfterApplyBackjump_2: assumes "InvariantConsistent (getM state)" "InvariantUniq (getM state)" "InvariantWatchListsContainOnlyClausesFromF (getWatchList state) (getF state)" and "InvariantWatchListsUniq (getWatchList state)" and "InvariantWatchListsCharacterization (getWatchList state) (getWatch1 state) (getWatch2 state)" "InvariantWatchesEl (getF state) (getWatch1 state) (getWatch2 state)" and "InvariantWatchesDiffer (getF state) (getWatch1 state) (getWatch2 state)" and "InvariantWatchCharacterization (getF state) (getWatch1 state) (getWatch2 state) (getM state)" and "InvariantUniqC (getC state)" "getC state \ [opposite (getCl state)]" "InvariantNoDecisionsWhenConflict (butlast (getF state)) (getM state) (currentLevel (getM state))" "getF state \ []" "last (getF state) = getC state" "getConflictFlag state" "InvariantCFalse (getConflictFlag state) (getM state) (getC state)" and "InvariantCEntailed (getConflictFlag state) F0 (getC state)" and "InvariantClCharacterization (getCl state) (getC state) (getM state)" and "InvariantCllCharacterization (getCl state) (getCll state) (getC state) (getM state)" and "InvariantClCurrentLevel (getCl state) (getM state)" "currentLevel (getM state) > 0" "isUIP (opposite (getCl state)) (getC state) (getM state)" shows "let state' = (applyBackjump state) in InvariantConflictFlagCharacterization (getConflictFlag state') (getF state') (getM state')" proof- let ?l = "getCl state" let ?level = "getBackjumpLevel state" let ?prefix = "prefixToLevel ?level (getM state)" let ?state' = "state\ getConflictFlag := False, getQ := [], getM := ?prefix \" let ?state'' = "setReason (opposite ?l) (length (getF state) - 1) ?state'" let ?stateB = "applyBackjump state" have "?level < elementLevel ?l (getM state)" using assms using isMinimalBackjumpLevelGetBackjumpLevel[of "state"] unfolding isMinimalBackjumpLevel_def unfolding isBackjumpLevel_def by (simp add: Let_def) hence "?level < currentLevel (getM state)" using elementLevelLeqCurrentLevel[of "?l" "getM state"] by simp hence "InvariantConflictFlagCharacterization (getConflictFlag ?state') (butlast (getF ?state')) (getM ?state')" using \InvariantNoDecisionsWhenConflict (butlast (getF state)) (getM state) (currentLevel (getM state))\ unfolding InvariantNoDecisionsWhenConflict_def unfolding InvariantConflictFlagCharacterization_def by simp moreover have "isBackjumpLevel (getBackjumpLevel state) (opposite (getCl state)) (getC state) (getM state)" using assms using isMinimalBackjumpLevelGetBackjumpLevel[of "state"] unfolding isMinimalBackjumpLevel_def by (simp add: Let_def) hence "isUnitClause (last (getF state)) (opposite ?l) (elements ?prefix)" using isBackjumpLevelEnsuresIsUnitInPrefix[of "getM state" "getC state" "getBackjumpLevel state" "opposite ?l"] using \InvariantUniq (getM state)\ using \InvariantConsistent (getM state)\ using \getConflictFlag state\ using \InvariantCFalse (getConflictFlag state) (getM state) (getC state)\ using \last (getF state) = getC state\ unfolding InvariantUniq_def unfolding InvariantConsistent_def unfolding InvariantCFalse_def by (simp add: Let_def) hence "\ clauseFalse (last (getF state)) (elements ?prefix)" unfolding isUnitClause_def by (auto simp add: clauseFalseIffAllLiteralsAreFalse) moreover from \getF state \ []\ have "butlast (getF state) @ [last (getF state)] = getF state" using append_butlast_last_id[of "getF state"] by simp hence "getF state = butlast (getF state) @ [last (getF state)]" by (rule sym) ultimately have "InvariantConflictFlagCharacterization (getConflictFlag ?state') (getF ?state') (getM ?state')" using set_append[of "butlast (getF state)" "[last (getF state)]"] unfolding InvariantConflictFlagCharacterization_def by (auto simp add: formulaFalseIffContainsFalseClause) moreover have "InvariantConsistent (?prefix @ [(opposite ?l, False)])" using assms using InvariantConsistentAfterApplyBackjump[of "state" "F0"] using assertLiteralEffect unfolding applyBackjump_def unfolding setReason_def by (auto simp add: Let_def split: if_split_asm) ultimately show ?thesis using InvariantConflictFlagCharacterizationAfterAssertLiteral[of "?state'"] using InvariantConflictFlagCharacterizationAfterAssertLiteral[of "?state''"] using InvariantWatchCharacterizationInBackjumpPrefix[of "state"] using assms using assertLiteralEffect unfolding applyBackjump_def unfolding setReason_def by (auto simp add: Let_def) qed lemma InvariantConflictClauseCharacterizationAfterApplyBackjump: assumes "InvariantWatchListsContainOnlyClausesFromF (getWatchList state) (getF state)" and "InvariantWatchListsUniq (getWatchList state)" and "InvariantWatchListsCharacterization (getWatchList state) (getWatch1 state) (getWatch2 state)" and "InvariantWatchesEl (getF state) (getWatch1 state) (getWatch2 state)" shows "let state' = applyBackjump state in InvariantConflictClauseCharacterization (getConflictFlag state') (getConflictClause state') (getF state') (getM state')" proof- let ?l = "getCl state" let ?level = "getBackjumpLevel state" let ?prefix = "prefixToLevel ?level (getM state)" let ?state' = "state\ getConflictFlag := False, getQ := [], getM := ?prefix \" let ?state'' = "if 0 < ?level then setReason (opposite ?l) (length (getF state) - 1) ?state' else ?state'" have "\ getConflictFlag ?state'" by simp hence "InvariantConflictClauseCharacterization (getConflictFlag ?state'') (getConflictClause ?state'') (getF ?state'') (getM ?state'')" unfolding InvariantConflictClauseCharacterization_def unfolding setReason_def by auto moreover have "getF ?state'' = getF state" "getWatchList ?state'' = getWatchList state" "getWatch1 ?state'' = getWatch1 state" "getWatch2 ?state'' = getWatch2 state" unfolding setReason_def by auto ultimately show ?thesis using assms using InvariantConflictClauseCharacterizationAfterAssertLiteral[of "?state''"] unfolding applyBackjump_def by (simp only: Let_def) qed lemma InvariantGetReasonIsReasonAfterApplyBackjump: assumes "InvariantConsistent (getM state)" "InvariantUniq (getM state)" "InvariantWatchesEl (getF state) (getWatch1 state) (getWatch2 state)" "InvariantWatchListsContainOnlyClausesFromF (getWatchList state) (getF state)" and "InvariantWatchListsUniq (getWatchList state)" and "InvariantWatchListsCharacterization (getWatchList state) (getWatch1 state) (getWatch2 state)" and "getConflictFlag state" "InvariantUniqC (getC state)" "InvariantCFalse (getConflictFlag state) (getM state) (getC state)" "InvariantCEntailed (getConflictFlag state) F0 (getC state)" "InvariantClCharacterization (getCl state) (getC state) (getM state)" "InvariantCllCharacterization (getCl state) (getCll state) (getC state) (getM state)" "InvariantClCurrentLevel (getCl state) (getM state)" "isUIP (opposite (getCl state)) (getC state) (getM state)" "0 < currentLevel (getM state)" "InvariantGetReasonIsReason (getReason state) (getF state) (getM state) (set (getQ state))" "getBackjumpLevel state > 0 \ getF state \ [] \ last (getF state) = getC state" shows "let state' = applyBackjump state in InvariantGetReasonIsReason (getReason state') (getF state') (getM state') (set (getQ state')) " proof- let ?l = "getCl state" let ?level = "getBackjumpLevel state" let ?prefix = "prefixToLevel ?level (getM state)" let ?state' = "state\ getConflictFlag := False, getQ := [], getM := ?prefix \" let ?state'' = "if 0 < ?level then setReason (opposite ?l) (length (getF state) - 1) ?state' else ?state'" let ?stateB = "applyBackjump state" have "InvariantGetReasonIsReason (getReason ?state') (getF ?state') (getM ?state') (set (getQ ?state'))" proof- { fix l::Literal assume *: "l el (elements ?prefix) \ \ l el (decisions ?prefix) \ elementLevel l ?prefix > 0" hence "l el (elements (getM state)) \ \ l el (decisions (getM state)) \ elementLevel l (getM state) > 0" using \InvariantUniq (getM state)\ unfolding InvariantUniq_def using isPrefixPrefixToLevel[of "?level" "(getM state)"] using isPrefixElements[of "?prefix" "getM state"] using prefixIsSubset[of "elements ?prefix" "elements (getM state)"] using markedElementsTrailMemPrefixAreMarkedElementsPrefix[of "getM state" "?prefix" "l"] using elementLevelPrefixElement[of "l" "getBackjumpLevel state" "getM state"] by auto with assms obtain reason where "reason < length (getF state)" "isReason (nth (getF state) reason) l (elements (getM state))" "getReason state l = Some reason" unfolding InvariantGetReasonIsReason_def by auto hence "\ reason. getReason state l = Some reason \ reason < length (getF state) \ isReason (nth (getF state) reason) l (elements ?prefix)" using isReasonHoldsInPrefix[of "l" "elements ?prefix" "elements (getM state)" "nth (getF state) reason"] using isPrefixPrefixToLevel[of "?level" "(getM state)"] using isPrefixElements[of "?prefix" "getM state"] using * by auto } thus ?thesis unfolding InvariantGetReasonIsReason_def by auto qed let ?stateM = "?state'' \ getM := getM ?state'' @ [(opposite ?l, False)] \" have **: "getM ?stateM = ?prefix @ [(opposite ?l, False)]" "getF ?stateM = getF state" "getQ ?stateM = []" "getWatchList ?stateM = getWatchList state" "getWatch1 ?stateM = getWatch1 state" "getWatch2 ?stateM = getWatch2 state" unfolding setReason_def by auto have "InvariantGetReasonIsReason (getReason ?stateM) (getF ?stateM) (getM ?stateM) (set (getQ ?stateM))" proof- { fix l::Literal assume *: "l el (elements (getM ?stateM)) \ \ l el (decisions (getM ?stateM)) \ elementLevel l (getM ?stateM) > 0" have "isPrefix ?prefix (getM ?stateM)" unfolding setReason_def unfolding isPrefix_def by auto have "\ reason. getReason ?stateM l = Some reason \ reason < length (getF ?stateM) \ isReason (nth (getF ?stateM) reason) l (elements (getM ?stateM))" proof (cases "l = opposite ?l") case False hence "l el (elements ?prefix)" using * using ** by auto moreover hence "\ l el (decisions ?prefix)" using elementLevelAppend[of "l" "?prefix" "[(opposite ?l, False)]"] using \isPrefix ?prefix (getM ?stateM)\ using markedElementsPrefixAreMarkedElementsTrail[of "?prefix" "getM ?stateM" "l"] using * using ** by auto moreover have "elementLevel l ?prefix = elementLevel l (getM ?stateM)" using \l el (elements ?prefix)\ using * using ** using elementLevelAppend[of "l" "?prefix" "[(opposite ?l, False)]"] by auto hence "elementLevel l ?prefix > 0" using * by simp ultimately obtain reason where "reason < length (getF state)" "isReason (nth (getF state) reason) l (elements ?prefix)" "getReason state l = Some reason" using \InvariantGetReasonIsReason (getReason ?state') (getF ?state') (getM ?state') (set (getQ ?state'))\ unfolding InvariantGetReasonIsReason_def by auto moreover have "getReason ?stateM l = getReason ?state' l" using False unfolding setReason_def by auto ultimately show ?thesis using isReasonAppend[of "nth (getF state) reason" "l" "elements ?prefix" "[opposite ?l]"] using ** by auto next case True show ?thesis proof (cases "?level = 0") case True hence "currentLevel (getM ?stateM) = 0" using currentLevelPrefixToLevel[of "0" "getM state"] using * unfolding currentLevel_def by (simp add: markedElementsAppend) hence "elementLevel l (getM ?stateM) = 0" using \?level = 0\ using elementLevelLeqCurrentLevel[of "l" "getM ?stateM"] by simp with * have False by simp thus ?thesis by simp next case False let ?reason = "length (getF state) - 1" have "getReason ?stateM l = Some ?reason" using \?level \ 0\ using \l = opposite ?l\ unfolding setReason_def by auto moreover have "(nth (getF state) ?reason) = (getC state)" using \?level \ 0\ using \getBackjumpLevel state > 0 \ getF state \ [] \ last (getF state) = getC state\ using last_conv_nth[of "getF state"] by simp hence "isUnitClause (nth (getF state) ?reason) l (elements ?prefix)" using assms using applyBackjumpEffect[of "state" "F0"] using \l = opposite ?l\ by (simp add: Let_def) hence "isReason (nth (getF state) ?reason) l (elements (getM ?stateM))" using ** using isUnitClauseIsReason[of "nth (getF state) ?reason" "l" "elements ?prefix" "[opposite ?l]"] using \l = opposite ?l\ by simp moreover have "?reason < length (getF state)" using \?level \ 0\ using \getBackjumpLevel state > 0 \ getF state \ [] \ last (getF state) = getC state\ by simp ultimately show ?thesis using \?level \ 0\ using \l = opposite ?l\ using ** by auto qed qed } thus ?thesis unfolding InvariantGetReasonIsReason_def unfolding setReason_def by auto qed thus ?thesis using InvariantGetReasonIsReasonAfterNotifyWatches[of "?stateM" "getWatchList ?stateM ?l" "?l" "?prefix" "False" "{}" "[]"] unfolding applyBackjump_def unfolding Let_def unfolding assertLiteral_def unfolding Let_def unfolding notifyWatches_def using ** using assms unfolding InvariantWatchListsCharacterization_def unfolding InvariantWatchListsUniq_def unfolding InvariantWatchListsContainOnlyClausesFromF_def by auto qed lemma InvariantsNoDecisionsWhenConflictNorUnitAfterApplyBackjump_1: assumes "InvariantConsistent (getM state)" "InvariantUniq (getM state)" "InvariantWatchesEl (getF state) (getWatch1 state) (getWatch2 state)" and "InvariantWatchListsContainOnlyClausesFromF (getWatchList state) (getF state)" and "InvariantUniqC (getC state)" "getC state = [opposite (getCl state)]" "InvariantNoDecisionsWhenConflict (getF state) (getM state) (currentLevel (getM state))" "InvariantNoDecisionsWhenUnit (getF state) (getM state) (currentLevel (getM state))" "InvariantCFalse (getConflictFlag state) (getM state) (getC state)" and "InvariantCEntailed (getConflictFlag state) F0 (getC state)" and "InvariantClCharacterization (getCl state) (getC state) (getM state)" and "InvariantCllCharacterization (getCl state) (getCll state) (getC state) (getM state)" and "InvariantClCurrentLevel (getCl state) (getM state)" "getConflictFlag state" "isUIP (opposite (getCl state)) (getC state) (getM state)" "currentLevel (getM state) > 0" shows "let state' = applyBackjump state in InvariantNoDecisionsWhenConflict (getF state') (getM state') (currentLevel (getM state')) \ InvariantNoDecisionsWhenUnit (getF state') (getM state') (currentLevel (getM state'))" proof- let ?l = "getCl state" let ?bClause = "getC state" let ?bLiteral = "opposite ?l" let ?level = "getBackjumpLevel state" let ?prefix = "prefixToLevel ?level (getM state)" let ?state' = "applyBackjump state" have "getM ?state' = ?prefix @ [(?bLiteral, False)]" "getF ?state' = getF state" using assms using applyBackjumpEffect[of "state"] by (auto simp add: Let_def) show ?thesis proof- have "?level < elementLevel ?l (getM state)" using assms using isMinimalBackjumpLevelGetBackjumpLevel[of "state"] unfolding isMinimalBackjumpLevel_def unfolding isBackjumpLevel_def by (simp add: Let_def) hence "?level < currentLevel (getM state)" using elementLevelLeqCurrentLevel[of "?l" "getM state"] by simp have "currentLevel (getM ?state') = currentLevel ?prefix" using \getM ?state' = ?prefix @ [(?bLiteral, False)]\ using markedElementsAppend[of "?prefix" "[(?bLiteral, False)]"] unfolding currentLevel_def by simp hence "currentLevel (getM ?state') \ ?level" using currentLevelPrefixToLevel[of "?level" "getM state"] by simp show ?thesis proof- { fix level assume "level < currentLevel (getM ?state')" hence "level < currentLevel ?prefix" using \currentLevel (getM ?state') = currentLevel ?prefix\ by simp hence "prefixToLevel level (getM (applyBackjump state)) = prefixToLevel level ?prefix" using \getM ?state' = ?prefix @ [(?bLiteral, False)]\ using prefixToLevelAppend[of "level" "?prefix" "[(?bLiteral, False)]"] by simp have "level < ?level" using \level < currentLevel ?prefix\ using \currentLevel (getM ?state') \ ?level\ using \currentLevel (getM ?state') = currentLevel ?prefix\ by simp have "prefixToLevel level (getM ?state') = prefixToLevel level ?prefix" using \getM ?state' = ?prefix @ [(?bLiteral, False)]\ using prefixToLevelAppend[of "level" "?prefix" "[(?bLiteral, False)]"] using \level < currentLevel ?prefix\ by simp hence "\ formulaFalse (getF ?state') (elements (prefixToLevel level (getM ?state')))" (is "?false") using \InvariantNoDecisionsWhenConflict (getF state) (getM state) (currentLevel (getM state))\ unfolding InvariantNoDecisionsWhenConflict_def using \level < ?level\ using \?level < currentLevel (getM state)\ using prefixToLevelPrefixToLevelHigherLevel[of "level" "?level" "getM state", THEN sym] using \getF ?state' = getF state\ using \prefixToLevel level (getM ?state') = prefixToLevel level ?prefix\ using prefixToLevelPrefixToLevelHigherLevel[of "level" "?level" "getM state", THEN sym] by (auto simp add: formulaFalseIffContainsFalseClause) moreover have "\ (\ clause literal. clause el (getF ?state') \ isUnitClause clause literal (elements (prefixToLevel level (getM ?state'))))" (is "?unit") using \InvariantNoDecisionsWhenUnit (getF state) (getM state) (currentLevel (getM state))\ unfolding InvariantNoDecisionsWhenUnit_def using \level < ?level\ using \?level < currentLevel (getM state)\ using \getF ?state' = getF state\ using \prefixToLevel level (getM ?state') = prefixToLevel level ?prefix\ using prefixToLevelPrefixToLevelHigherLevel[of "level" "?level" "getM state", THEN sym] by simp ultimately have "?false \ ?unit" by simp } thus ?thesis unfolding InvariantNoDecisionsWhenConflict_def unfolding InvariantNoDecisionsWhenUnit_def by (auto simp add: Let_def) qed qed qed lemma InvariantsNoDecisionsWhenConflictNorUnitAfterApplyBackjump_2: assumes "InvariantConsistent (getM state)" "InvariantUniq (getM state)" "InvariantWatchesEl (getF state) (getWatch1 state) (getWatch2 state)" and "InvariantWatchListsContainOnlyClausesFromF (getWatchList state) (getF state)" and "InvariantUniqC (getC state)" "getC state \ [opposite (getCl state)]" "InvariantNoDecisionsWhenConflict (butlast (getF state)) (getM state) (currentLevel (getM state))" "InvariantNoDecisionsWhenUnit (butlast (getF state)) (getM state) (currentLevel (getM state))" "getF state \ []" "last (getF state) = getC state" "InvariantNoDecisionsWhenConflict [getC state] (getM state) (getBackjumpLevel state)" "InvariantNoDecisionsWhenUnit [getC state] (getM state) (getBackjumpLevel state)" "getConflictFlag state" "InvariantCFalse (getConflictFlag state) (getM state) (getC state)" and "InvariantCEntailed (getConflictFlag state) F0 (getC state)" and "InvariantClCharacterization (getCl state) (getC state) (getM state)" and "InvariantCllCharacterization (getCl state) (getCll state) (getC state) (getM state)" and "InvariantClCurrentLevel (getCl state) (getM state)" "isUIP (opposite (getCl state)) (getC state) (getM state)" "currentLevel (getM state) > 0" shows "let state' = applyBackjump state in InvariantNoDecisionsWhenConflict (getF state') (getM state') (currentLevel (getM state')) \ InvariantNoDecisionsWhenUnit (getF state') (getM state') (currentLevel (getM state'))" proof- let ?l = "getCl state" let ?bClause = "getC state" let ?bLiteral = "opposite ?l" let ?level = "getBackjumpLevel state" let ?prefix = "prefixToLevel ?level (getM state)" let ?state' = "applyBackjump state" have "getM ?state' = ?prefix @ [(?bLiteral, False)]" "getF ?state' = getF state" using assms using applyBackjumpEffect[of "state"] by (auto simp add: Let_def) show ?thesis proof- have "?level < elementLevel ?l (getM state)" using assms using isMinimalBackjumpLevelGetBackjumpLevel[of "state"] unfolding isMinimalBackjumpLevel_def unfolding isBackjumpLevel_def by (simp add: Let_def) hence "?level < currentLevel (getM state)" using elementLevelLeqCurrentLevel[of "?l" "getM state"] by simp have "currentLevel (getM ?state') = currentLevel ?prefix" using \getM ?state' = ?prefix @ [(?bLiteral, False)]\ using markedElementsAppend[of "?prefix" "[(?bLiteral, False)]"] unfolding currentLevel_def by simp hence "currentLevel (getM ?state') \ ?level" using currentLevelPrefixToLevel[of "?level" "getM state"] by simp show ?thesis proof- { fix level assume "level < currentLevel (getM ?state')" hence "level < currentLevel ?prefix" using \currentLevel (getM ?state') = currentLevel ?prefix\ by simp hence "prefixToLevel level (getM (applyBackjump state)) = prefixToLevel level ?prefix" using \getM ?state' = ?prefix @ [(?bLiteral, False)]\ using prefixToLevelAppend[of "level" "?prefix" "[(?bLiteral, False)]"] by simp have "level < ?level" using \level < currentLevel ?prefix\ using \currentLevel (getM ?state') \ ?level\ using \currentLevel (getM ?state') = currentLevel ?prefix\ by simp have "prefixToLevel level (getM ?state') = prefixToLevel level ?prefix" using \getM ?state' = ?prefix @ [(?bLiteral, False)]\ using prefixToLevelAppend[of "level" "?prefix" "[(?bLiteral, False)]"] using \level < currentLevel ?prefix\ by simp have "\ formulaFalse (butlast (getF ?state')) (elements (prefixToLevel level (getM ?state')))" using \getF ?state' = getF state\ using \InvariantNoDecisionsWhenConflict (butlast (getF state)) (getM state) (currentLevel (getM state))\ using \level < ?level\ using \?level < currentLevel (getM state)\ using \prefixToLevel level (getM ?state') = prefixToLevel level ?prefix\ using prefixToLevelPrefixToLevelHigherLevel[of "level" "?level" "getM state", THEN sym] unfolding InvariantNoDecisionsWhenConflict_def by (auto simp add: formulaFalseIffContainsFalseClause) moreover have "\ clauseFalse (last (getF ?state')) (elements (prefixToLevel level (getM ?state')))" using \getF ?state' = getF state\ using \InvariantNoDecisionsWhenConflict [getC state] (getM state) (getBackjumpLevel state)\ using \last (getF state) = getC state\ using \level < ?level\ using \prefixToLevel level (getM ?state') = prefixToLevel level ?prefix\ using prefixToLevelPrefixToLevelHigherLevel[of "level" "?level" "getM state", THEN sym] unfolding InvariantNoDecisionsWhenConflict_def by (simp add: formulaFalseIffContainsFalseClause) moreover from \getF state \ []\ have "butlast (getF state) @ [last (getF state)] = getF state" using append_butlast_last_id[of "getF state"] by simp hence "getF state = butlast (getF state) @ [last (getF state)]" by (rule sym) ultimately have "\ formulaFalse (getF ?state') (elements (prefixToLevel level (getM ?state')))" (is "?false") using \getF ?state' = getF state\ using set_append[of "butlast (getF state)" "[last (getF state)]"] by (auto simp add: formulaFalseIffContainsFalseClause) have "\ (\ clause literal. clause el (butlast (getF ?state')) \ isUnitClause clause literal (elements (prefixToLevel level (getM ?state'))))" using \InvariantNoDecisionsWhenUnit (butlast (getF state)) (getM state) (currentLevel (getM state))\ unfolding InvariantNoDecisionsWhenUnit_def using \level < ?level\ using \?level < currentLevel (getM state)\ using \getF ?state' = getF state\ using \prefixToLevel level (getM ?state') = prefixToLevel level ?prefix\ using prefixToLevelPrefixToLevelHigherLevel[of "level" "?level" "getM state", THEN sym] by simp moreover have "\ (\ l. isUnitClause (last (getF ?state')) l (elements (prefixToLevel level (getM ?state'))))" using \getF ?state' = getF state\ using \InvariantNoDecisionsWhenUnit [getC state] (getM state) (getBackjumpLevel state)\ using \last (getF state) = getC state\ using \level < ?level\ using \prefixToLevel level (getM ?state') = prefixToLevel level ?prefix\ using prefixToLevelPrefixToLevelHigherLevel[of "level" "?level" "getM state", THEN sym] unfolding InvariantNoDecisionsWhenUnit_def by simp moreover from \getF state \ []\ have "butlast (getF state) @ [last (getF state)] = getF state" using append_butlast_last_id[of "getF state"] by simp hence "getF state = butlast (getF state) @ [last (getF state)]" by (rule sym) ultimately have "\ (\ clause literal. clause el (getF ?state') \ isUnitClause clause literal (elements (prefixToLevel level (getM ?state'))))" (is ?unit) using \getF ?state' = getF state\ using set_append[of "butlast (getF state)" "[last (getF state)]"] by auto have "?false \ ?unit" using \?false\ \?unit\ by simp } thus ?thesis unfolding InvariantNoDecisionsWhenConflict_def unfolding InvariantNoDecisionsWhenUnit_def by (auto simp add: Let_def) qed qed qed lemma InvariantEquivalentZLAfterApplyBackjump: assumes "InvariantConsistent (getM state)" "InvariantUniq (getM state)" "InvariantWatchesEl (getF state) (getWatch1 state) (getWatch2 state)" and "InvariantWatchListsContainOnlyClausesFromF (getWatchList state) (getF state)" and "getConflictFlag state" "InvariantUniqC (getC state)" "InvariantCFalse (getConflictFlag state) (getM state) (getC state)" and "InvariantCEntailed (getConflictFlag state) F0 (getC state)" and "InvariantClCharacterization (getCl state) (getC state) (getM state)" and "InvariantCllCharacterization (getCl state) (getCll state) (getC state) (getM state)" and "InvariantClCurrentLevel (getCl state) (getM state)" "InvariantEquivalentZL (getF state) (getM state) F0" "isUIP (opposite (getCl state)) (getC state) (getM state)" "currentLevel (getM state) > 0" shows "let state' = applyBackjump state in InvariantEquivalentZL (getF state') (getM state') F0 " proof- let ?l = "getCl state" let ?bClause = "getC state" let ?bLiteral = "opposite ?l" let ?level = "getBackjumpLevel state" let ?prefix = "prefixToLevel ?level (getM state)" let ?state' = "applyBackjump state" have "formulaEntailsClause F0 ?bClause" "isUnitClause ?bClause ?bLiteral (elements ?prefix)" "getM ?state' = ?prefix @ [(?bLiteral, False)] " "getF ?state' = getF state" using assms using applyBackjumpEffect[of "state" "F0"] by (auto simp add: Let_def) note * = this show ?thesis proof (cases "?level = 0") case False have "?level < elementLevel ?l (getM state)" using assms using isMinimalBackjumpLevelGetBackjumpLevel[of "state"] unfolding isMinimalBackjumpLevel_def unfolding isBackjumpLevel_def by (simp add: Let_def) hence "?level < currentLevel (getM state)" using elementLevelLeqCurrentLevel[of "?l" "getM state"] by simp hence "prefixToLevel 0 (getM ?state') = prefixToLevel 0 ?prefix" using * using prefixToLevelAppend[of "0" "?prefix" "[(?bLiteral, False)]"] using \?level \ 0\ using currentLevelPrefixToLevelEq[of "?level" "getM state"] by simp hence "prefixToLevel 0 (getM ?state') = prefixToLevel 0 (getM state)" using \?level \ 0\ using prefixToLevelPrefixToLevelHigherLevel[of "0" "?level" "getM state"] by simp thus ?thesis using * using \InvariantEquivalentZL (getF state) (getM state) F0\ unfolding InvariantEquivalentZL_def by (simp add: Let_def) next case True hence "prefixToLevel 0 (getM ?state') = ?prefix @ [(?bLiteral, False)]" using * using prefixToLevelAppend[of "0" "?prefix" "[(?bLiteral, False)]"] using currentLevelPrefixToLevel[of "0" "getM state"] by simp let ?FM = "getF state @ val2form (elements (prefixToLevel 0 (getM state)))" let ?FM' = "getF ?state' @ val2form (elements (prefixToLevel 0 (getM ?state')))" have "formulaEntailsValuation F0 (elements ?prefix)" using \?level = 0\ using val2formIsEntailed[of "getF state" "elements (prefixToLevel 0 (getM state))" "[]"] using \InvariantEquivalentZL (getF state) (getM state) F0\ unfolding formulaEntailsValuation_def unfolding InvariantEquivalentZL_def unfolding equivalentFormulae_def unfolding formulaEntailsLiteral_def by auto have "formulaEntailsLiteral (F0 @ val2form (elements ?prefix)) ?bLiteral" using * using unitLiteralIsEntailed [of "?bClause" "?bLiteral" "elements ?prefix" "F0"] by simp have "formulaEntailsLiteral F0 ?bLiteral" proof- { fix valuation::Valuation assume "model valuation F0" hence "formulaTrue (val2form (elements ?prefix)) valuation" using \formulaEntailsValuation F0 (elements ?prefix)\ using val2formFormulaTrue[of "elements ?prefix" "valuation"] unfolding formulaEntailsValuation_def unfolding formulaEntailsLiteral_def by simp hence "formulaTrue (F0 @ (val2form (elements ?prefix))) valuation" using \model valuation F0\ by (simp add: formulaTrueAppend) hence "literalTrue ?bLiteral valuation" using \model valuation F0\ using \formulaEntailsLiteral (F0 @ val2form (elements ?prefix)) ?bLiteral\ unfolding formulaEntailsLiteral_def by auto } thus ?thesis unfolding formulaEntailsLiteral_def by simp qed hence "formulaEntailsClause F0 [?bLiteral]" unfolding formulaEntailsLiteral_def unfolding formulaEntailsClause_def by (auto simp add: clauseTrueIffContainsTrueLiteral) hence "formulaEntailsClause ?FM [?bLiteral]" using \InvariantEquivalentZL (getF state) (getM state) F0\ unfolding InvariantEquivalentZL_def unfolding equivalentFormulae_def unfolding formulaEntailsClause_def by auto have "?FM' = ?FM @ [[?bLiteral]]" using * using \?level = 0\ using \prefixToLevel 0 (getM ?state') = ?prefix @ [(?bLiteral, False)]\ by (auto simp add: val2formAppend) show ?thesis using \InvariantEquivalentZL (getF state) (getM state) F0\ using \?FM' = ?FM @ [[?bLiteral]]\ using \formulaEntailsClause ?FM [?bLiteral]\ unfolding InvariantEquivalentZL_def using extendEquivalentFormulaWithEntailedClause[of "F0" "?FM" "[?bLiteral]"] by (simp add: equivalentFormulaeSymmetry) qed qed lemma InvariantsVarsAfterApplyBackjump: assumes "InvariantConsistent (getM state)" "InvariantUniq (getM state)" "InvariantWatchesEl (getF state) (getWatch1 state) (getWatch2 state)" and "InvariantWatchListsContainOnlyClausesFromF (getWatchList state) (getF state)" and "InvariantWatchListsUniq (getWatchList state)" "InvariantWatchListsCharacterization (getWatchList state) (getWatch1 state) (getWatch2 state)" "InvariantWatchesEl (getF state) (getWatch1 state) (getWatch2 state)" "InvariantWatchesDiffer (getF state) (getWatch1 state) (getWatch2 state)" "InvariantWatchCharacterization (getF state) (getWatch1 state) (getWatch2 state) (getM state)" and "getConflictFlag state" "InvariantCFalse (getConflictFlag state) (getM state) (getC state)" and "InvariantUniqC (getC state)" and "InvariantCEntailed (getConflictFlag state) F0' (getC state)" and "InvariantClCharacterization (getCl state) (getC state) (getM state)" and "InvariantCllCharacterization (getCl state) (getCll state) (getC state) (getM state)" and "InvariantClCurrentLevel (getCl state) (getM state)" "InvariantEquivalentZL (getF state) (getM state) F0'" "isUIP (opposite (getCl state)) (getC state) (getM state)" "currentLevel (getM state) > 0" "vars F0' \ vars F0" "InvariantVarsM (getM state) F0 Vbl" "InvariantVarsF (getF state) F0 Vbl" "InvariantVarsQ (getQ state) F0 Vbl" shows "let state' = applyBackjump state in InvariantVarsM (getM state') F0 Vbl \ InvariantVarsF (getF state') F0 Vbl \ InvariantVarsQ (getQ state') F0 Vbl " proof- let ?l = "getCl state" let ?bClause = "getC state" let ?bLiteral = "opposite ?l" let ?level = "getBackjumpLevel state" let ?prefix = "prefixToLevel ?level (getM state)" let ?state' = "state\ getConflictFlag := False, getQ := [], getM := ?prefix \" let ?state'' = "setReason (opposite (getCl state)) (length (getF state) - 1) ?state'" let ?stateB = "applyBackjump state" have "formulaEntailsClause F0' ?bClause" "isUnitClause ?bClause ?bLiteral (elements ?prefix)" "getM ?stateB = ?prefix @ [(?bLiteral, False)] " "getF ?stateB = getF state" using assms using applyBackjumpEffect[of "state" "F0'"] by (auto simp add: Let_def) note * = this have "var ?bLiteral \ vars F0 \ Vbl" proof- have "vars (getC state) \ vars (elements (getM state))" using \getConflictFlag state\ using \InvariantCFalse (getConflictFlag state) (getM state) (getC state)\ using valuationContainsItsFalseClausesVariables[of "getC state" "elements (getM state)"] unfolding InvariantCFalse_def by simp moreover have "?bLiteral el (getC state)" using \InvariantClCharacterization (getCl state) (getC state) (getM state)\ unfolding InvariantClCharacterization_def unfolding isLastAssertedLiteral_def using literalElListIffOppositeLiteralElOppositeLiteralList[of "?bLiteral" "getC state"] by simp ultimately show ?thesis using \InvariantVarsM (getM state) F0 Vbl\ using \vars F0' \ vars F0\ unfolding InvariantVarsM_def using clauseContainsItsLiteralsVariable[of "?bLiteral" "getC state"] by auto qed hence "InvariantVarsM (getM ?stateB) F0 Vbl" using \InvariantVarsM (getM state) F0 Vbl\ using InvariantVarsMAfterBackjump[of "getM state" "F0" "Vbl" "?prefix" "?bLiteral" "getM ?stateB"] using * by (simp add: isPrefixPrefixToLevel) moreover have "InvariantConsistent (prefixToLevel (getBackjumpLevel state) (getM state) @ [(opposite (getCl state), False)])" "InvariantUniq (prefixToLevel (getBackjumpLevel state) (getM state) @ [(opposite (getCl state), False)])" "InvariantWatchCharacterization (getF state) (getWatch1 state) (getWatch2 state) (prefixToLevel (getBackjumpLevel state) (getM state))" using assms using InvariantConsistentAfterApplyBackjump[of "state" "F0'"] using InvariantUniqAfterApplyBackjump[of "state" "F0'"] using * using InvariantWatchCharacterizationInBackjumpPrefix[of "state"] by (auto simp add: Let_def) hence "InvariantVarsQ (getQ ?stateB) F0 Vbl" using \InvariantVarsF (getF state) F0 Vbl\ using \InvariantWatchListsContainOnlyClausesFromF (getWatchList state) (getF state)\ using \InvariantWatchListsUniq (getWatchList state)\ using \InvariantWatchListsCharacterization (getWatchList state) (getWatch1 state) (getWatch2 state)\ using \InvariantWatchesEl (getF state) (getWatch1 state) (getWatch2 state)\ using \InvariantWatchesDiffer (getF state) (getWatch1 state) (getWatch2 state)\ using InvariantVarsQAfterAssertLiteral[of "if ?level > 0 then ?state'' else ?state'" "?bLiteral" "False" "F0" "Vbl"] unfolding applyBackjump_def unfolding InvariantVarsQ_def unfolding setReason_def by (auto simp add: Let_def) moreover have "InvariantVarsF (getF ?stateB) F0 Vbl" using assms using assertLiteralEffect[of "if ?level > 0 then ?state'' else ?state'" "?bLiteral" "False"] using \InvariantVarsF (getF state) F0 Vbl\ unfolding applyBackjump_def unfolding setReason_def by (simp add: Let_def) ultimately show ?thesis by (simp add: Let_def) qed end diff --git a/thys/SPARCv8/SparcModel_MMU/Sparc_Properties.thy b/thys/SPARCv8/SparcModel_MMU/Sparc_Properties.thy --- a/thys/SPARCv8/SparcModel_MMU/Sparc_Properties.thy +++ b/thys/SPARCv8/SparcModel_MMU/Sparc_Properties.thy @@ -1,8589 +1,8589 @@ (* * Copyright 2016, NTU * * This software may be distributed and modified according to the terms of * the BSD 2-Clause license. Note that NO WARRANTY is provided. * See "LICENSE_BSD2.txt" for details. * * Author: Zhe Hou. *) theory Sparc_Properties imports Main Sparc_Execution begin (*********************************************************************) section\Single step theorem\ (*********************************************************************) text \The following shows that, if the pre-state satisfies certain conditions called \good_context\, there must be a defined post-state after a single step execution.\ method save_restore_proof = ((simp add: save_restore_instr_def), (simp add: Let_def simpler_gets_def bind_def h1_def h2_def), (simp add: case_prod_unfold), (simp add: raise_trap_def simpler_modify_def), (simp add: simpler_gets_def bind_def h1_def h2_def), (simp add: save_retore_sub1_def), (simp add: write_cpu_def simpler_modify_def), (simp add: write_reg_def simpler_modify_def), (simp add: get_curr_win_def), (simp add: simpler_gets_def bind_def h1_def h2_def)) method select_trap_proof0 = ((simp add: select_trap_def exec_gets return_def), (simp add: DetMonad.bind_def h1_def h2_def simpler_modify_def), (simp add: write_cpu_tt_def write_cpu_def), (simp add: DetMonad.bind_def h1_def h2_def simpler_modify_def), (simp add: return_def simpler_gets_def)) method select_trap_proof1 = ((simp add: select_trap_def exec_gets return_def), (simp add: DetMonad.bind_def h1_def h2_def simpler_modify_def), (simp add: write_cpu_tt_def write_cpu_def), (simp add: DetMonad.bind_def h1_def h2_def simpler_modify_def), (simp add: return_def simpler_gets_def), (simp add: emp_trap_set_def err_mode_val_def cpu_reg_mod_def)) method dispatch_instr_proof1 = ((simp add: dispatch_instruction_def), (simp add: simpler_gets_def bind_def h1_def h2_def), (simp add: Let_def)) method exe_proof_to_decode = ((simp add: execute_instruction_def), (simp add: exec_gets bind_def h1_def h2_def Let_def return_def), clarsimp, (simp add: simpler_gets_def bind_def h1_def h2_def Let_def simpler_modify_def), (simp add: return_def)) method exe_proof_dispatch_rett = ((simp add: dispatch_instruction_def), (simp add: simpler_gets_def bind_def h1_def h2_def Let_def), (simp add: rett_instr_def), (simp add: simpler_gets_def bind_def h1_def h2_def Let_def)) lemma write_cpu_result: "snd (write_cpu w r s) = False" by (simp add: write_cpu_def simpler_modify_def) lemma set_annul_result: "snd (set_annul b s) = False" by (simp add: set_annul_def simpler_modify_def) lemma raise_trap_result : "snd (raise_trap t s) = False" by (simp add: raise_trap_def simpler_modify_def) context includes bit_operations_syntax begin lemma rett_instr_result: "(fst i) = ctrl_type RETT \ (get_ET (cpu_reg_val PSR s) \ 1 \ (((get_S (cpu_reg_val PSR s)))::word1) \ 0 \ (get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR s))) + 1) mod NWINDOWS)) (cpu_reg_val WIM s)) = 0 \ ((AND) (get_addr (snd i) s) (0b00000000000000000000000000000011::word32)) = 0) \ snd (rett_instr i s) = False" apply (simp add: rett_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: write_cpu_def simpler_modify_def) apply (auto simp add: Let_def return_def) done lemma call_instr_result: "(fst i) = call_type CALL \ snd (call_instr i s) = False" apply (simp add: call_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def case_prod_unfold) apply (simp add: write_cpu_def write_reg_def) apply (simp add: get_curr_win_def get_CWP_def) by (simp add: simpler_modify_def simpler_gets_def) lemma branch_instr_result: "(fst i) \ {bicc_type BE,bicc_type BNE,bicc_type BGU, bicc_type BLE,bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA,bicc_type BN} \ snd (branch_instr i s) = False" proof (cases "eval_icc (fst i) (get_icc_N ((cpu_reg s) PSR)) (get_icc_Z ((cpu_reg s) PSR)) (get_icc_V ((cpu_reg s) PSR)) (get_icc_C ((cpu_reg s) PSR)) = 1") case True then have f1: "eval_icc (fst i) (get_icc_N ((cpu_reg s) PSR)) (get_icc_Z ((cpu_reg s) PSR)) (get_icc_V ((cpu_reg s) PSR)) (get_icc_C ((cpu_reg s) PSR)) = 1" by auto then show ?thesis proof (cases "(fst i) = bicc_type BA \ get_operand_flag ((snd i)!0) = 1") case True then show ?thesis using f1 apply (simp add: branch_instr_def) apply (simp add: Let_def simpler_gets_def bind_def h1_def h2_def) apply (simp add: set_annul_def case_prod_unfold) apply (simp add: write_cpu_def simpler_modify_def) by (simp add: return_def) next case False then have f2: "\ (fst i = bicc_type BA \ get_operand_flag (snd i ! 0) = 1)" by auto then show ?thesis using f1 apply (simp add: branch_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: write_cpu_def simpler_modify_def) apply (simp add: branch_instr_sub1_def) apply (simp add: Let_def) apply auto apply (simp add: write_cpu_def simpler_modify_def) by (simp add: write_cpu_def simpler_modify_def) qed next case False then show ?thesis apply (simp add: branch_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: write_cpu_def simpler_modify_def) apply (simp add: branch_instr_sub1_def) apply (simp add: Let_def) apply auto apply (simp add: Let_def bind_def h1_def h2_def) apply (simp add: write_cpu_def simpler_modify_def) apply (simp add: cpu_reg_mod_def set_annul_def simpler_modify_def) by (simp add: write_cpu_def simpler_modify_def) qed lemma nop_instr_result: "(fst i) = nop_type NOP \ snd (nop_instr i s) = False" apply (simp add: nop_instr_def) by (simp add: returnOk_def return_def) lemma sethi_instr_result: "(fst i) = sethi_type SETHI \ snd (sethi_instr i s) = False" apply (simp add: sethi_instr_def) apply (simp add: Let_def) apply (simp add: get_curr_win_def get_CWP_def cpu_reg_val_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: write_reg_def simpler_modify_def) by (simp add: return_def) lemma jmpl_instr_result: "(fst i) = ctrl_type JMPL \ snd (jmpl_instr i s) = False" apply (simp add: jmpl_instr_def) apply (simp add: get_curr_win_def get_CWP_def cpu_reg_val_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: write_reg_def simpler_modify_def) apply (simp add: write_cpu_def simpler_modify_def) by (simp add: raise_trap_def simpler_modify_def) lemma save_restore_instr_result: "(fst i) \ {ctrl_type SAVE,ctrl_type RESTORE} \ snd (save_restore_instr i s) = False" proof (cases "(fst i) = ctrl_type SAVE") case True then show ?thesis by save_restore_proof next case False then show ?thesis by save_restore_proof qed lemma flush_instr_result: "(fst i) = load_store_type FLUSH \ snd (flush_instr i s) = False" apply (simp add: flush_instr_def) by (simp add: simpler_gets_def bind_def h1_def h2_def simpler_modify_def) lemma read_state_reg_instr_result: "(fst i) \ {sreg_type RDY,sreg_type RDPSR, sreg_type RDWIM,sreg_type RDTBR} \ snd (read_state_reg_instr i s) = False" apply (simp add: read_state_reg_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: simpler_gets_def bind_def) apply (simp add: write_reg_def simpler_modify_def) apply (simp add: raise_trap_def simpler_modify_def return_def) apply (simp add: bind_def h1_def h2_def) by (simp add: get_curr_win_def simpler_gets_def) lemma write_state_reg_instr_result: "(fst i) \ {sreg_type WRY,sreg_type WRPSR, sreg_type WRWIM,sreg_type WRTBR} \ snd (write_state_reg_instr i s) = False" apply (simp add: write_state_reg_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: simpler_modify_def) apply (simp add: raise_trap_def simpler_modify_def return_def) apply (simp add: bind_def h1_def h2_def) apply (simp add: simpler_gets_def) apply (simp add: write_cpu_def simpler_modify_def) by (simp add: get_curr_win_def simpler_gets_def) lemma logical_instr_result: "(fst i) \ {logic_type ANDs,logic_type ANDcc, logic_type ANDN,logic_type ANDNcc,logic_type ORs,logic_type ORcc, logic_type ORN,logic_type XORs,logic_type XNOR} \ snd (logical_instr i s) = False" apply (simp add: logical_instr_def) apply (simp add: Let_def simpler_gets_def) apply (simp add: write_reg_def simpler_modify_def) apply (simp add: bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: logical_instr_sub1_def) apply (simp add: return_def) apply (simp add: write_cpu_def simpler_modify_def) apply (simp add: bind_def h1_def h2_def) apply (simp add: case_prod_unfold) apply (simp add: simpler_gets_def) by (simp add: get_curr_win_def simpler_gets_def) lemma shift_instr_result: "(fst i) \ {shift_type SLL,shift_type SRL,shift_type SRA} \ snd (shift_instr i s) = False" apply (simp add: shift_instr_def) apply (simp add: Let_def) apply (simp add: get_curr_win_def simpler_gets_def bind_def h1_def h2_def) apply (simp add: return_def) apply (simp add: bind_def h1_def h2_def) by (simp add: write_reg_def simpler_modify_def) method add_sub_instr_proof = ((simp add: Let_def), auto, (simp add: write_reg_def simpler_modify_def), (simp add: simpler_gets_def bind_def), (simp add: get_curr_win_def simpler_gets_def), (simp add: write_reg_def write_cpu_def simpler_modify_def), (simp add: bind_def), (simp add: case_prod_unfold), (simp add: simpler_gets_def), (simp add: get_curr_win_def simpler_gets_def), (simp add: write_reg_def simpler_modify_def), (simp add: simpler_gets_def bind_def), (simp add: get_curr_win_def simpler_gets_def)) lemma add_instr_result: "(fst i) \ {arith_type ADD,arith_type ADDcc,arith_type ADDX} \ snd (add_instr i s) = False" apply (simp add: add_instr_def) apply (simp add: Let_def) apply auto apply (simp add: add_instr_sub1_def) apply (simp add: write_reg_def simpler_modify_def) apply (simp add: bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: simpler_gets_def) apply (simp add: get_curr_win_def simpler_gets_def) apply (simp add: add_instr_sub1_def) apply (simp add: write_reg_def simpler_modify_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: get_curr_win_def simpler_gets_def) apply (simp add: write_cpu_def simpler_modify_def) apply (simp add: add_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: get_curr_win_def simpler_gets_def) by (simp add: write_reg_def simpler_modify_def) lemma sub_instr_result: "(fst i) \ {arith_type SUB,arith_type SUBcc, arith_type SUBX} \ snd (sub_instr i s) = False" apply (simp add: sub_instr_def) apply (simp add: Let_def) apply auto apply (simp add: sub_instr_sub1_def) apply (simp add: write_reg_def simpler_modify_def) apply (simp add: bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: simpler_gets_def) apply (simp add: get_curr_win_def simpler_gets_def) apply (simp add: sub_instr_sub1_def) apply (simp add: write_reg_def simpler_modify_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: get_curr_win_def simpler_gets_def) apply (simp add: write_cpu_def simpler_modify_def) apply (simp add: sub_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: get_curr_win_def simpler_gets_def) by (simp add: write_reg_def simpler_modify_def) lemma mul_instr_result: "(fst i) \ {arith_type UMUL,arith_type SMUL, arith_type SMULcc} \ snd (mul_instr i s) = False" apply (simp add: mul_instr_def) apply (simp add: Let_def) apply auto apply (simp add: mul_instr_sub1_def) apply (simp add: write_reg_def simpler_modify_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: get_curr_win_def simpler_gets_def) apply (simp add: write_reg_def write_cpu_def simpler_modify_def) apply (simp add: mul_instr_sub1_def) apply (simp add: simpler_gets_def) apply (simp add: write_cpu_def write_reg_def simpler_modify_def) apply (simp add: bind_def h1_def h2_def Let_def) apply (simp add: get_curr_win_def simpler_gets_def) apply (simp add: mul_instr_sub1_def) apply (simp add: simpler_gets_def) apply (simp add: write_cpu_def write_reg_def simpler_modify_def) apply (simp add: bind_def h1_def h2_def) by (simp add: get_curr_win_def simpler_gets_def) lemma div_write_new_val_result: "snd (div_write_new_val i result temp_V s) = False" apply (simp add: div_write_new_val_def) apply (simp add: return_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) by (simp add: write_cpu_def simpler_modify_def) lemma div_result: "snd (div_comp instr rs1 rd operand2 s) = False" apply (simp add: div_comp_def) apply (simp add: simpler_gets_def) apply (simp add: bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: write_reg_def simpler_modify_def) apply (simp add: get_curr_win_def simpler_gets_def) by (simp add: div_write_new_val_result) lemma div_instr_result: "(fst i) \ {arith_type UDIV,arith_type UDIVcc, arith_type SDIV} \ snd (div_instr i s) = False" apply (simp add: div_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: raise_trap_def simpler_modify_def) apply (simp add: return_def bind_def) by (simp add: div_result) lemma load_sub2_result: "snd (load_sub2 address asi rd curr_win word0 s) = False" apply (simp add: load_sub2_def) apply (simp add: write_reg_def simpler_modify_def) apply (simp add: bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: raise_trap_def simpler_modify_def) apply (simp add: bind_def h1_def h2_def) apply (simp add: write_reg_def simpler_modify_def) by (simp add: simpler_gets_def) lemma load_sub3_result: "snd (load_sub3 instr curr_win rd asi address s) = False" apply (simp add: load_sub3_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: case_prod_unfold) apply (simp add: simpler_modify_def bind_def h1_def h2_def Let_def) apply (simp add: write_reg_def simpler_modify_def) apply (simp add: load_sub2_result) by (simp add: raise_trap_def simpler_modify_def) lemma load_sub1_result: "snd (load_sub1 i rd s_val s) = False" apply (simp add: load_sub1_def) apply (simp add: bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: raise_trap_def simpler_modify_def) apply (simp add: get_curr_win_def simpler_gets_def) by (simp add: load_sub3_result) lemma load_instr_result: "(fst i) \ {load_store_type LDSB,load_store_type LDUB, load_store_type LDUBA,load_store_type LDUH,load_store_type LD, load_store_type LDA,load_store_type LDD} \ snd (load_instr i s) = False" apply (simp add: load_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: raise_trap_def simpler_modify_def) apply (simp add: return_def) by (simp add: load_sub1_result) lemma store_sub2_result: "snd (store_sub2 instr curr_win rd asi address s) = False" apply (simp add: store_sub2_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: raise_trap_def simpler_modify_def) apply (simp add: return_def) apply (simp add: raise_trap_def simpler_modify_def) by (simp add: bind_def h1_def h2_def) lemma store_sub1_result: "snd (store_sub1 instr rd s_val s) = False" apply (simp add: store_sub1_def) apply (simp add: bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: raise_trap_def simpler_modify_def) apply (simp add: get_curr_win_def) apply (simp add: simpler_gets_def) by (simp add: store_sub2_result) lemma store_instr_result: "(fst i) \ {load_store_type STB,load_store_type STH, load_store_type ST,load_store_type STA,load_store_type STD} \ snd (store_instr i s) = False" apply (simp add: store_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: raise_trap_def simpler_modify_def) apply (simp add: return_def) by (simp add: store_sub1_result) lemma supported_instr_set: "supported_instruction i = True \ i \ {load_store_type LDSB,load_store_type LDUB,load_store_type LDUBA, load_store_type LDUH,load_store_type LD,load_store_type LDA, load_store_type LDD, load_store_type STB,load_store_type STH,load_store_type ST, load_store_type STA,load_store_type STD, sethi_type SETHI, nop_type NOP, logic_type ANDs,logic_type ANDcc,logic_type ANDN,logic_type ANDNcc, logic_type ORs,logic_type ORcc,logic_type ORN,logic_type XORs, logic_type XNOR, shift_type SLL,shift_type SRL,shift_type SRA, arith_type ADD,arith_type ADDcc,arith_type ADDX, arith_type SUB,arith_type SUBcc,arith_type SUBX, arith_type UMUL,arith_type SMUL,arith_type SMULcc, arith_type UDIV,arith_type UDIVcc,arith_type SDIV, ctrl_type SAVE,ctrl_type RESTORE, call_type CALL, ctrl_type JMPL, ctrl_type RETT, sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" apply (simp add: supported_instruction_def) by presburger lemma dispatch_instr_result: assumes a1: "supported_instruction (fst i) = True \ (fst i) \ ctrl_type RETT" shows "snd (dispatch_instruction i s) = False" proof (cases "get_trap_set s = {}") case True then have f1: "get_trap_set s = {}" by auto then show ?thesis proof (cases "(fst i) \ {load_store_type LDSB,load_store_type LDUB, load_store_type LDUBA,load_store_type LDUH,load_store_type LD, load_store_type LDA,load_store_type LDD}") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (simp add: load_instr_result) next case False then have f2: "(fst i) \ {load_store_type STB,load_store_type STH,load_store_type ST, load_store_type STA,load_store_type STD, sethi_type SETHI, nop_type NOP, logic_type ANDs,logic_type ANDcc,logic_type ANDN,logic_type ANDNcc, logic_type ORs,logic_type ORcc,logic_type ORN,logic_type XORs, logic_type XNOR, shift_type SLL,shift_type SRL,shift_type SRA, arith_type ADD,arith_type ADDcc,arith_type ADDX, arith_type SUB,arith_type SUBcc,arith_type SUBX, arith_type UMUL,arith_type SMUL,arith_type SMULcc, arith_type UDIV,arith_type UDIVcc,arith_type SDIV, ctrl_type SAVE,ctrl_type RESTORE, call_type CALL, ctrl_type JMPL, sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using a1 apply (simp add: supported_instruction_def) by presburger then show ?thesis proof (cases "(fst i) \ {load_store_type STB,load_store_type STH, load_store_type ST, load_store_type STA,load_store_type STD}") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (auto simp add: store_instr_result) next case False then have f3: "(fst i) \ {sethi_type SETHI, nop_type NOP, logic_type ANDs,logic_type ANDcc,logic_type ANDN,logic_type ANDNcc, logic_type ORs,logic_type ORcc,logic_type ORN,logic_type XORs, logic_type XNOR, shift_type SLL,shift_type SRL,shift_type SRA, arith_type ADD,arith_type ADDcc,arith_type ADDX, arith_type SUB,arith_type SUBcc,arith_type SUBX, arith_type UMUL,arith_type SMUL,arith_type SMULcc, arith_type UDIV,arith_type UDIVcc,arith_type SDIV, ctrl_type SAVE,ctrl_type RESTORE, call_type CALL, ctrl_type JMPL, sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f2 by auto then show ?thesis proof (cases "(fst i) = sethi_type SETHI") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (simp add: sethi_instr_result) next case False then have f4: "(fst i) \ {nop_type NOP, logic_type ANDs,logic_type ANDcc,logic_type ANDN,logic_type ANDNcc, logic_type ORs,logic_type ORcc,logic_type ORN,logic_type XORs, logic_type XNOR, shift_type SLL,shift_type SRL,shift_type SRA, arith_type ADD,arith_type ADDcc,arith_type ADDX, arith_type SUB,arith_type SUBcc,arith_type SUBX, arith_type UMUL,arith_type SMUL,arith_type SMULcc, arith_type UDIV,arith_type UDIVcc,arith_type SDIV, ctrl_type SAVE,ctrl_type RESTORE, call_type CALL, ctrl_type JMPL, sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f3 by auto then show ?thesis proof (cases "fst i = nop_type NOP") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (simp add: nop_instr_result) next case False then have f5: "(fst i) \ {logic_type ANDs,logic_type ANDcc, logic_type ANDN,logic_type ANDNcc, logic_type ORs,logic_type ORcc,logic_type ORN,logic_type XORs, logic_type XNOR, shift_type SLL,shift_type SRL,shift_type SRA, arith_type ADD,arith_type ADDcc,arith_type ADDX, arith_type SUB,arith_type SUBcc,arith_type SUBX, arith_type UMUL,arith_type SMUL,arith_type SMULcc, arith_type UDIV,arith_type UDIVcc,arith_type SDIV, ctrl_type SAVE,ctrl_type RESTORE, call_type CALL, ctrl_type JMPL, sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f4 by auto then show ?thesis proof (cases "(fst i) \ {logic_type ANDs,logic_type ANDcc, logic_type ANDN,logic_type ANDNcc, logic_type ORs,logic_type ORcc,logic_type ORN,logic_type XORs, logic_type XNOR}") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (auto simp add: logical_instr_result) next case False then have f6: "(fst i) \ {shift_type SLL,shift_type SRL, shift_type SRA, arith_type ADD,arith_type ADDcc,arith_type ADDX, arith_type SUB,arith_type SUBcc,arith_type SUBX, arith_type UMUL,arith_type SMUL,arith_type SMULcc, arith_type UDIV,arith_type UDIVcc,arith_type SDIV, ctrl_type SAVE,ctrl_type RESTORE, call_type CALL, ctrl_type JMPL, sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f5 by auto then show ?thesis proof (cases "(fst i) \ {shift_type SLL,shift_type SRL, shift_type SRA}") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (auto simp add: shift_instr_result) next case False then have f7: "(fst i) \ {arith_type ADD,arith_type ADDcc, arith_type ADDX, arith_type SUB,arith_type SUBcc,arith_type SUBX, arith_type UMUL,arith_type SMUL,arith_type SMULcc, arith_type UDIV,arith_type UDIVcc,arith_type SDIV, ctrl_type SAVE,ctrl_type RESTORE, call_type CALL, ctrl_type JMPL, sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f6 by auto then show ?thesis proof (cases "(fst i) \ {arith_type ADD,arith_type ADDcc, arith_type ADDX}") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (auto simp add: add_instr_result) next case False then have f8: "(fst i) \ {arith_type SUB,arith_type SUBcc, arith_type SUBX, arith_type UMUL,arith_type SMUL,arith_type SMULcc, arith_type UDIV,arith_type UDIVcc,arith_type SDIV, ctrl_type SAVE,ctrl_type RESTORE, call_type CALL, ctrl_type JMPL, sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f7 by auto then show ?thesis proof (cases "(fst i) \ {arith_type SUB,arith_type SUBcc, arith_type SUBX}") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (auto simp add: sub_instr_result) next case False then have f9: "(fst i) \ {arith_type UMUL,arith_type SMUL, arith_type SMULcc, arith_type UDIV,arith_type UDIVcc,arith_type SDIV, ctrl_type SAVE,ctrl_type RESTORE, call_type CALL, ctrl_type JMPL, sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f8 by auto then show ?thesis proof (cases "(fst i) \ {arith_type UMUL,arith_type SMUL, arith_type SMULcc}") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (auto simp add: mul_instr_result) next case False then have f10: "(fst i) \ {arith_type UDIV,arith_type UDIVcc, arith_type SDIV, ctrl_type SAVE,ctrl_type RESTORE, call_type CALL, ctrl_type JMPL, sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f9 by auto then show ?thesis proof (cases "(fst i) \ {arith_type UDIV,arith_type UDIVcc, arith_type SDIV}") case True then show ?thesis apply dispatch_instr_proof1 using f1 by (auto simp add: div_instr_result) next case False then have f11: "(fst i) \ {ctrl_type SAVE,ctrl_type RESTORE, call_type CALL, ctrl_type JMPL, sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f10 by auto then show ?thesis proof (cases "(fst i) \ {ctrl_type SAVE,ctrl_type RESTORE}") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (auto simp add: save_restore_instr_result) next case False then have f12: "(fst i) \ {call_type CALL, ctrl_type JMPL, sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f11 by auto then show ?thesis proof (cases "(fst i) = call_type CALL") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (auto simp add: call_instr_result) next case False then have f13: "(fst i) \ {ctrl_type JMPL, sreg_type RDY,sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f12 by auto then show ?thesis proof (cases "(fst i) = ctrl_type JMPL") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (auto simp add: jmpl_instr_result) next case False then have f14: "(fst i) \ { sreg_type RDY, sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR, sreg_type WRY,sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f13 by auto then show ?thesis proof (cases "(fst i) \ {sreg_type RDY, sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR}") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (auto simp add: read_state_reg_instr_result) next case False then have f15: "(fst i) \ { sreg_type WRY, sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR, load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f14 by auto then show ?thesis proof (cases "(fst i) \ {sreg_type WRY, sreg_type WRPSR,sreg_type WRWIM,sreg_type WRTBR}") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (auto simp add: write_state_reg_instr_result) next case False then have f16: "(fst i) \ { load_store_type FLUSH, bicc_type BE,bicc_type BNE,bicc_type BGU,bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f15 by auto then show ?thesis proof (cases "(fst i) = load_store_type FLUSH") case True then show ?thesis using f1 apply dispatch_instr_proof1 by (auto simp add: flush_instr_result) next case False then have f17: "(fst i) \ { bicc_type BE,bicc_type BNE,bicc_type BGU, bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA, bicc_type BN}" using f16 by auto then show ?thesis using f1 proof (cases "(fst i) \ {bicc_type BE, bicc_type BNE,bicc_type BGU, bicc_type BLE, bicc_type BL,bicc_type BGE,bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC,bicc_type BA }") case True then show ?thesis using f1 apply dispatch_instr_proof1 apply auto by (auto simp add: branch_instr_result) next case False then have f18: "(fst i) \ {bicc_type BN}" using f17 by auto then show ?thesis using f1 apply dispatch_instr_proof1 apply auto by (auto simp add: branch_instr_result) qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed next case False then show ?thesis apply (simp add: dispatch_instruction_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: Let_def) by (simp add: returnOk_def return_def) qed lemma dispatch_instr_result_rett: assumes a1: "(fst i) = ctrl_type RETT \ (get_ET (cpu_reg_val PSR s) \ 1 \ (((get_S (cpu_reg_val PSR s)))::word1) \ 0 \ (get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR s))) + 1) mod NWINDOWS)) (cpu_reg_val WIM s)) = 0 \ ((AND) (get_addr (snd i) s) (0b00000000000000000000000000000011::word32)) = 0)" shows "snd (dispatch_instruction i s) = False" proof (cases "get_trap_set s = {}") case True then show ?thesis using a1 apply (simp add: dispatch_instruction_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (simp add: rett_instr_result) next case False then show ?thesis using a1 apply (simp add: dispatch_instruction_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (simp add: return_def) qed lemma execute_instr_sub1_result: "snd (execute_instr_sub1 i s) = False" proof (cases "get_trap_set s = {} \ (fst i) \ {call_type CALL,ctrl_type RETT, ctrl_type JMPL}") case True then show ?thesis apply (simp add: execute_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: write_cpu_def simpler_modify_def) apply auto by (auto simp add: return_def) next case False then show ?thesis apply (simp add: execute_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: write_cpu_def simpler_modify_def) by (auto simp add: return_def) qed lemma next_match : "snd (execute_instruction () s) = False \ NEXT s = Some (snd (fst (execute_instruction () s)))" apply (simp add: NEXT_def) by (simp add: case_prod_unfold) lemma exec_ss1 : "\s'. (execute_instruction () s = (s', False)) \ \s''. (execute_instruction() s = (s'', False))" proof - assume "\s'. (execute_instruction () s = (s', False))" hence "(snd (execute_instruction() s)) = False" by (auto simp add: execute_instruction_def case_prod_unfold) hence "(execute_instruction() s) = ((fst (execute_instruction() s)),False)" by (metis (full_types) prod.collapse) hence "\s''. (execute_instruction() s = (s'', False))" by blast thus ?thesis by assumption qed lemma exec_ss2 : "snd (execute_instruction() s) = False \ snd (execute_instruction () s) = False" proof - assume "snd (execute_instruction() s) = False" hence "snd (execute_instruction () s) = False" by (auto simp add:execute_instruction_def) thus ?thesis by assumption qed lemma good_context_1 : "good_context s \ s' = s \ (get_trap_set s') \ {} \ (reset_trap_val s') = False \ get_ET (cpu_reg_val PSR s') = 0 \ False" proof - assume asm: "good_context s \ s' = s \ (get_trap_set s') \ {} \ (reset_trap_val s') = False \ get_ET (cpu_reg_val PSR s') = 0" then have "(get_trap_set s') \ {} \ (reset_trap_val s') = False \ get_ET (cpu_reg_val PSR s') = 0 \ False" by (simp add: good_context_def get_ET_def cpu_reg_val_def) then show ?thesis using asm by auto qed lemma fetch_instr_result_1 : "\ (\e. fetch_instruction s' = Inl e) \ (\v. fetch_instruction s' = Inr v)" by (meson sumE) lemma fetch_instr_result_2 : "(\v. fetch_instruction s' = Inr v) \ \ (\e. fetch_instruction s' = Inl e)" by force lemma fetch_instr_result_3 : "(\e. fetch_instruction s' = Inl e) \ \ (\v. fetch_instruction s' = Inr v)" by auto lemma decode_instr_result_1 : "\(\v2. ((decode_instruction v1)::(Exception list + instruction)) = Inr v2) \ (\e. ((decode_instruction v1)::(Exception list + instruction)) = Inl e)" by (meson sumE) lemma decode_instr_result_2 : "(\e. ((decode_instruction v1)::(Exception list + instruction)) = Inl e) \ \(\v2. ((decode_instruction v1)::(Exception list + instruction)) = Inr v2)" by force lemma decode_instr_result_3 : "x = decode_instruction v1 \ y = decode_instruction v2 \ v1 = v2 \ x = y" by auto lemma decode_instr_result_4 : "\ (\e. ((decode_instruction v1)::(Exception list + instruction)) = Inl e) \ (\v2. ((decode_instruction v1)::(Exception list + instruction)) = Inr v2)" by (meson sumE) lemma good_context_2 : "good_context (s::(('a::len) sparc_state)) \ fetch_instruction (delayed_pool_write s) = Inr v1 \ \(\v2. (decode_instruction v1::(Exception list + instruction)) = Inr v2) \ False" proof - assume "good_context s \ fetch_instruction (delayed_pool_write s) = Inr v1 \ \(\v2. ((decode_instruction v1)::(Exception list + instruction)) = Inr v2)" hence fact1: "good_context s \ fetch_instruction (delayed_pool_write s) = Inr v1 \ (\e. ((decode_instruction v1)::(Exception list + instruction)) = Inl e)" using decode_instr_result_1 by auto hence fact2: "\(\e. fetch_instruction (delayed_pool_write s) = Inl e)" using fetch_instr_result_2 by auto then have "fetch_instruction (delayed_pool_write s) = Inr v1 \ (\e. ((decode_instruction v1)::(Exception list + instruction)) = Inl e) \ False" proof (cases "(get_trap_set s) \ {} \ (reset_trap_val s) = False \ get_ET (cpu_reg_val PSR s) = 0") case True from this fact1 show ?thesis using good_context_1 by blast next case False then have fact3: "(get_trap_set s) = {} \ (reset_trap_val s) \ False \ get_ET (cpu_reg_val PSR s) \ 0" by auto then show ?thesis using fact1 decode_instr_result_3 by (metis (no_types, lifting) good_context_def sum.case(1) sum.case(2)) qed thus ?thesis using fact1 by auto qed lemma good_context_3 : "good_context (s::(('a::len) sparc_state)) \ s'' = delayed_pool_write s \ fetch_instruction s'' = Inr v1 \ (decode_instruction v1::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = False \ False" proof - assume asm: "good_context (s::(('a::len) sparc_state)) \ s'' = delayed_pool_write s \ fetch_instruction s'' = Inr v1 \ (decode_instruction v1::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = False" then have "annul_val s'' = False \ supported_instruction (fst v2) = False \ False" proof (cases "(get_trap_set s) \ {} \ (reset_trap_val s) = False \ get_ET (cpu_reg_val PSR s) = 0") case True from this asm show ?thesis using good_context_1 by blast next case False then have fact3: "(get_trap_set s) = {} \ (reset_trap_val s) \ False \ get_ET (cpu_reg_val PSR s) \ 0" by auto thus ?thesis using asm by (auto simp add: good_context_def) qed thus ?thesis using asm by auto qed lemma good_context_4 : "good_context (s::(('a::len) sparc_state)) \ s'' = delayed_pool_write s \ fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = True \ \ \This line is redundant\ (fst v2) = ctrl_type RETT \ get_ET (cpu_reg_val PSR s'') \ 1 \ (((get_S (cpu_reg_val PSR s'')))::word1) = 0 \ False" proof - assume asm: "good_context (s::(('a::len) sparc_state)) \ s'' = delayed_pool_write s \ fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = True \ \ \This line is redundant\ (fst v2) = ctrl_type RETT \ get_ET (cpu_reg_val PSR s'') \ 1 \ (((get_S (cpu_reg_val PSR s'')))::word1) = 0" then have "(fst v2) = ctrl_type RETT \ get_ET (cpu_reg_val PSR s'') \ 1 \ (((get_S (cpu_reg_val PSR s'')))::word1) = 0 \ False" proof (cases "(get_trap_set s) \ {} \ (reset_trap_val s) = False \ get_ET (cpu_reg_val PSR s) = 0") case True from this asm show ?thesis using good_context_1 by blast next case False then have fact3: "(get_trap_set s) = {} \ (reset_trap_val s) \ False \ get_ET (cpu_reg_val PSR s) \ 0" by auto thus ?thesis using asm by (auto simp add: good_context_def) qed thus ?thesis using asm by auto qed lemma good_context_5 : "good_context (s::(('a::len) sparc_state)) \ s'' = delayed_pool_write s \ fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = True \ \ \This line is redundant\ (fst v2) = ctrl_type RETT \ get_ET (cpu_reg_val PSR s'') \ 1 \ (((get_S (cpu_reg_val PSR s'')))::word1) \ 0 \ (get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR s''))) + 1) mod NWINDOWS)) (cpu_reg_val WIM s'')) \ 0 \ False" proof - assume asm: "good_context (s::(('a::len) sparc_state)) \ s'' = delayed_pool_write s \ fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = True \ \ \This line is redundant\ (fst v2) = ctrl_type RETT \ get_ET (cpu_reg_val PSR s'') \ 1 \ (((get_S (cpu_reg_val PSR s'')))::word1) \ 0 \ (get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR s''))) + 1) mod NWINDOWS)) (cpu_reg_val WIM s'')) \ 0" then have "(fst v2) = ctrl_type RETT \ get_ET (cpu_reg_val PSR s'') \ 1 \ (((get_S (cpu_reg_val PSR s'')))::word1) \ 0 \ (get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR s''))) + 1) mod NWINDOWS)) (cpu_reg_val WIM s'')) \ 0 \ False" proof (cases "(get_trap_set s) \ {} \ (reset_trap_val s) = False \ get_ET (cpu_reg_val PSR s) = 0") case True from this asm show ?thesis using good_context_1 by blast next case False then have fact3: "(get_trap_set s) = {} \ (reset_trap_val s) \ False \ get_ET (cpu_reg_val PSR s) \ 0" by auto thus ?thesis using asm by (auto simp add: good_context_def) qed thus ?thesis using asm by auto qed lemma good_context_6 : "good_context (s::(('a::len) sparc_state)) \ s'' = delayed_pool_write s \ fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = True \ \ \This line is redundant\ (fst v2) = ctrl_type RETT \ get_ET (cpu_reg_val PSR s'') \ 1 \ (((get_S (cpu_reg_val PSR s'')))::word1) \ 0 \ (get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR s''))) + 1) mod NWINDOWS)) (cpu_reg_val WIM s'')) = 0 \ ((AND) (get_addr (snd v2) s'') (0b00000000000000000000000000000011::word32)) \ 0 \ False" proof - assume asm: "good_context (s::(('a::len) sparc_state)) \ s'' = delayed_pool_write s \ fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = True \ \ \This line is redundant\ (fst v2) = ctrl_type RETT \ get_ET (cpu_reg_val PSR s'') \ 1 \ (((get_S (cpu_reg_val PSR s'')))::word1) \ 0 \ (get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR s''))) + 1) mod NWINDOWS)) (cpu_reg_val WIM s'')) = 0 \ ((AND) (get_addr (snd v2) s'') (0b00000000000000000000000000000011::word32)) \ 0" then have "(fst v2) = ctrl_type RETT \ get_ET (cpu_reg_val PSR s'') \ 1 \ (((get_S (cpu_reg_val PSR s'')))::word1) \ 0 \ (get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR s''))) + 1) mod NWINDOWS)) (cpu_reg_val WIM s'')) = 0 \ ((AND) (get_addr (snd v2) s'') (0b00000000000000000000000000000011::word32)) \ 0 \ False" proof (cases "(get_trap_set s) \ {} \ (reset_trap_val s) = False \ get_ET (cpu_reg_val PSR s) = 0") case True from this asm show ?thesis using good_context_1 by blast next case False then have fact3: "(get_trap_set s) = {} \ (reset_trap_val s) \ False \ get_ET (cpu_reg_val PSR s) \ 0" by auto thus ?thesis using asm by (auto simp add: good_context_def) qed thus ?thesis using asm by auto qed lemma good_context_all : "good_context (s::(('a::len) sparc_state)) \ s'' = delayed_pool_write s \ (get_trap_set s = {} \ (reset_trap_val s) \ False \ get_ET (cpu_reg_val PSR s) \ 0) \ ((\e. fetch_instruction s'' = Inl e) \ (\v1 v2. fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ (annul_val s'' = True \ (annul_val s'' = False \ (\v1' v2'. fetch_instruction s'' = Inr v1' \ ((decode_instruction v1')::(Exception list + instruction)) = Inr v2' \ supported_instruction (fst v2') = True) \ ((fst v2) \ ctrl_type RETT \ ((fst v2) = ctrl_type RETT \ (get_ET (cpu_reg_val PSR s'') = 1 \ (get_ET (cpu_reg_val PSR s'') \ 1 \ (((get_S (cpu_reg_val PSR s'')))::word1) \ 0 \ (get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR s''))) + 1) mod NWINDOWS)) (cpu_reg_val WIM s'')) = 0 \ ((AND) (get_addr (snd v2) s'') (0b00000000000000000000000000000011::word32)) = 0))))))))" proof - assume asm: "good_context s \ s'' = delayed_pool_write s" from asm have "(get_trap_set s) \ {} \ (reset_trap_val s) = False \ get_ET (cpu_reg_val PSR s) = 0 \ False" using good_context_1 by blast hence fact1: "(get_trap_set s = {} \ (reset_trap_val s) \ False \ get_ET (cpu_reg_val PSR s) \ 0)" by auto have fact2: "\(\e. fetch_instruction s'' = Inl e) \ \ (\v1. fetch_instruction s'' = Inr v1) \ False" using fetch_instr_result_1 by blast from asm have fact3: "\v1. fetch_instruction s'' = Inr v1 \ \(\v2.((decode_instruction v1)::(Exception list + instruction)) = Inr v2) \ False" using good_context_2 by blast from asm have fact4: "\v1 v2. fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = False \ False" using good_context_3 by blast from asm have fact5: "\v1 v2. fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = True \ (fst v2) = ctrl_type RETT \ get_ET (cpu_reg_val PSR s'') \ 1 \ (((get_S (cpu_reg_val PSR s'')))::word1) = 0 \ False" using good_context_4 by blast from asm have fact6: "\v1 v2. fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = True \ (fst v2) = ctrl_type RETT \ get_ET (cpu_reg_val PSR s'') \ 1 \ (((get_S (cpu_reg_val PSR s'')))::word1) \ 0 \ (get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR s''))) + 1) mod NWINDOWS)) (cpu_reg_val WIM s'')) \ 0 \ False" using good_context_5 by blast from asm have fact7: "\v1 v2. fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = True \ (fst v2) = ctrl_type RETT \ get_ET (cpu_reg_val PSR s'') \ 1 \ (((get_S (cpu_reg_val PSR s'')))::word1) \ 0 \ (get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR s''))) + 1) mod NWINDOWS)) (cpu_reg_val WIM s'')) = 0 \ ((AND) (get_addr (snd v2) s'') (0b00000000000000000000000000000011::word32)) \ 0 \ False" using good_context_6 by blast from asm show ?thesis proof (cases "(\e. fetch_instruction s'' = Inl e)") case True then show ?thesis using fact1 by auto next case False then have fact8: "\v1. fetch_instruction s'' = Inr v1 \ (\v2.((decode_instruction v1)::(Exception list + instruction)) = Inr v2)" using fact2 fact3 by auto then show ?thesis proof (cases "annul_val s'' = True") case True then show ?thesis using fact1 fact8 by auto next case False then have fact9: "\v1 v2. fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = True" using fact4 fact8 by blast then show ?thesis proof (cases "\v1 v2. fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ (fst v2) \ ctrl_type RETT") case True then show ?thesis using fact1 fact9 by auto next case False then have fact10: "\v1 v2. fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ annul_val s'' = False \ supported_instruction (fst v2) = True \ (fst v2) = ctrl_type RETT" using fact9 by auto then show ?thesis proof (cases "get_ET (cpu_reg_val PSR s'') = 1") case True then show ?thesis using fact1 fact9 by auto next case False then have fact11: "get_ET (cpu_reg_val PSR s'') \ 1 \ (((get_S (cpu_reg_val PSR s'')))::word1) \ 0" using fact10 fact5 by auto then have fact12: "(get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR s''))) + 1) mod NWINDOWS)) (cpu_reg_val WIM s'')) = 0" using fact10 fact6 by auto then have fact13: "\v1 v2. fetch_instruction s'' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ ((AND) (get_addr (snd v2) s'') (0b00000000000000000000000000000011::word32)) = 0" using fact10 fact11 fact7 by blast thus ?thesis using fact1 fact10 fact11 fact12 by auto qed qed qed qed qed lemma select_trap_result1 : "(reset_trap_val s) = True \ snd (select_trap() s) = False" apply (simp add: select_trap_def exec_gets return_def) by (simp add: bind_def h1_def h2_def simpler_modify_def) lemma select_trap_result2 : assumes a1: "\(reset_trap_val s = False \ get_ET (cpu_reg_val PSR s) = 0)" shows "snd (select_trap() s) = False" proof (cases "reset_trap_val s = True") case True then show ?thesis using select_trap_result1 by blast next case False then have f1: "reset_trap_val s = False \ get_ET (cpu_reg_val PSR s) \ 0" using a1 by auto then show ?thesis proof (cases "data_store_error \ get_trap_set s") case True then show ?thesis using f1 by select_trap_proof0 next case False then have f2: "data_store_error \ get_trap_set s" by auto then show ?thesis proof (cases "instruction_access_error \ get_trap_set s") case True then show ?thesis using f1 f2 by select_trap_proof0 next case False then have f3: "instruction_access_error \ get_trap_set s" by auto then show ?thesis proof (cases "r_register_access_error \ get_trap_set s") case True then show ?thesis using f1 f2 f3 by select_trap_proof0 next case False then have f4: "r_register_access_error \ get_trap_set s" by auto then show ?thesis proof (cases "instruction_access_exception \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 by select_trap_proof0 next case False then have f5: "instruction_access_exception \ get_trap_set s" by auto then show ?thesis proof (cases "privileged_instruction \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 by select_trap_proof0 next case False then have f6: "privileged_instruction \ get_trap_set s" by auto then show ?thesis proof (cases "illegal_instruction \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 by select_trap_proof0 next case False then have f7: "illegal_instruction \ get_trap_set s" by auto then show ?thesis proof (cases "fp_disabled \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 by select_trap_proof0 next case False then have f8: "fp_disabled \ get_trap_set s" by auto then show ?thesis proof (cases "cp_disabled \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 by select_trap_proof0 next case False then have f9: "cp_disabled \ get_trap_set s" by auto then show ?thesis proof (cases "unimplemented_FLUSH \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 by select_trap_proof0 next case False then have f10: "unimplemented_FLUSH \ get_trap_set s" by auto then show ?thesis proof (cases "window_overflow \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 by select_trap_proof0 next case False then have f11: "window_overflow \ get_trap_set s" by auto then show ?thesis proof (cases "window_underflow \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 by select_trap_proof0 next case False then have f12: "window_underflow \ get_trap_set s" by auto then show ?thesis proof (cases "mem_address_not_aligned \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 by select_trap_proof0 next case False then have f13: "mem_address_not_aligned \ get_trap_set s" by auto then show ?thesis proof (cases "fp_exception \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 by select_trap_proof0 next case False then have f14: "fp_exception \ get_trap_set s" by auto then show ?thesis proof (cases "cp_exception \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 by select_trap_proof0 next case False then have f15: "cp_exception \ get_trap_set s" by auto then show ?thesis proof (cases "data_access_error \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 by select_trap_proof0 next case False then have f16: "data_access_error \ get_trap_set s" by auto then show ?thesis proof (cases "data_access_exception \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 by select_trap_proof0 next case False then have f17: "data_access_exception \ get_trap_set s" by auto then show ?thesis proof (cases "tag_overflow \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 by select_trap_proof0 next case False then have f18: "tag_overflow \ get_trap_set s" by auto then show ?thesis proof (cases "division_by_zero \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 by select_trap_proof0 next case False then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 apply (simp add: select_trap_def exec_gets return_def) apply (simp add: DetMonad.bind_def h1_def h2_def simpler_modify_def) apply (simp add: return_def simpler_gets_def) apply (simp add: case_prod_unfold) apply (simp add: return_def) apply (simp add: write_cpu_tt_def write_cpu_def) by (simp add: simpler_gets_def bind_def h1_def h2_def simpler_modify_def) qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed lemma emp_trap_set_err_mode : "err_mode_val s = err_mode_val (emp_trap_set s)" by (auto simp add: emp_trap_set_def err_mode_val_def) lemma write_cpu_tt_err_mode : "err_mode_val s = err_mode_val (snd (fst (write_cpu_tt w s)))" apply (simp add: write_cpu_tt_def err_mode_val_def write_cpu_def) apply (simp add: exec_gets return_def) apply (simp add: bind_def simpler_modify_def) by (simp add: cpu_reg_mod_def) lemma select_trap_monad : "snd (select_trap() s) = False \ err_mode_val s = err_mode_val (snd (fst (select_trap () s)))" proof - assume a1: "snd (select_trap() s) = False" then have f0: "reset_trap_val s = False \ get_ET (cpu_reg_val PSR s) = 0 \ False" apply (simp add: select_trap_def exec_gets return_def) apply (simp add: bind_def h1_def h2_def simpler_modify_def) by (simp add: fail_def split_def) then show ?thesis proof (cases "reset_trap_val s = True") case True from a1 f0 this show ?thesis apply (simp add: select_trap_def exec_gets return_def) apply (simp add: bind_def h1_def h2_def simpler_modify_def) by (simp add: emp_trap_set_def err_mode_val_def) next case False then have f1: "reset_trap_val s = False \ get_ET (cpu_reg_val PSR s) \ 0" using f0 by auto then show ?thesis using f1 a1 proof (cases "data_store_error \ get_trap_set s") case True then show ?thesis using f1 a1 by select_trap_proof1 next case False then have f2: "data_store_error \ get_trap_set s" by auto then show ?thesis proof (cases "instruction_access_error \ get_trap_set s") case True then show ?thesis using f1 f2 a1 by select_trap_proof1 next case False then have f3: "instruction_access_error \ get_trap_set s" by auto then show ?thesis proof (cases "r_register_access_error \ get_trap_set s") case True then show ?thesis using f1 f2 f3 a1 by select_trap_proof1 next case False then have f4: "r_register_access_error \ get_trap_set s" by auto then show ?thesis proof (cases "instruction_access_exception \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 a1 by select_trap_proof1 next case False then have f5: "instruction_access_exception \ get_trap_set s" by auto then show ?thesis proof (cases "privileged_instruction \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 a1 by select_trap_proof1 next case False then have f6: "privileged_instruction \ get_trap_set s" by auto then show ?thesis proof (cases "illegal_instruction \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 a1 by select_trap_proof1 next case False then have f7: "illegal_instruction \ get_trap_set s" by auto then show ?thesis proof (cases "fp_disabled \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 a1 by select_trap_proof1 next case False then have f8: "fp_disabled \ get_trap_set s" by auto then show ?thesis proof (cases "cp_disabled \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 a1 by select_trap_proof1 next case False then have f9: "cp_disabled \ get_trap_set s" by auto then show ?thesis proof (cases "unimplemented_FLUSH \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 a1 by select_trap_proof1 next case False then have f10: "unimplemented_FLUSH \ get_trap_set s" by auto then show ?thesis proof (cases "window_overflow \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 a1 by select_trap_proof1 next case False then have f11: "window_overflow \ get_trap_set s" by auto then show ?thesis proof (cases "window_underflow \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 a1 by select_trap_proof1 next case False then have f12: "window_underflow \ get_trap_set s" by auto then show ?thesis proof (cases "mem_address_not_aligned \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 a1 by select_trap_proof1 next case False then have f13: "mem_address_not_aligned \ get_trap_set s" by auto then show ?thesis proof (cases "fp_exception \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 a1 by select_trap_proof1 next case False then have f14: "fp_exception \ get_trap_set s" by auto then show ?thesis proof (cases "cp_exception \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 a1 by select_trap_proof1 next case False then have f15: "cp_exception \ get_trap_set s" by auto then show ?thesis proof (cases "data_access_error \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 a1 by select_trap_proof1 next case False then have f16: "data_access_error \ get_trap_set s" by auto then show ?thesis proof (cases "data_access_exception \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 a1 by select_trap_proof1 next case False then have f17: "data_access_exception \ get_trap_set s" by auto then show ?thesis proof (cases "tag_overflow \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 a1 by select_trap_proof1 next case False then have f18: "tag_overflow \ get_trap_set s" by auto then show ?thesis proof (cases "division_by_zero \ get_trap_set s") case True then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 a1 by select_trap_proof1 next case False then show ?thesis using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 a1 apply (simp add: select_trap_def exec_gets return_def) apply (simp add: bind_def h1_def h2_def simpler_modify_def) apply (simp add: return_def simpler_gets_def) apply (simp add: emp_trap_set_def err_mode_val_def cpu_reg_mod_def) apply (simp add: case_prod_unfold) apply (simp add: return_def) apply clarsimp apply (simp add: write_cpu_tt_def write_cpu_def write_tt_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: simpler_modify_def) by (simp add: cpu_reg_val_def cpu_reg_mod_def) qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed lemma exe_trap_st_pc_result : "snd (exe_trap_st_pc() s) = False" proof (cases "annul_val s = True") case True then show ?thesis apply (simp add: exe_trap_st_pc_def get_curr_win_def) apply (simp add: exec_gets return_def) apply (simp add: DetMonad.bind_def h1_def h2_def) by (simp add: set_annul_def write_reg_def simpler_modify_def) next case False then show ?thesis apply (simp add: exe_trap_st_pc_def get_curr_win_def) apply (simp add: exec_gets return_def) apply (simp add: DetMonad.bind_def h1_def h2_def) by (simp add: write_reg_def simpler_modify_def) qed lemma exe_trap_wr_pc_result : "snd (exe_trap_wr_pc() s) = False" proof (cases "reset_trap_val s = True") case True then show ?thesis apply (simp add: exe_trap_wr_pc_def get_curr_win_def) apply (simp add: exec_gets return_def) apply (simp add: DetMonad.bind_def h1_def h2_def) apply (simp add: write_cpu_def simpler_modify_def) apply (simp add: simpler_gets_def) apply (simp add: cpu_reg_val_def update_S_def cpu_reg_mod_def reset_trap_val_def) apply (simp add: write_cpu_def simpler_modify_def DetMonad.bind_def h1_def h2_def) apply (simp add: return_def) by (simp add: set_reset_trap_def simpler_modify_def DetMonad.bind_def h1_def h2_def return_def) next case False then show ?thesis apply (simp add: exe_trap_wr_pc_def get_curr_win_def) apply (simp add: exec_gets return_def) apply (simp add: DetMonad.bind_def h1_def h2_def) apply (simp add: write_cpu_def simpler_modify_def) apply (simp add: simpler_gets_def) apply (simp add: cpu_reg_val_def update_S_def cpu_reg_mod_def reset_trap_val_def) apply (simp add: write_cpu_def simpler_modify_def DetMonad.bind_def h1_def h2_def) by (simp add: return_def) qed lemma execute_trap_result : "\(reset_trap_val s = False \ get_ET (cpu_reg_val PSR s) = 0) \ snd (execute_trap() s) = False" proof - assume "\(reset_trap_val s = False \ get_ET (cpu_reg_val PSR s) = 0)" then have fact1: "snd (select_trap() s) = False" using select_trap_result2 by blast then show ?thesis proof (cases "err_mode_val s = True") case True then show ?thesis using fact1 apply (simp add: execute_trap_def exec_gets return_def) apply (simp add: DetMonad.bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) by (simp add: in_gets return_def select_trap_monad simpler_gets_def) next case False then show ?thesis using fact1 select_trap_monad apply (simp add: execute_trap_def exec_gets return_def) apply (simp add: DetMonad.bind_def h1_def h2_def) apply (simp add: case_prod_unfold) apply (simp add: simpler_gets_def) apply (auto simp add: select_trap_monad) apply (simp add: DetMonad.bind_def h1_def h2_def get_curr_win_def) apply (simp add: get_CWP_def cpu_reg_val_def) apply (simp add: simpler_gets_def return_def write_cpu_def) apply (simp add: simpler_modify_def DetMonad.bind_def h1_def h2_def) apply (simp add: exe_trap_st_pc_result) by (simp add: case_prod_unfold exe_trap_wr_pc_result) qed qed lemma execute_trap_result2 : "\(reset_trap_val s = False \ get_ET (cpu_reg_val PSR s) = 0) \ snd (execute_trap() s) = False" using execute_trap_result by blast lemma exe_instr_all : "good_context (s::(('a::len) sparc_state)) \ snd (execute_instruction() s) = False" proof - assume asm1: "good_context s" let ?s' = "delayed_pool_write s" from asm1 have f1 : "(get_trap_set s = {} \ (reset_trap_val s) \ False \ get_ET (cpu_reg_val PSR s) \ 0) \ ((\e. fetch_instruction ?s' = Inl e) \ (\v1 v2. fetch_instruction ?s' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ (annul_val ?s' = True \ (annul_val ?s' = False \ (\v1' v2'. fetch_instruction ?s' = Inr v1' \ ((decode_instruction v1')::(Exception list + instruction)) = Inr v2' \ supported_instruction (fst v2') = True) \ ((fst v2) \ ctrl_type RETT \ ((fst v2) = ctrl_type RETT \ (get_ET (cpu_reg_val PSR ?s') = 1 \ (get_ET (cpu_reg_val PSR ?s') \ 1 \ (((get_S (cpu_reg_val PSR ?s')))::word1) \ 0 \ (get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR ?s'))) + 1) mod NWINDOWS)) (cpu_reg_val WIM ?s')) = 0 \ ((AND) (get_addr (snd v2) ?s') (0b00000000000000000000000000000011::word32)) = 0))))))))" using good_context_all by blast from f1 have f2: "get_trap_set s \ {} \ (reset_trap_val s) \ False \ get_ET (cpu_reg_val PSR s) \ 0" by auto show ?thesis proof (cases "get_trap_set s = {}") case True then have f3: "get_trap_set s = {}" by auto then show ?thesis proof (cases "exe_mode_val s = True") case True then have f4: "exe_mode_val s = True" by auto then show ?thesis proof (cases "\e1. fetch_instruction ?s' = Inl e1") case True then show ?thesis using f3 apply exe_proof_to_decode apply (simp add: raise_trap_def simpler_modify_def) by (simp add: bind_def h1_def h2_def return_def) next case False then have f5: "\ v1. fetch_instruction ?s' = Inr v1" using fetch_instr_result_1 by blast then have f6: "\v1 v2. fetch_instruction ?s' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2" using f1 fetch_instr_result_2 by blast then show ?thesis proof (cases "annul_val ?s' = True") case True then show ?thesis using f3 f4 f6 apply exe_proof_to_decode apply (simp add: set_annul_def annul_mod_def simpler_modify_def bind_def h1_def h2_def) apply (simp add: return_def simpler_gets_def) by (simp add: write_cpu_def simpler_modify_def) next case False then have f7: "\v1 v2. fetch_instruction ?s' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ (\v1' v2'. fetch_instruction ?s' = Inr v1' \ ((decode_instruction v1')::(Exception list + instruction)) = Inr v2' \ supported_instruction (fst v2') = True) \ annul_val ?s' = False" using f1 f6 fetch_instr_result_2 by auto then have f7': "\v1 v2. fetch_instruction ?s' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ supported_instruction (fst v2) = True \ annul_val ?s' = False" by auto then show ?thesis proof (cases "\v1 v2. fetch_instruction ?s' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ (fst v2) = ctrl_type RETT") case True then have f8: "\v1 v2. fetch_instruction ?s' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ (fst v2) = ctrl_type RETT" by auto then show ?thesis proof (cases "get_trap_set ?s' = {}") case True then have f9: "get_trap_set ?s' = {}" by auto then show ?thesis proof (cases "get_ET (cpu_reg_val PSR ?s') = 1") case True then have f10: "get_ET (cpu_reg_val PSR ?s') = 1" by auto then show ?thesis proof (cases "(((get_S (cpu_reg_val PSR ?s')))::word1) = 0") case True then show ?thesis using f3 f4 f7 f8 f9 f10 apply exe_proof_to_decode apply exe_proof_dispatch_rett apply (simp add: raise_trap_def simpler_modify_def) apply (auto simp add: execute_instr_sub1_result return_def) by (simp add: case_prod_unfold) next case False then show ?thesis using f3 f4 f7 f8 f9 f10 apply exe_proof_to_decode apply exe_proof_dispatch_rett apply (simp add: raise_trap_def simpler_modify_def) apply (auto simp add: execute_instr_sub1_result return_def) by (simp add: case_prod_unfold) qed next case False then have f11: "\v1 v2. fetch_instruction ?s' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ annul_val ?s' = False \ (fst v2) = ctrl_type RETT \ (get_ET (cpu_reg_val PSR ?s') \ 1 \ (((get_S (cpu_reg_val PSR ?s')))::word1) \ 0 \ (get_WIM_bit (nat (((uint (get_CWP (cpu_reg_val PSR ?s'))) + 1) mod NWINDOWS)) (cpu_reg_val WIM ?s')) = 0 \ ((AND) (get_addr (snd v2) ?s') (0b00000000000000000000000000000011::word32)) = 0)" using f1 fetch_instr_result_2 f7' f8 by auto then show ?thesis using f3 f4 proof (cases "get_trap_set ?s' = {}") case True then show ?thesis using f3 f4 f11 apply (simp add: execute_instruction_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def simpler_modify_def) apply clarsimp apply (simp add: return_def) apply (simp add: bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply auto apply (simp add: execute_instr_sub1_result) apply (simp add: dispatch_instruction_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (simp add: rett_instr_result) next case False then show ?thesis using f3 f4 f11 apply (simp add: execute_instruction_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def simpler_modify_def) apply clarsimp apply (simp add: return_def) apply (simp add: bind_def h1_def h2_def) apply (simp add: case_prod_unfold) apply (simp add: execute_instr_sub1_result) apply (simp add: dispatch_instruction_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (simp add: return_def) qed qed next case False then show ?thesis using f3 f4 f7 f8 apply exe_proof_to_decode apply (simp add: dispatch_instruction_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: case_prod_unfold) by (auto simp add: execute_instr_sub1_result return_def Let_def) qed next case False \ \Instruction is not \RETT\.\ then have "\v1 v2. fetch_instruction ?s' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ (fst v2) \ ctrl_type RETT" using f7 by auto then have "\v1 v2. fetch_instruction ?s' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ (fst v2) \ ctrl_type RETT \ supported_instruction (fst v2) = True \ annul_val ?s' = False" using f7 by auto then have "\v1 v2. fetch_instruction ?s' = Inr v1 \ ((decode_instruction v1)::(Exception list + instruction)) = Inr v2 \ (fst v2) \ ctrl_type RETT \ supported_instruction (fst v2) = True \ annul_val ?s' = False \ snd (dispatch_instruction v2 ?s') = False" by (auto simp add: dispatch_instr_result) then show ?thesis using f3 f4 apply exe_proof_to_decode apply (simp add: bind_def h1_def h2_def) apply (simp add: case_prod_unfold) by (simp add: execute_instr_sub1_result) qed qed qed next case False then show ?thesis using f3 apply (simp add: execute_instruction_def) by (simp add: exec_gets return_def) qed next case False then have "get_trap_set s \ {} \ ((reset_trap_val s) \ False \ get_ET (cpu_reg_val PSR s) \ 0)" using f2 by auto then show ?thesis apply (simp add: execute_instruction_def exec_gets) by (simp add: execute_trap_result2) qed qed lemma dispatch_fail: "snd (execute_instruction() (s::(('a::len) sparc_state))) = False \ get_trap_set s = {} \ exe_mode_val s \ fetch_instruction (delayed_pool_write s) = Inr v \ ((decode_instruction v)::(Exception list + instruction)) = Inl e \ False" using decode_instr_result_2 apply (simp add: execute_instruction_def) apply (simp add: exec_gets bind_def) apply clarsimp apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: simpler_modify_def return_def) by (simp add: fail_def) lemma no_error : "good_context s \ snd (execute_instruction () s) = False" proof - assume "good_context s" hence "snd (execute_instruction() s) = False" using exe_instr_all by auto hence "snd (execute_instruction () s) = False" by (simp add: exec_ss2) thus ?thesis by assumption qed theorem single_step : "good_context s \ NEXT s = Some (snd (fst (execute_instruction () s)))" by (simp add: no_error next_match) (*********************************************************************) section \Privilege safty\ (*********************************************************************) text \The following shows that, if the pre-state is under user mode, then after a singel step execution, the post-state is aslo under user mode.\ lemma write_cpu_pc_privilege: "s' = snd (fst (write_cpu w PC s)) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: write_cpu_def simpler_modify_def) apply (simp add: cpu_reg_mod_def) by (simp add: cpu_reg_val_def) lemma write_cpu_npc_privilege: "s' = snd (fst (write_cpu w nPC s)) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: write_cpu_def simpler_modify_def) apply (simp add: cpu_reg_mod_def) by (simp add: cpu_reg_val_def) lemma write_cpu_y_privilege: "s' = snd (fst (write_cpu w Y s)) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: write_cpu_def simpler_modify_def) apply (simp add: cpu_reg_mod_def) by (simp add: cpu_reg_val_def) lemma cpu_reg_mod_y_privilege: "s' = cpu_reg_mod w Y s \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" by (simp add: cpu_reg_mod_def cpu_reg_val_def) lemma cpu_reg_mod_asr_privilege: "s' = cpu_reg_mod w (ASR r) s \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" by (simp add: cpu_reg_mod_def cpu_reg_val_def) lemma global_reg_mod_privilege: "s' = global_reg_mod w1 n w2 s \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (induction n arbitrary:s) apply (clarsimp) apply (auto) apply (simp add: Let_def) by (simp add: cpu_reg_val_def) lemma out_reg_mod_privilege: "s' = out_reg_mod a w r s \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: out_reg_mod_def Let_def) by (simp add: cpu_reg_val_def) lemma in_reg_mod_privilege: "s' = in_reg_mod a w r s \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: in_reg_mod_def Let_def) by (simp add: cpu_reg_val_def) lemma user_reg_mod_privilege: assumes a1: " s' = user_reg_mod d (w::(('a::len) window_size)) r (s::(('a::len) sparc_state)) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "r = 0") case True then show ?thesis using a1 by (simp add: user_reg_mod_def) next case False then have f1: "r \ 0" by auto then show ?thesis proof (cases "0 < r \ r < 8") case True then show ?thesis using a1 f1 apply (simp add: user_reg_mod_def) by (auto intro: global_reg_mod_privilege) next case False then have f2: "\(0 < r \ r < 8)" by auto then show ?thesis proof (cases "7 < r \ r < 16") case True then show ?thesis using a1 f1 f2 apply (simp add: user_reg_mod_def) by (auto intro: out_reg_mod_privilege) next case False then have f3: "\ (7 < r \ r < 16)" by auto then show ?thesis proof (cases "15 < r \ r < 24") case True then show ?thesis using a1 f1 f2 f3 apply (simp add: user_reg_mod_def) by (simp add: cpu_reg_val_def) next case False then show ?thesis using a1 f1 f2 f3 apply (simp add: user_reg_mod_def) by (auto intro: in_reg_mod_privilege) qed qed qed qed lemma write_reg_privilege: "s' = snd (fst (write_reg w1 w2 w3 (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: write_reg_def simpler_modify_def) by (auto intro: user_reg_mod_privilege) lemma set_annul_privilege: "s' = snd (fst (set_annul b s)) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: set_annul_def simpler_modify_def) apply (simp add: annul_mod_def write_annul_def) by (simp add: cpu_reg_val_def) lemma set_reset_trap_privilege: "s' = snd (fst (set_reset_trap b s)) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: set_reset_trap_def simpler_modify_def) apply (simp add: reset_trap_mod_def write_annul_def) by (simp add: cpu_reg_val_def) lemma empty_delayed_pool_write_privilege: "get_delayed_pool s = [] \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ s' = delayed_pool_write s \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: delayed_pool_write_def) by (simp add: get_delayed_write_def delayed_write_all_def delayed_pool_rm_list_def) lemma raise_trap_privilege: "(((get_S (cpu_reg_val PSR s)))::word1) = 0 \ s' = snd (fst (raise_trap t s)) \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: raise_trap_def) apply (simp add: simpler_modify_def add_trap_set_def) by (simp add: cpu_reg_val_def) lemma write_cpu_tt_privilege: "s' = snd (fst (write_cpu_tt w s)) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: write_cpu_tt_def) apply (simp add: exec_gets) apply (simp add: write_cpu_def cpu_reg_mod_def write_tt_def) apply (simp add: simpler_modify_def) by (simp add: cpu_reg_val_def) lemma emp_trap_set_privilege: "s' = emp_trap_set s \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: emp_trap_set_def) by (simp add: cpu_reg_val_def) lemma sys_reg_mod_privilege: "s' = sys_reg_mod w r s \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: sys_reg_mod_def) by (simp add: cpu_reg_val_def) lemma mem_mod_privilege: assumes a1: "s' = mem_mod a1 a2 v s \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "(uint a1) = 8 \ (uint a1) = 10") case True then show ?thesis using a1 apply (simp add: mem_mod_def) apply (simp add: Let_def) by (simp add: cpu_reg_val_def) next case False then have f1: "\((uint a1) = 8 \ (uint a1) = 10)" by auto then show ?thesis proof (cases "(uint a1) = 9 \ (uint a1) = 11") case True then show ?thesis using a1 f1 apply (simp add: mem_mod_def) apply (simp add: Let_def) by (simp add: cpu_reg_val_def) next case False then show ?thesis using a1 f1 apply (simp add: mem_mod_def) by (simp add: cpu_reg_val_def) qed qed lemma mem_mod_w32_privilege: "s' = mem_mod_w32 a1 a2 b d s \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: mem_mod_w32_def) apply (simp add: Let_def) by (auto intro: mem_mod_privilege) lemma add_instr_cache_privilege: "s' = add_instr_cache s addr y m \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: add_instr_cache_def) apply (simp add: Let_def) by (simp add: icache_mod_def cpu_reg_val_def) lemma add_data_cache_privilege: "s' = add_data_cache s addr y m \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: add_data_cache_def) apply (simp add: Let_def) by (simp add: dcache_mod_def cpu_reg_val_def) lemma memory_read_privilege: assumes a1: "s' = snd (memory_read asi addr s) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "uint asi = 1") case True then show ?thesis using a1 apply (simp add: memory_read_def) by (simp add: Let_def) next case False then have f1: "uint asi \ 1" by auto then show ?thesis proof (cases "uint asi = 2") case True then show ?thesis using a1 f1 by (simp add: memory_read_def) next case False then have f2: "uint asi \ 2" by auto then show ?thesis proof (cases "uint asi \ {8,9}") case True then have f3: "uint asi \ {8,9}" by auto then show ?thesis proof (cases "load_word_mem s addr asi = None") case True then have f4: "load_word_mem s addr asi = None" by auto then show ?thesis using a1 f1 f2 f3 f4 by (simp add: memory_read_def) next case False then show ?thesis using a1 f1 f2 f3 apply (simp add: memory_read_def) apply auto apply (simp add: add_instr_cache_privilege) by (simp add: add_instr_cache_privilege) qed next case False then have f5: "uint asi \ {8, 9}" by auto then show ?thesis proof (cases "uint asi \ {10,11}") case True then have f6: "uint asi \ {10,11}" by auto then show ?thesis proof (cases "load_word_mem s addr asi = None") case True then have f7: "load_word_mem s addr asi = None" by auto then show ?thesis using a1 f1 f2 f5 f6 f7 by (simp add: memory_read_def) next case False then show ?thesis using a1 f1 f2 f5 f6 apply (simp add: memory_read_def) apply auto apply (simp add: add_data_cache_privilege) by (simp add: add_data_cache_privilege) qed next case False then have f8: "uint asi \ {10,11}" by auto then show ?thesis proof (cases "uint asi = 13") case True then have f9: "uint asi = 13" by auto then show ?thesis proof (cases "read_instr_cache s addr = None") case True then show ?thesis using a1 f1 f2 f5 f8 f9 by (simp add: memory_read_def) next case False then show ?thesis using a1 f1 f2 f5 f8 f9 apply (simp add: memory_read_def) by auto qed next case False then have f10: "uint asi \ 13" by auto then show ?thesis proof (cases "uint asi = 15") case True then show ?thesis using a1 f1 f2 f5 f8 f10 apply (simp add: memory_read_def) apply (cases "read_data_cache s addr = None") by auto next case False then show ?thesis using a1 f1 f2 f5 f8 f10 apply (simp add: memory_read_def) \ \The rest cases are easy.\ by (simp add: Let_def) qed qed qed qed qed qed lemma get_curr_win_privilege: "s' = snd (fst (get_curr_win() s)) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: get_curr_win_def) by (simp add: simpler_gets_def) lemma load_sub2_privilege: assumes a1: "s' = snd (fst (load_sub2 addr asi r win w s)) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "fst (memory_read asi (addr + 4) (snd (fst (write_reg w win (r AND 30) s)))) = None") case True then show ?thesis using a1 apply (simp add: load_sub2_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) by (auto intro: raise_trap_privilege write_reg_privilege) next case False then show ?thesis using a1 apply (simp add: load_sub2_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: case_prod_unfold) apply clarsimp apply (simp add: simpler_modify_def bind_def h1_def h2_def Let_def) by (auto intro: write_reg_privilege memory_read_privilege) qed lemma load_sub3_privilege: assumes a1: "s' = snd (fst (load_sub3 instr curr_win rd asi address s)) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "fst (memory_read asi address s) = None") case True then show ?thesis using a1 apply (simp add: load_sub3_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: case_prod_unfold) by (auto intro: raise_trap_privilege) next case False then have f1: "fst (memory_read asi address s) \ None " by auto then show ?thesis proof (cases "rd \ 0 \ (fst instr = load_store_type LD \ fst instr = load_store_type LDA \ fst instr = load_store_type LDUH \ fst instr = load_store_type LDSB \ fst instr = load_store_type LDUB \ fst instr = load_store_type LDUBA \ fst instr = load_store_type LDSH \ fst instr = load_store_type LDSHA \ fst instr = load_store_type LDUHA \ fst instr = load_store_type LDSBA)") case True then show ?thesis using a1 f1 apply (simp add: load_sub3_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: case_prod_unfold) apply clarsimp apply (simp add: simpler_modify_def bind_def h1_def h2_def Let_def) by (auto intro: write_reg_privilege memory_read_privilege) next case False then show ?thesis using a1 f1 apply (simp add: load_sub3_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: case_prod_unfold) apply auto apply (simp add: simpler_modify_def bind_def h1_def h2_def) apply (auto intro: load_sub2_privilege memory_read_privilege) apply (simp add: simpler_modify_def bind_def h1_def h2_def) by (auto intro: load_sub2_privilege memory_read_privilege) qed qed lemma load_sub1_privilege: assumes a1: "s' = snd (fst (load_sub1 instr rd s_val s)) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: load_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply auto by (auto intro: get_curr_win_privilege raise_trap_privilege load_sub3_privilege) lemma load_instr_privilege: "s' = snd (fst (load_instr i s)) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: load_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: Let_def) apply clarsimp by (auto intro: get_curr_win_privilege raise_trap_privilege load_sub1_privilege) lemma store_barrier_pending_mod_privilege: "s' = store_barrier_pending_mod b s \ (((get_S (cpu_reg_val PSR s)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: store_barrier_pending_mod_def) apply (simp add: write_store_barrier_pending_def) by (simp add: cpu_reg_val_def) lemma store_word_mem_privilege: assumes a1: "store_word_mem s addr data byte_mask asi = Some s' \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: store_word_mem_def) apply (case_tac "virt_to_phys addr (mmu s) (mem s) = None") apply auto apply (case_tac "mmu_writable (get_acc_flag b) asi") apply auto by (simp add: mem_mod_w32_privilege) lemma flush_instr_cache_privilege: "(((get_S (cpu_reg_val PSR s)))::word1) = 0 \ s' = flush_instr_cache s \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: flush_instr_cache_def) by (simp add: cpu_reg_val_def) lemma flush_data_cache_privilege: "(((get_S (cpu_reg_val PSR s)))::word1) = 0 \ s' = flush_data_cache s \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: flush_data_cache_def) by (simp add: cpu_reg_val_def) lemma flush_cache_all_privilege: "(((get_S (cpu_reg_val PSR s)))::word1) = 0 \ s' = flush_cache_all s \ (((get_S (cpu_reg_val PSR s')))::word1) = 0" apply (simp add: flush_cache_all_def) by (simp add: cpu_reg_val_def) lemma memory_write_asi_privilege: assumes a1: "r = memory_write_asi asi addr byte_mask data s \ r = Some s' \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "uint asi = 1") case True then show ?thesis using a1 apply (simp add: memory_write_asi_def) by (auto intro: store_word_mem_privilege) next case False then have f1: "uint asi \ 1" by auto then show ?thesis proof (cases "uint asi = 2") case True then have f01: "uint asi = 2" by auto then show ?thesis proof (cases "uint addr = 0") case True then show ?thesis using a1 f1 f01 apply (simp add: memory_write_asi_def) apply (simp add: ccr_flush_def) apply (simp add: Let_def) apply auto apply (metis flush_data_cache_privilege flush_instr_cache_privilege sys_reg_mod_privilege) apply (metis flush_instr_cache_privilege sys_reg_mod_privilege) apply (metis flush_data_cache_privilege sys_reg_mod_privilege) by (simp add: sys_reg_mod_privilege) next case False then show ?thesis using a1 f1 f01 apply (simp add: memory_write_asi_def) apply clarsimp by (metis option.distinct(1) option.sel sys_reg_mod_privilege) qed next case False then have f2: "uint asi \ 2" by auto then show ?thesis proof (cases "uint asi \ {8,9}") case True then show ?thesis using a1 f1 f2 apply (simp add: memory_write_asi_def) using store_word_mem_privilege add_instr_cache_privilege by blast next case False then have f3: "uint asi \ {8,9}" by auto then show ?thesis proof (cases "uint asi \ {10,11}") case True then show ?thesis using a1 f1 f2 f3 apply (simp add: memory_write_asi_def) using store_word_mem_privilege add_data_cache_privilege by blast next case False then have f4: "uint asi \ {10,11}" by auto then show ?thesis proof (cases "uint asi = 13") case True then show ?thesis using a1 f1 f2 f3 f4 apply (simp add: memory_write_asi_def) by (auto simp add: add_instr_cache_privilege) next case False then have f5: "uint asi \ 13" by auto then show ?thesis proof (cases "uint asi = 15") case True then show ?thesis using a1 f1 f2 f3 f4 f5 apply (simp add: memory_write_asi_def) by (auto simp add: add_data_cache_privilege) next case False then have f6: "uint asi \ 15" by auto then show ?thesis proof (cases "uint asi = 16") case True then show ?thesis using a1 apply (simp add: memory_write_asi_def) by (auto simp add: flush_instr_cache_privilege) next case False then have f7: "uint asi \ 16" by auto then show ?thesis proof (cases "uint asi = 17") case True then show ?thesis using a1 apply (simp add: memory_write_asi_def) by (auto simp add: flush_data_cache_privilege) next case False then have f8: "uint asi \ 17" by auto then show ?thesis proof (cases "uint asi = 24") case True then show ?thesis using a1 apply (simp add: memory_write_asi_def) by (auto simp add: flush_cache_all_privilege) next case False then have f9: "uint asi \ 24" by auto then show ?thesis proof (cases "uint asi = 25") case True then show ?thesis using a1 apply (simp add: memory_write_asi_def) apply (case_tac "mmu_reg_mod (mmu s) addr data = None") apply auto by (simp add: cpu_reg_val_def) next case False then have f10: "uint asi \ 25" by auto then show ?thesis proof (cases "uint asi = 28") case True then show ?thesis using a1 apply (simp add: memory_write_asi_def) by (auto simp add: mem_mod_w32_privilege) next case False \ \The remaining cases are easy.\ then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 apply (simp add: memory_write_asi_def) apply (auto simp add: Let_def) apply (case_tac "uint asi = 20 \ uint asi = 21") by auto qed qed qed qed qed qed qed qed qed qed qed lemma memory_write_privilege: assumes a1: "r = memory_write asi addr byte_mask data (s::(('a::len) sparc_state)) \ r = Some s' \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR (s'::(('a::len) sparc_state)))))::word1) = 0" proof - have "\x. Some x \ None" by auto then have "r \ None" using a1 by (simp add: \r = memory_write asi addr byte_mask data s \ r = Some s' \ (get_S (cpu_reg_val PSR s)) = 0\) then have "\s''. r = Some (store_barrier_pending_mod False s'')" using a1 by (metis (no_types, lifting) memory_write_def option.case_eq_if) then have "\s''. s' = store_barrier_pending_mod False s''" using a1 by blast then have "\s''. memory_write_asi asi addr byte_mask data s = Some s'' \ s' = store_barrier_pending_mod False s''" by (metis (no_types, lifting) assms memory_write_def not_None_eq option.case_eq_if option.sel) then show ?thesis using a1 using memory_write_asi_privilege store_barrier_pending_mod_privilege by blast qed lemma store_sub2_privilege: assumes a1: "s' = snd (fst (store_sub2 instr curr_win rd asi address s)) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "memory_write asi address (st_byte_mask instr address) (st_data0 instr curr_win rd address s) s = None") case True then show ?thesis using a1 apply (simp add: store_sub2_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) by (metis fst_conv raise_trap_privilege return_def snd_conv) next case False then have f1: "\(memory_write asi address (st_byte_mask instr address) (st_data0 instr curr_win rd address s) s = None)" by auto then show ?thesis proof (cases "(fst instr) \ {load_store_type STD,load_store_type STDA}") case True then have f2: "(fst instr) \ {load_store_type STD,load_store_type STDA}" by auto then show ?thesis using a1 f1 apply (simp add: store_sub2_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: simpler_modify_def bind_def h1_def h2_def Let_def) apply (simp add: return_def) apply (simp add: bind_def case_prod_unfold) apply (simp add: simpler_modify_def) apply clarsimp apply (simp add: case_prod_unfold bind_def h1_def h2_def Let_def simpler_modify_def) apply (simp add: simpler_gets_def) apply auto using memory_write_privilege raise_trap_privilege apply blast apply (simp add: simpler_modify_def simpler_gets_def bind_def) apply (meson memory_write_privilege) using memory_write_privilege raise_trap_privilege apply blast by (meson memory_write_privilege) next case False then show ?thesis using a1 f1 apply (simp add: store_sub2_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply clarsimp apply (simp add: simpler_modify_def return_def) by (auto intro: memory_write_privilege) qed qed lemma store_sub1_privilege: assumes a1: "s' = snd (fst (store_sub1 instr rd s_val (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR (s'::(('a::len) sparc_state)))))::word1) = 0" proof (cases "(fst instr = load_store_type STH \ fst instr = load_store_type STHA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s)))))::word1) \ 0") case True then show ?thesis using a1 apply (simp add: store_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) using get_curr_win_privilege raise_trap_privilege by blast next case False then have f1: "\((fst instr = load_store_type STH \ fst instr = load_store_type STHA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s)))))::word1) \ 0)" by auto then show ?thesis proof (cases "(fst instr \ {load_store_type ST,load_store_type STA}) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s)))))::word2) \ 0") case True then show ?thesis using a1 f1 apply (simp add: store_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) using get_curr_win_privilege raise_trap_privilege by blast next case False then have f2: "\((fst instr \ {load_store_type ST,load_store_type STA}) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s)))))::word2) \ 0)" by auto then show ?thesis proof (cases "(fst instr \ {load_store_type STD,load_store_type STDA}) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s)))))::word3) \ 0") case True then show ?thesis using a1 f1 f2 apply (simp add: store_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) using get_curr_win_privilege raise_trap_privilege by blast next case False then show ?thesis using a1 f1 f2 apply (simp add: store_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) by (meson get_curr_win_privilege store_sub2_privilege) qed qed qed lemma store_instr_privilege: assumes a1: "s' = snd (fst (store_instr instr (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR (s'::(('a::len) sparc_state)))))::word1) = 0" using a1 apply (simp add: store_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: Let_def) using raise_trap_privilege store_sub1_privilege by blast lemma sethi_instr_privilege: assumes a1: "s' = snd (fst (sethi_instr instr (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: sethi_instr_def) apply (simp add: Let_def) apply auto apply (simp add: bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) using get_curr_win_privilege write_reg_privilege apply blast by (simp add: return_def) lemma nop_instr_privilege: assumes a1: "s' = snd (fst (nop_instr instr (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: nop_instr_def) by (simp add: return_def) lemma ucast_0: "(((get_S w))::word1) = 0 \ get_S w = 0" by simp lemma ucast_02: "get_S w = 0 \ (((get_S w))::word1) = 0" by simp lemma ucast_s: "(((get_S w))::word1) = 0 \ (AND) w (0b00000000000000000000000010000000::word32) = 0" by (simp add: get_S_def split: if_splits) lemma ucast_s2: "(AND) w 0b00000000000000000000000010000000 = 0 \ (((get_S w))::word1) = 0" by (simp add: get_S_def) lemma update_PSR_icc_1: "w' = (AND) w (0b11111111000011111111111111111111::word32) \ (((get_S w))::word1) = 0 \ (((get_S w'))::word1) = 0" by (simp add: get_S_def word_bw_assocs(1)) lemma and_num_1048576_128: "(AND) (0b00000000000100000000000000000000::word32) (0b00000000000000000000000010000000::word32) = 0" by simp lemma and_num_2097152_128: "(AND) (0b00000000001000000000000000000000::word32) (0b00000000000000000000000010000000::word32) = 0" by simp lemma and_num_4194304_128: "(AND) (0b00000000010000000000000000000000::word32) (0b00000000000000000000000010000000::word32) = 0" by simp lemma and_num_8388608_128: "(AND) (0b00000000100000000000000000000000::word32) (0b00000000000000000000000010000000::word32) = 0" by simp lemma or_and_s: "(AND) w1 (0b00000000000000000000000010000000::word32) = 0 \ (AND) w2 (0b00000000000000000000000010000000::word32) = 0 \ (AND) ((OR) w1 w2) (0b00000000000000000000000010000000::word32) = 0" by (simp add: word_ao_dist) lemma and_or_s: assumes "(((get_S w1))::word1) = 0 \ (AND) w2 (0b00000000000000000000000010000000::word32) = 0" shows "(((get_S ((OR) ((AND) w1 (0b11111111000011111111111111111111::word32)) w2)))::word1) = 0" proof - from assms have "w1 AND 128 = 0" using ucast_s by blast then have "(w1 AND 4279238655 OR w2) AND 128 = 0" using assms by (metis word_ao_absorbs(6) word_ao_dist word_bw_comms(2)) then show ?thesis using ucast_s2 by blast qed lemma and_or_or_s: assumes a1: "(((get_S w1))::word1) = 0 \ (AND) w2 (0b00000000000000000000000010000000::word32) = 0 \ (AND) w3 (0b00000000000000000000000010000000::word32) = 0" shows "(((get_S ((OR) ((OR) ((AND) w1 (0b11111111000011111111111111111111::word32)) w2) w3)))::word1) = 0" using and_or_s assms or_and_s ucast_s ucast_s2 by blast lemma and_or_or_or_s: assumes a1: "(((get_S w1))::word1) = 0 \ (AND) w2 (0b00000000000000000000000010000000::word32) = 0 \ (AND) w3 (0b00000000000000000000000010000000::word32) = 0 \ (AND) w4 (0b00000000000000000000000010000000::word32) = 0" shows "(((get_S ((OR) ((OR) ((OR) ((AND) w1 (0b11111111000011111111111111111111::word32)) w2) w3) w4)))::word1) = 0" using and_or_or_s assms or_and_s ucast_s ucast_s2 by (simp add: and_or_or_s assms or_and_s ucast_s ucast_s2) lemma and_or_or_or_or_s: assumes a1: "(((get_S w1))::word1) = 0 \ (AND) w2 (0b00000000000000000000000010000000::word32) = 0 \ (AND) w3 (0b00000000000000000000000010000000::word32) = 0 \ (AND) w4 (0b00000000000000000000000010000000::word32) = 0 \ (AND) w5 (0b00000000000000000000000010000000::word32) = 0" shows "(((get_S ((OR) ((OR) ((OR) ((OR) ((AND) w1 (0b11111111000011111111111111111111::word32)) w2) w3) w4) w5)))::word1) = 0" using and_or_or_or_s assms or_and_s ucast_s ucast_s2 by (simp add: and_or_or_or_s assms or_and_s ucast_s ucast_s2) lemma write_cpu_PSR_icc_privilege: assumes a1: "s' = snd (fst (write_cpu (update_PSR_icc n_val z_val v_val c_val (cpu_reg_val PSR s)) PSR (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: write_cpu_def) apply (simp add: simpler_modify_def) apply (simp add: cpu_reg_mod_def update_PSR_icc_def) apply (simp add: cpu_reg_val_def) apply auto using update_PSR_icc_1 apply blast using update_PSR_icc_1 and_num_1048576_128 and_or_s apply blast using update_PSR_icc_1 and_num_2097152_128 and_or_s apply blast using update_PSR_icc_1 and_num_1048576_128 and_num_2097152_128 and_or_or_s apply blast using update_PSR_icc_1 and_num_4194304_128 and_or_s apply blast using update_PSR_icc_1 and_num_1048576_128 and_num_4194304_128 and_or_or_s apply blast using update_PSR_icc_1 and_num_2097152_128 and_num_4194304_128 and_or_or_s apply blast using update_PSR_icc_1 and_num_1048576_128 and_num_2097152_128 and_num_4194304_128 and_or_or_or_s apply blast using update_PSR_icc_1 and_num_8388608_128 and_or_s apply blast using update_PSR_icc_1 and_num_1048576_128 and_num_8388608_128 and_or_or_s apply blast using update_PSR_icc_1 and_num_2097152_128 and_num_8388608_128 and_or_or_s apply blast using update_PSR_icc_1 and_num_1048576_128 and_num_2097152_128 and_num_8388608_128 and_or_or_or_s apply blast using update_PSR_icc_1 and_num_4194304_128 and_num_8388608_128 and_or_or_s apply blast using update_PSR_icc_1 and_num_1048576_128 and_num_4194304_128 and_num_8388608_128 and_or_or_or_s apply blast using update_PSR_icc_1 and_num_2097152_128 and_num_4194304_128 and_num_8388608_128 and_or_or_or_s apply blast using update_PSR_icc_1 and_num_1048576_128 and_num_2097152_128 and_num_4194304_128 and_num_8388608_128 and_or_or_or_or_s by blast lemma and_num_4294967167_128: "(AND) (0b11111111111111111111111101111111::word32) (0b00000000000000000000000010000000::word32) = 0" by simp lemma s_0_word: "(((get_S ((AND) w (0b11111111111111111111111101111111::word32))))::word1) = 0" apply (simp add: get_S_def) using and_num_4294967167_128 by (simp add: ac_simps) lemma update_PSR_CWP_1: "w' = (AND) w (0b11111111111111111111111111100000::word32) \ (((get_S w))::word1) = 0 \ (((get_S w'))::word1) = 0" by (simp add: get_S_def word_bw_assocs(1)) lemma write_cpu_PSR_CWP_privilege: assumes a1: "s' = snd (fst (write_cpu (update_CWP cwp_val (cpu_reg_val PSR s)) PSR (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: write_cpu_def) apply (simp add: simpler_modify_def) apply (simp add: cpu_reg_mod_def) apply (simp add: update_CWP_def) apply (simp add: Let_def) apply auto apply (simp add: cpu_reg_val_def) using s_0_word by blast lemma logical_instr_sub1_privilege: assumes a1: "s' = snd (fst (logical_instr_sub1 instr_name result (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "instr_name = logic_type ANDcc \ instr_name = logic_type ANDNcc \ instr_name = logic_type ORcc \ instr_name = logic_type ORNcc \ instr_name = logic_type XORcc \ instr_name = logic_type XNORcc") case True then show ?thesis using a1 apply (simp add: logical_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: logical_new_psr_val_def) using write_cpu_PSR_icc_privilege by blast next case False then show ?thesis using a1 apply (simp add: logical_instr_sub1_def) by (simp add: return_def) qed lemma logical_instr_privilege: assumes a1: "s' = snd (fst (logical_instr instr (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: logical_instr_def) apply (simp add: Let_def simpler_gets_def bind_def h1_def h2_def) apply (simp add: case_prod_unfold) apply auto apply (meson get_curr_win_privilege logical_instr_sub1_privilege write_reg_privilege) by (meson get_curr_win_privilege logical_instr_sub1_privilege write_reg_privilege) method shift_instr_privilege_proof = ( (simp add: shift_instr_def), (simp add: Let_def), (simp add: simpler_gets_def), (simp add: bind_def h1_def h2_def Let_def case_prod_unfold), auto, (blast intro: get_curr_win_privilege write_reg_privilege), (blast intro: get_curr_win_privilege write_reg_privilege) ) lemma shift_instr_privilege: assumes a1: "s' = snd (fst (shift_instr instr (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "(fst instr = shift_type SLL) \ (get_operand_w5 ((snd instr)!3) \ 0)") case True then show ?thesis using a1 by shift_instr_privilege_proof next case False then have f1: "\((fst instr = shift_type SLL) \ (get_operand_w5 ((snd instr)!3) \ 0))" by auto then show ?thesis proof (cases "(fst instr = shift_type SRL) \ (get_operand_w5 ((snd instr)!3) \ 0)") case True then show ?thesis using a1 f1 by shift_instr_privilege_proof next case False then have f2: "\((fst instr = shift_type SRL) \ (get_operand_w5 ((snd instr)!3) \ 0))" by auto then show ?thesis proof (cases "(fst instr = shift_type SRA) \ (get_operand_w5 ((snd instr)!3) \ 0)") case True then show ?thesis using a1 f1 f2 by shift_instr_privilege_proof next case False then show ?thesis using a1 f1 f2 apply (simp add: shift_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def) apply (simp add: bind_def h1_def h2_def Let_def case_prod_unfold) apply (simp add: return_def) using get_curr_win_privilege by blast qed qed qed lemma add_instr_sub1_privilege: assumes a1: "s' = snd (fst (add_instr_sub1 instr_name result rs1_val operand2 (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "instr_name = arith_type ADDcc \ instr_name = arith_type ADDXcc") case True then show ?thesis using a1 apply (simp add: add_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (blast intro: write_cpu_PSR_icc_privilege) next case False then show ?thesis using a1 apply (simp add: add_instr_sub1_def) by (simp add: return_def) qed lemma add_instr_privilege: assumes a1: "s' = snd (fst (add_instr instr (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: add_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) by (meson add_instr_sub1_privilege get_curr_win_privilege write_reg_privilege) lemma sub_instr_sub1_privilege: assumes a1: "s' = snd (fst (sub_instr_sub1 instr_name result rs1_val operand2 (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "instr_name = arith_type SUBcc \ instr_name = arith_type SUBXcc") case True then show ?thesis using a1 apply (simp add: sub_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (blast intro: write_cpu_PSR_icc_privilege) next case False then show ?thesis using a1 apply (simp add: sub_instr_sub1_def) by (simp add: return_def) qed lemma sub_instr_privilege: assumes a1: "s' = snd (fst (sub_instr instr (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: sub_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) by (meson sub_instr_sub1_privilege get_curr_win_privilege write_reg_privilege) lemma mul_instr_sub1_privilege: assumes a1: "s' = snd (fst (mul_instr_sub1 instr_name result (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "instr_name \ {arith_type SMULcc,arith_type UMULcc}") case True then show ?thesis using a1 apply (simp add: mul_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (blast intro: write_cpu_PSR_icc_privilege) next case False then show ?thesis using a1 apply (simp add: mul_instr_sub1_def) by (simp add: return_def) qed lemma mul_instr_privilege: assumes a1: "s' = snd (fst (mul_instr instr (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: mul_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) by (meson get_curr_win_privilege mul_instr_sub1_privilege write_cpu_y_privilege write_reg_privilege) lemma div_write_new_val_privilege: assumes a1: "s' = snd (fst (div_write_new_val i result temp_V (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "(fst i) \ {arith_type UDIVcc,arith_type SDIVcc}") case True then show ?thesis using a1 apply (simp add: div_write_new_val_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (blast intro: write_cpu_PSR_icc_privilege) next case False then show ?thesis using a1 apply (simp add: div_write_new_val_def) by (simp add: return_def) qed lemma div_comp_privilege: assumes a1: "s' = snd (fst (div_comp instr rs1 rd operand2 (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: div_comp_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) by (meson get_curr_win_privilege div_write_new_val_privilege write_reg_privilege) lemma div_instr_privilege: assumes a1: "s' = snd (fst (div_instr instr (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: div_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: return_def) apply auto using raise_trap_privilege apply blast using div_comp_privilege by blast lemma save_retore_sub1_privilege: assumes a1: "s' = snd (fst (save_retore_sub1 result new_cwp rd (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: save_retore_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) using write_cpu_PSR_CWP_privilege write_reg_privilege by blast method save_restore_instr_privilege_proof = ( (simp add: save_restore_instr_def), (simp add: Let_def), (simp add: simpler_gets_def bind_def h1_def h2_def Let_def), (simp add: case_prod_unfold), auto, (blast intro: get_curr_win_privilege raise_trap_privilege), (simp add: simpler_gets_def bind_def h1_def h2_def Let_def case_prod_unfold), (blast intro: get_curr_win_privilege save_retore_sub1_privilege) ) lemma save_restore_instr_privilege: assumes a1: "s' = snd (fst (save_restore_instr instr (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "fst instr = ctrl_type SAVE") case True then have f1: "fst instr = ctrl_type SAVE" by auto then show ?thesis using a1 by save_restore_instr_privilege_proof next case False then show ?thesis using a1 by save_restore_instr_privilege_proof qed lemma call_instr_privilege: assumes a1: "s' = snd (fst (call_instr instr (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: call_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) by (meson get_curr_win_privilege write_cpu_npc_privilege write_cpu_pc_privilege write_reg_privilege) lemma jmpl_instr_privilege: assumes a1: "s' = snd (fst (jmpl_instr instr (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: jmpl_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply auto using get_curr_win_privilege raise_trap_privilege apply blast apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) by (meson get_curr_win_privilege write_cpu_npc_privilege write_cpu_pc_privilege write_reg_privilege) lemma rett_instr_privilege: assumes a1: "snd (rett_instr i s) = False \ s' = snd (fst (rett_instr instr (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: rett_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply auto apply (simp add: case_prod_unfold) apply (simp add: return_def) apply (blast intro: raise_trap_privilege) apply (simp add: bind_def h1_def h2_def Let_def) by (simp add: case_prod_unfold fail_def) method read_state_reg_instr_privilege_proof = ( (simp add: read_state_reg_instr_def), (simp add: Let_def), (simp add: simpler_gets_def bind_def h1_def h2_def Let_def), (simp add: case_prod_unfold) ) lemma read_state_reg_instr_privilege: assumes a1: "s' = snd (fst (read_state_reg_instr instr (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "(fst instr \ {sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR} \ (fst instr = sreg_type RDASR \ privileged_ASR (get_operand_w5 ((snd instr)!0))))") case True then have "(fst instr \ {sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR} \ (fst instr = sreg_type RDASR \ privileged_ASR (get_operand_w5 ((snd instr)!0)))) \ (((get_S (cpu_reg_val PSR (snd (fst (get_curr_win () s))))))::word1) = 0" by (metis assms get_curr_win_privilege) then show ?thesis using a1 apply read_state_reg_instr_privilege_proof by (blast intro: raise_trap_privilege get_curr_win_privilege) next case False then have f1: "\((fst instr = sreg_type RDPSR \ fst instr = sreg_type RDWIM \ fst instr = sreg_type RDTBR \ fst instr = sreg_type RDASR \ privileged_ASR (get_operand_w5 (snd instr ! 0))) \ (get_S (cpu_reg_val PSR (snd (fst (get_curr_win () s))))) = 0)" by blast then show ?thesis proof (cases "illegal_instruction_ASR (get_operand_w5 ((snd instr)!0))") case True then show ?thesis using a1 f1 apply read_state_reg_instr_privilege_proof by (simp add: illegal_instruction_ASR_def) next case False then have f2: "\(illegal_instruction_ASR (get_operand_w5 ((snd instr)!0)))" by auto then show ?thesis proof (cases "(get_operand_w5 ((snd instr)!1)) \ 0") case True then have f3: "(get_operand_w5 ((snd instr)!1)) \ 0" by auto then show ?thesis proof (cases "fst instr = sreg_type RDY") case True then show ?thesis using a1 f1 f2 f3 apply (simp add: read_state_reg_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) by (blast intro: get_curr_win_privilege write_reg_privilege) next case False then have f4: "\(fst instr = sreg_type RDY)" by auto then show ?thesis proof (cases "fst instr = sreg_type RDASR") case True then show ?thesis using a1 f1 f2 f3 f4 apply read_state_reg_instr_privilege_proof apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (blast intro: get_curr_win_privilege write_reg_privilege) next case False then have f5: "\(fst instr = sreg_type RDASR)" by auto then show ?thesis proof (cases "fst instr = sreg_type RDPSR") case True then show ?thesis using a1 f1 f2 f3 f4 f5 apply read_state_reg_instr_privilege_proof by (blast intro: get_curr_win_privilege write_reg_privilege) next case False then show ?thesis using a1 f1 f2 f3 f4 f5 apply read_state_reg_instr_privilege_proof apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (blast intro: get_curr_win_privilege write_reg_privilege) qed qed qed next case False then show ?thesis using a1 apply read_state_reg_instr_privilege_proof apply (simp add: return_def) using f1 f2 get_curr_win_privilege by blast qed qed qed method write_state_reg_instr_privilege_proof = ( (simp add: write_state_reg_instr_def), (simp add: Let_def), (simp add: simpler_gets_def bind_def h1_def h2_def Let_def), (simp add: case_prod_unfold) ) lemma write_state_reg_instr_privilege: assumes a1: "s' = snd (fst (write_state_reg_instr instr (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "fst instr = sreg_type WRY") case True then show ?thesis using a1 apply write_state_reg_instr_privilege_proof apply (simp add: simpler_modify_def) apply (simp add: delayed_pool_add_def DELAYNUM_def) by (blast intro: cpu_reg_mod_y_privilege get_curr_win_privilege) next case False then have f1: "\(fst instr = sreg_type WRY)" by auto then show ?thesis proof (cases "fst instr = sreg_type WRASR") case True then show ?thesis using a1 f1 apply write_state_reg_instr_privilege_proof apply (simp add: simpler_modify_def) apply auto using illegal_instruction_ASR_def apply blast using illegal_instruction_ASR_def apply blast using illegal_instruction_ASR_def apply blast using raise_trap_privilege get_curr_win_privilege apply blast apply (simp add: simpler_modify_def delayed_pool_add_def DELAYNUM_def) using cpu_reg_mod_asr_privilege get_curr_win_privilege apply blast apply (simp add: simpler_modify_def delayed_pool_add_def DELAYNUM_def) using cpu_reg_mod_asr_privilege get_curr_win_privilege by blast next case False then have f2: "\(fst instr = sreg_type WRASR)" by auto have f3: "get_S (cpu_reg_val PSR (snd (fst (get_curr_win () s)))) = 0" using get_curr_win_privilege a1 by (metis ucast_id) then show ?thesis proof (cases "fst instr = sreg_type WRPSR") case True then show ?thesis using a1 f1 f2 f3 apply write_state_reg_instr_privilege_proof by (metis raise_trap_privilege ucast_0) next case False then have f4: "\(fst instr = sreg_type WRPSR)" by auto then show ?thesis proof (cases "fst instr = sreg_type WRWIM") case True then show ?thesis using a1 f1 f2 f3 f4 apply write_state_reg_instr_privilege_proof by (metis raise_trap_privilege ucast_0) next case False then have f5: "\(fst instr = sreg_type WRWIM)" by auto then show ?thesis using a1 f1 f2 f3 f4 f5 apply write_state_reg_instr_privilege_proof by (metis raise_trap_privilege ucast_0) qed qed qed qed lemma flush_instr_privilege: assumes a1: "s' = snd (fst (flush_instr instr (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: flush_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def simpler_modify_def) by (auto simp add: flush_cache_all_privilege) lemma branch_instr_privilege: assumes a1: "s' = snd (fst (branch_instr instr (s::(('a::len) sparc_state)))) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" using a1 apply (simp add: branch_instr_def) apply (simp add: Let_def simpler_gets_def bind_def h1_def h2_def) apply (simp add: case_prod_unfold return_def) by (meson set_annul_privilege write_cpu_npc_privilege write_cpu_pc_privilege) method dispath_instr_privilege_proof = ( (simp add: dispatch_instruction_def), (simp add: simpler_gets_def bind_def h1_def h2_def Let_def), (simp add: Let_def) ) lemma dispath_instr_privilege: assumes a1: "snd (dispatch_instruction instr s) = False \ s' = snd (fst (dispatch_instruction instr s)) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "get_trap_set s = {}") case True then have f1: "get_trap_set s = {}" by auto show ?thesis proof (cases "fst instr \ {load_store_type LDSB,load_store_type LDUB, load_store_type LDUBA,load_store_type LDUH,load_store_type LD, load_store_type LDA,load_store_type LDD}") case True then show ?thesis using a1 f1 apply dispath_instr_privilege_proof by (blast intro: load_instr_privilege) next case False then have f2: "\(fst instr \ {load_store_type LDSB,load_store_type LDUB, load_store_type LDUBA,load_store_type LDUH,load_store_type LD, load_store_type LDA,load_store_type LDD})" by auto then show ?thesis proof (cases "fst instr \ {load_store_type STB,load_store_type STH, load_store_type ST,load_store_type STA,load_store_type STD}") case True then show ?thesis using a1 f1 f2 apply dispath_instr_privilege_proof by (blast intro: store_instr_privilege) next case False then have f3: "\(fst instr \ {load_store_type STB,load_store_type STH, load_store_type ST,load_store_type STA,load_store_type STD})" by auto then show ?thesis proof (cases "fst instr \ {sethi_type SETHI}") case True then show ?thesis using a1 f1 f2 f3 apply dispath_instr_privilege_proof by (blast intro: sethi_instr_privilege) next case False then have f4: "\(fst instr \ {sethi_type SETHI})" by auto then show ?thesis proof (cases "fst instr \ {nop_type NOP}") case True then show ?thesis using a1 f1 f2 f3 f4 apply dispath_instr_privilege_proof by (blast intro: nop_instr_privilege) next case False then have f5: "\(fst instr \ {nop_type NOP})" by auto then show ?thesis proof (cases "fst instr \ {logic_type ANDs,logic_type ANDcc,logic_type ANDN, logic_type ANDNcc,logic_type ORs,logic_type ORcc,logic_type ORN, logic_type XORs,logic_type XNOR}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 apply dispath_instr_privilege_proof by (blast intro: logical_instr_privilege) next case False then have f6: "\(fst instr \ {logic_type ANDs,logic_type ANDcc,logic_type ANDN, logic_type ANDNcc,logic_type ORs,logic_type ORcc,logic_type ORN, logic_type XORs,logic_type XNOR})" by auto show ?thesis proof (cases "fst instr \ {shift_type SLL,shift_type SRL,shift_type SRA}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 f6 apply dispath_instr_privilege_proof by (blast intro: shift_instr_privilege) next case False then have f7: "\(fst instr \ {shift_type SLL,shift_type SRL,shift_type SRA})" by auto then show ?thesis proof (cases "fst instr \ {arith_type ADD,arith_type ADDcc,arith_type ADDX}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 apply dispath_instr_privilege_proof by (blast intro: add_instr_privilege) next case False then have f8: "\(fst instr \ {arith_type ADD,arith_type ADDcc,arith_type ADDX})" by auto then show ?thesis proof (cases "fst instr \ {arith_type SUB,arith_type SUBcc,arith_type SUBX}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 f8 apply dispath_instr_privilege_proof by (blast intro: sub_instr_privilege) next case False then have f9: "\(fst instr \ {arith_type SUB,arith_type SUBcc,arith_type SUBX})" by auto then show ?thesis proof (cases "fst instr \ {arith_type UMUL,arith_type SMUL,arith_type SMULcc}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 f8 f9 apply dispath_instr_privilege_proof by (blast intro: mul_instr_privilege) next case False then have f10: "\(fst instr \ {arith_type UMUL,arith_type SMUL, arith_type SMULcc})" by auto then show ?thesis proof (cases "fst instr \ {arith_type UDIV,arith_type UDIVcc,arith_type SDIV}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 apply dispath_instr_privilege_proof by (blast intro: div_instr_privilege) next case False then have f11: "\(fst instr \ {arith_type UDIV, arith_type UDIVcc,arith_type SDIV})" by auto then show ?thesis proof (cases "fst instr \ {ctrl_type SAVE,ctrl_type RESTORE}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 apply dispath_instr_privilege_proof by (blast intro: save_restore_instr_privilege) next case False then have f12: "\(fst instr \ {ctrl_type SAVE,ctrl_type RESTORE})" by auto then show ?thesis proof (cases "fst instr \ {call_type CALL}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 apply dispath_instr_privilege_proof by (blast intro: call_instr_privilege) next case False then have f13: "\(fst instr \ {call_type CALL})" by auto then show ?thesis proof (cases "fst instr \ {ctrl_type JMPL}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 apply dispath_instr_privilege_proof by (blast intro: jmpl_instr_privilege) next case False then have f14: "\(fst instr \ {ctrl_type JMPL})" by auto then show ?thesis proof (cases "fst instr \ {ctrl_type RETT}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 apply dispath_instr_privilege_proof by (blast intro: rett_instr_privilege) next case False then have f15: "\(fst instr \ {ctrl_type RETT})" by auto then show ?thesis proof (cases "fst instr \ {sreg_type RDY,sreg_type RDPSR, sreg_type RDWIM, sreg_type RDTBR}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 apply dispath_instr_privilege_proof by (blast intro: read_state_reg_instr_privilege) next case False then have f16: "\(fst instr \ {sreg_type RDY,sreg_type RDPSR, sreg_type RDWIM, sreg_type RDTBR})" by auto then show ?thesis proof (cases "fst instr \ {sreg_type WRY,sreg_type WRPSR, sreg_type WRWIM, sreg_type WRTBR}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 apply dispath_instr_privilege_proof by (blast intro: write_state_reg_instr_privilege) next case False then have f17: "\(fst instr \ {sreg_type WRY,sreg_type WRPSR, sreg_type WRWIM, sreg_type WRTBR})" by auto then show ?thesis proof (cases "fst instr \ {load_store_type FLUSH}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 apply dispath_instr_privilege_proof by (blast intro: flush_instr_privilege) next case False then have f18: "\(fst instr \ {load_store_type FLUSH})" by auto then show ?thesis proof (cases "fst instr \ {bicc_type BE,bicc_type BNE, bicc_type BGU,bicc_type BLE,bicc_type BL,bicc_type BGE, bicc_type BNEG,bicc_type BG,bicc_type BCS,bicc_type BLEU, bicc_type BCC,bicc_type BA,bicc_type BN}") case True then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 apply dispath_instr_privilege_proof by (blast intro: branch_instr_privilege) next case False then show ?thesis using a1 f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 f18 apply dispath_instr_privilege_proof by (simp add: fail_def) qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed next case False then show ?thesis using a1 apply (simp add: dispatch_instruction_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: Let_def) by (simp add: return_def) qed lemma execute_instr_sub1_privilege: assumes a1: "snd (execute_instr_sub1 i s) = False \ s' = snd (fst (execute_instr_sub1 i s)) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "get_trap_set s = {} \ fst i \ {call_type CALL,ctrl_type RETT,ctrl_type JMPL, bicc_type BE,bicc_type BNE,bicc_type BGU, bicc_type BLE,bicc_type BL,bicc_type BGE, bicc_type BNEG,bicc_type BG, bicc_type BCS,bicc_type BLEU,bicc_type BCC, bicc_type BA,bicc_type BN}") case True then show ?thesis using a1 apply (simp add: execute_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold return_def) by (auto intro: write_cpu_pc_privilege write_cpu_npc_privilege) next case False then show ?thesis using a1 apply (simp add: execute_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold return_def) by auto qed text \ Assume that there is no \delayed_write\ and there is no traps to be executed. If an instruction is executed as a user, the privilege will not be changed to supervisor after the execution.\ theorem safe_privilege : assumes a1: "get_delayed_pool s = [] \ get_trap_set s = {} \ snd (execute_instruction() s) = False \ s' = snd (fst (execute_instruction() s)) \ (((get_S (cpu_reg_val PSR s)))::word1) = 0" shows "(((get_S (cpu_reg_val PSR s')))::word1) = 0" proof (cases "exe_mode_val s") case True then have f2: "exe_mode_val s = True" by auto then show ?thesis proof (cases "\e. fetch_instruction (delayed_pool_write s) = Inl e") case True then have f3: "\e. fetch_instruction (delayed_pool_write s) = Inl e" by auto then have f4: "\ (\v. fetch_instruction (delayed_pool_write s) = Inr v)" using fetch_instr_result_3 by auto then show ?thesis using a1 f2 f3 raise_trap_result empty_delayed_pool_write_privilege raise_trap_privilege apply (simp add: execute_instruction_def) apply (simp add: exec_gets return_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: simpler_modify_def) apply clarsimp apply (simp add: case_prod_unfold) by (blast intro: empty_delayed_pool_write_privilege raise_trap_privilege) next case False then have f5: "\v. fetch_instruction (delayed_pool_write s) = Inr v" using fetch_instr_result_1 by blast then have f6: "\v. fetch_instruction (delayed_pool_write s) = Inr v \ \ (\e. ((decode_instruction v)::(Exception list + instruction)) = Inl e)" using a1 f2 dispatch_fail by blast then have f7: "\v. fetch_instruction (delayed_pool_write s) = Inr v \ (\v1. ((decode_instruction v)::(Exception list + instruction)) = Inr v1)" using decode_instr_result_4 by auto then show ?thesis proof (cases "annul_val (delayed_pool_write s)") case True then show ?thesis using a1 f2 f7 apply (simp add: execute_instruction_def) apply (simp add: exec_gets return_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: simpler_modify_def) apply clarsimp apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) by (auto intro: empty_delayed_pool_write_privilege set_annul_privilege write_cpu_npc_privilege write_cpu_pc_privilege) next case False then show ?thesis using a1 f2 f7 apply (simp add: execute_instruction_def) apply (simp add: exec_gets return_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: simpler_modify_def) apply clarsimp apply (simp add: bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: simpler_modify_def return_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) by (auto intro: empty_delayed_pool_write_privilege dispath_instr_privilege execute_instr_sub1_privilege) qed qed next case False then show ?thesis using a1 apply (simp add: execute_instruction_def) by (simp add: simpler_gets_def bind_def h1_def h2_def Let_def return_def) qed (*********************************************************************) section \Single step non-interference property.\ (*********************************************************************) definition user_accessible:: "('a::len) sparc_state \ phys_address \ bool" where "user_accessible s pa \ \va p. (virt_to_phys va (mmu s) (mem s)) = Some p \ mmu_readable (get_acc_flag (snd p)) 10 \ (fst p) = pa" \ \Passing \asi = 8\ is the same.\ lemma user_accessible_8: assumes a1: "mmu_readable (get_acc_flag (snd p)) 8" shows "mmu_readable (get_acc_flag (snd p)) 10" using a1 by (simp add: mmu_readable_def) definition mem_equal:: "('a) sparc_state \ ('a) sparc_state \ phys_address \ bool" where "mem_equal s1 s2 pa \ (mem s1) 8 (pa AND 68719476732) = (mem s2) 8 (pa AND 68719476732) \ (mem s1) 8 ((pa AND 68719476732) + 1) = (mem s2) 8 ((pa AND 68719476732) + 1) \ (mem s1) 8 ((pa AND 68719476732) + 2) = (mem s2) 8 ((pa AND 68719476732) + 2) \ (mem s1) 8 ((pa AND 68719476732) + 3) = (mem s2) 8 ((pa AND 68719476732) + 3) \ (mem s1) 9 (pa AND 68719476732) = (mem s2) 9 (pa AND 68719476732) \ (mem s1) 9 ((pa AND 68719476732) + 1) = (mem s2) 9 ((pa AND 68719476732) + 1) \ (mem s1) 9 ((pa AND 68719476732) + 2) = (mem s2) 9 ((pa AND 68719476732) + 2) \ (mem s1) 9 ((pa AND 68719476732) + 3) = (mem s2) 9 ((pa AND 68719476732) + 3) \ (mem s1) 10 (pa AND 68719476732) = (mem s2) 10 (pa AND 68719476732) \ (mem s1) 10 ((pa AND 68719476732) + 1) = (mem s2) 10 ((pa AND 68719476732) + 1) \ (mem s1) 10 ((pa AND 68719476732) + 2) = (mem s2) 10 ((pa AND 68719476732) + 2) \ (mem s1) 10 ((pa AND 68719476732) + 3) = (mem s2) 10 ((pa AND 68719476732) + 3) \ (mem s1) 11 (pa AND 68719476732) = (mem s2) 11 (pa AND 68719476732) \ (mem s1) 11 ((pa AND 68719476732) + 1) = (mem s2) 11 ((pa AND 68719476732) + 1) \ (mem s1) 11 ((pa AND 68719476732) + 2) = (mem s2) 11 ((pa AND 68719476732) + 2) \ (mem s1) 11 ((pa AND 68719476732) + 3) = (mem s2) 11 ((pa AND 68719476732) + 3)" text \\low_equal\ defines the equivalence relation over two sparc states that is an analogy to the \=\<^sub>L\ relation over memory contexts in the traditional non-interference theorem.\ definition low_equal:: "('a::len) sparc_state \ ('a) sparc_state \ bool" where "low_equal s1 s2 \ (cpu_reg s1) = (cpu_reg s2) \ (user_reg s1) = (user_reg s2) \ (sys_reg s1) = (sys_reg s2) \ (\va. (virt_to_phys va (mmu s1) (mem s1)) = (virt_to_phys va (mmu s2) (mem s2))) \ (\pa. (user_accessible s1 pa) \ mem_equal s1 s2 pa) \ (mmu s1) = (mmu s2) \ (state_var s1) = (state_var s2) \ (traps s1) = (traps s2) \ (undef s1) = (undef s2) " lemma low_equal_com: "low_equal s1 s2 \ low_equal s2 s1" apply (simp add: low_equal_def) apply (simp add: mem_equal_def user_accessible_def) by metis lemma non_exe_mode_equal: "exe_mode_val s = False \ get_trap_set s = {} \ Some t = NEXT s \ t = s" apply (simp add: NEXT_def execute_instruction_def) apply auto by (simp add: simpler_gets_def bind_def h1_def h2_def Let_def return_def) lemma exe_mode_low_equal: assumes a1: "low_equal s1 s2" shows " exe_mode_val s1 = exe_mode_val s2" using a1 apply (simp add: low_equal_def) by (simp add: exe_mode_val_def) lemma mem_val_mod_state: "mem_val_alt asi a s = mem_val_alt asi a (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\)" apply (simp add: mem_val_alt_def) by (simp add: Let_def) lemma mem_val_w32_mod_state: "mem_val_w32 asi a s = mem_val_w32 asi a (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\)" apply (simp add: mem_val_w32_def) apply (simp add: Let_def) by (metis mem_val_mod_state) lemma load_word_mem_mod_state: "load_word_mem s addr asi = load_word_mem (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr asi" apply (simp add: load_word_mem_def) apply (case_tac "virt_to_phys addr (mmu s) (mem s) = None") apply auto by (auto simp add: mem_val_w32_mod_state) lemma load_word_mem2_mod_state: "fst (case load_word_mem s addr asi of None \ (None, s) | Some w \ (Some w, add_data_cache s addr w 15)) = fst (case load_word_mem (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr asi of None \ (None, (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\)) | Some w \ (Some w, add_data_cache (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr w 15))" proof (cases "load_word_mem s addr asi = None") case True then have "load_word_mem s addr asi = None \ load_word_mem (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr asi = None" using load_word_mem_mod_state by metis then show ?thesis by auto next case False then have "\w. load_word_mem s addr asi = Some w" by auto then have "\w. load_word_mem s addr asi = Some w \ load_word_mem (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr asi = Some w" using load_word_mem_mod_state by metis then show ?thesis by auto qed lemma load_word_mem3_mod_state: "fst (case load_word_mem s addr asi of None \ (None, s) | Some w \ (Some w, add_instr_cache s addr w 15)) = fst (case load_word_mem (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr asi of None \ (None, (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\)) | Some w \ (Some w, add_instr_cache (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr w 15))" proof (cases "load_word_mem s addr asi = None") case True then have "load_word_mem s addr asi = None \ load_word_mem (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr asi = None" using load_word_mem_mod_state by metis then show ?thesis by auto next case False then have "\w. load_word_mem s addr asi = Some w" by auto then have "\w. load_word_mem s addr asi = Some w \ load_word_mem (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr asi = Some w" using load_word_mem_mod_state by metis then show ?thesis by auto qed lemma read_dcache_mod_state: "read_data_cache s addr = read_data_cache (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr" apply (simp add: read_data_cache_def) by (simp add: dcache_val_def) lemma read_dcache2_mod_state: "fst (case read_data_cache s addr of None \ (None, s) | Some w \ (Some w, s)) = fst (case read_data_cache (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr of None \ (None, (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\)) | Some w \ (Some w, (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\)))" proof (cases "read_data_cache s addr = None") case True then have "read_data_cache s addr = None \ read_data_cache (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr = None" using read_dcache_mod_state by metis then show ?thesis by auto next case False then have "\w. read_data_cache s addr = Some w" by auto then have "\w. read_data_cache s addr = Some w \ read_data_cache (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr = Some w" using read_dcache_mod_state by metis then show ?thesis by auto qed lemma read_icache_mod_state: "read_instr_cache s addr = read_instr_cache (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr" apply (simp add: read_instr_cache_def) by (simp add: icache_val_def) lemma read_icache2_mod_state: "fst (case read_instr_cache s addr of None \ (None, s) | Some w \ (Some w, s)) = fst (case read_instr_cache (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr of None \ (None, (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\)) | Some w \ (Some w, (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\)))" proof (cases "read_instr_cache s addr = None") case True then have "read_instr_cache s addr = None \ read_instr_cache (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr = None" using read_icache_mod_state by metis then show ?thesis by auto next case False then have "\w. read_instr_cache s addr = Some w" by auto then have "\w. read_instr_cache s addr = Some w \ read_instr_cache (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) addr = Some w" using read_icache_mod_state by metis then show ?thesis by auto qed lemma mem_read_mod_state: "fst (memory_read asi addr s) = fst (memory_read asi addr (s\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\))" apply (simp add: memory_read_def) apply (case_tac "uint asi = 1") apply (simp add: Let_def) apply (metis load_word_mem_mod_state option.distinct(1)) apply (case_tac "uint asi = 2") apply (simp add: Let_def) apply (simp add: sys_reg_val_def) apply (case_tac "uint asi \ {8,9}") apply (simp add: Let_def) apply (simp add: load_word_mem3_mod_state) apply (simp add: load_word_mem_mod_state) apply (case_tac "uint asi \ {10,11}") apply (simp add: Let_def) apply (simp add: load_word_mem2_mod_state) apply (simp add: load_word_mem_mod_state) apply (case_tac "uint asi = 13") apply (simp add: Let_def) apply (simp add: read_icache2_mod_state) apply (case_tac "uint asi = 15") apply (simp add: Let_def) apply (simp add: read_dcache2_mod_state) apply (case_tac "uint asi = 25") apply (simp add: Let_def) apply (case_tac "uint asi = 28") apply (simp add: Let_def) apply (simp add: mem_val_w32_mod_state) by (simp add: Let_def) lemma insert_trap_mem: "fst (memory_read asi addr s) = fst (memory_read asi addr (s\traps := new_traps\))" proof - have "fst (memory_read asi addr s) = fst (memory_read asi addr (s\cpu_reg := (cpu_reg s), user_reg := (user_reg s), dwrite := (dwrite s), state_var := (state_var s), traps := new_traps, undef := (undef s)\))" using mem_read_mod_state by blast then show ?thesis by auto qed lemma cpu_reg_mod_mem: "fst (memory_read asi addr s) = fst (memory_read asi addr (s\cpu_reg := new_cpu_reg\))" proof - have "fst (memory_read asi addr s) = fst (memory_read asi addr (s\cpu_reg := new_cpu_reg, user_reg := (user_reg s), dwrite := (dwrite s), state_var := (state_var s), traps := (traps s), undef := (undef s)\))" using mem_read_mod_state by blast then show ?thesis by auto qed lemma user_reg_mod_mem: "fst (memory_read asi addr s) = fst (memory_read asi addr (s\user_reg := new_user_reg\))" proof - have "fst (memory_read asi addr s) = fst (memory_read asi addr (s\cpu_reg := (cpu_reg s), user_reg := new_user_reg, dwrite := (dwrite s), state_var := (state_var s), traps := (traps s), undef := (undef s)\))" using mem_read_mod_state by blast then show ?thesis by auto qed lemma annul_mem: "fst (memory_read asi addr s) = fst (memory_read asi addr (s\state_var := new_state_var, cpu_reg := new_cpu_reg\))" proof - have "fst (memory_read asi addr s) = fst (memory_read asi addr (s\cpu_reg := new_cpu_reg, user_reg := (user_reg s), dwrite := (dwrite s), state_var := new_state_var, traps := (traps s), undef := (undef s)\))" using mem_read_mod_state by blast then have "fst (memory_read asi addr s) = fst (memory_read asi addr (s\cpu_reg := new_cpu_reg, state_var := new_state_var\))" by auto then show ?thesis by (metis Sparc_State.sparc_state.surjective Sparc_State.sparc_state.update_convs(1) Sparc_State.sparc_state.update_convs(8)) qed lemma state_var_mod_mem: "fst (memory_read asi addr s) = fst (memory_read asi addr (s\state_var := new_state_var\))" proof - have "fst (memory_read asi addr s) = fst (memory_read asi addr (s\cpu_reg := (cpu_reg s), user_reg := (user_reg s), dwrite := (dwrite s), state_var := new_state_var, traps := (traps s), undef := (undef s)\))" using mem_read_mod_state by blast then show ?thesis by auto qed lemma mod_state_low_equal: "low_equal s1 s2 \ t1 = (s1\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) \ t2 = (s2\cpu_reg := new_cpu_reg, user_reg := new_user_reg, dwrite := new_dwrite, state_var := new_state_var, traps := new_traps, undef := new_undef\) \ low_equal t1 t2" apply (simp add: low_equal_def) apply clarsimp apply (simp add: mem_equal_def) by (simp add: user_accessible_def) lemma user_reg_state_mod_low_equal: assumes a1: "low_equal s1 s2 \ t1 = (s1\user_reg := new_user_reg\) \ t2 = (s2\user_reg := new_user_reg\)" shows "low_equal t1 t2" proof - have "low_equal s1 s2 \ t1 = (s1\cpu_reg := (cpu_reg s1), user_reg := new_user_reg, dwrite := (dwrite s1), state_var := (state_var s1), traps := (traps s1), undef := (undef s1)\) \ t2 = (s2\cpu_reg := (cpu_reg s2), user_reg := new_user_reg, dwrite := (dwrite s2), state_var := (state_var s2), traps := (traps s2), undef := (undef s2)\) \ low_equal t1 t2" using mod_state_low_equal apply (simp add: low_equal_def) apply (simp add: user_accessible_def mem_equal_def) by clarsimp then show ?thesis using a1 by clarsimp qed lemma mod_trap_low_equal: assumes a1: "low_equal s1 s2 \ t1 = (s1\traps := new_traps\) \ t2 = (s2\traps := new_traps\)" shows "low_equal t1 t2" proof - have "low_equal s1 s2 \ t1 = (s1\cpu_reg := (cpu_reg s1), user_reg := (user_reg s1), dwrite := (dwrite s1), state_var := (state_var s1), traps := new_traps, undef := (undef s1)\) \ t2 = (s2\cpu_reg := (cpu_reg s2), user_reg := (user_reg s2), dwrite := (dwrite s2), state_var := (state_var s2), traps := new_traps, undef := (undef s2)\) \ low_equal t1 t2" using mod_state_low_equal apply (simp add: low_equal_def) apply (simp add: user_accessible_def mem_equal_def) by clarsimp then show ?thesis using a1 by clarsimp qed lemma state_var_low_equal: "low_equal s1 s2 \ state_var s1 = state_var s2" by (simp add: low_equal_def) lemma state_var2_low_equal: assumes a1: "low_equal s1 s2 \ t1 = (s1\state_var := new_state_var\) \ t2 = (s2\state_var := new_state_var\)" shows "low_equal t1 t2" proof - have "low_equal s1 s2 \ t1 = (s1\cpu_reg := (cpu_reg s1), user_reg := (user_reg s1), dwrite := (dwrite s1), state_var := new_state_var, traps := (traps s1), undef := (undef s1)\) \ t2 = (s2\cpu_reg := (cpu_reg s2), user_reg := (user_reg s2), dwrite := (dwrite s2), state_var := new_state_var, traps := (traps s2), undef := (undef s2)\) \ low_equal t1 t2" using mod_state_low_equal apply (simp add: low_equal_def) apply (simp add: user_accessible_def mem_equal_def) by clarsimp then show ?thesis using a1 by clarsimp qed lemma traps_low_equal: "low_equal s1 s2 \ traps s1 = traps s2" by (simp add: low_equal_def) lemma s_low_equal: "low_equal s1 s2 \ (get_S (cpu_reg_val PSR s1)) = (get_S (cpu_reg_val PSR s2))" by (simp add: low_equal_def cpu_reg_val_def) lemma cpu_reg_val_low_equal: "low_equal s1 s2 \ (cpu_reg_val cr s1) = (cpu_reg_val cr s2)" by (simp add: cpu_reg_val_def low_equal_def) lemma get_curr_win_low_equal: "low_equal s1 s2 \ (fst (fst (get_curr_win () s1))) = (fst (fst (get_curr_win () s2)))" apply (simp add: low_equal_def) apply (simp add: get_curr_win_def cpu_reg_val_def get_CWP_def) by (simp add: simpler_gets_def) lemma get_curr_win2_low_equal: "low_equal s1 s2 \ t1 = (snd (fst (get_curr_win () s1))) \ t2 = (snd (fst (get_curr_win () s2))) \ low_equal t1 t2" apply (simp add: low_equal_def) apply (simp add: get_curr_win_def cpu_reg_val_def get_CWP_def) by (auto simp add: simpler_gets_def) lemma get_curr_win3_low_equal: "low_equal s1 s2 \ (traps (snd (fst (get_curr_win () s1)))) = (traps (snd (fst (get_curr_win () s2))))" using low_equal_def get_curr_win2_low_equal by blast lemma get_addr_low_equal: "low_equal s1 s2 \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word3) = ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word3) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word2) = ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word2) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word1) = ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word1)" apply (simp add: low_equal_def) apply (simp add: get_curr_win_def cpu_reg_val_def get_CWP_def) apply (simp add: simpler_gets_def get_addr_def user_reg_val_def) apply (simp add: Let_def ) apply (simp add: get_CWP_def cpu_reg_val_def get_operand2_def) by (simp add: user_reg_val_def) lemma get_addr2_low_equal: "low_equal s1 s2 \ get_addr (snd instr) (snd (fst (get_curr_win () s1))) = get_addr (snd instr) (snd (fst (get_curr_win () s2)))" apply (simp add: low_equal_def) apply (simp add: get_curr_win_def cpu_reg_val_def get_CWP_def) apply (simp add: simpler_gets_def get_addr_def user_reg_val_def) apply (simp add: Let_def ) apply (simp add: get_CWP_def cpu_reg_val_def get_operand2_def) by (simp add: user_reg_val_def) lemma sys_reg_low_equal: "low_equal s1 s2 \ sys_reg s1 = sys_reg s2" by (simp add: low_equal_def) lemma user_reg_low_equal: "low_equal s1 s2 \ user_reg s1 = user_reg s2" by (simp add: low_equal_def) lemma user_reg_val_low_equal: "low_equal s1 s2 \ user_reg_val win ur s1 = user_reg_val win ur s2" apply (simp add: user_reg_val_def) by (simp add: user_reg_low_equal) lemma get_operand2_low_equal: "low_equal s1 s2 \ get_operand2 op_list s1 = get_operand2 op_list s2" apply (simp add: get_operand2_def) apply (simp add: cpu_reg_val_low_equal) apply auto apply (simp add: user_reg_val_def) using user_reg_low_equal by fastforce lemma mem_val_mod_cache: "mem_val_alt asi a s = mem_val_alt asi a (s\cache := new_cache\)" apply (simp add: mem_val_alt_def) by (simp add: Let_def) lemma mem_val_w32_mod_cache: "mem_val_w32 asi a s = mem_val_w32 asi a (s\cache := new_cache\)" apply (simp add: mem_val_w32_def) apply (simp add: Let_def) by (metis mem_val_mod_cache) lemma load_word_mem_mod_cache: "load_word_mem s addr asi = load_word_mem (s\cache := new_cache\) addr asi" apply (simp add: load_word_mem_def) apply (case_tac "virt_to_phys addr (mmu s) (mem s) = None") apply auto by (simp add: mem_val_w32_mod_cache) lemma memory_read_8_mod_cache: "fst (memory_read 8 addr s) = fst (memory_read 8 addr (s\cache := new_cache\))" apply (simp add: memory_read_def) apply (case_tac "sys_reg s CCR AND 1 \ 0") apply auto apply (simp add: option.case_eq_if load_word_mem_mod_cache) apply (auto intro: load_word_mem_mod_cache) apply (metis load_word_mem_mod_cache option.distinct(1)) by (metis load_word_mem_mod_cache option.distinct(1)) lemma memory_read_10_mod_cache: "fst (memory_read 10 addr s) = fst (memory_read 10 addr (s\cache := new_cache\))" apply (simp add: memory_read_def) apply (case_tac "sys_reg s CCR AND 1 \ 0") apply auto apply (simp add: option.case_eq_if load_word_mem_mod_cache) apply (auto intro: load_word_mem_mod_cache) apply (metis load_word_mem_mod_cache option.distinct(1)) by (metis load_word_mem_mod_cache option.distinct(1)) lemma mem_equal_mod_cache: "mem_equal s1 s2 pa \ mem_equal (s1\cache := new_cache1\) (s2\cache := new_cache2\) pa" by (simp add: mem_equal_def) lemma user_accessible_mod_cache: "user_accessible (s\cache := new_cache\) pa = user_accessible s pa" by (simp add: user_accessible_def) lemma mem_equal_mod_user_reg: "mem_equal s1 s2 pa \ mem_equal (s1\user_reg := new_user_reg1\) (s2\user_reg := user_reg2\) pa" by (simp add: mem_equal_def) lemma user_accessible_mod_user_reg: "user_accessible (s\user_reg := new_user_reg\) pa = user_accessible s pa" by (simp add: user_accessible_def) lemma mem_equal_mod_cpu_reg: "mem_equal s1 s2 pa \ mem_equal (s1\cpu_reg := new_cpu1\) (s2\cpu_reg := cpu_reg2\) pa" by (simp add: mem_equal_def) lemma user_accessible_mod_cpu_reg: "user_accessible (s\cpu_reg := new_cpu_reg\) pa = user_accessible s pa" by (simp add: user_accessible_def) lemma mem_equal_mod_trap: "mem_equal s1 s2 pa \ mem_equal (s1\traps := new_traps1\) (s2\traps := traps2\) pa" by (simp add: mem_equal_def) lemma user_accessible_mod_trap: "user_accessible (s\traps := new_traps\) pa = user_accessible s pa" by (simp add: user_accessible_def) lemma mem_equal_annul: "mem_equal s1 s2 pa \ mem_equal (s1\state_var := new_state_var, cpu_reg := new_cpu_reg\) (s2\state_var := new_state_var2, cpu_reg := new_cpu_reg2\) pa" by (simp add: mem_equal_def) lemma user_accessible_annul: "user_accessible (s\state_var := new_state_var, cpu_reg := new_cpu_reg\) pa = user_accessible s pa" by (simp add: user_accessible_def) lemma mem_val_alt_10_mem_equal_0: "mem_equal s1 s2 pa \ mem_val_alt 10 (pa AND 68719476732) s1 = mem_val_alt 10 (pa AND 68719476732) s2" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_10_mem_equal_1: "mem_equal s1 s2 pa \ mem_val_alt 10 ((pa AND 68719476732) + 1) s1 = mem_val_alt 10 ((pa AND 68719476732) + 1) s2" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_10_mem_equal_2: "mem_equal s1 s2 pa \ mem_val_alt 10 ((pa AND 68719476732) + 2) s1 = mem_val_alt 10 ((pa AND 68719476732) + 2) s2" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_10_mem_equal_3: "mem_equal s1 s2 pa \ mem_val_alt 10 ((pa AND 68719476732) + 3) s1 = mem_val_alt 10 ((pa AND 68719476732) + 3) s2" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_10_mem_equal: assumes a1: "mem_equal s1 s2 pa" shows "mem_val_alt 10 (pa AND 68719476732) s1 = mem_val_alt 10 (pa AND 68719476732) s2 \ mem_val_alt 10 ((pa AND 68719476732) + 1) s1 = mem_val_alt 10 ((pa AND 68719476732) + 1) s2 \ mem_val_alt 10 ((pa AND 68719476732) + 2) s1 = mem_val_alt 10 ((pa AND 68719476732) + 2) s2 \ mem_val_alt 10 ((pa AND 68719476732) + 3) s1 = mem_val_alt 10 ((pa AND 68719476732) + 3) s2" using mem_val_alt_10_mem_equal_0 mem_val_alt_10_mem_equal_1 mem_val_alt_10_mem_equal_2 mem_val_alt_10_mem_equal_3 a1 by blast lemma mem_val_w32_10_mem_equal: assumes a1: "mem_equal s1 s2 a" shows "mem_val_w32 10 a s1 = mem_val_w32 10 a s2" apply (simp add: mem_val_w32_def) apply (simp add: Let_def) using mem_val_alt_10_mem_equal a1 apply auto apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce by fastforce lemma mem_val_alt_8_mem_equal_0: "mem_equal s1 s2 pa \ mem_val_alt 8 (pa AND 68719476732) s1 = mem_val_alt 8 (pa AND 68719476732) s2" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_8_mem_equal_1: "mem_equal s1 s2 pa \ mem_val_alt 8 ((pa AND 68719476732) + 1) s1 = mem_val_alt 8 ((pa AND 68719476732) + 1) s2" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_8_mem_equal_2: "mem_equal s1 s2 pa \ mem_val_alt 8 ((pa AND 68719476732) + 2) s1 = mem_val_alt 8 ((pa AND 68719476732) + 2) s2" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_8_mem_equal_3: "mem_equal s1 s2 pa \ mem_val_alt 8 ((pa AND 68719476732) + 3) s1 = mem_val_alt 8 ((pa AND 68719476732) + 3) s2" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_8_mem_equal: assumes a1: "mem_equal s1 s2 pa" shows "mem_val_alt 8 (pa AND 68719476732) s1 = mem_val_alt 8 (pa AND 68719476732) s2 \ mem_val_alt 8 ((pa AND 68719476732) + 1) s1 = mem_val_alt 8 ((pa AND 68719476732) + 1) s2 \ mem_val_alt 8 ((pa AND 68719476732) + 2) s1 = mem_val_alt 8 ((pa AND 68719476732) + 2) s2 \ mem_val_alt 8 ((pa AND 68719476732) + 3) s1 = mem_val_alt 8 ((pa AND 68719476732) + 3) s2" using mem_val_alt_8_mem_equal_0 mem_val_alt_8_mem_equal_1 mem_val_alt_8_mem_equal_2 mem_val_alt_8_mem_equal_3 a1 by blast lemma mem_val_w32_8_mem_equal: assumes a1: "mem_equal s1 s2 a" shows "mem_val_w32 8 a s1 = mem_val_w32 8 a s2" apply (simp add: mem_val_w32_def) apply (simp add: Let_def) using mem_val_alt_8_mem_equal a1 apply auto apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce by fastforce lemma load_word_mem_10_low_equal: assumes a1: "low_equal s1 s2" shows "load_word_mem s1 address 10 = load_word_mem s2 address 10" using a1 apply (simp add: low_equal_def load_word_mem_def) apply clarsimp apply (case_tac "virt_to_phys address (mmu s2) (mem s2) = None") apply auto apply (simp add: user_accessible_def) using mem_val_w32_10_mem_equal apply blast apply (simp add: user_accessible_def) using mem_val_w32_10_mem_equal by blast lemma load_word_mem_8_low_equal: assumes a1: "low_equal s1 s2" shows "load_word_mem s1 address 8 = load_word_mem s2 address 8" using a1 apply (simp add: low_equal_def load_word_mem_def) apply clarsimp apply (case_tac "virt_to_phys address (mmu s2) (mem s2) = None") apply auto apply (simp add: user_accessible_def) using mem_val_w32_8_mem_equal user_accessible_8 apply fastforce apply (simp add: user_accessible_def) using mem_val_w32_8_mem_equal user_accessible_8 by fastforce lemma mem_read_low_equal: assumes a1: "low_equal s1 s2 \ asi \ {8,10}" shows "fst (memory_read asi address s1) = fst (memory_read asi address s2)" proof (cases "asi = 8") case True then show ?thesis using a1 apply (simp add: low_equal_def) apply (simp add: memory_read_def) using a1 load_word_mem_8_low_equal apply auto apply (simp add: option.case_eq_if) by (simp add: option.case_eq_if) next case False then have "asi = 10" using a1 by auto then show ?thesis using a1 apply (simp add: low_equal_def) apply (simp add: memory_read_def) using a1 load_word_mem_10_low_equal apply auto apply (simp add: option.case_eq_if) by (simp add: option.case_eq_if) qed lemma read_mem_pc_low_equal: assumes a1: "low_equal s1 s2" shows "fst (memory_read 8 (cpu_reg_val PC s1) s1) = fst (memory_read 8 (cpu_reg_val PC s2) s2)" proof - have f2: "cpu_reg_val PC s1 = cpu_reg_val PC s2" using a1 by (simp add: low_equal_def cpu_reg_val_def) then show ?thesis using a1 f2 mem_read_low_equal by auto qed lemma dcache_mod_low_equal: assumes a1: "low_equal s1 s2 \ t1 = dcache_mod c v s1 \ t2 = dcache_mod c v s2" shows "low_equal t1 t2" using a1 apply (simp add: low_equal_def) apply (simp add: dcache_mod_def) apply auto apply (simp add: user_accessible_mod_cache mem_equal_mod_cache) by (simp add: user_accessible_mod_cache mem_equal_mod_cache) lemma add_data_cache_low_equal: assumes a1: "low_equal s1 s2 \ t1 = add_data_cache s1 address w bm \ t2 = add_data_cache s2 address w bm" shows "low_equal t1 t2" using a1 apply (simp add: add_data_cache_def) apply (case_tac "bm AND 8 >> 3 = 1") apply auto apply (case_tac "bm AND 4 >> 2 = 1") apply auto apply (case_tac "bm AND 2 >> Suc 0 = 1") apply auto apply (case_tac "bm AND 1 = 1") apply auto apply (meson dcache_mod_low_equal) apply (meson dcache_mod_low_equal) apply (case_tac "bm AND 1 = 1") apply auto apply (meson dcache_mod_low_equal) apply (meson dcache_mod_low_equal) apply (case_tac "bm AND 2 >> Suc 0 = 1") apply auto apply (case_tac "bm AND 1 = 1") apply auto apply (meson dcache_mod_low_equal) apply (meson dcache_mod_low_equal) apply (case_tac "bm AND 1 = 1") apply auto apply (meson dcache_mod_low_equal) apply (meson dcache_mod_low_equal) apply (case_tac "bm AND 4 >> 2 = 1") apply auto apply (case_tac "bm AND 2 >> Suc 0 = 1") apply auto apply (case_tac "bm AND 1 = 1") apply auto apply (meson dcache_mod_low_equal) apply (meson dcache_mod_low_equal) apply (case_tac "bm AND 1 = 1") apply auto apply (meson dcache_mod_low_equal) apply (meson dcache_mod_low_equal) apply (case_tac "bm AND 2 >> Suc 0 = 1") apply auto apply (case_tac "bm AND 1 = 1") apply auto apply (meson dcache_mod_low_equal) apply (meson dcache_mod_low_equal) by (meson dcache_mod_low_equal) lemma mem_read2_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (memory_read (10::word8) address s1) \ t2 = snd (memory_read (10::word8) address s2)" shows "low_equal t1 t2" using a1 apply (simp add: memory_read_def) using a1 apply (auto simp add: sys_reg_low_equal mod_2_eq_odd) using a1 apply (simp add: load_word_mem_10_low_equal) apply (auto split: option.splits) using add_data_cache_low_equal apply force using add_data_cache_low_equal apply force done lemma mem_read_delayed_write_low_equal: assumes a1: "low_equal s1 s2 \ get_delayed_pool s1 = [] \ get_delayed_pool s2 = []" shows "fst (memory_read 8 (cpu_reg_val PC (delayed_pool_write s1)) (delayed_pool_write s1)) = fst (memory_read 8 (cpu_reg_val PC (delayed_pool_write s2)) (delayed_pool_write s2))" using a1 apply (simp add: delayed_pool_write_def) apply (simp add: Let_def) apply (simp add: get_delayed_write_def) by (simp add: read_mem_pc_low_equal) lemma global_reg_mod_low_equal: assumes a1: "low_equal s1 s2\ t1 = (global_reg_mod w n rd s1) \ t2 = (global_reg_mod w n rd s2)" shows "low_equal t1 t2" using a1 apply (induction n arbitrary: s1 s2) apply clarsimp apply auto apply (simp add: Let_def) apply (simp add: user_reg_low_equal) using user_reg_state_mod_low_equal by blast lemma out_reg_mod_low_equal: assumes a1: "low_equal s1 s2\ t1 = (out_reg_mod w curr_win rd s1) \ t2 = (out_reg_mod w curr_win rd s2)" shows "low_equal t1 t2" using a1 apply (simp add: out_reg_mod_def Let_def) apply auto apply (simp add: user_reg_low_equal) using user_reg_state_mod_low_equal apply fastforce apply (simp add: user_reg_low_equal) using user_reg_state_mod_low_equal by blast lemma in_reg_mod_low_equal: assumes a1: "low_equal s1 s2\ t1 = (in_reg_mod w curr_win rd s1) \ t2 = (in_reg_mod w curr_win rd s2)" shows "low_equal t1 t2" using a1 apply (simp add: in_reg_mod_def Let_def) apply auto apply (simp add: user_reg_low_equal) using user_reg_state_mod_low_equal apply fastforce apply (simp add: user_reg_low_equal) using user_reg_state_mod_low_equal by blast lemma user_reg_mod_low_equal: assumes a1: "low_equal s1 s2 \ t1 = user_reg_mod w curr_win rd s1 \ t2 = user_reg_mod w curr_win rd s2" shows "low_equal t1 t2" proof (cases "rd = 0") case True then show ?thesis using a1 by (simp add: user_reg_mod_def) next case False then have f1: "rd \ 0" by auto then show ?thesis proof (cases "0 < rd \ rd < 8") case True then show ?thesis using a1 f1 apply (simp add: user_reg_mod_def) using global_reg_mod_low_equal by blast next case False then have f2: "\ (0 < rd \ rd < 8)" by auto then show ?thesis proof (cases "7 < rd \ rd < 16") case True then show ?thesis using a1 f1 f2 apply (simp add: user_reg_mod_def) by (auto intro: out_reg_mod_low_equal) next case False then have f3: "\ (7 < rd \ rd < 16)" by auto then show ?thesis proof (cases "15 < rd \ rd < 24") case True then show ?thesis using a1 f1 f2 f3 apply (simp add: user_reg_mod_def) apply (simp add: low_equal_def) apply clarsimp by (simp add: user_accessible_mod_user_reg mem_equal_mod_user_reg) next case False then show ?thesis using a1 f1 f2 f3 apply (simp add: user_reg_mod_def) by (auto intro: in_reg_mod_low_equal) qed qed qed qed lemma virt_to_phys_low_equal: "low_equal s1 s2 \ virt_to_phys addr (mmu s1) (mem s1) = virt_to_phys addr (mmu s2) (mem s2)" by (auto simp add: low_equal_def) lemma write_reg_low_equal: assumes a1: "low_equal s1 s2 \ t1 = (snd (fst (write_reg w curr_win rd s1))) \ t2 = (snd (fst (write_reg w curr_win rd s2)))" shows "low_equal t1 t2" using a1 apply (simp add: write_reg_def) apply (simp add: simpler_modify_def) by (auto intro: user_reg_mod_low_equal) lemma write_cpu_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (write_cpu w cr s1)) \ t2 = (snd (fst (write_cpu w cr s2)))" shows "low_equal t1 t2" using a1 apply (simp add: write_cpu_def simpler_modify_def) apply (simp add: cpu_reg_mod_def) apply (simp add: low_equal_def) using user_accessible_mod_cpu_reg mem_equal_mod_cpu_reg by metis lemma cpu_reg_mod_low_equal: assumes a1: "low_equal s1 s2 \ t1 = cpu_reg_mod w cr s1 \ t2 = cpu_reg_mod w cr s2" shows "low_equal t1 t2" using a1 apply (simp add: cpu_reg_mod_def) apply (simp add: low_equal_def) using user_accessible_mod_cpu_reg mem_equal_mod_cpu_reg by metis lemma load_sub2_low_equal: assumes a1: "low_equal s1 s2 \ t1 = (snd (fst (load_sub2 address 10 rd curr_win w s1))) \ t2 = (snd (fst (load_sub2 address 10 rd curr_win w s2)))" shows "low_equal t1 t2" proof (cases "fst (memory_read 10 (address + 4) (snd (fst (write_reg w curr_win (rd AND 30) s1)))) = None") case True then have f0: "fst (memory_read 10 (address + 4) (snd (fst (write_reg w curr_win (rd AND 30) s1)))) = None" by auto have f1: "low_equal (snd (fst (write_reg w curr_win (rd AND 30) s1))) (snd (fst (write_reg w curr_win (rd AND 30) s2)))" using a1 by (auto intro: write_reg_low_equal) then have "fst (memory_read 10 (address + 4) (snd (fst (write_reg w curr_win (rd AND 30) s1)))) = None \ fst (memory_read 10 (address + 4) (snd (fst (write_reg w curr_win (rd AND 30) s1)))) = fst (memory_read 10 (address + 4) (snd (fst (write_reg w curr_win (rd AND 30) s2))))" using f0 by (blast intro: mem_read_low_equal) then have "fst (memory_read 10 (address + 4) (snd (fst (write_reg w curr_win (rd AND 30) s1)))) = None \ fst (memory_read 10 (address + 4) (snd (fst (write_reg w curr_win (rd AND 30) s2)))) = None" by auto then show ?thesis using a1 apply (simp add: load_sub2_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: raise_trap_def add_trap_set_def) apply (simp add: simpler_modify_def) using f1 apply (simp add: traps_low_equal) using f1 by (auto intro: mod_trap_low_equal) next case False then have f2: "fst (memory_read 10 (address + 4) (snd (fst (write_reg w curr_win (rd AND 30) s1)))) \ None" by auto have f3: "low_equal (snd (fst (write_reg w curr_win (rd AND 30) s1))) (snd (fst (write_reg w curr_win (rd AND 30) s2)))" using a1 by (auto intro: write_reg_low_equal) then have f4: "fst (memory_read 10 (address + 4) (snd (fst (write_reg w curr_win (rd AND 30) s1)))) = fst (memory_read 10 (address + 4) (snd (fst (write_reg w curr_win (rd AND 30) s2))))" using f2 by (blast intro: mem_read_low_equal) then have "fst (memory_read 10 (address + 4) (snd (fst (write_reg w curr_win (rd AND 30) s1)))) \ None \ fst (memory_read 10 (address + 4) (snd (fst (write_reg w curr_win (rd AND 30) s2)))) \ None" using f2 by auto then show ?thesis using a1 apply (simp add: load_sub2_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply clarsimp apply (simp add: simpler_modify_def bind_def h1_def h2_def Let_def) using f4 apply clarsimp using f3 by (auto intro: mem_read2_low_equal write_reg_low_equal) qed lemma load_sub3_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (load_sub3 instr curr_win rd (10::word8) address s1)) \ t2 = snd (fst (load_sub3 instr curr_win rd (10::word8) address s2))" shows "low_equal t1 t2" proof (cases "fst (memory_read 10 address s1) = None") case True then have "fst (memory_read 10 address s1) = None \ fst (memory_read 10 address s2) = None" using a1 by (auto simp add: mem_read_low_equal) then show ?thesis using a1 apply (simp add: load_sub3_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: case_prod_unfold) apply (simp add: raise_trap_def add_trap_set_def) apply (simp add: simpler_modify_def) apply (auto simp add: traps_low_equal) by (auto intro: mod_trap_low_equal) next case False then have f1: "fst (memory_read 10 address s1) \ None \ fst (memory_read 10 address s2) \ None" using a1 by (auto simp add: mem_read_low_equal) then show ?thesis proof (cases "rd \ 0 \ (fst instr = load_store_type LD \ fst instr = load_store_type LDA \ fst instr = load_store_type LDUH \ fst instr = load_store_type LDSB \ fst instr = load_store_type LDUB \ fst instr = load_store_type LDUBA \ fst instr = load_store_type LDSH \ fst instr = load_store_type LDSHA \ fst instr = load_store_type LDUHA \ fst instr = load_store_type LDSBA)") case True then show ?thesis using a1 f1 apply (simp add: load_sub3_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: case_prod_unfold) apply clarsimp apply (simp add: simpler_modify_def bind_def h1_def h2_def Let_def) apply (simp add: mem_read_low_equal) by (meson mem_read2_low_equal write_reg_low_equal) next case False then show ?thesis using a1 f1 apply (simp add: load_sub3_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: case_prod_unfold) apply clarsimp apply (simp add: simpler_modify_def bind_def h1_def h2_def Let_def) apply (simp add: mem_read_low_equal) by (meson load_sub2_low_equal mem_read2_low_equal) qed qed lemma ld_asi_user: "(fst instr = load_store_type LDSB \ fst instr = load_store_type LDUB \ fst instr = load_store_type LDUH \ fst instr = load_store_type LD \ fst instr = load_store_type LDD) \ ld_asi instr 0 = 10" apply (simp add: ld_asi_def) by auto lemma load_sub1_low_equal: assumes a1: "low_equal s1 s2 \ (fst instr = load_store_type LDSB \ fst instr = load_store_type LDUB \ fst instr = load_store_type LDUH \ fst instr = load_store_type LD \ fst instr = load_store_type LDD) \ t1 = snd (fst (load_sub1 instr rd 0 s1)) \ t2 = snd (fst (load_sub1 instr rd 0 s2))" shows "low_equal t1 t2" proof (cases "(fst instr = load_store_type LDD \ fst instr = load_store_type LDDA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word3) \ 0 \ (fst instr = load_store_type LD \ fst instr = load_store_type LDA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word2) \ 0 \ (fst instr = load_store_type LDUH \ fst instr = load_store_type LDUHA \ fst instr = load_store_type LDSH \ fst instr = load_store_type LDSHA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word1) \ 0") case True then have "((fst instr = load_store_type LDD \ fst instr = load_store_type LDDA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word3) \ 0 \ (fst instr = load_store_type LD \ fst instr = load_store_type LDA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word2) \ 0 \ (fst instr = load_store_type LDUH \ fst instr = load_store_type LDUHA \ fst instr = load_store_type LDSH \ fst instr = load_store_type LDSHA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word1) \ 0) \ ((fst instr = load_store_type LDD \ fst instr = load_store_type LDDA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word3) \ 0 \ (fst instr = load_store_type LD \ fst instr = load_store_type LDA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word2) \ 0 \ (fst instr = load_store_type LDUH \ fst instr = load_store_type LDUHA \ fst instr = load_store_type LDSH \ fst instr = load_store_type LDSHA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word1) \ 0)" by (metis (mono_tags, lifting) assms get_addr_low_equal) then show ?thesis using a1 apply (simp add: load_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: raise_trap_def add_trap_set_def) apply (simp add: simpler_modify_def) apply clarsimp apply (simp add: get_curr_win3_low_equal) by (auto intro: get_curr_win2_low_equal mod_trap_low_equal) next case False then have f1: "\ ((fst instr = load_store_type LDD \ fst instr = load_store_type LDDA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word3) \ 0 \ (fst instr = load_store_type LD \ fst instr = load_store_type LDA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word2) \ 0 \ (fst instr = load_store_type LDUH \ fst instr = load_store_type LDUHA \ fst instr = load_store_type LDSH \ fst instr = load_store_type LDSHA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word1) \ 0) \ \ ((fst instr = load_store_type LDD \ fst instr = load_store_type LDDA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word3) \ 0 \ (fst instr = load_store_type LD \ fst instr = load_store_type LDA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word2) \ 0 \ (fst instr = load_store_type LDUH \ fst instr = load_store_type LDUHA \ fst instr = load_store_type LDSH \ fst instr = load_store_type LDSHA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word1) \ 0)" by (metis assms get_addr_low_equal) show ?thesis proof - have "low_equal s1 s2 \ low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2)))" using get_curr_win2_low_equal by auto then have f2: "low_equal s1 s2 \ low_equal (snd (fst (load_sub3 instr (fst (fst (get_curr_win () s2))) rd 10 (get_addr (snd instr) (snd (fst (get_curr_win () s2)))) (snd (fst (get_curr_win () s1)))))) (snd (fst (load_sub3 instr (fst (fst (get_curr_win () s2))) rd 10 (get_addr (snd instr) (snd (fst (get_curr_win () s2)))) (snd (fst (get_curr_win () s2))))))" using load_sub3_low_equal by blast show ?thesis using a1 unfolding load_sub1_def simpler_gets_def bind_def h1_def h2_def Let_def case_prod_unfold using f1 f2 apply clarsimp by (simp add: get_addr2_low_equal get_curr_win_low_equal ld_asi_user) qed qed lemma load_instr_low_equal: assumes a1: "low_equal s1 s2 \ (fst instr = load_store_type LDSB \ fst instr = load_store_type LDUB \ fst instr = load_store_type LDUBA \ fst instr = load_store_type LDUH \ fst instr = load_store_type LD \ fst instr = load_store_type LDA \ fst instr = load_store_type LDD) \ (((get_S (cpu_reg_val PSR s1)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s2)))::word1) = 0 \ t1 = snd (fst (load_instr instr s1)) \ t2 = snd (fst (load_instr instr s2))" shows "low_equal t1 t2" proof - have "get_S (cpu_reg_val PSR s1) = 0 \ get_S (cpu_reg_val PSR s2) = 0" using a1 by (simp add: ucast_id) then show ?thesis using a1 apply (simp add: load_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: Let_def) apply clarsimp apply (simp add: raise_trap_def add_trap_set_def) apply (simp add: simpler_modify_def) apply (simp add: traps_low_equal) by (auto intro: mod_trap_low_equal load_sub1_low_equal) qed lemma st_data0_low_equal: "low_equal s1 s2 \ st_data0 instr curr_win rd addr s1 = st_data0 instr curr_win rd addr s2" apply (simp add: st_data0_def) by (simp add: user_reg_val_def low_equal_def) lemma store_word_mem_low_equal_none: "low_equal s1 s2 \ store_word_mem (add_data_cache s1 addr data bm) addr data bm 10 = None \ store_word_mem (add_data_cache s2 addr data bm) addr data bm 10 = None" apply (simp add: store_word_mem_def) proof - assume a1: "low_equal s1 s2" assume a2: "(case virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm)) of None \ None | Some pair \ if mmu_writable (get_acc_flag (snd pair)) 10 then Some (mem_mod_w32 10 (fst pair) bm data (add_data_cache s1 addr data bm)) else None) = None" have f3: "(if mmu_writable (get_acc_flag (snd (v1_2 (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))))) 10 then Some (mem_mod_w32 10 (fst (v1_2 (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm))))) bm data (add_data_cache s2 addr data bm)) else None) = (case Some (v1_2 (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))) of None \ if mmu_writable (get_acc_flag (snd (v1_2 (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))))) 10 then Some (mem_mod_w32 10 (fst (v1_2 (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm))))) bm data (add_data_cache s1 addr data bm)) else None | Some p \ if mmu_writable (get_acc_flag (snd p)) 10 then Some (mem_mod_w32 10 (fst p) bm data (add_data_cache s2 addr data bm)) else None)" by auto obtain pp :: "(word36 \ word8) option \ word36 \ word8" where f4: "virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm)) = None \ virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm)) = Some (pp (virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm))))" by (metis (no_types) option.exhaust) have f5: "virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm)) = virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm))" using a1 by (meson add_data_cache_low_equal virt_to_phys_low_equal) { assume "Some (mem_mod_w32 10 (fst (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm))))) bm data (add_data_cache s1 addr data bm)) \ (case Some (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))) of None \ None | Some p \ if mmu_writable (get_acc_flag (snd p)) 10 then Some (mem_mod_w32 10 (fst p) bm data (add_data_cache s1 addr data bm)) else None)" then have "None = (if mmu_writable (get_acc_flag (snd (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))))) 10 then Some (mem_mod_w32 10 (fst (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm))))) bm data (add_data_cache s2 addr data bm)) else None)" by fastforce moreover { assume "(if mmu_writable (get_acc_flag (snd (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))))) 10 then Some (mem_mod_w32 10 (fst (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm))))) bm data (add_data_cache s2 addr data bm)) else None) \ (case virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)) of None \ None | Some p \ if mmu_writable (get_acc_flag (snd p)) 10 then Some (mem_mod_w32 10 (fst p) bm data (add_data_cache s2 addr data bm)) else None)" then have "(case Some (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))) of None \ if mmu_writable (get_acc_flag (snd (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))))) 10 then Some (mem_mod_w32 10 (fst (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm))))) bm data (add_data_cache s1 addr data bm)) else None | Some p \ if mmu_writable (get_acc_flag (snd p)) 10 then Some (mem_mod_w32 10 (fst p) bm data (add_data_cache s2 addr data bm)) else None) \ (case virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)) of None \ None | Some p \ if mmu_writable (get_acc_flag (snd p)) 10 then Some (mem_mod_w32 10 (fst p) bm data (add_data_cache s2 addr data bm)) else None)" using f3 by simp then have "Some (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))) \ virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)) \ (if mmu_writable (get_acc_flag (snd (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))))) 10 then Some (mem_mod_w32 10 (fst (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm))))) bm data (add_data_cache s1 addr data bm)) else None) \ None" proof - have "(case virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)) of None \ if mmu_writable (get_acc_flag (snd (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))))) 10 then Some (mem_mod_w32 10 (fst (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm))))) bm data (add_data_cache s1 addr data bm)) else None | Some p \ if mmu_writable (get_acc_flag (snd p)) 10 then Some (mem_mod_w32 10 (fst p) bm data (add_data_cache s2 addr data bm)) else None) = (case virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)) of None \ None | Some p \ if mmu_writable (get_acc_flag (snd p)) 10 then Some (mem_mod_w32 10 (fst p) bm data (add_data_cache s2 addr data bm)) else None) \ Some (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))) \ virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)) \ (if mmu_writable (get_acc_flag (snd (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))))) 10 then Some (mem_mod_w32 10 (fst (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm))))) bm data (add_data_cache s1 addr data bm)) else None) \ None" by simp then show ?thesis using \(case Some (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))) of None \ if mmu_writable (get_acc_flag (snd (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))))) 10 then Some (mem_mod_w32 10 (fst (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm))))) bm data (add_data_cache s1 addr data bm)) else None | Some p \ if mmu_writable (get_acc_flag (snd p)) 10 then Some (mem_mod_w32 10 (fst p) bm data (add_data_cache s2 addr data bm)) else None) \ (case virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)) of None \ None | Some p \ if mmu_writable (get_acc_flag (snd p)) 10 then Some (mem_mod_w32 10 (fst p) bm data (add_data_cache s2 addr data bm)) else None)\ by force qed moreover { assume "Some (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))) \ virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm))" then have "virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm)) \ Some (pp (virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm))))" using f5 by simp } ultimately have "virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm)) \ Some (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))) \ virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm)) \ Some (pp (virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm))))" using a2 by force } ultimately have "virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm)) = Some (pp (virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm)))) \ virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm)) = Some (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))) \ (case virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)) of None \ None | Some p \ if mmu_writable (get_acc_flag (snd p)) 10 then Some (mem_mod_w32 10 (fst p) bm data (add_data_cache s2 addr data bm)) else None) = None" by fastforce } then have "virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm)) = Some (pp (virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm)))) \ virt_to_phys addr (mmu (add_data_cache s1 addr data bm)) (mem (add_data_cache s1 addr data bm)) = Some (pp (virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)))) \ (case virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)) of None \ None | Some p \ if mmu_writable (get_acc_flag (snd p)) 10 then Some (mem_mod_w32 10 (fst p) bm data (add_data_cache s2 addr data bm)) else None) = None" using a2 by force then show "(case virt_to_phys addr (mmu (add_data_cache s2 addr data bm)) (mem (add_data_cache s2 addr data bm)) of None \ None | Some p \ if mmu_writable (get_acc_flag (snd p)) 10 then Some (mem_mod_w32 10 (fst p) bm data (add_data_cache s2 addr data bm)) else None) = None" using f5 f4 by force qed lemma memory_write_asi_low_equal_none: "low_equal s1 s2 \ memory_write_asi 10 addr bm data s1 = None \ memory_write_asi 10 addr bm data s2 = None" apply (simp add: memory_write_asi_def) by (simp add: store_word_mem_low_equal_none) lemma memory_write_low_equal_none: "low_equal s1 s2 \ memory_write 10 addr bm data s1 = None \ memory_write 10 addr bm data s2 = None" apply (simp add: memory_write_def) by (metis map_option_case memory_write_asi_low_equal_none option.map_disc_iff) lemma memory_write_low_equal_none2: "low_equal s1 s2 \ memory_write 10 addr bm data s2 = None \ memory_write 10 addr bm data s1 = None" apply (simp add: memory_write_def) by (metis low_equal_com memory_write_def memory_write_low_equal_none) lemma mem_context_val_9_unchanged: "mem_context_val 9 addr1 (mem s1) = mem_context_val 9 addr1 - ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None)))" + ((mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None)))" apply (simp add: mem_context_val_def) by (simp add: Let_def) lemma mem_context_val_w32_9_unchanged: "mem_context_val_w32 9 addr1 (mem s1) = mem_context_val_w32 9 addr1 - ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None)))" + ((mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None)))" apply (simp add: mem_context_val_w32_def) apply (simp add: Let_def) by (metis mem_context_val_9_unchanged) lemma ptd_lookup_unchanged_4: "ptd_lookup va ptp (mem s1) 4 = -ptd_lookup va ptp ((mem s1)(10 := mem s1 10(addr \ val), +ptd_lookup va ptp ((mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) 4" by auto lemma ptd_lookup_unchanged_3: "ptd_lookup va ptp (mem s1) 3 = -ptd_lookup va ptp ((mem s1)(10 := mem s1 10(addr \ val), +ptd_lookup va ptp ((mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) 3" proof (cases "mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 12))::word6))::word32)))::word36) (mem s1) = None") case True then have "mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 12))::word6))::word32)))::word36) (mem s1) = None \ mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 12))::word6))::word32)))::word36) - ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = None" + ((mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) = None" using mem_context_val_w32_9_unchanged by metis then show ?thesis by auto next case False then have "mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 12))::word6))::word32)))::word36) (mem s1) \ None \ mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 12))::word6))::word32)))::word36) - ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) \ None" + ((mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) \ None" using mem_context_val_w32_9_unchanged by metis then have "mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 12))::word6))::word32)))::word36) (mem s1) \ None \ mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 12))::word6))::word32)))::word36) - ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) \ None \ + ((mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) \ None \ (\y. (mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 12))::word6))::word32)))::word36) (mem s1) = Some y) \ (mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 12))::word6))::word32)))::word36) - ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None)))= Some y))" + ((mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None)))= Some y))" using mem_context_val_w32_9_unchanged by metis then show ?thesis apply auto by (simp add: Let_def) qed lemma ptd_lookup_unchanged_2: "ptd_lookup va ptp (mem s1) 2 = -ptd_lookup va ptp ((mem s1)(10 := mem s1 10(addr \ val), +ptd_lookup va ptp ((mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) 2" proof (cases "mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 18))::word6))::word32)))::word36) (mem s1) = None") case True then have "mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 18))::word6))::word32)))::word36) (mem s1) = None \ mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 18))::word6))::word32)))::word36) - ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = None" + ((mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) = None" using mem_context_val_w32_9_unchanged by metis then show ?thesis by auto next case False then have "mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 18))::word6))::word32)))::word36) (mem s1) \ None \ mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 18))::word6))::word32)))::word36) - ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) \ None" + ((mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) \ None" using mem_context_val_w32_9_unchanged by metis then have "mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 18))::word6))::word32)))::word36) (mem s1) \ None \ mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 18))::word6))::word32)))::word36) - ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) \ None \ + ((mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) \ None \ (\y. (mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 18))::word6))::word32)))::word36) (mem s1) = Some y) \ (mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 18))::word6))::word32)))::word36) - ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None)))= Some y))" + ((mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None)))= Some y))" using mem_context_val_w32_9_unchanged by metis then show ?thesis apply auto using ptd_lookup_unchanged_3 unfolding Let_def by auto qed lemma ptd_lookup_unchanged_1: "ptd_lookup va ptp (mem s1) 1 = -ptd_lookup va ptp ((mem s1)(10 := mem s1 10(addr \ val), +ptd_lookup va ptp ((mem s1)(10 := (mem s1 10) (addr \ val), 11 := (mem s1 11)(addr := None))) 1" proof (cases "mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 24))::word8))::word32)))::word36) (mem s1) = None") case True then have "mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 24))::word8))::word32)))::word36) (mem s1) = None \ mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 24))::word8))::word32)))::word36) - ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = None" + ((mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) = None" using mem_context_val_w32_9_unchanged by metis then show ?thesis by auto next case False then have "mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 24))::word8))::word32)))::word36) (mem s1) \ None \ mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 24))::word8))::word32)))::word36) - ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) \ None" + ((mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) \ None" using mem_context_val_w32_9_unchanged by metis then have "mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 24))::word8))::word32)))::word36) (mem s1) \ None \ mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 24))::word8))::word32)))::word36) - ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) \ None \ + ((mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) \ None \ (\y. (mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 24))::word8))::word32)))::word36) (mem s1) = Some y) \ (mem_context_val_w32 9 ((ucast (ptp OR ((ucast ((ucast (va >> 24))::word8))::word32)))::word36) - ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None)))= Some y))" + ((mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None)))= Some y))" using mem_context_val_w32_9_unchanged by metis then show ?thesis apply auto using ptd_lookup_unchanged_2 unfolding Let_def proof - fix y :: word32 - have "(y AND 3 \ 0 \ y AND 3 = 0 \ (y AND 3 \ 1 \ ptd_lookup va (y AND 4294967292) (mem s1) (Suc 0 + 1) = None) \ (y AND 3 = 1 \ y AND 3 \ 2 \ None = Some ((ucast (ucast (y >> 8)::word24) << 12) OR (ucast (ucast va::word12)::word36), ucast y::word8))) \ (y AND 3 = 0 \ (y AND 3 \ 1 \ (y AND 3 \ 0 \ ptd_lookup va (y AND 4294967292) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) (Suc 0 + 1) = None) \ (y AND 3 = 0 \ (y AND 3 \ 1 \ ptd_lookup va (y AND 4294967292) (mem s1) (Suc 0 + 1) = ptd_lookup va (y AND 4294967292) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) (Suc 0 + 1)) \ (y AND 3 = 1 \ (y AND 3 \ 2 \ ptd_lookup va (y AND 4294967292) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) (Suc 0 + 1) = Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y)) \ (y AND 3 = 2 \ ptd_lookup va (y AND 4294967292) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) (Suc 0 + 1) = None)))) \ (y AND 3 = 1 \ (y AND 3 \ 2 \ (y AND 3 \ 0 \ None = Some ((ucast (ucast (y >> 8)::word24) << 12) OR (ucast (ucast va::word12)::word36), ucast y::word8)) \ (y AND 3 = 0 \ (y AND 3 \ 1 \ ptd_lookup va (y AND 4294967292) (mem s1) (Suc 0 + 1) = Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y)) \ (y AND 3 = 1 \ y AND 3 = 2 \ None = Some ((ucast (ucast (y >> 8)::word24) << 12) OR (ucast (ucast va::word12)::word36), ucast y::word8)))) \ (y AND 3 = 2 \ y AND 3 = 0 \ (y AND 3 \ 1 \ ptd_lookup va (y AND 4294967292) (mem s1) (Suc 0 + 1) = None) \ (y AND 3 = 1 \ y AND 3 \ 2 \ None = Some ((ucast (ucast (y >> 8)::word24) << 12) OR (ucast (ucast va::word12)::word36), ucast y::word8))))) \ (\w. mem s1 w = ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) w)" + have "(y AND 3 \ 0 \ y AND 3 = 0 \ (y AND 3 \ 1 \ ptd_lookup va (y AND 4294967292) (mem s1) (Suc 0 + 1) = None) \ (y AND 3 = 1 \ y AND 3 \ 2 \ None = Some ((ucast (ucast (y >> 8)::word24) << 12) OR (ucast (ucast va::word12)::word36), ucast y::word8))) \ (y AND 3 = 0 \ (y AND 3 \ 1 \ (y AND 3 \ 0 \ ptd_lookup va (y AND 4294967292) ((mem s1) (10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) (Suc 0 + 1) = None) \ (y AND 3 = 0 \ (y AND 3 \ 1 \ ptd_lookup va (y AND 4294967292) (mem s1) (Suc 0 + 1) = ptd_lookup va (y AND 4294967292) ((mem s1) (10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) (Suc 0 + 1)) \ (y AND 3 = 1 \ (y AND 3 \ 2 \ ptd_lookup va (y AND 4294967292) ((mem s1) (10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) (Suc 0 + 1) = Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y)) \ (y AND 3 = 2 \ ptd_lookup va (y AND 4294967292) ((mem s1) (10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) (Suc 0 + 1) = None)))) \ (y AND 3 = 1 \ (y AND 3 \ 2 \ (y AND 3 \ 0 \ None = Some ((ucast (ucast (y >> 8)::word24) << 12) OR (ucast (ucast va::word12)::word36), ucast y::word8)) \ (y AND 3 = 0 \ (y AND 3 \ 1 \ ptd_lookup va (y AND 4294967292) (mem s1) (Suc 0 + 1) = Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y)) \ (y AND 3 = 1 \ y AND 3 = 2 \ None = Some ((ucast (ucast (y >> 8)::word24) << 12) OR (ucast (ucast va::word12)::word36), ucast y::word8)))) \ (y AND 3 = 2 \ y AND 3 = 0 \ (y AND 3 \ 1 \ ptd_lookup va (y AND 4294967292) (mem s1) (Suc 0 + 1) = None) \ (y AND 3 = 1 \ y AND 3 \ 2 \ None = Some ((ucast (ucast (y >> 8)::word24) << 12) OR (ucast (ucast va::word12)::word36), ucast y::word8))))) \ (\w. mem s1 w = ((mem s1) (10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) w)" by (metis (no_types) One_nat_def Suc_1 Suc_eq_plus1 ptd_lookup_unchanged_2) - then show "(if y AND 3 = 0 then None else if y AND 3 = 1 then ptd_lookup va (y AND 4294967292) (mem s1) (Suc 0 + 1) else if y AND 3 = 2 then Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y) else None) = (if y AND 3 = 0 then None else if y AND 3 = 1 then ptd_lookup va (y AND 4294967292) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) (Suc 0 + 1) else if y AND 3 = 2 then Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y) else None)" + then show "(if y AND 3 = 0 then None else if y AND 3 = 1 then ptd_lookup va (y AND 4294967292) (mem s1) (Suc 0 + 1) else if y AND 3 = 2 then Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y) else None) = (if y AND 3 = 0 then None else if y AND 3 = 1 then ptd_lookup va (y AND 4294967292) ((mem s1) (10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) (Suc 0 + 1) else if y AND 3 = 2 then Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y) else None)" proof - have f1: "2 = Suc 0 + 1" by (metis One_nat_def Suc_1 Suc_eq_plus1) { assume "y AND 3 = 1" moreover - { assume "y AND 3 = 1 \ (if y AND 3 = 0 then None else if y AND 3 = 1 then ptd_lookup va (y AND 4294967292) (mem s1) (Suc 0 + 1) else if y AND 3 = 2 then Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y) else None) \ (if y AND 3 = 0 then None else if y AND 3 = 1 then ptd_lookup va (y AND 4294967292) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) (Suc 0 + 1) else if y AND 3 = 2 then Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y) else None)" - have "y AND 3 = 1 \ (if y AND 3 = 0 then None else if y AND 3 = 1 then ptd_lookup va (y AND 4294967292) (mem s1) (Suc 0 + 1) else if y AND 3 = 2 then Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y) else None) \ ptd_lookup va (y AND 4294967292) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) (Suc 0 + 1) \ (if y AND 3 = 0 then None else if y AND 3 = 1 then ptd_lookup va (y AND 4294967292) (mem s1) (Suc 0 + 1) else if y AND 3 = 2 then Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y) else None) = (if y AND 3 = 0 then None else if y AND 3 = 1 then ptd_lookup va (y AND 4294967292) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) (Suc 0 + 1) else if y AND 3 = 2 then Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y) else None)" + { assume "y AND 3 = 1 \ (if y AND 3 = 0 then None else if y AND 3 = 1 then ptd_lookup va (y AND 4294967292) (mem s1) (Suc 0 + 1) else if y AND 3 = 2 then Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y) else None) \ (if y AND 3 = 0 then None else if y AND 3 = 1 then ptd_lookup va (y AND 4294967292) ((mem s1) (10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) (Suc 0 + 1) else if y AND 3 = 2 then Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y) else None)" + have "y AND 3 = 1 \ (if y AND 3 = 0 then None else if y AND 3 = 1 then ptd_lookup va (y AND 4294967292) (mem s1) (Suc 0 + 1) else if y AND 3 = 2 then Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y) else None) \ ptd_lookup va (y AND 4294967292) ((mem s1) (10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) (Suc 0 + 1) \ (if y AND 3 = 0 then None else if y AND 3 = 1 then ptd_lookup va (y AND 4294967292) (mem s1) (Suc 0 + 1) else if y AND 3 = 2 then Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y) else None) = (if y AND 3 = 0 then None else if y AND 3 = 1 then ptd_lookup va (y AND 4294967292) ((mem s1) (10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) (Suc 0 + 1) else if y AND 3 = 2 then Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y) else None)" by presburger moreover - { assume "y AND 3 = 1 \ (if y AND 3 = 0 then None else if y AND 3 = 1 then ptd_lookup va (y AND 4294967292) (mem s1) (Suc 0 + 1) else if y AND 3 = 2 then Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y) else None) \ ptd_lookup va (y AND 4294967292) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) (Suc 0 + 1)" + { assume "y AND 3 = 1 \ (if y AND 3 = 0 then None else if y AND 3 = 1 then ptd_lookup va (y AND 4294967292) (mem s1) (Suc 0 + 1) else if y AND 3 = 2 then Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y) else None) \ ptd_lookup va (y AND 4294967292) ((mem s1) (10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) (Suc 0 + 1)" then have "y AND 3 = 1 \ (if y AND 3 = 0 then None else if y AND 3 = 1 then ptd_lookup va (y AND 4294967292) (mem s1) (Suc 0 + 1) else if y AND 3 = 2 then Some ((ucast (ucast (y >> 8)::word24) << 12) OR ucast (ucast va::word12), ucast y) else None) \ ptd_lookup va (y AND 4294967292) (mem s1) 2" by (metis One_nat_def Suc_1 Suc_eq_plus1 ptd_lookup_unchanged_2) then have ?thesis using f1 by auto } ultimately have ?thesis by blast } ultimately have ?thesis by blast } then show ?thesis by presburger qed qed qed lemma virt_to_phys_unchanged_sub1: assumes a1: "(let context_table_entry = (v1 >> 11 << 11) OR (v2 AND 511 << 2) in Let (mem_context_val_w32 (word_of_int 9) (ucast context_table_entry) (mem s1)) (case_option None (\lvl1_page_table. ptd_lookup va lvl1_page_table (mem s1) 1))) = (let context_table_entry = (v1 >> 11 << 11) OR (v2 AND 511 << 2) in Let (mem_context_val_w32 (word_of_int 9) (ucast context_table_entry) (mem s2)) (case_option None (\lvl1_page_table. ptd_lookup va lvl1_page_table (mem s2) 1)))" shows "(let context_table_entry = (v1 >> 11 << 11) OR (v2 AND 511 << 2) in Let (mem_context_val_w32 (word_of_int 9) (ucast context_table_entry) - ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None)))) + ((mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None)))) (case_option None (\lvl1_page_table. ptd_lookup va lvl1_page_table - ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) 1))) = + ((mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) 1))) = (let context_table_entry = (v1 >> 11 << 11) OR (v2 AND 511 << 2) in Let (mem_context_val_w32 (word_of_int 9) (ucast context_table_entry) - ((mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None)))) + ((mem s2)(10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None)))) (case_option None (\lvl1_page_table. ptd_lookup va lvl1_page_table - ((mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) 1)))" + ((mem s2)(10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))) 1)))" proof - from a1 have "(case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) (mem s1) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table (mem s1) 1) = (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) (mem s2) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table (mem s2) 1)" unfolding Let_def by auto then have "(case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) - ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) of + ((mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table (mem s1) 1) = (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) - ((mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) of + ((mem s2)(10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table (mem s2) 1)" using mem_context_val_w32_9_unchanged by (metis word_numeral_alt) then have "(case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) - ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) of + ((mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table - ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) 1) = + ((mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) 1) = (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) - ((mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) of + ((mem s2)(10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table - ((mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) 1)" + ((mem s2)(10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))) 1)" using ptd_lookup_unchanged_1 proof - obtain ww :: "word32 option \ word32" where f1: "\z. (z = None \ z = Some (ww z)) \ (z \ None \ (\w. z \ Some w))" by moura - then have f2: "(mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = None \ mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = Some (ww (mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None)))))) \ (mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) \ None \ (\w. mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) \ Some w))" + then have f2: "(mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) = None \ mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) = Some (ww (mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None)))))) \ (mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) \ None \ (\w. mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) \ Some w))" by blast - then have f3: "(case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) of None \ None | Some w \ ptd_lookup va w ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) 1) \ None \ (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) of None \ None | Some w \ ptd_lookup va w ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) 1) = (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) of None \ None | Some w \ ptd_lookup va w (mem s1) 1)" - by (metis (no_types) \\val va s1 ptp addr. ptd_lookup va ptp (mem s1) 1 = ptd_lookup va ptp ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) 1\ option.case(2) option.simps(4)) - have f4: "mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) = Some (ww (mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))))) \ mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = Some (ww (mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))))) \ (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) of None \ None | Some w \ ptd_lookup va w ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) 1) = (case Some (ww (mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))))) of None \ None | Some w \ ptd_lookup va w ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) 1)" - by (metis (no_types) \(case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table (mem s1) 1) = (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table (mem s2) 1)\ \\val va s1 ptp addr. ptd_lookup va ptp (mem s1) 1 = ptd_lookup va ptp ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) 1\ option.case(2)) - have f5: "(mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) = None \ mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) = Some (ww (mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None)))))) \ (mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) \ None \ (\w. mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) \ Some w))" + then have f3: "(case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) of None \ None | Some w \ ptd_lookup va w ((mem s1) (10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) 1) \ None \ (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) of None \ None | Some w \ ptd_lookup va w ((mem s1) (10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) 1) = (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) of None \ None | Some w \ ptd_lookup va w (mem s1) 1)" + by (metis (no_types) \\val va s1 ptp addr. ptd_lookup va ptp (mem s1) 1 = ptd_lookup va ptp ((mem s1) (10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) 1\ option.case(2) option.simps(4)) + have f4: "mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))) = Some (ww (mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))))) \ mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) = Some (ww (mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))))) \ (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) of None \ None | Some w \ ptd_lookup va w ((mem s1) (10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) 1) = (case Some (ww (mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))))) of None \ None | Some w \ ptd_lookup va w ((mem s2) (10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))) 1)" + by (metis (no_types) \(case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table (mem s1) 1) = (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table (mem s2) 1)\ \\val va s1 ptp addr. ptd_lookup va ptp (mem s1) 1 = ptd_lookup va ptp ((mem s1) (10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) 1\ option.case(2)) + have f5: "(mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))) = None \ mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))) = Some (ww (mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None)))))) \ (mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))) \ None \ (\w. mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))) \ Some w))" using f1 by blast - { assume "(case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) of None \ None | Some w \ ptd_lookup va w ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) 1) \ (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) of None \ None | Some w \ ptd_lookup va w ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) 1)" - { assume "(case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) of None \ None | Some w \ ptd_lookup va w (mem s2) 1) \ None \ (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) of None \ None | Some w \ ptd_lookup va w (mem s2) 1) \ None" - then have "(case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) of None \ None | Some w \ ptd_lookup va w (mem s2) 1) \ None \ mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) \ None" - by (metis (no_types) \(case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table (mem s1) 1) = (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table (mem s2) 1)\ option.simps(4)) + { assume "(case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) of None \ None | Some w \ ptd_lookup va w ((mem s1) (10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) 1) \ (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))) of None \ None | Some w \ ptd_lookup va w ((mem s2) (10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))) 1)" + { assume "(case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))) of None \ None | Some w \ ptd_lookup va w (mem s2) 1) \ None \ (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))) of None \ None | Some w \ ptd_lookup va w (mem s2) 1) \ None" + then have "(case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))) of None \ None | Some w \ ptd_lookup va w (mem s2) 1) \ None \ mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) \ None" + by (metis (no_types) \(case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table (mem s1) 1) = (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table (mem s2) 1)\ option.simps(4)) then have ?thesis using f5 f4 f2 by force } then have ?thesis - using f5 f3 by (metis (no_types) \(case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table (mem s1) 1) = (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table (mem s2) 1)\ \\val va s1 ptp addr. ptd_lookup va ptp (mem s1) 1 = ptd_lookup va ptp ((mem s1) (10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) 1\ option.case(2) option.simps(4)) } + using f5 f3 by (metis (no_types) \(case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s1) (10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table (mem s1) 1) = (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) (10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table (mem s2) 1)\ \\val va s1 ptp addr. ptd_lookup va ptp (mem s1) 1 = ptd_lookup va ptp ((mem s1) (10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) 1\ option.case(2) option.simps(4)) } then show ?thesis by blast qed then show ?thesis unfolding Let_def by auto qed lemma virt_to_phys_unchanged: assumes a1: "(\va. virt_to_phys va (mmu s2) (mem s1) = virt_to_phys va (mmu s2) (mem s2))" -shows "(\va. virt_to_phys va (mmu s2) ((mem s1)(10 := mem s1 10(addr \ val), +shows "(\va. virt_to_phys va (mmu s2) ((mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) = - virt_to_phys va (mmu s2) ((mem s2)(10 := mem s2 10(addr \ val), + virt_to_phys va (mmu s2) ((mem s2)(10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))))" proof (cases "registers (mmu s2) CR AND 1 \ 0") case True then have f1: "registers (mmu s2) CR AND 1 \ 0" by auto then show ?thesis proof (cases "mmu_reg_val (mmu s2) 256 = None") case True then show ?thesis by (simp add: virt_to_phys_def) next case False then have f2: "mmu_reg_val (mmu s2) 256 \ None" by auto then show ?thesis proof (cases "mmu_reg_val (mmu s2) 512 = None") case True then show ?thesis using f1 f2 apply (simp add: virt_to_phys_def) by auto next case False then show ?thesis using f1 f2 a1 apply (simp add: virt_to_phys_def) apply clarify using virt_to_phys_unchanged_sub1 by fastforce qed qed next case False then show ?thesis by (simp add: virt_to_phys_def) qed lemma virt_to_phys_unchanged2_sub1: "(case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) (mem s2) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table (mem s2) 1) = (case mem_context_val_w32 (word_of_int 9) (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) - (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) of + (10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))) of None \ None | Some lvl1_page_table \ ptd_lookup va lvl1_page_table ((mem s2) - (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) 1)" + (10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))) 1)" proof (cases "mem_context_val_w32 9 (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) (mem s2) = None") case True then have "mem_context_val_w32 9 (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) (mem s2) = None \ mem_context_val_w32 9 (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) - (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) = None" + (10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))) = None" using mem_context_val_w32_9_unchanged by metis then show ?thesis by auto next case False then have "mem_context_val_w32 9 (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) (mem s2) \ None \ (\y. mem_context_val_w32 9 (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) (mem s2) = Some y \ mem_context_val_w32 9 (ucast ((v1 >> 11 << 11) OR (v2 AND 511 << 2))) ((mem s2) - (10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))) = Some y)" + (10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))) = Some y)" using mem_context_val_w32_9_unchanged by metis then show ?thesis using ptd_lookup_unchanged_1 by fastforce qed lemma virt_to_phys_unchanged2: "virt_to_phys va (mmu s2) (mem s2) = -virt_to_phys va (mmu s2) ((mem s2)(10 := mem s2 10(addr \ val), +virt_to_phys va (mmu s2) ((mem s2)(10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None)))" proof (cases "registers (mmu s2) CR AND 1 \ 0") case True then have f1: "registers (mmu s2) CR AND 1 \ 0" by auto then show ?thesis proof (cases "mmu_reg_val (mmu s2) 256 = None") case True then show ?thesis by (simp add: virt_to_phys_def) next case False then have f2: "mmu_reg_val (mmu s2) 256 \ None" by auto then show ?thesis proof (cases "mmu_reg_val (mmu s2) 512 = None") case True then show ?thesis using f1 f2 apply (simp add: virt_to_phys_def) by auto next case False then show ?thesis using f1 f2 apply (simp add: virt_to_phys_def) apply clarify unfolding Let_def using virt_to_phys_unchanged2_sub1 by auto qed qed next case False then show ?thesis by (simp add: virt_to_phys_def) qed lemma virt_to_phys_unchanged_low_equal: assumes a1: "low_equal s1 s2" -shows "(\va. virt_to_phys va (mmu s2) ((mem s1)(10 := mem s1 10(addr \ val), +shows "(\va. virt_to_phys va (mmu s2) ((mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) = - virt_to_phys va (mmu s2) ((mem s2)(10 := mem s2 10(addr \ val), + virt_to_phys va (mmu s2) ((mem s2)(10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))))" using a1 apply (simp add: low_equal_def) using virt_to_phys_unchanged by metis lemma mmu_low_equal: "low_equal s1 s2 \ mmu s1 = mmu s2" by (simp add: low_equal_def) lemma mem_val_alt_8_unchanged0: assumes a1: "mem_equal s1 s2 pa" -shows "mem_val_alt 8 (pa AND 68719476732) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), +shows "mem_val_alt 8 (pa AND 68719476732) (s1\mem := (mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))\) = - mem_val_alt 8 (pa AND 68719476732) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), + mem_val_alt 8 (pa AND 68719476732) (s2\mem := (mem s2)(10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))\)" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) using a1 apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_8_unchanged1: assumes a1: "mem_equal s1 s2 pa" -shows "mem_val_alt 8 ((pa AND 68719476732) + 1) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), +shows "mem_val_alt 8 ((pa AND 68719476732) + 1) (s1\mem := (mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))\) = - mem_val_alt 8 ((pa AND 68719476732) + 1) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), + mem_val_alt 8 ((pa AND 68719476732) + 1) (s2\mem := (mem s2)(10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))\)" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) using a1 apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_8_unchanged2: assumes a1: "mem_equal s1 s2 pa" -shows "mem_val_alt 8 ((pa AND 68719476732) + 2) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), +shows "mem_val_alt 8 ((pa AND 68719476732) + 2) (s1\mem := (mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))\) = - mem_val_alt 8 ((pa AND 68719476732) + 2) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), + mem_val_alt 8 ((pa AND 68719476732) + 2) (s2\mem := (mem s2)(10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))\)" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) using a1 apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_8_unchanged3: assumes a1: "mem_equal s1 s2 pa" -shows "mem_val_alt 8 ((pa AND 68719476732) + 3) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), +shows "mem_val_alt 8 ((pa AND 68719476732) + 3) (s1\mem := (mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))\) = - mem_val_alt 8 ((pa AND 68719476732) + 3) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), + mem_val_alt 8 ((pa AND 68719476732) + 3) (s2\mem := (mem s2)(10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))\)" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) using a1 apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_8_unchanged: assumes a1: "mem_equal s1 s2 pa" -shows "mem_val_alt 8 (pa AND 68719476732) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), - 11 := (mem s1 11)(addr := None))\) = - mem_val_alt 8 (pa AND 68719476732) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), - 11 := (mem s2 11)(addr := None))\) \ - mem_val_alt 8 ((pa AND 68719476732) + 1) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), +shows "mem_val_alt 8 (pa AND 68719476732) (s1\mem := (mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))\) = - mem_val_alt 8 ((pa AND 68719476732) + 1) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), + mem_val_alt 8 (pa AND 68719476732) (s2\mem := (mem s2)(10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))\) \ - mem_val_alt 8 ((pa AND 68719476732) + 2) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), + mem_val_alt 8 ((pa AND 68719476732) + 1) (s1\mem := (mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))\) = - mem_val_alt 8 ((pa AND 68719476732) + 2) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), + mem_val_alt 8 ((pa AND 68719476732) + 1) (s2\mem := (mem s2)(10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))\) \ - mem_val_alt 8 ((pa AND 68719476732) + 3) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), + mem_val_alt 8 ((pa AND 68719476732) + 2) (s1\mem := (mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))\) = - mem_val_alt 8 ((pa AND 68719476732) + 3) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), + mem_val_alt 8 ((pa AND 68719476732) + 2) (s2\mem := (mem s2)(10 := (mem s2 10)(addr \ val), + 11 := (mem s2 11)(addr := None))\) \ + mem_val_alt 8 ((pa AND 68719476732) + 3) (s1\mem := (mem s1)(10 := (mem s1 10)(addr \ val), + 11 := (mem s1 11)(addr := None))\) = + mem_val_alt 8 ((pa AND 68719476732) + 3) (s2\mem := (mem s2)(10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))\)" using a1 mem_val_alt_8_unchanged0 mem_val_alt_8_unchanged1 mem_val_alt_8_unchanged2 mem_val_alt_8_unchanged3 by blast lemma mem_val_w32_8_unchanged: assumes a1: "mem_equal s1 s2 a" -shows "mem_val_w32 8 a (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), +shows "mem_val_w32 8 a (s1\mem := (mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))\) = -mem_val_w32 8 a (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), +mem_val_w32 8 a (s2\mem := (mem s2)(10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))\)" apply (simp add: mem_val_w32_def) apply (simp add: Let_def) using mem_val_alt_8_unchanged a1 apply auto apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce by fastforce lemma load_word_mem_8_unchanged: assumes a1: "low_equal s1 s2 \ load_word_mem s1 addra 8 = load_word_mem s2 addra 8" -shows "load_word_mem (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), +shows "load_word_mem (s1\mem := (mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))\) addra 8 = - load_word_mem (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), + load_word_mem (s2\mem := (mem s2)(10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))\) addra 8" -proof (cases "virt_to_phys addra (mmu s1) ((mem s1)(10 := mem s1 10(addr \ val), +proof (cases "virt_to_phys addra (mmu s1) ((mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) = None") case True - then have "virt_to_phys addra (mmu s1) ((mem s1)(10 := mem s1 10(addr \ val), + then have "virt_to_phys addra (mmu s1) ((mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) = None \ - virt_to_phys addra (mmu s2) ((mem s2)(10 := mem s2 10(addr \ val), + virt_to_phys addra (mmu s2) ((mem s2)(10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))) = None" using a1 apply (auto simp add: mmu_low_equal) using a1 virt_to_phys_unchanged_low_equal by metis then show ?thesis by (simp add: load_word_mem_def) next case False - then have "\p. virt_to_phys addra (mmu s1) ((mem s1)(10 := mem s1 10(addr \ val), + then have "\p. virt_to_phys addra (mmu s1) ((mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) = Some p \ - virt_to_phys addra (mmu s2) ((mem s2)(10 := mem s2 10(addr \ val), + virt_to_phys addra (mmu s2) ((mem s2)(10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))) = Some p" using a1 apply (auto simp add: mmu_low_equal) using a1 virt_to_phys_unchanged_low_equal by metis - then have "\p. virt_to_phys addra (mmu s1) ((mem s1)(10 := mem s1 10(addr \ val), + then have "\p. virt_to_phys addra (mmu s1) ((mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) = Some p \ - virt_to_phys addra (mmu s2) ((mem s2)(10 := mem s2 10(addr \ val), + virt_to_phys addra (mmu s2) ((mem s2)(10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))) = Some p \ virt_to_phys addra (mmu s1) (mem s1) = Some p \ virt_to_phys addra (mmu s2) (mem s2) = Some p" using virt_to_phys_unchanged2 by metis then show ?thesis using a1 apply (simp add: load_word_mem_def) apply auto apply (simp add: low_equal_def) apply (simp add: user_accessible_def) using mem_val_w32_8_unchanged a1 user_accessible_8 by (metis snd_conv) qed lemma load_word_mem_select_8: assumes a1: "fst (case load_word_mem s1 addra 8 of None \ (None, s1) | Some w \ (Some w, add_instr_cache s1 addra w 15)) = fst (case load_word_mem s2 addra 8 of None \ (None, s2) | Some w \ (Some w, add_instr_cache s2 addra w 15))" shows "load_word_mem s1 addra 8 = load_word_mem s2 addra 8" using a1 by (metis (mono_tags, lifting) fst_conv not_None_eq option.simps(4) option.simps(5)) lemma memory_read_8_unchanged: assumes a1: "low_equal s1 s2 \ fst (memory_read 8 addra s1) = fst (memory_read 8 addra s2)" shows "fst (memory_read 8 addra - (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), + (s1\mem := (mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))\)) = fst (memory_read 8 addra - (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), + (s2\mem := (mem s2)(10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))\))" proof (cases "sys_reg s1 CCR AND 1 = 0") case True then have "sys_reg s1 CCR AND 1 = 0 \ sys_reg s2 CCR AND 1 = 0" using a1 sys_reg_low_equal by fastforce then show ?thesis using a1 apply (simp add: memory_read_def) using load_word_mem_8_unchanged by blast next case False then have f1: "sys_reg s1 CCR AND 1 \ 0 \ sys_reg s2 CCR AND 1 \ 0" using a1 sys_reg_low_equal by fastforce then show ?thesis using a1 - proof (cases "load_word_mem (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), + proof (cases "load_word_mem (s1\mem := (mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))\) addra 8 = None") case True - then have "load_word_mem (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), + then have "load_word_mem (s1\mem := (mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))\) addra 8 = None \ - load_word_mem (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), + load_word_mem (s2\mem := (mem s2)(10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))\) addra 8 = None" using a1 f1 apply (simp add: memory_read_def) apply clarsimp using load_word_mem_select_8 load_word_mem_8_unchanged by fastforce then show ?thesis by (simp add: memory_read_def) next case False - then have "\y. load_word_mem (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), + then have "\y. load_word_mem (s1\mem := (mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))\) addra 8 = Some y" by auto - then have "\y. load_word_mem (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), + then have "\y. load_word_mem (s1\mem := (mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))\) addra 8 = Some y \ - load_word_mem (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), + load_word_mem (s2\mem := (mem s2)(10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))\) addra 8 = Some y" using a1 f1 apply (simp add: memory_read_def) apply clarsimp using load_word_mem_select_8 load_word_mem_8_unchanged by fastforce then show ?thesis using a1 f1 apply (simp add: memory_read_def) by auto qed qed lemma mem_val_alt_mod: assumes a1: "addr1 \ addr2" shows "mem_val_alt 10 addr1 s = -mem_val_alt 10 addr1 (s\mem := (mem s)(10 := mem s 10(addr2 \ val), +mem_val_alt 10 addr1 (s\mem := (mem s)(10 := (mem s 10)(addr2 \ val), 11 := (mem s 11)(addr2 := None))\)" using a1 apply (simp add: mem_val_alt_def) by (simp add: Let_def) lemma mem_val_alt_mod2: -"mem_val_alt 10 addr (s\mem := (mem s)(10 := mem s 10(addr \ val), +"mem_val_alt 10 addr (s\mem := (mem s)(10 := (mem s 10)(addr \ val), 11 := (mem s 11)(addr := None))\) = Some val" by (simp add: mem_val_alt_def) lemma mem_val_alt_10_unchanged0: assumes a1: "mem_equal s1 s2 pa" -shows "mem_val_alt 10 (pa AND 68719476732) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), +shows "mem_val_alt 10 (pa AND 68719476732) (s1\mem := (mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))\) = - mem_val_alt 10 (pa AND 68719476732) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), + mem_val_alt 10 (pa AND 68719476732) (s2\mem := (mem s2)(10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))\)" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) using a1 apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_10_unchanged1: assumes a1: "mem_equal s1 s2 pa" -shows "mem_val_alt 10 ((pa AND 68719476732) + 1) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), +shows "mem_val_alt 10 ((pa AND 68719476732) + 1) (s1\mem := (mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))\) = - mem_val_alt 10 ((pa AND 68719476732) + 1) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), + mem_val_alt 10 ((pa AND 68719476732) + 1) (s2\mem := (mem s2)(10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))\)" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) using a1 apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_10_unchanged2: assumes a1: "mem_equal s1 s2 pa" -shows "mem_val_alt 10 ((pa AND 68719476732) + 2) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), +shows "mem_val_alt 10 ((pa AND 68719476732) + 2) (s1\mem := (mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))\) = - mem_val_alt 10 ((pa AND 68719476732) + 2) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), + mem_val_alt 10 ((pa AND 68719476732) + 2) (s2\mem := (mem s2)(10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))\)" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) using a1 apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_10_unchanged3: assumes a1: "mem_equal s1 s2 pa" -shows "mem_val_alt 10 ((pa AND 68719476732) + 3) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), +shows "mem_val_alt 10 ((pa AND 68719476732) + 3) (s1\mem := (mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))\) = - mem_val_alt 10 ((pa AND 68719476732) + 3) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), + mem_val_alt 10 ((pa AND 68719476732) + 3) (s2\mem := (mem s2)(10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))\)" apply (simp add: mem_val_alt_def) apply (simp add: Let_def) using a1 apply (simp add: mem_equal_def) by (metis option.distinct(1)) lemma mem_val_alt_10_unchanged: assumes a1: "mem_equal s1 s2 pa" -shows "mem_val_alt 10 (pa AND 68719476732) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), - 11 := (mem s1 11)(addr := None))\) = - mem_val_alt 10 (pa AND 68719476732) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), - 11 := (mem s2 11)(addr := None))\) \ - mem_val_alt 10 ((pa AND 68719476732) + 1) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), +shows "mem_val_alt 10 (pa AND 68719476732) (s1\mem := (mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))\) = - mem_val_alt 10 ((pa AND 68719476732) + 1) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), + mem_val_alt 10 (pa AND 68719476732) (s2\mem := (mem s2)(10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))\) \ - mem_val_alt 10 ((pa AND 68719476732) + 2) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), + mem_val_alt 10 ((pa AND 68719476732) + 1) (s1\mem := (mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))\) = - mem_val_alt 10 ((pa AND 68719476732) + 2) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), + mem_val_alt 10 ((pa AND 68719476732) + 1) (s2\mem := (mem s2)(10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))\) \ - mem_val_alt 10 ((pa AND 68719476732) + 3) (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), + mem_val_alt 10 ((pa AND 68719476732) + 2) (s1\mem := (mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))\) = - mem_val_alt 10 ((pa AND 68719476732) + 3) (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), + mem_val_alt 10 ((pa AND 68719476732) + 2) (s2\mem := (mem s2)(10 := (mem s2 10)(addr \ val), + 11 := (mem s2 11)(addr := None))\) \ + mem_val_alt 10 ((pa AND 68719476732) + 3) (s1\mem := (mem s1)(10 := (mem s1 10)(addr \ val), + 11 := (mem s1 11)(addr := None))\) = + mem_val_alt 10 ((pa AND 68719476732) + 3) (s2\mem := (mem s2)(10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))\)" using a1 mem_val_alt_10_unchanged0 mem_val_alt_10_unchanged1 mem_val_alt_10_unchanged2 mem_val_alt_10_unchanged3 by blast lemma mem_val_w32_10_unchanged: assumes a1: "mem_equal s1 s2 a" -shows "mem_val_w32 10 a (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), +shows "mem_val_w32 10 a (s1\mem := (mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))\) = -mem_val_w32 10 a (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), +mem_val_w32 10 a (s2\mem := (mem s2)(10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))\)" apply (simp add: mem_val_w32_def) apply (simp add: Let_def) using mem_val_alt_10_unchanged a1 apply auto apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce apply fastforce by fastforce lemma is_accessible: "low_equal s1 s2 \ virt_to_phys addra (mmu s1) (mem s1) = Some (a, b) \ virt_to_phys addra (mmu s2) (mem s2) = Some (a, b) \ mmu_readable (get_acc_flag b) 10 \ mem_equal s1 s2 a" apply (simp add: low_equal_def) apply (simp add: user_accessible_def) by fastforce lemma load_word_mem_10_unchanged: assumes a1: "low_equal s1 s2 \ load_word_mem s1 addra 10 = load_word_mem s2 addra 10" -shows "load_word_mem (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), +shows "load_word_mem (s1\mem := (mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))\) addra 10 = - load_word_mem (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), + load_word_mem (s2\mem := (mem s2)(10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))\) addra 10" -proof (cases "virt_to_phys addra (mmu s1) ((mem s1)(10 := mem s1 10(addr \ val), +proof (cases "virt_to_phys addra (mmu s1) ((mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) = None") case True - then have "virt_to_phys addra (mmu s1) ((mem s1)(10 := mem s1 10(addr \ val), + then have "virt_to_phys addra (mmu s1) ((mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) = None \ - virt_to_phys addra (mmu s2) ((mem s2)(10 := mem s2 10(addr \ val), + virt_to_phys addra (mmu s2) ((mem s2)(10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))) = None" using a1 apply (auto simp add: mmu_low_equal) using a1 virt_to_phys_unchanged_low_equal by metis then show ?thesis by (simp add: load_word_mem_def) next case False - then have "\p. virt_to_phys addra (mmu s1) ((mem s1)(10 := mem s1 10(addr \ val), + then have "\p. virt_to_phys addra (mmu s1) ((mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) = Some p \ - virt_to_phys addra (mmu s2) ((mem s2)(10 := mem s2 10(addr \ val), + virt_to_phys addra (mmu s2) ((mem s2)(10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))) = Some p" using a1 apply (auto simp add: mmu_low_equal) using a1 virt_to_phys_unchanged_low_equal by metis - then have "\p. virt_to_phys addra (mmu s1) ((mem s1)(10 := mem s1 10(addr \ val), + then have "\p. virt_to_phys addra (mmu s1) ((mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) = Some p \ - virt_to_phys addra (mmu s2) ((mem s2)(10 := mem s2 10(addr \ val), + virt_to_phys addra (mmu s2) ((mem s2)(10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))) = Some p \ virt_to_phys addra (mmu s1) (mem s1) = Some p \ virt_to_phys addra (mmu s2) (mem s2) = Some p" using virt_to_phys_unchanged2 by metis then show ?thesis using a1 apply (simp add: load_word_mem_def) apply auto apply (simp add: low_equal_def) apply (simp add: user_accessible_def) using mem_val_w32_10_unchanged a1 by metis qed lemma load_word_mem_select_10: assumes a1: "fst (case load_word_mem s1 addra 10 of None \ (None, s1) | Some w \ (Some w, add_data_cache s1 addra w 15)) = fst (case load_word_mem s2 addra 10 of None \ (None, s2) | Some w \ (Some w, add_data_cache s2 addra w 15))" shows "load_word_mem s1 addra 10 = load_word_mem s2 addra 10" using a1 by (metis (mono_tags, lifting) fst_conv not_None_eq option.simps(4) option.simps(5)) lemma memory_read_10_unchanged: assumes a1: "low_equal s1 s2 \ fst (memory_read 10 addra s1) = fst (memory_read 10 addra s2)" shows "fst (memory_read 10 addra - (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), + (s1\mem := (mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))\)) = fst (memory_read 10 addra - (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), + (s2\mem := (mem s2)(10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))\))" proof (cases "sys_reg s1 CCR AND 1 = 0") case True then have "sys_reg s1 CCR AND 1 = 0 \ sys_reg s2 CCR AND 1 = 0" using a1 sys_reg_low_equal by fastforce then show ?thesis using a1 apply (simp add: memory_read_def) using load_word_mem_10_unchanged by blast next case False then have f1: "sys_reg s1 CCR AND 1 \ 0 \ sys_reg s2 CCR AND 1 \ 0" using a1 sys_reg_low_equal by fastforce then show ?thesis using a1 - proof (cases "load_word_mem (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), + proof (cases "load_word_mem (s1\mem := (mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))\) addra 10 = None") case True - then have "load_word_mem (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), + then have "load_word_mem (s1\mem := (mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))\) addra 10 = None \ - load_word_mem (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), + load_word_mem (s2\mem := (mem s2)(10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))\) addra 10 = None" using a1 f1 apply (simp add: memory_read_def) apply clarsimp using load_word_mem_select_10 load_word_mem_10_unchanged by fastforce then show ?thesis by (simp add: memory_read_def) next case False - then have "\y. load_word_mem (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), + then have "\y. load_word_mem (s1\mem := (mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))\) addra 10 = Some y" by auto - then have "\y. load_word_mem (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), + then have "\y. load_word_mem (s1\mem := (mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))\) addra 10 = Some y \ - load_word_mem (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), + load_word_mem (s2\mem := (mem s2)(10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))\) addra 10 = Some y" using a1 f1 apply (simp add: memory_read_def) apply clarsimp using load_word_mem_select_10 load_word_mem_10_unchanged by fastforce then show ?thesis using a1 f1 apply (simp add: memory_read_def) by auto qed qed lemma state_mem_mod_1011_low_equal_sub1: assumes a1: "(\va. virt_to_phys va (mmu s2) (mem s1) = virt_to_phys va (mmu s2) (mem s2)) \ (\pa. (\va b. virt_to_phys va (mmu s2) (mem s2) = Some (pa, b) \ mmu_readable (get_acc_flag b) 10) \ mem_equal s1 s2 pa) \ mmu s1 = mmu s2 \ virt_to_phys va (mmu s2) - ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = + ((mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) = Some (pa, b) \ mmu_readable (get_acc_flag b) 10" shows "mem_equal s1 s2 pa" proof - have "virt_to_phys va (mmu s1) - ((mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))) = + ((mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))) = Some (pa, b)" using a1 by auto then have "virt_to_phys va (mmu s1) (mem s1) = Some (pa, b)" using virt_to_phys_unchanged2 by metis then have "virt_to_phys va (mmu s2) (mem s2) = Some (pa, b)" using a1 by auto then show ?thesis using a1 by auto qed lemma mem_equal_unchanged: assumes a1: "mem_equal s1 s2 pa" -shows "mem_equal (s1\mem := (mem s1)(10 := mem s1 10(addr \ val), +shows "mem_equal (s1\mem := (mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))\) - (s2\mem := (mem s2)(10 := mem s2 10(addr \ val), + (s2\mem := (mem s2)(10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))\) pa" using a1 apply (simp add: mem_equal_def) by auto lemma state_mem_mod_1011_low_equal: assumes a1: "low_equal s1 s2 \ -t1 = s1\mem := (mem s1)(10 := mem s1 10(addr \ val), 11 := (mem s1 11)(addr := None))\ \ -t2 = s2\mem := (mem s2)(10 := mem s2 10(addr \ val), 11 := (mem s2 11)(addr := None))\" +t1 = s1\mem := (mem s1)(10 := (mem s1 10)(addr \ val), 11 := (mem s1 11)(addr := None))\ \ +t2 = s2\mem := (mem s2)(10 := (mem s2 10)(addr \ val), 11 := (mem s2 11)(addr := None))\" shows "low_equal t1 t2" using a1 apply (simp add: low_equal_def) apply (simp add: user_accessible_def) apply auto apply (simp add: assms virt_to_phys_unchanged_low_equal) using state_mem_mod_1011_low_equal_sub1 mem_equal_unchanged apply metis apply (metis virt_to_phys_unchanged2) using state_mem_mod_1011_low_equal_sub1 mem_equal_unchanged by metis lemma mem_mod_low_equal: assumes a1: "low_equal s1 s2 \ t1 = (mem_mod 10 addr val s1) \ t2 = (mem_mod 10 addr val s2)" shows "low_equal t1 t2" using a1 apply (simp add: mem_mod_def) by (auto intro: state_mem_mod_1011_low_equal) lemma mem_mod_w32_low_equal: assumes a1: "low_equal s1 s2 \ t1 = mem_mod_w32 10 a bm data s1 \ t2 = mem_mod_w32 10 a bm data s2" shows "low_equal t1 t2" using a1 apply (simp add: mem_mod_w32_def) apply (simp add: Let_def) by (meson mem_mod_low_equal) lemma store_word_mem_low_equal: assumes a1: "low_equal s1 s2 \ Some t1 = store_word_mem s1 addr data bm 10 \ Some t2 = store_word_mem s2 addr data bm 10" shows "low_equal t1 t2" using a1 apply (simp add: store_word_mem_def) apply (auto simp add: virt_to_phys_low_equal) apply (case_tac "virt_to_phys addr (mmu s2) (mem s2) = None") apply auto apply (case_tac "mmu_writable (get_acc_flag b) 10") apply auto using mem_mod_w32_low_equal by blast lemma memory_write_asi_low_equal: assumes a1: "low_equal s1 s2 \ Some t1 = memory_write_asi 10 addr bm data s1 \ Some t2 = memory_write_asi 10 addr bm data s2" shows "low_equal t1 t2" using a1 apply (simp add: memory_write_asi_def) by (meson add_data_cache_low_equal store_word_mem_low_equal) lemma store_barrier_pending_mod_low_equal: assumes a1: "low_equal s1 s2 \ t1 = store_barrier_pending_mod False s1 \ t2 = store_barrier_pending_mod False s2" shows "low_equal t1 t2" using a1 apply (simp add: store_barrier_pending_mod_def) apply clarsimp using a1 apply (auto simp add: state_var_low_equal) by (auto intro: state_var2_low_equal) lemma memory_write_low_equal: assumes a1: "low_equal s1 s2 \ Some t1 = memory_write 10 addr bm data s1 \ Some t2 = memory_write 10 addr bm data s2" shows "low_equal t1 t2" apply (case_tac "memory_write_asi 10 addr bm data s1 = None") using a1 apply (simp add: memory_write_def) apply (case_tac "memory_write_asi 10 addr bm data s2 = None") apply (meson assms low_equal_com memory_write_asi_low_equal_none) using a1 apply (simp add: memory_write_def) apply auto by (metis memory_write_asi_low_equal store_barrier_pending_mod_low_equal) lemma memory_write_low_equal2: assumes a1: "low_equal s1 s2 \ Some t1 = memory_write 10 addr bm data s1" shows "\t2. Some t2 = memory_write 10 addr bm data s2" using a1 apply (simp add: memory_write_def) apply auto by (metis (full_types) memory_write_def memory_write_low_equal_none2 not_None_eq) lemma store_sub2_low_equal_sub1: assumes a1: "low_equal s1 s2 \ memory_write 10 addr (st_byte_mask instr addr) (st_data0 instr curr_win rd addr s2) s1 = Some y \ memory_write 10 addr (st_byte_mask instr addr) (st_data0 instr curr_win rd addr s2) s2 = Some ya" shows "low_equal (y\traps := insert data_access_exception (traps y)\) (ya\traps := insert data_access_exception (traps ya)\)" proof - from a1 have f1: "low_equal y ya" using memory_write_low_equal by metis then have "traps y = traps ya" by (simp add: low_equal_def) then show ?thesis using f1 mod_trap_low_equal by fastforce qed lemma store_sub2_low_equal_sub2: assumes a1: "low_equal s1 s2 \ memory_write 10 addr (st_byte_mask instr addr) (st_data0 instr curr_win rd addr s2) s1 = Some y \ memory_write 10 addr (st_byte_mask instr addr) (st_data0 instr curr_win rd addr s2) s2 = Some ya \ memory_write 10 (addr + 4) 15 (user_reg_val curr_win (rd OR 1) y) y = None \ memory_write 10 (addr + 4) 15 (user_reg_val curr_win (rd OR 1) ya) ya = Some yb" shows "False" proof - from a1 have f1: "low_equal y ya" using memory_write_low_equal by metis then have "(user_reg_val curr_win (rd OR 1) y) = (user_reg_val curr_win (rd OR 1) ya)" by (simp add: low_equal_def user_reg_val_def) then show ?thesis using a1 using f1 memory_write_low_equal_none by fastforce qed lemma store_sub2_low_equal_sub3: assumes a1: "low_equal s1 s2 \ memory_write 10 addr (st_byte_mask instr addr) (st_data0 instr curr_win rd addr s2) s1 = Some y \ memory_write 10 addr (st_byte_mask instr addr) (st_data0 instr curr_win rd addr s2) s2 = Some ya \ memory_write 10 (addr + 4) 15 (user_reg_val curr_win (rd OR 1) y) y = Some yb \ memory_write 10 (addr + 4) 15 (user_reg_val curr_win (rd OR 1) ya) ya = None" shows "False" proof - from a1 have f1: "low_equal y ya" using memory_write_low_equal by metis then have "(user_reg_val curr_win (rd OR 1) y) = (user_reg_val curr_win (rd OR 1) ya)" by (simp add: low_equal_def user_reg_val_def) then show ?thesis using a1 using f1 memory_write_low_equal_none2 by fastforce qed lemma store_sub2_low_equal_sub4: assumes a1: "low_equal s1 s2 \ memory_write 10 addr (st_byte_mask instr addr) (st_data0 instr curr_win rd addr s2) s1 = Some y \ memory_write 10 addr (st_byte_mask instr addr) (st_data0 instr curr_win rd addr s2) s2 = Some ya \ memory_write 10 (addr + 4) 15 (user_reg_val curr_win (rd OR 1) y) y = Some yb \ memory_write 10 (addr + 4) 15 (user_reg_val curr_win (rd OR 1) ya) ya = Some yc" shows "low_equal yb yc" proof - from a1 have f1: "low_equal y ya" using memory_write_low_equal by metis then have "(user_reg_val curr_win (rd OR 1) y) = (user_reg_val curr_win (rd OR 1) ya)" by (simp add: low_equal_def user_reg_val_def) then show ?thesis using a1 f1 by (metis memory_write_low_equal) qed lemma store_sub2_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (store_sub2 instr curr_win rd 10 addr s1)) \ t2 = snd (fst (store_sub2 instr curr_win rd 10 addr s2))" shows "low_equal t1 t2" proof (cases "memory_write 10 addr (st_byte_mask instr addr) (st_data0 instr curr_win rd addr s1) s1 = None") case True then have "memory_write 10 addr (st_byte_mask instr addr) (st_data0 instr curr_win rd addr s1) s1 = None \ memory_write 10 addr (st_byte_mask instr addr) (st_data0 instr curr_win rd addr s2) s2 = None" using a1 by (metis memory_write_low_equal_none st_data0_low_equal) then show ?thesis using a1 apply (simp add: store_sub2_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold return_def) apply (simp add: raise_trap_def add_trap_set_def) apply (simp add: simpler_modify_def) using mod_trap_low_equal traps_low_equal by fastforce next case False then have f1: "memory_write 10 addr (st_byte_mask instr addr) (st_data0 instr curr_win rd addr s1) s1 \ None \ memory_write 10 addr (st_byte_mask instr addr) (st_data0 instr curr_win rd addr s2) s2 \ None" using a1 by (metis memory_write_low_equal_none2 st_data0_low_equal) then show ?thesis proof (cases "(fst instr) \ {load_store_type STD,load_store_type STDA}") case True then show ?thesis using a1 f1 apply (simp add: store_sub2_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: simpler_modify_def bind_def h1_def h2_def Let_def) apply (simp add: return_def) apply (simp add: bind_def case_prod_unfold) apply (simp add: simpler_modify_def) apply clarsimp apply (simp add: case_prod_unfold bind_def h1_def h2_def Let_def simpler_modify_def) apply (simp add: simpler_gets_def) apply auto apply (simp add: raise_trap_def add_trap_set_def) apply (simp add: simpler_modify_def) apply (simp add: st_data0_low_equal) apply (simp add: store_sub2_low_equal_sub1) apply (simp add: st_data0_low_equal) using store_sub2_low_equal_sub2 apply blast apply (simp add: st_data0_low_equal) using store_sub2_low_equal_sub3 apply blast apply (simp add: st_data0_low_equal) using store_sub2_low_equal_sub4 apply blast apply (simp add: st_data0_low_equal) apply (simp add: raise_trap_def add_trap_set_def) apply (simp add: simpler_modify_def) using store_sub2_low_equal_sub1 apply blast apply (simp add: st_data0_low_equal) using store_sub2_low_equal_sub2 apply blast apply (simp add: st_data0_low_equal) using store_sub2_low_equal_sub3 apply blast apply (simp add: st_data0_low_equal) using store_sub2_low_equal_sub4 by blast next case False then show ?thesis using a1 f1 apply (simp add: store_sub2_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: simpler_modify_def bind_def h1_def h2_def Let_def) apply (simp add: return_def) apply (simp add: bind_def case_prod_unfold) apply clarsimp apply (simp add: simpler_modify_def) apply (simp add: st_data0_low_equal) using memory_write_low_equal by metis qed qed lemma store_sub1_low_equal: assumes a1: "low_equal s1 s2 \ (fst instr = load_store_type STB \ fst instr = load_store_type STH \ fst instr = load_store_type ST \ fst instr = load_store_type STD) \ t1 = snd (fst (store_sub1 instr rd 0 s1)) \ t2 = snd (fst (store_sub1 instr rd 0 s2))" shows "low_equal t1 t2" proof (cases "(fst instr = load_store_type STH \ fst instr = load_store_type STHA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word1) \ 0") case True then have "((fst instr = load_store_type STH \ fst instr = load_store_type STHA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word1) \ 0) \ ((fst instr = load_store_type STH \ fst instr = load_store_type STHA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word1) \ 0)" by (metis (mono_tags, lifting) assms get_addr_low_equal) then show ?thesis using a1 apply (simp add: store_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: raise_trap_def add_trap_set_def) apply (simp add: simpler_modify_def) apply clarsimp apply (simp add: get_curr_win3_low_equal) by (auto intro: get_curr_win2_low_equal mod_trap_low_equal) next case False then have f1: "\ ((fst instr = load_store_type STH \ fst instr = load_store_type STHA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word1) \ 0) \ \ ((fst instr = load_store_type STH \ fst instr = load_store_type STHA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word1) \ 0)" by (metis (mono_tags, lifting) assms get_addr_low_equal) then show ?thesis proof (cases "(fst instr \ {load_store_type ST,load_store_type STA}) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word2) \ 0") case True then have "(fst instr \ {load_store_type ST,load_store_type STA}) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word2) \ 0 \ (fst instr \ {load_store_type ST,load_store_type STA}) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word2) \ 0" by (metis (mono_tags, lifting) assms get_addr_low_equal) then show ?thesis using a1 f1 apply (simp add: store_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: raise_trap_def add_trap_set_def) apply (simp add: simpler_modify_def) apply clarsimp apply (simp add: get_curr_win3_low_equal) by (auto intro: get_curr_win2_low_equal mod_trap_low_equal) next case False then have "\((fst instr \ {load_store_type ST,load_store_type STA}) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word2) \ 0) \ \((fst instr \ {load_store_type ST,load_store_type STA}) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word2) \ 0)" by (metis (mono_tags, lifting) assms get_addr_low_equal) then have f2: "\((fst instr = load_store_type ST \ fst instr = load_store_type STA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word2) \ 0) \ \((fst instr = load_store_type ST \ fst instr = load_store_type STA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word2) \ 0)" by auto then show ?thesis proof (cases "(fst instr \ {load_store_type STD,load_store_type STDA}) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word3) \ 0") case True then have "(fst instr \ {load_store_type STD,load_store_type STDA}) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word3) \ 0 \ (fst instr \ {load_store_type STD,load_store_type STDA}) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word3) \ 0" by (metis (mono_tags, lifting) assms get_addr_low_equal) then show ?thesis using a1 apply (simp add: store_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply auto apply (simp add: case_prod_unfold) apply (simp add: raise_trap_def add_trap_set_def) apply (simp add: simpler_modify_def) apply (simp add: get_curr_win3_low_equal) by (auto intro: get_curr_win2_low_equal mod_trap_low_equal) next case False then have "\ (fst instr \ {load_store_type STD, load_store_type STDA} \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word3) \ 0) \ \ (fst instr \ {load_store_type STD, load_store_type STDA} \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word3) \ 0)" by (metis (mono_tags, lifting) assms get_addr_low_equal) then have f3: "\ ((fst instr = load_store_type STD \ fst instr = load_store_type STDA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s1)))))::word3) \ 0) \ \ ((fst instr = load_store_type STD \ fst instr = load_store_type STDA) \ ((ucast (get_addr (snd instr) (snd (fst (get_curr_win () s2)))))::word3) \ 0)" by auto show ?thesis using a1 apply (simp add: store_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (unfold case_prod_beta) apply (simp add: f1 f2 f3) apply (simp_all add: st_asi_def) using a1 apply clarsimp apply (simp add: get_curr_win_low_equal get_addr2_low_equal) by (metis store_sub2_low_equal get_curr_win2_low_equal) qed qed qed lemma store_instr_low_equal: assumes a1: "low_equal s1 s2 \ (fst instr = load_store_type STB \ fst instr = load_store_type STH \ fst instr = load_store_type ST \ fst instr = load_store_type STA \ fst instr = load_store_type STD) \ (((get_S (cpu_reg_val PSR s1)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s2)))::word1) = 0 \ t1 = snd (fst (store_instr instr s1)) \ t2 = snd (fst (store_instr instr s2))" shows "low_equal t1 t2" proof - have "get_S (cpu_reg_val PSR s1) = 0 \ get_S (cpu_reg_val PSR s2) = 0" using a1 by (simp add: ucast_id) then show ?thesis using a1 apply (simp add: store_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: Let_def) apply clarsimp apply (simp add: raise_trap_def add_trap_set_def) apply (simp add: simpler_modify_def) apply (simp add: traps_low_equal) by (auto intro: mod_trap_low_equal store_sub1_low_equal) qed lemma sethi_low_equal: "low_equal s1 s2 \ t1 = snd (fst (sethi_instr instr s1)) \ t2 = snd (fst (sethi_instr instr s2)) \ low_equal t1 t2" apply (simp add: sethi_instr_def) apply (simp add: Let_def) apply (case_tac "get_operand_w5 (snd instr ! Suc 0) \ 0") apply auto apply (simp add: bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: get_curr_win_low_equal) using get_curr_win2_low_equal write_reg_low_equal apply metis by (simp add: return_def) lemma nop_low_equal: "low_equal s1 s2 \ t1 = snd (fst (nop_instr instr s1)) \ t2 = snd (fst (nop_instr instr s2)) \ low_equal t1 t2" apply (simp add: nop_instr_def) by (simp add: return_def) lemma logical_instr_sub1_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (logical_instr_sub1 instr_name result s1)) \ t2 = snd (fst (logical_instr_sub1 instr_name result s2))" shows "low_equal t1 t2" proof (cases "instr_name = logic_type ANDcc \ instr_name = logic_type ANDNcc \ instr_name = logic_type ORcc \ instr_name = logic_type ORNcc \ instr_name = logic_type XORcc \ instr_name = logic_type XNORcc") case True then show ?thesis using a1 apply (simp add: logical_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: logical_new_psr_val_def) using write_cpu_low_equal cpu_reg_val_low_equal by fastforce next case False then show ?thesis using a1 apply (simp add: logical_instr_sub1_def) by (simp add: return_def) qed lemma logical_instr_low_equal: "low_equal s1 s2 \ t1 = snd (fst (logical_instr instr s1)) \ t2 = snd (fst (logical_instr instr s2)) \ low_equal t1 t2" apply (simp add: logical_instr_def) apply (simp add: Let_def simpler_gets_def bind_def h1_def h2_def) apply (simp add: case_prod_unfold) apply auto apply (simp_all add: get_curr_win_low_equal) apply (simp_all add: get_operand2_low_equal) using logical_instr_sub1_low_equal get_operand2_low_equal get_curr_win2_low_equal write_reg_low_equal user_reg_val_low_equal proof - assume a1: "low_equal s1 s2" assume "t2 = snd (fst (logical_instr_sub1 (fst instr) (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2)) (snd (fst (write_reg (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))))" assume "t1 = snd (fst (logical_instr_sub1 (fst instr) (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s2)) (snd (fst (write_reg (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s2)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))))" have "\w wa. user_reg_val w wa (snd (fst (get_curr_win () s2))) = user_reg_val w wa (snd (fst (get_curr_win () s1)))" using a1 by (metis (no_types) get_curr_win2_low_equal user_reg_val_low_equal) then show "low_equal (snd (fst (logical_instr_sub1 (fst instr) (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s2)) (snd (fst (write_reg (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s2)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))))) (snd (fst (logical_instr_sub1 (fst instr) (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2)) (snd (fst (write_reg (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))))))" using a1 by (metis (no_types) get_curr_win2_low_equal logical_instr_sub1_low_equal write_reg_low_equal) next assume a2: "low_equal s1 s2" assume "t1 = snd (fst (logical_instr_sub1 (fst instr) (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s2)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1))))))))" assume "t2 = snd (fst (logical_instr_sub1 (fst instr) (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))))))))" have "\w wa. user_reg_val w wa (snd (fst (get_curr_win () s2))) = user_reg_val w wa (snd (fst (get_curr_win () s1)))" using a2 by (metis (no_types) get_curr_win2_low_equal user_reg_val_low_equal) then show "low_equal (snd (fst (logical_instr_sub1 (fst instr) (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s2)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1))))))))) (snd (fst (logical_instr_sub1 (fst instr) (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))))))))" proof - have "low_equal (snd (fst (logical_instr_sub1 (fst instr) (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s2)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1))))))))) (snd (fst (logical_instr_sub1 (fst instr) (logical_result (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s2)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))))))))" by (meson a2 get_curr_win2_low_equal logical_instr_sub1_low_equal write_reg_low_equal) then show ?thesis using \\wa w. user_reg_val w wa (snd (fst (get_curr_win () s2))) = user_reg_val w wa (snd (fst (get_curr_win () s1)))\ by presburger qed qed lemma shift_instr_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (shift_instr instr s1)) \ t2 = snd (fst (shift_instr instr s2))" shows "low_equal t1 t2" proof (cases "(fst instr = shift_type SLL) \ (get_operand_w5 ((snd instr)!3) \ 0)") case True then show ?thesis using a1 apply (simp add: shift_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def) apply (simp add: bind_def h1_def h2_def Let_def case_prod_unfold) apply auto apply (simp_all add: get_curr_win_low_equal) proof - assume a1: "low_equal s1 s2" assume "t2 = snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) << unat (ucast (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 2)) (snd (fst (get_curr_win () s2))))::word5)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))" assume "t1 = snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) << unat (ucast (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 2)) (snd (fst (get_curr_win () s1))))::word5)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))" have "\w wa wb. low_equal (snd (fst (write_reg w wa wb s1))) (snd (fst (write_reg w wa wb s2)))" using a1 by (metis write_reg_low_equal) then show "low_equal (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) << unat (ucast (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 2)) (snd (fst (get_curr_win () s1))))::word5)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) << unat (ucast (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 2)) (snd (fst (get_curr_win () s2))))::word5)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))" using a1 by (simp add: get_curr_win_def simpler_gets_def user_reg_val_low_equal) next assume a2: "low_equal s1 s2" assume "t1 = snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) << unat (get_operand_w5 (snd instr ! 2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))" assume "t2 = snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) << unat (get_operand_w5 (snd instr ! 2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))" have "\w wa wb. low_equal (snd (fst (write_reg w wa wb s1))) (snd (fst (write_reg w wa wb s2)))" using a2 by (metis write_reg_low_equal) then show "low_equal (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) << unat (get_operand_w5 (snd instr ! 2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) << unat (get_operand_w5 (snd instr ! 2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))" proof - assume a1: "\w wa wb. low_equal (snd (fst (write_reg w wa wb s1))) (snd (fst (write_reg w wa wb s2)))" have "\u s. fst (get_curr_win u s) = (ucast (get_CWP (cpu_reg_val PSR s))::'a word, s)" by (simp add: get_curr_win_def simpler_gets_def) then show ?thesis using a1 assms user_reg_val_low_equal by fastforce qed qed next case False then have f1: "\((fst instr = shift_type SLL) \ (get_operand_w5 ((snd instr)!3) \ 0))" by auto then show ?thesis proof (cases "(fst instr = shift_type SRL) \ (get_operand_w5 ((snd instr)!3) \ 0)") case True then show ?thesis using a1 f1 apply (simp add: shift_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def) apply (simp add: bind_def h1_def h2_def Let_def case_prod_unfold) apply auto apply (simp_all add: get_curr_win_low_equal) proof - assume a1: "low_equal s1 s2" assume "t2 = snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) >> unat (ucast (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 2)) (snd (fst (get_curr_win () s2))))::word5)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))" assume "t1 = snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) >> unat (ucast (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 2)) (snd (fst (get_curr_win () s1))))::word5)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))" have "\u s. fst (get_curr_win u s) = (ucast (get_CWP (cpu_reg_val PSR s))::'a word, s)" by (simp add: get_curr_win_def simpler_gets_def) then show "low_equal (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) >> unat (ucast (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 2)) (snd (fst (get_curr_win () s1))))::word5)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) >> unat (ucast (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 2)) (snd (fst (get_curr_win () s2))))::word5)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))" using a1 user_reg_val_low_equal write_reg_low_equal by fastforce next assume a2: "low_equal s1 s2" assume "t1 = snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) >> unat (get_operand_w5 (snd instr ! 2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))" assume "t2 = snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) >> unat (get_operand_w5 (snd instr ! 2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))" have "\u s. fst (get_curr_win u s) = (ucast (get_CWP (cpu_reg_val PSR s))::'a word, s)" by (simp add: get_curr_win_def simpler_gets_def) then show "low_equal (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) >> unat (get_operand_w5 (snd instr ! 2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) >> unat (get_operand_w5 (snd instr ! 2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))" using a2 user_reg_val_low_equal write_reg_low_equal by fastforce qed next case False then have f2: "\((fst instr = shift_type SRL) \ (get_operand_w5 ((snd instr)!3) \ 0))" by auto then show ?thesis proof (cases "(fst instr = shift_type SRA) \ (get_operand_w5 ((snd instr)!3) \ 0)") case True then show ?thesis using a1 f1 f2 apply (simp add: shift_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def) apply (simp add: bind_def h1_def h2_def Let_def case_prod_unfold) apply auto apply (simp_all add: get_curr_win_low_equal) proof - assume a1: "low_equal s1 s2" assume "t1 = snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) >>> unat (ucast (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 2)) (snd (fst (get_curr_win () s1))))::word5)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))" assume "t2 = snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) >>> unat (ucast (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 2)) (snd (fst (get_curr_win () s2))))::word5)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))" have "\w wa. user_reg_val wa w (snd (fst (get_curr_win () s1))) = user_reg_val wa w (snd (fst (get_curr_win () s2)))" using a1 by (meson get_curr_win2_low_equal user_reg_val_low_equal) then show "low_equal (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) >>> unat (ucast (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 2)) (snd (fst (get_curr_win () s1))))::word5)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) >>> unat (ucast (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 2)) (snd (fst (get_curr_win () s2))))::word5)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))" using a1 by (metis (no_types) get_curr_win2_low_equal write_reg_low_equal) next assume a2: "low_equal s1 s2" assume "t1 = snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) >>> unat (get_operand_w5 (snd instr ! 2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))" assume "t2 = snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) >>> unat (get_operand_w5 (snd instr ! 2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))" have "\w wa. user_reg_val wa w (snd (fst (get_curr_win () s1))) = user_reg_val wa w (snd (fst (get_curr_win () s2)))" using a2 by (meson get_curr_win2_low_equal user_reg_val_low_equal) then show "low_equal (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) >>> unat (get_operand_w5 (snd instr ! 2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) >>> unat (get_operand_w5 (snd instr ! 2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))" using a2 get_curr_win2_low_equal write_reg_low_equal by fastforce qed next case False then show ?thesis using a1 f1 f2 apply (simp add: shift_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def) apply (simp add: bind_def h1_def h2_def Let_def case_prod_unfold) apply (simp add: return_def) using get_curr_win2_low_equal by blast qed qed qed lemma add_instr_sub1_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (add_instr_sub1 instr_name result rs1_val operand2 s1)) \ t2 = snd (fst (add_instr_sub1 instr_name result rs1_val operand2 s2))" shows "low_equal t1 t2" proof (cases "instr_name = arith_type ADDcc \ instr_name = arith_type ADDXcc") case True then show ?thesis using a1 apply (simp add: add_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (clarsimp simp add: cpu_reg_val_low_equal) using write_cpu_low_equal by blast next case False then show ?thesis using a1 apply (simp add: add_instr_sub1_def) by (simp add: return_def) qed lemma add_instr_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (add_instr instr s1)) \ t2 = snd (fst (add_instr instr s2))" shows "low_equal t1 t2" proof - have f1: "low_equal s1 s2 \ t1 = snd (fst (add_instr_sub1 (fst instr) (if fst instr = arith_type ADD \ fst instr = arith_type ADDcc then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1 else user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) \ 0 then if fst instr = arith_type ADD \ fst instr = arith_type ADDcc then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1 else user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1))))) else user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))))) \ t2 = snd (fst (add_instr_sub1 (fst instr) (if fst instr = arith_type ADD \ fst instr = arith_type ADDcc then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2 else user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) \ 0 then if fst instr = arith_type ADD \ fst instr = arith_type ADDcc then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2 else user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2))))) else user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))))" using a1 apply (simp add: add_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (simp add: case_prod_unfold) then show ?thesis proof (cases "get_operand_w5 (snd instr ! 3) \ 0") case True then have f2: "get_operand_w5 (snd instr ! 3) \ 0" by auto then show ?thesis proof (cases "fst instr = arith_type ADD \ fst instr = arith_type ADDcc") case True then show ?thesis using f1 f2 apply clarsimp proof - assume a1: "low_equal s1 s2" assume "t1 = snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))))" assume a2: "t2 = snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))))" have f3: "\is. get_operand2 is s1 = get_operand2 is s2" using a1 by (metis get_operand2_low_equal) have f4: "fst (fst (get_curr_win () s1)) = fst (fst (get_curr_win () s2))" using a1 by (meson get_curr_win_low_equal) have "\s. snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) s + get_operand2 (snd instr) s2) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) s) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) s + get_operand2 (snd instr) s2) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))))) = t2 \ \ low_equal s (snd (fst (get_curr_win () s2)))" using a2 user_reg_val_low_equal by fastforce then show "low_equal (snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))))) (snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))))))" using f4 f3 a2 a1 by (metis (no_types) add_instr_sub1_low_equal get_curr_win2_low_equal write_reg_low_equal) qed next case False then show ?thesis using f1 f2 apply clarsimp proof - assume a1: "low_equal s1 s2" have f2: "\s sa sb w wa wb sc. (\ low_equal s sa \ sb \ snd (fst (write_reg w (wa::'a word) wb s)) \ sc \ snd (fst (write_reg w wa wb sa))) \ low_equal sb sc" by (meson write_reg_low_equal) have f3: "gets (\s. ucast (get_CWP (cpu_reg_val PSR s))::'a word) = get_curr_win ()" by (simp add: get_curr_win_def) then have "((ucast (get_CWP (cpu_reg_val PSR s1)), s1), False) = (fst (get_curr_win () s1), snd (get_curr_win () s1))" by (metis (no_types) prod.collapse simpler_gets_def) then have "(ucast (get_CWP (cpu_reg_val PSR s1)), s1) = fst (get_curr_win () s1) \ \ snd (get_curr_win () s1)" by blast then have f4: "ucast (get_CWP (cpu_reg_val PSR s1)) = fst (fst (get_curr_win () s1)) \ s1 = snd (fst (get_curr_win () s1))" by (metis (no_types) prod.collapse prod.simps(1)) have "((ucast (get_CWP (cpu_reg_val PSR s2)), s2), False) = (fst (get_curr_win () s2), snd (get_curr_win () s2))" using f3 by (metis (no_types) prod.collapse simpler_gets_def) then have "(ucast (get_CWP (cpu_reg_val PSR s2)), s2) = fst (get_curr_win () s2) \ \ snd (get_curr_win () s2)" by blast then have f5: "ucast (get_CWP (cpu_reg_val PSR s2)) = fst (fst (get_curr_win () s2)) \ s2 = snd (fst (get_curr_win () s2))" by (metis prod.collapse prod.simps(1)) then have f6: "low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2))) = low_equal s1 s2" using f4 by presburger have f7: "fst (fst (get_curr_win () s1)) = ucast (get_CWP (cpu_reg_val PSR s1))" using f4 by presburger have f8: "cpu_reg_val PSR s1 = cpu_reg_val PSR s2" using a1 by (meson cpu_reg_val_low_equal) have f9: "user_reg_val (ucast (get_CWP (cpu_reg_val PSR s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) = user_reg_val (ucast (get_CWP (cpu_reg_val PSR s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))" using f6 a1 by (meson user_reg_val_low_equal) have f10: "ucast (get_CWP (cpu_reg_val PSR s2)) = fst (fst (get_curr_win () s2))" using f5 by meson have f11: "\s sa is. \ low_equal (s::'a sparc_state) sa \ get_operand2 is s = get_operand2 is sa" using get_operand2_low_equal by blast then have f12: "user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1))))) = user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))" using f9 f8 f5 f4 a1 by auto then have "low_equal (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))" using f10 f8 f6 f4 f2 a1 by simp then show "low_equal (snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))))) (snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))))))" using f12 f11 f10 f9 f8 f7 a1 add_instr_sub1_low_equal by fastforce qed qed next case False then have f3: "\ get_operand_w5 (snd instr ! 3) \ 0" by auto then show ?thesis proof (cases "fst instr = arith_type ADD \ fst instr = arith_type ADDcc") case True then show ?thesis using f1 f3 apply clarsimp proof - assume a1: "low_equal s1 s2" assume "t1 = snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1))))))))" assume "t2 = snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))))))))" have f2: "\is. get_operand2 is s1 = get_operand2 is s2" using a1 by (meson get_operand2_low_equal) have f3: "fst (fst (get_curr_win () s1)) = fst (fst (get_curr_win () s2))" using a1 by (meson get_curr_win_low_equal) have "\w wa. user_reg_val wa w (snd (fst (get_curr_win () s1))) = user_reg_val wa w (snd (fst (get_curr_win () s2)))" using a1 by (meson get_curr_win2_low_equal user_reg_val_low_equal) then show "low_equal (snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1))))))))) (snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))))))))" using f3 f2 a1 by (metis (no_types) add_instr_sub1_low_equal get_curr_win2_low_equal write_reg_low_equal) qed next case False then show ?thesis using f1 f3 apply clarsimp proof - assume a1: "low_equal s1 s2" have f2: "gets (\s. ucast (get_CWP (cpu_reg_val PSR s))::'a word) = get_curr_win ()" by (simp add: get_curr_win_def) then have "((ucast (get_CWP (cpu_reg_val PSR s1)), s1), False) = (fst (get_curr_win () s1), snd (get_curr_win () s1))" by (metis (no_types) prod.collapse simpler_gets_def) then have "(ucast (get_CWP (cpu_reg_val PSR s1)), s1) = fst (get_curr_win () s1) \ \ snd (get_curr_win () s1)" by fastforce then have f3: "ucast (get_CWP (cpu_reg_val PSR s1)) = fst (fst (get_curr_win () s1)) \ s1 = snd (fst (get_curr_win () s1))" by (metis prod.collapse prod.simps(1)) have "((ucast (get_CWP (cpu_reg_val PSR s2)), s2), False) = (fst (get_curr_win () s2), snd (get_curr_win () s2))" using f2 by (metis (no_types) prod.collapse simpler_gets_def) then have "(ucast (get_CWP (cpu_reg_val PSR s2)), s2) = fst (get_curr_win () s2) \ \ snd (get_curr_win () s2)" by fastforce then have f4: "ucast (get_CWP (cpu_reg_val PSR s2)) = fst (fst (get_curr_win () s2)) \ s2 = snd (fst (get_curr_win () s2))" by (metis (no_types) prod.collapse prod.simps(1)) then have f5: "low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2))) = low_equal s1 s2" using f3 by presburger have f6: "fst (fst (get_curr_win () s1)) = ucast (get_CWP (cpu_reg_val PSR s1))" using f3 by auto have f7: "cpu_reg_val PSR s1 = cpu_reg_val PSR s2" using a1 by (meson cpu_reg_val_low_equal) have f8: "\s sa w wa. \ low_equal s sa \ user_reg_val (w::'a word) wa s = user_reg_val w wa sa" by (meson user_reg_val_low_equal) have f9: "ucast (get_CWP (cpu_reg_val PSR s2)) = fst (fst (get_curr_win () s2))" using f4 by meson have "\s sa is. \ low_equal (s::'a sparc_state) sa \ get_operand2 is s = get_operand2 is sa" using get_operand2_low_equal by blast then have f10: "get_operand2 (snd instr) s1 = get_operand2 (snd instr) s2" using a1 by meson have f11: "cpu_reg_val PSR (snd (fst (get_curr_win () s2))) = cpu_reg_val PSR s1" using f4 a1 by (simp add: cpu_reg_val_low_equal) have f12: "user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1))) = 0" by (meson user_reg_val_def) have "user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))) = 0" by (meson user_reg_val_def) then have "low_equal (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))))))" using f12 f9 f7 f5 f3 a1 write_reg_low_equal by fastforce then have "low_equal (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))))) \ snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))))))) = snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (if get_operand_w5 (snd instr ! Suc 0) = 0 then 0 else user_reg (snd (fst (get_curr_win () s2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))))))) \ snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))))))) = snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (if get_operand_w5 (snd instr ! Suc 0) = 0 then 0 else user_reg (snd (fst (get_curr_win () s2))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))))))))" using f11 f10 f9 f8 f7 f6 f5 f3 a1 by (simp add: user_reg_val_def) then show "low_equal (snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) + get_operand2 (snd instr) s1 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1))))))))) (snd (fst (add_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) + get_operand2 (snd instr) s2 + ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))))))))" using add_instr_sub1_low_equal by blast qed qed qed qed lemma sub_instr_sub1_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (sub_instr_sub1 instr_name result rs1_val operand2 s1)) \ t2 = snd (fst (sub_instr_sub1 instr_name result rs1_val operand2 s2))" shows "low_equal t1 t2" proof (cases "instr_name = arith_type SUBcc \ instr_name = arith_type SUBXcc") case True then show ?thesis using a1 apply (simp add: sub_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (clarsimp simp add: cpu_reg_val_low_equal) using write_cpu_low_equal by blast next case False then show ?thesis using a1 apply (simp add: sub_instr_sub1_def) by (simp add: return_def) qed lemma sub_instr_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (sub_instr instr s1)) \ t2 = snd (fst (sub_instr instr s2))" shows "low_equal t1 t2" proof - have f1: "low_equal s1 s2 \ t1 = snd (fst (sub_instr_sub1 (fst instr) (if fst instr = arith_type SUB \ fst instr = arith_type SUBcc then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 else user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) \ 0 then if fst instr = arith_type SUB \ fst instr = arith_type SUBcc then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 else user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1))))) else user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))))) \ t2 = snd (fst (sub_instr_sub1 (fst instr) (if fst instr = arith_type SUB \ fst instr = arith_type SUBcc then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2 else user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) \ 0 then if fst instr = arith_type SUB \ fst instr = arith_type SUBcc then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2 else user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2))))) else user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))))" using a1 apply (simp add: sub_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (simp add: case_prod_unfold) then show ?thesis proof (cases "get_operand_w5 (snd instr ! 3) \ 0") case True then have f2: "get_operand_w5 (snd instr ! 3) \ 0" by auto then show ?thesis proof (cases "fst instr = arith_type SUB \ fst instr = arith_type SUBcc") case True then show ?thesis using f1 f2 apply clarsimp proof - assume a1: "low_equal s1 s2" assume a2: "t1 = snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))))" assume a3: "t2 = snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))))" then have f4: "snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s1) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s1) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))))) = t2" using a1 by (simp add: get_operand2_low_equal) have "\s. \ low_equal (snd (fst (get_curr_win () s1))) s \ snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) s - get_operand2 (snd instr) s1) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) s) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) s - get_operand2 (snd instr) s1) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))))) = t1" using a2 a1 by (simp add: get_curr_win_low_equal user_reg_val_low_equal) then show "low_equal (snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))))) (snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))))))" using f4 a3 a2 a1 by (metis (no_types) get_curr_win2_low_equal sub_instr_sub1_low_equal write_reg_low_equal) qed next case False then show ?thesis using f1 f2 apply clarsimp proof - assume a1: "low_equal s1 s2" have f2: "fst (get_curr_win () s1) = (ucast (get_CWP (cpu_reg_val PSR s1)), s1)" by (simp add: get_curr_win_def simpler_gets_def) have f3: "cpu_reg_val PSR s1 = cpu_reg_val PSR s2" using a1 by (meson cpu_reg_val_low_equal) then have f4: "user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) = user_reg_val (ucast (get_CWP (cpu_reg_val PSR s2))) (get_operand_w5 (snd instr ! Suc 0)) s1" using f2 by simp have f5: "\s sa is. \ low_equal (s::'a sparc_state) sa \ get_operand2 is s = get_operand2 is sa" using get_operand2_low_equal by blast then have f6: "sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (user_reg_val (ucast (get_CWP (cpu_reg_val PSR s2))) (get_operand_w5 (snd instr ! Suc 0)) s2) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))) = sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))" using f4 a1 by (simp add: user_reg_val_low_equal) have f7: "fst (get_curr_win () s2) = (ucast (get_CWP (cpu_reg_val PSR s2)), s2)" by (simp add: get_curr_win_def simpler_gets_def) then have f8: "user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2))))) = user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))" using f5 f2 a1 by (simp add: cpu_reg_val_low_equal user_reg_val_low_equal) then have f9: "sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (user_reg_val (ucast (get_CWP (cpu_reg_val PSR s2))) (get_operand_w5 (snd instr ! Suc 0)) s2) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))) = sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))" using f7 by fastforce have "write_reg (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (ucast (get_CWP (cpu_reg_val PSR s2))) (get_operand_w5 (snd instr ! 3)) s2 = write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))" using f8 f7 by simp then have "low_equal (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))" using f3 f2 a1 by (metis (no_types) prod.sel(1) prod.sel(2) write_reg_low_equal) then show "low_equal (snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))))) (snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))))))" using f9 f6 by (metis (no_types) sub_instr_sub1_low_equal) qed qed next case False then have f3: "\ get_operand_w5 (snd instr ! 3) \ 0" by auto then show ?thesis proof (cases "fst instr = arith_type SUB \ fst instr = arith_type SUBcc") case True then show ?thesis using f1 f3 apply clarsimp proof - assume a1: "low_equal s1 s2" assume "t1 = snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1))))))))" assume "t2 = snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))))))))" have f2: "\is. get_operand2 is s1 = get_operand2 is s2" using a1 get_operand2_low_equal by blast have f3: "fst (fst (get_curr_win () s1)) = fst (fst (get_curr_win () s2))" using a1 by (meson get_curr_win_low_equal) have "\w wa. user_reg_val wa w (snd (fst (get_curr_win () s1))) = user_reg_val wa w (snd (fst (get_curr_win () s2)))" using a1 by (metis (no_types) get_curr_win2_low_equal user_reg_val_low_equal) then show "low_equal (snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1))))))))) (snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))))))))" using f3 f2 a1 by (metis (no_types) get_curr_win2_low_equal sub_instr_sub1_low_equal write_reg_low_equal) qed next case False then show ?thesis using f1 f3 apply clarsimp proof - assume a1: "low_equal s1 s2" have f2: "\s sa sb w wa wb sc. (\ low_equal s sa \ sb \ snd (fst (write_reg w (wa::'a word) wb s)) \ sc \ snd (fst (write_reg w wa wb sa))) \ low_equal sb sc" by (meson write_reg_low_equal) have "((ucast (get_CWP (cpu_reg_val PSR s1)), s1), False) = get_curr_win () s1" by (simp add: get_curr_win_def simpler_gets_def) then have f3: "ucast (get_CWP (cpu_reg_val PSR s1)) = fst (fst (get_curr_win () s1)) \ s1 = snd (fst (get_curr_win () s1))" by (metis (no_types) prod.collapse prod.simps(1)) have "((ucast (get_CWP (cpu_reg_val PSR s2)), s2), False) = get_curr_win () s2" by (simp add: get_curr_win_def simpler_gets_def) then have f4: "ucast (get_CWP (cpu_reg_val PSR s2)) = fst (fst (get_curr_win () s2)) \ s2 = snd (fst (get_curr_win () s2))" by (metis (no_types) prod.collapse prod.simps(1)) have f5: "\s sa sb sc w wa wb sd. (\ low_equal (s::'a sparc_state) sa \ sb \ snd (fst (sub_instr_sub1 sc w wa wb s)) \ sd \ snd (fst (sub_instr_sub1 sc w wa wb sa))) \ low_equal sb sd" by (meson sub_instr_sub1_low_equal) have "low_equal (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))))))" using f4 f3 f2 a1 by (simp add: cpu_reg_val_low_equal user_reg_val_low_equal) then show "low_equal (snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) - get_operand2 (snd instr) s1 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))))) (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) (get_operand2 (snd instr) s1) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (get_curr_win () s1))))))))) (snd (fst (sub_instr_sub1 (fst instr) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) - get_operand2 (snd instr) s2 - ucast (get_icc_C (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))))) (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) (get_operand2 (snd instr) s2) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))))))))" using f5 f4 f3 a1 by (simp add: cpu_reg_val_low_equal get_operand2_low_equal user_reg_val_low_equal) qed qed qed qed lemma mul_instr_sub1_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (mul_instr_sub1 instr_name result s1)) \ t2 = snd (fst (mul_instr_sub1 instr_name result s2))" shows "low_equal t1 t2" proof (cases "instr_name \ {arith_type SMULcc,arith_type UMULcc}") case True then show ?thesis using a1 apply (simp add: mul_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (clarsimp simp add: cpu_reg_val_low_equal) using write_cpu_low_equal by blast next case False then show ?thesis using a1 apply (simp add: mul_instr_sub1_def) by (simp add: return_def) qed lemma mul_instr_low_equal: \low_equal t1 t2\ if \low_equal s1 s2 \ t1 = snd (fst (mul_instr instr s1)) \ t2 = snd (fst (mul_instr instr s2))\ proof - from that have \low_equal s1 s2\ and t1: \t1 = snd (fst (mul_instr instr s1))\ and t2: \t2 = snd (fst (mul_instr instr s2))\ by simp_all have f2: "\s sa sb sc w sd. \ low_equal (s::'a sparc_state) sa \ sb \ snd (fst (mul_instr_sub1 sc w s)) \ sd \ snd (fst (mul_instr_sub1 sc w sa)) \ low_equal sb sd" using mul_instr_sub1_low_equal by blast have f3: "\s sa sb w wa wb sc. \ low_equal s sa \ sb \ snd (fst (write_reg w (wa::'a word) wb s)) \ sc \ snd (fst (write_reg w wa wb sa)) \ low_equal sb sc" by (meson write_reg_low_equal) have f4: "\s sa sb w c sc. \ low_equal (s::'a sparc_state) sa \ sb \ snd (fst (write_cpu w c s)) \ sc \ snd (fst (write_cpu w c sa)) \ low_equal sb sc" by (meson write_cpu_low_equal) have f6: "((ucast (get_CWP (cpu_reg_val PSR s1)), s1), False) = (fst (get_curr_win () s1), snd (get_curr_win () s1))" by (simp add: get_curr_win_def simpler_gets_def) have f7: "fst (fst (get_curr_win () s1)) = fst (fst (get_curr_win () s2))" using \low_equal s1 s2\ by (meson get_curr_win_low_equal) have "((ucast (get_CWP (cpu_reg_val PSR s2)), s2), False) = (fst (get_curr_win () s2), snd (get_curr_win () s2))" by (simp add: get_curr_win_def simpler_gets_def) then have f8: "ucast (get_CWP (cpu_reg_val PSR s2)) = fst (fst (get_curr_win () s2)) \ s2 = snd (fst (get_curr_win () s2))" by (metis prod.collapse prod.simps(1)) then have f9: "low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2)))" using f6 \low_equal s1 s2\ by (metis (no_types) prod.collapse prod.simps(1)) have f10: "\s sa w wa. \ low_equal s sa \ user_reg_val (w::'a word) wa s = user_reg_val w wa sa" using user_reg_val_low_equal by blast have f11: "get_operand2 (snd instr) s1 = get_operand2 (snd instr) (snd (fst (get_curr_win () s2)))" using f9 f6 by (metis (no_types) get_operand2_low_equal prod.collapse prod.simps(1)) then have f12: "uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2) = uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1)" using f10 f9 f8 f7 by presburger then have f13: "(word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) \ (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1)) else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) \ (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) \ (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2)) else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) \ low_equal (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2))))))" using f9 f4 by presburger have "get_operand_w5 (snd instr ! 3) = 0 \ low_equal (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) \ write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) = write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1))))))" using f10 f7 by force then have f14: "get_operand_w5 (snd instr ! 3) \ 0 \ low_equal (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1))))))))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2))))))))) \ \ low_equal (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2))))))" using f3 by metis then have f15: "low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))) \ fst instr \ arith_type UMULcc \ get_operand_w5 (snd instr ! 3) \ 0 \ (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) \ (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1)) else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) \ (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) \ (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2)) else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))" using f13 f12 f2 by fastforce have f16: "user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) = user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))" using f10 f9 f7 by presburger { assume "fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))" moreover { assume "\ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))" moreover { assume "low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))) \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))" moreover { assume "mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))))) \ mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))" then have "write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))) \ write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) \ fst instr \ arith_type UMULcc" by fastforce } moreover { assume "mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))) \ mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))" then have "write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))) \ write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) \ fst instr \ arith_type UMULcc" by fastforce } ultimately have "write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))) \ write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) \ write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))) \ write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) \ fst instr \ arith_type UMULcc" by force } ultimately have "write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))) \ write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) \ write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))) \ write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) \ \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))) \ fst instr \ arith_type UMULcc" by fastforce } ultimately have "fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))) \ write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))) = write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) \ write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))) = write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) \ get_operand_w5 (snd instr ! 3) = 0 \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))))" by blast } moreover { assume "\ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))" moreover { assume "\ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))) \ snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1))))))))))) = snd (fst (mul_instr_sub1 (arith_type UMULcc) (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1))))))))))) \ snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))) = snd (fst (mul_instr_sub1 (arith_type UMULcc) (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))" then have "\ low_equal (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1))))))))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))" using f2 by blast moreover { assume "(if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) \ (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))))" moreover { assume "(if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) \ ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))" then have "(if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) = (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1)) else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) \ get_operand_w5 (snd instr ! 3) = 0" by (metis f11 f16 f8) } ultimately have "(if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) = (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1)) else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) \ get_operand_w5 (snd instr ! 3) = 0" by fastforce } ultimately have "fst instr = arith_type UMULcc \ (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) = (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1)) else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) \ get_operand_w5 (snd instr ! 3) = 0" using f13 f7 f3 by fastforce } moreover { assume "mul_instr_sub1 (arith_type UMULcc) (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2))))))))) \ mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))" moreover { assume "(if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) \ ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))" then have "(if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) = (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1)) else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) \ get_operand_w5 (snd instr ! 3) = 0" by (metis f11 f16 f8) } ultimately have "fst instr = arith_type UMULcc \ (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) = (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1)) else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) \ get_operand_w5 (snd instr ! 3) = 0" by fastforce } ultimately have "fst instr = arith_type UMULcc \ get_operand_w5 (snd instr ! 3) = 0" using f12 by fastforce } moreover { assume "write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))) \ write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2))))))" then have "fst instr = arith_type UMULcc \ get_operand_w5 (snd instr ! 3) = 0" by presburger } ultimately have "fst instr = arith_type UMULcc \ get_operand_w5 (snd instr ! 3) = 0 \ get_operand_w5 (snd instr ! 3) = 0 \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))))" by force moreover { assume "fst instr \ arith_type UMULcc" { assume "fst instr \ arith_type UMULcc \ low_equal (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1))))))))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))" moreover { assume "fst instr \ arith_type UMULcc \ low_equal (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1))))))))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2))))))))) \ snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))) = snd (fst (mul_instr_sub1 (arith_type UMUL) (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))" then have "(fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ fst instr \ arith_type UMUL \ fst instr \ arith_type UMULcc \ (get_operand_w5 (snd instr ! 3) = 0 \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))))) \ get_operand_w5 (snd instr ! 3) = 0" using f2 by presburger } ultimately have "fst instr \ arith_type UMULcc \ (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) \ ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ fst instr \ arith_type UMUL \ fst instr \ arith_type UMULcc \ (get_operand_w5 (snd instr ! 3) = 0 \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))))) \ get_operand_w5 (snd instr ! 3) = 0" by fastforce } then have "(get_operand_w5 (snd instr ! 3) = 0 \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))))) \ get_operand_w5 (snd instr ! 3) = 0" using f16 f11 f9 f8 f7 f4 f3 f2 by force } moreover { assume "get_operand_w5 (snd instr ! 3) = 0" moreover { assume "(fst instr = arith_type UMUL \ fst instr = arith_type UMULcc) \ get_operand_w5 (snd instr ! 3) = 0" moreover { assume "((fst instr = arith_type UMUL \ fst instr = arith_type UMULcc) \ get_operand_w5 (snd instr ! 3) = 0) \ \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))" moreover { assume "((fst instr = arith_type UMUL \ fst instr = arith_type UMULcc) \ get_operand_w5 (snd instr ! 3) = 0) \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))) \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))" then have "((fst instr = arith_type UMUL \ fst instr = arith_type UMULcc) \ get_operand_w5 (snd instr ! 3) = 0) \ arith_type UMUL \ arith_type UMULcc \ fst instr \ arith_type UMULcc \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc) \ get_operand_w5 (snd instr ! 3) = 0" by fastforce then have "((fst instr = arith_type UMUL \ fst instr = arith_type UMULcc) \ get_operand_w5 (snd instr ! 3) = 0) \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ fst instr \ arith_type UMULcc \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc) \ get_operand_w5 (snd instr ! 3) = 0" by force } ultimately have "((fst instr = arith_type UMUL \ fst instr = arith_type UMULcc) \ get_operand_w5 (snd instr ! 3) = 0) \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ ((fst instr = arith_type UMUL \ fst instr = arith_type UMULcc) \ get_operand_w5 (snd instr ! 3) = 0) \ \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))) \ fst instr \ arith_type UMULcc \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc) \ get_operand_w5 (snd instr ! 3) = 0" by simp } ultimately have "((fst instr = arith_type UMUL \ fst instr = arith_type UMULcc) \ get_operand_w5 (snd instr ! 3) = 0) \ \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) = 0 then user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))) else ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))) \ fst instr \ arith_type UMULcc \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc) \ get_operand_w5 (snd instr ! 3) = 0 \ (get_operand_w5 (snd instr ! 3) = 0 \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))))) \ (get_operand_w5 (snd instr ! 3) \ 0 \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))))" by auto then have "fst instr \ arith_type UMULcc \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc) \ get_operand_w5 (snd instr ! 3) = 0 \ (get_operand_w5 (snd instr ! 3) = 0 \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))))) \ (get_operand_w5 (snd instr ! 3) \ 0 \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))))" using f15 by presburger then have "(get_operand_w5 (snd instr ! 3) = 0 \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))))) \ (get_operand_w5 (snd instr ! 3) \ 0 \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))))" using f14 f13 f12 f2 by force } ultimately have "(get_operand_w5 (snd instr ! 3) = 0 \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))))) \ (get_operand_w5 (snd instr ! 3) \ 0 \ (fst instr \ arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMUL \ fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))))" using f16 f14 f11 f9 f8 f4 f2 by fastforce } ultimately have "(get_operand_w5 (snd instr ! 3) \ 0 \ (fst instr = arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMUL \ fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))))) \ (get_operand_w5 (snd instr ! 3) = 0 \ (fst instr = arith_type UMUL \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMUL) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr = arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (arith_type UMULcc) (ucast (word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))) \ (fst instr \ arith_type UMUL \ fst instr \ arith_type UMULcc \ low_equal (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))::word64) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (mul_instr_sub1 (fst instr) (ucast (word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64)) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (write_cpu (ucast ((word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))::word64) >> 32)) Y (snd (fst (get_curr_win () s2))))))))))))))" by blast moreover from t1 have \t1 = snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) \ 0 then ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) else user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1))))))) (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * uint (get_operand2 (snd instr) s1))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s1))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))) * sint (get_operand2 (snd instr) s1))) >> 32)) Y (snd (fst (get_curr_win () s1)))))))))))\ by (simp add: mul_instr_def) (simp add: simpler_gets_def bind_def h1_def h2_def Let_def case_prod_unfold) moreover from t2 have \t2 = snd (fst (mul_instr_sub1 (fst instr) (ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2)))) (snd (fst (write_reg (if get_operand_w5 (snd instr ! 3) \ 0 then ucast (if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) else user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2))))))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (write_cpu (ucast ((if fst instr = arith_type UMUL \ fst instr = arith_type UMULcc then word_of_int (uint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * uint (get_operand2 (snd instr) s2))::word64 else word_of_int (sint (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))) * sint (get_operand2 (snd instr) s2))) >> 32)) Y (snd (fst (get_curr_win () s2)))))))))))\ by (simp add: mul_instr_def) (simp add: simpler_gets_def bind_def h1_def h2_def Let_def case_prod_unfold) ultimately show ?thesis by (simp add: mul_instr_def) qed lemma div_write_new_val_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (div_write_new_val i result temp_V s1)) \ t2 = snd (fst (div_write_new_val i result temp_V s2))" shows "low_equal t1 t2" proof (cases "(fst i) \ {arith_type UDIVcc,arith_type SDIVcc}") case True then show ?thesis using a1 apply (simp add: div_write_new_val_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (clarsimp simp add: cpu_reg_val_low_equal) using write_cpu_low_equal by blast next case False then show ?thesis using a1 apply (simp add: div_write_new_val_def) by (simp add: return_def) qed lemma div_comp_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (div_comp instr rs1 rd operand2 s1)) \ t2 = snd (fst (div_comp instr rs1 rd operand2 s2))" shows "low_equal t1 t2" using a1 apply (simp add: div_comp_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (clarsimp simp add: get_curr_win_low_equal) proof - assume a1: "low_equal s1 s2" have f2: "\s sa sb w wa wb sc. \ low_equal s sa \ sb \ snd (fst (write_reg w (wa::'a word) wb s)) \ sc \ snd (fst (write_reg w wa wb sa)) \ low_equal sb sc" by (meson write_reg_low_equal) have f3: "gets (\s. ucast (get_CWP (cpu_reg_val PSR s))::'a word) = get_curr_win ()" by (simp add: get_curr_win_def) then have "((ucast (get_CWP (cpu_reg_val PSR s1)), s1), False) = (fst (get_curr_win () s1), snd (get_curr_win () s1))" by (metis (no_types) prod.collapse simpler_gets_def) then have f4: "ucast (get_CWP (cpu_reg_val PSR s1)) = fst (fst (get_curr_win () s1)) \ s1 = snd (fst (get_curr_win () s1))" by (metis prod.collapse prod.simps(1)) have "((ucast (get_CWP (cpu_reg_val PSR s2)), s2), False) = (fst (get_curr_win () s2), snd (get_curr_win () s2))" using f3 by (metis (no_types) prod.collapse simpler_gets_def) then have f5: "ucast (get_CWP (cpu_reg_val PSR s2)) = fst (fst (get_curr_win () s2)) \ s2 = snd (fst (get_curr_win () s2))" by (metis (no_types) prod.collapse prod.simps(1)) then have f6: "low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2)))" using f4 a1 by presburger have f7: "\s sa sb p w wa sc. \ low_equal (s::'a sparc_state) sa \ sb \ snd (fst (div_write_new_val p w wa s)) \ sc \ snd (fst (div_write_new_val p w wa sa)) \ low_equal sb sc" by (meson div_write_new_val_low_equal) have f8: "cpu_reg_val PSR s2 = cpu_reg_val PSR s1" using a1 by (simp add: cpu_reg_val_def low_equal_def) then have "fst (fst (get_curr_win () s2)) = ucast (get_CWP (cpu_reg_val PSR s1))" using f5 by presburger then have f9: "fst (fst (get_curr_win () s2)) = fst (fst (get_curr_win () s1))" using f4 by presburger have f10: "fst (fst (get_curr_win () s1)) = fst (fst (get_curr_win () s2))" using f8 f5 f4 by presburger have f11: "(word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))::word64) = word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))" using f5 f4 a1 by (metis (no_types) cpu_reg_val_def low_equal_def user_reg_val_low_equal) have f12: "ucast (get_CWP (cpu_reg_val PSR s1)) = fst (fst (get_curr_win () s2))" using f8 f5 by presburger then have "rd = 0 \ (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) = user_reg_val (ucast (get_CWP (cpu_reg_val PSR s1))) 0 (snd (fst (get_curr_win () s1)))" using f6 user_reg_val_low_equal by fastforce then have f13: "rd = 0 \ write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (ucast (get_CWP (cpu_reg_val PSR s1))) 0 (snd (fst (get_curr_win () s1))) = write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s1))) rd (snd (fst (get_curr_win () s1))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s1)))" using f12 f10 by presburger have f14: "write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (ucast (get_CWP (cpu_reg_val PSR s1))) rd (snd (fst (get_curr_win () s2))) = write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2)))" using f12 f11 by auto have "write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (ucast (get_CWP (cpu_reg_val PSR s1))) rd (snd (fst (get_curr_win () s1))) = write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s1))) rd (snd (fst (get_curr_win () s1))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s1))) \ write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (ucast (get_CWP (cpu_reg_val PSR s1))) rd (snd (fst (get_curr_win () s2))) = write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))) \ low_equal (snd (fst (write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s1))) rd (snd (fst (get_curr_win () s1))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))))))" using f6 f2 by metis moreover { assume "low_equal (snd (fst (write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s1))) rd (snd (fst (get_curr_win () s1))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))))))" then have "low_equal (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (snd (fst (write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s1))) rd (snd (fst (get_curr_win () s1))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s1))))))))) (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (snd (fst (write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2)))))))))" using f11 f9 f7 by metis moreover { assume "low_equal (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (snd (fst (write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s1))) rd (snd (fst (get_curr_win () s1))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s1))))))))) (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (snd (fst (write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))))))))) \ low_equal (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (snd (fst (write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s1))))))))) (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (snd (fst (write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2)))))))))" then have "div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2) = (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s1))) rd (snd (fst (get_curr_win () s1))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) \ rd = 0" by fastforce } ultimately have "div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2) = (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s1))) rd (snd (fst (get_curr_win () s1))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) \ rd = 0 \ (rd \ 0 \ low_equal (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (snd (fst (write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s1))))))))) (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (snd (fst (write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2)))))))))) \ (rd = 0 \ low_equal (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1))))))))) (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))))))))))" by fastforce } moreover { assume "write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (ucast (get_CWP (cpu_reg_val PSR s1))) rd (snd (fst (get_curr_win () s1))) \ write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s1))) rd (snd (fst (get_curr_win () s1))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s1)))" then have "div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2) \ (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s1))) rd (snd (fst (get_curr_win () s1))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2))" using f12 f9 by fastforce } moreover { assume "write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (ucast (get_CWP (cpu_reg_val PSR s1))) rd (snd (fst (get_curr_win () s2))) \ write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2)))" then have "rd = 0" using f14 by presburger } moreover { assume "rd = 0" then have "rd = 0 \ low_equal (snd (fst (write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s1))) rd (snd (fst (get_curr_win () s1))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))))))" using f13 f12 f6 f2 by metis then have "rd = 0 \ low_equal (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (snd (fst (write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s1))) rd (snd (fst (get_curr_win () s1))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s1))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s1))))))))) (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (snd (fst (write_reg (if rd = 0 then user_reg_val (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2))) else div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2)))))))))" using f11 f9 f7 by metis then have "(rd \ 0 \ low_equal (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (snd (fst (write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s1))))))))) (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (snd (fst (write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2)))))))))) \ (rd = 0 \ low_equal (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1))))))))) (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))))))))))" using f10 by fastforce } ultimately show "(rd \ 0 \ low_equal (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (snd (fst (write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s1))))))))) (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (snd (fst (write_reg (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (fst (fst (get_curr_win () s2))) rd (snd (fst (get_curr_win () s2)))))))))) \ (rd = 0 \ low_equal (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s1))))) operand2 >> 31))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1))))))))) (snd (fst (div_write_new_val instr (div_comp_result instr (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2)) (div_comp_temp_V instr (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 32)) (ucast (div_comp_temp_64bit instr (word_cat (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (user_reg_val (fst (fst (get_curr_win () s2))) rs1 (snd (fst (get_curr_win () s2))))) operand2 >> 31))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))))))))))" using f9 by fastforce qed lemma div_instr_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (div_instr instr s1)) \ t2 = snd (fst (div_instr instr s2))" shows "low_equal t1 t2" using a1 apply (simp add: div_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: return_def) apply (auto simp add: get_operand2_low_equal) apply (simp add: raise_trap_def add_trap_set_def simpler_modify_def) apply (auto simp add: traps_low_equal) apply (blast intro: mod_trap_low_equal) using div_comp_low_equal by blast lemma get_curr_win_traps_low_equal: assumes a1: "low_equal s1 s2" shows "low_equal (snd (fst (get_curr_win () s1)) \traps := insert some_trap (traps (snd (fst (get_curr_win () s1))))\) (snd (fst (get_curr_win () s2)) \traps := insert some_trap (traps (snd (fst (get_curr_win () s2))))\)" proof - from a1 have f1: "low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2)))" using get_curr_win2_low_equal by auto then have f2: "(traps (snd (fst (get_curr_win () s1)))) = (traps (snd (fst (get_curr_win () s2))))" using traps_low_equal by auto then show ?thesis using f1 f2 mod_trap_low_equal by fastforce qed lemma save_restore_instr_sub1_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (save_retore_sub1 result new_cwp rd s1)) \ t2 = snd (fst (save_retore_sub1 result new_cwp rd s2))" shows "low_equal t1 t2" using a1 apply (simp add: save_retore_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (auto simp add: cpu_reg_val_low_equal) using write_cpu_low_equal write_reg_low_equal by fastforce lemma get_WIM_bit_low_equal: \get_WIM_bit (nat ((uint (fst (fst (get_curr_win () s1))) - 1) mod NWINDOWS)) (cpu_reg_val WIM (snd (fst (get_curr_win () s1)))) = get_WIM_bit (nat ((uint (fst (fst (get_curr_win () s2))) - 1) mod NWINDOWS)) (cpu_reg_val WIM (snd (fst (get_curr_win () s2))))\ if \low_equal s1 s2\ proof - from that have f1: "low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2)))" using get_curr_win2_low_equal by blast then have f2: "(cpu_reg_val WIM (snd (fst (get_curr_win () s1)))) = (cpu_reg_val WIM (snd (fst (get_curr_win () s2))))" using cpu_reg_val_low_equal by auto from that have "(fst (fst (get_curr_win () s1))) = (fst (fst (get_curr_win () s2)))" using get_curr_win_low_equal by auto then show ?thesis using f1 f2 by auto qed lemma get_WIM_bit_low_equal2: \get_WIM_bit (nat ((uint (fst (fst (get_curr_win () s1))) + 1) mod NWINDOWS)) (cpu_reg_val WIM (snd (fst (get_curr_win () s1)))) = get_WIM_bit (nat ((uint (fst (fst (get_curr_win () s2))) + 1) mod NWINDOWS)) (cpu_reg_val WIM (snd (fst (get_curr_win () s2))))\ if \low_equal s1 s2\ proof - from that have f1: "low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2)))" using get_curr_win2_low_equal by blast then have f2: "(cpu_reg_val WIM (snd (fst (get_curr_win () s1)))) = (cpu_reg_val WIM (snd (fst (get_curr_win () s2))))" using cpu_reg_val_low_equal by auto from that have "(fst (fst (get_curr_win () s1))) = (fst (fst (get_curr_win () s2)))" using get_curr_win_low_equal by auto then show ?thesis using f1 f2 by auto qed lemma take_bit_5_mod_NWINDOWS_eq [simp]: \take_bit 5 (k mod NWINDOWS) = k mod NWINDOWS\ by (simp add: NWINDOWS_def take_bit_eq_mod) lemma save_restore_instr_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (save_restore_instr instr s1)) \ t2 = snd (fst (save_restore_instr instr s2))" shows "low_equal t1 t2" proof (cases "fst instr = ctrl_type SAVE") case True then have f1: "fst instr = ctrl_type SAVE" by auto then show ?thesis using a1 apply (simp add: save_restore_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (auto simp add: unsigned_of_int) apply (simp add: raise_trap_def add_trap_set_def simpler_modify_def) apply (simp add: get_curr_win_traps_low_equal) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: get_WIM_bit_low_equal) apply (simp add: get_WIM_bit_low_equal) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: get_curr_win_low_equal) using get_curr_win2_low_equal save_restore_instr_sub1_low_equal get_addr2_low_equal apply metis done next case False then show ?thesis using a1 apply (simp add: save_restore_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (auto simp add: unsigned_of_int) apply (simp add: raise_trap_def add_trap_set_def simpler_modify_def) apply (simp add: get_curr_win_traps_low_equal) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: get_WIM_bit_low_equal2) apply (simp add: get_WIM_bit_low_equal2) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: get_curr_win_low_equal) using get_curr_win2_low_equal save_restore_instr_sub1_low_equal get_addr2_low_equal apply metis done qed lemma call_instr_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (call_instr instr s1)) \ t2 = snd (fst (call_instr instr s2))" shows "low_equal t1 t2" using a1 apply (simp add: call_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (auto simp add: get_curr_win_low_equal) using cpu_reg_val_low_equal get_curr_win2_low_equal write_cpu_low_equal write_reg_low_equal proof - assume a1: "low_equal s1 s2" assume "t1 = snd (fst (write_cpu (cpu_reg_val PC (snd (fst (get_curr_win () s1))) + (ucast (get_operand_w30 (snd instr ! 0)) << 2)) nPC (snd (fst (write_cpu (cpu_reg_val nPC (snd (fst (get_curr_win () s1)))) PC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 15 (snd (fst (get_curr_win () s1)))))))))))" assume "t2 = snd (fst (write_cpu (cpu_reg_val PC (snd (fst (get_curr_win () s2))) + (ucast (get_operand_w30 (snd instr ! 0)) << 2)) nPC (snd (fst (write_cpu (cpu_reg_val nPC (snd (fst (get_curr_win () s2)))) PC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 15 (snd (fst (get_curr_win () s2)))))))))))" have "\c. cpu_reg_val c (snd (fst (get_curr_win () s1))) = cpu_reg_val c (snd (fst (get_curr_win () s2)))" using a1 by (meson cpu_reg_val_low_equal get_curr_win2_low_equal) then show "low_equal (snd (fst (write_cpu (cpu_reg_val PC (snd (fst (get_curr_win () s1))) + (ucast (get_operand_w30 (snd instr ! 0)) << 2)) nPC (snd (fst (write_cpu (cpu_reg_val nPC (snd (fst (get_curr_win () s1)))) PC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 15 (snd (fst (get_curr_win () s1)))))))))))) (snd (fst (write_cpu (cpu_reg_val PC (snd (fst (get_curr_win () s2))) + (ucast (get_operand_w30 (snd instr ! 0)) << 2)) nPC (snd (fst (write_cpu (cpu_reg_val nPC (snd (fst (get_curr_win () s2)))) PC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 15 (snd (fst (get_curr_win () s2))))))))))))" using a1 by (metis (no_types) get_curr_win2_low_equal write_cpu_low_equal write_reg_low_equal) qed lemma jmpl_instr_low_equal_sub1: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (write_cpu (get_addr (snd instr) (snd (fst (get_curr_win () s2)))) nPC (snd (fst (write_cpu (cpu_reg_val nPC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))) PC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))))))) \ t2 = snd (fst (write_cpu (get_addr (snd instr) (snd (fst (get_curr_win () s2)))) nPC (snd (fst (write_cpu (cpu_reg_val nPC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))) PC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))))))))" shows "low_equal t1 t2" proof - from a1 have f1: "low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2)))" using get_curr_win2_low_equal by blast then have f2: "(cpu_reg_val PC (snd (fst (get_curr_win () s1)))) = (cpu_reg_val PC (snd (fst (get_curr_win () s2))))" using cpu_reg_val_low_equal by blast then have f3: "low_equal (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))" using f1 write_reg_low_equal by fastforce then have "(cpu_reg_val nPC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))) = (cpu_reg_val nPC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))))" using cpu_reg_val_low_equal by auto then have f4: "low_equal (snd (fst (write_cpu (cpu_reg_val nPC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))) PC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s1))))))))) (snd (fst (write_cpu (cpu_reg_val nPC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2))))))) PC (snd (fst (write_reg (cpu_reg_val PC (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! 3)) (snd (fst (get_curr_win () s2)))))))))" using f3 write_cpu_low_equal by fastforce then show ?thesis using write_cpu_low_equal using assms by blast qed lemma jmpl_instr_low_equal_sub2: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (write_cpu (get_addr (snd instr) (snd (fst (get_curr_win () s2)))) nPC (snd (fst (write_cpu (cpu_reg_val nPC (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1))))))) PC (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1))))))))))) \ t2 = snd (fst (write_cpu (get_addr (snd instr) (snd (fst (get_curr_win () s2)))) nPC (snd (fst (write_cpu (cpu_reg_val nPC (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))))))) PC (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))))))))))" shows "low_equal t1 t2" proof - from a1 have f1: "low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2)))" using get_curr_win2_low_equal by blast then have f2: "(user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) = (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))))" using user_reg_val_low_equal by blast then have f3: "low_equal (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))))))" using f1 write_reg_low_equal by fastforce then have "(cpu_reg_val nPC (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1))))))) = (cpu_reg_val nPC (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))))))" using cpu_reg_val_low_equal by blast then have "low_equal (snd (fst (write_cpu (cpu_reg_val nPC (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1))))))) PC (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s1))))))))) (snd (fst (write_cpu (cpu_reg_val nPC (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2))))))) PC (snd (fst (write_reg (user_reg_val (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) 0 (snd (fst (get_curr_win () s2)))))))))" using f1 f2 f3 write_cpu_low_equal by fastforce then show ?thesis using write_cpu_low_equal using assms by blast qed lemma jmpl_instr_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (jmpl_instr instr s1)) \ t2 = snd (fst (jmpl_instr instr s2))" shows "low_equal t1 t2" using a1 apply (simp add: jmpl_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply auto apply (simp add: raise_trap_def add_trap_set_def simpler_modify_def) apply (simp add: get_curr_win_traps_low_equal) apply (simp add: get_addr2_low_equal) apply (simp add: get_addr2_low_equal) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp_all add: get_addr2_low_equal) apply (simp_all add: get_curr_win_low_equal) apply (case_tac "get_operand_w5 (snd instr ! 3) \ 0") apply auto using jmpl_instr_low_equal_sub1 apply blast apply (simp_all add: get_curr_win_low_equal) using jmpl_instr_low_equal_sub2 by blast lemma rett_instr_low_equal: assumes a1: "low_equal s1 s2 \ \ snd (rett_instr instr s1) \ \ snd (rett_instr instr s2) \ (((get_S (cpu_reg_val PSR s1)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s2)))::word1) = 0 \ t1 = snd (fst (rett_instr instr s1)) \ t2 = snd (fst (rett_instr instr s2))" shows "low_equal t1 t2" using a1 apply (simp add: rett_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply auto apply (simp add: raise_trap_def add_trap_set_def simpler_modify_def) apply (simp add: return_def) using mod_trap_low_equal traps_low_equal apply fastforce using cpu_reg_val_low_equal apply fastforce using cpu_reg_val_low_equal apply fastforce apply (simp add: bind_def h1_def h2_def Let_def) by (simp add: case_prod_unfold fail_def) lemma read_state_reg_low_equal: assumes a1: "low_equal s1 s2 \ (((get_S (cpu_reg_val PSR s1)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s2)))::word1) = 0 \ t1 = snd (fst (read_state_reg_instr instr s1)) \ t2 = snd (fst (read_state_reg_instr instr s2))" shows "low_equal t1 t2" proof (cases "(fst instr \ {sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR} \ (fst instr = sreg_type RDASR \ privileged_ASR (get_operand_w5 ((snd instr)!0))))") case True then have "(fst instr \ {sreg_type RDPSR,sreg_type RDWIM,sreg_type RDTBR} \ (fst instr = sreg_type RDASR \ privileged_ASR (get_operand_w5 ((snd instr)!0)))) \ (((get_S (cpu_reg_val PSR (snd (fst (get_curr_win () s1))))))::word1) = 0 \ (((get_S (cpu_reg_val PSR (snd (fst (get_curr_win () s2))))))::word1) = 0" by (metis assms get_curr_win_privilege) then show ?thesis using a1 apply (simp add: read_state_reg_instr_def) apply (simp add: Let_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: raise_trap_def add_trap_set_def simpler_modify_def) apply clarsimp using get_curr_win_traps_low_equal by auto next case False then have f1: "\((fst instr = sreg_type RDPSR \ fst instr = sreg_type RDWIM \ fst instr = sreg_type RDTBR \ fst instr = sreg_type RDASR \ privileged_ASR (get_operand_w5 (snd instr ! 0))))" by blast then show ?thesis proof (cases "illegal_instruction_ASR (get_operand_w5 ((snd instr)!0))") case True then show ?thesis using a1 f1 apply read_state_reg_instr_privilege_proof by (simp add: illegal_instruction_ASR_def) next case False then have f2: "\(illegal_instruction_ASR (get_operand_w5 ((snd instr)!0)))" by auto then show ?thesis proof (cases "(get_operand_w5 ((snd instr)!1)) \ 0") case True then have f3: "(get_operand_w5 ((snd instr)!1)) \ 0" by auto then show ?thesis proof (cases "fst instr = sreg_type RDY") case True then show ?thesis using a1 f1 f2 f3 apply (simp add: read_state_reg_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (auto simp add: get_curr_win_low_equal) using cpu_reg_val_low_equal get_curr_win2_low_equal write_reg_low_equal proof - assume "low_equal s1 s2" then have "low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2)))" by (meson get_curr_win2_low_equal) then show "low_equal (snd (fst (write_reg (cpu_reg_val Y (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (cpu_reg_val Y (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))))))" using cpu_reg_val_low_equal write_reg_low_equal by fastforce qed next case False then have f4: "\(fst instr = sreg_type RDY)" by auto then show ?thesis proof (cases "fst instr = sreg_type RDASR") case True then show ?thesis using a1 f1 f2 f3 f4 apply read_state_reg_instr_privilege_proof apply (clarsimp simp add: get_curr_win_low_equal) using cpu_reg_val_low_equal get_curr_win2_low_equal write_reg_low_equal proof - assume a1: "low_equal s1 s2" then have "cpu_reg_val (ASR (get_operand_w5 (snd instr ! 0))) (snd (fst (get_curr_win () s1))) = cpu_reg_val (ASR (get_operand_w5 (snd instr ! 0))) (snd (fst (get_curr_win () s2)))" by (meson cpu_reg_val_low_equal get_curr_win2_low_equal) then show "low_equal (snd (fst (write_reg (cpu_reg_val (ASR (get_operand_w5 (snd instr ! 0))) (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (cpu_reg_val (ASR (get_operand_w5 (snd instr ! 0))) (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))))))" using a1 by (metis (no_types) get_curr_win2_low_equal write_reg_low_equal) qed next case False then have f5: "\(fst instr = sreg_type RDASR)" by auto then show ?thesis using a1 f1 f2 f3 f4 f5 apply read_state_reg_instr_privilege_proof apply (clarsimp simp add: get_curr_win_low_equal) using cpu_reg_val_low_equal get_curr_win2_low_equal write_reg_low_equal proof - assume a1: "low_equal s1 s2" assume a2: "t1 = snd (fst (write_reg (cpu_reg_val TBR (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))))" assume "t2 = snd (fst (write_reg (cpu_reg_val TBR (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2)))))" have "\s. \ low_equal (snd (fst (get_curr_win () s1))) s \ snd (fst (write_reg (cpu_reg_val TBR s) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))))) = t1" using a2 by (simp add: cpu_reg_val_low_equal) then show "low_equal (snd (fst (write_reg (cpu_reg_val TBR (snd (fst (get_curr_win () s1)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1)))))) (snd (fst (write_reg (cpu_reg_val TBR (snd (fst (get_curr_win () s2)))) (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))))))" using a2 a1 by (metis (no_types) get_curr_win2_low_equal write_reg_low_equal) qed qed qed next case False then show ?thesis using a1 f1 f2 apply (simp add: read_state_reg_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: return_def) apply clarsimp apply (simp add: case_prod_unfold) using get_curr_win2_low_equal by auto qed qed qed lemma get_s_get_curr_win: assumes a1: "low_equal s1 s2" shows "get_S (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))) = get_S (cpu_reg_val PSR (snd (fst (get_curr_win () s2))))" proof - from a1 have "low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2)))" using get_curr_win2_low_equal by blast then show ?thesis using cpu_reg_val_low_equal by fastforce qed lemma write_state_reg_low_equal: assumes a1: "low_equal s1 s2 \ (((get_S (cpu_reg_val PSR s1)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s2)))::word1) = 0 \ t1 = snd (fst (write_state_reg_instr instr s1)) \ t2 = snd (fst (write_state_reg_instr instr s2))" shows "low_equal t1 t2" proof (cases "fst instr = sreg_type WRY") case True then show ?thesis using a1 apply write_state_reg_instr_privilege_proof apply (simp add: simpler_modify_def) apply (simp add: delayed_pool_add_def DELAYNUM_def) apply (auto simp add: get_curr_win_low_equal) using get_curr_win2_low_equal cpu_reg_mod_low_equal user_reg_val_low_equal get_operand2_low_equal proof - assume a1: "low_equal s1 s2" assume "t2 = cpu_reg_mod (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s2)))) Y (snd (fst (get_curr_win () s2)))" assume "t1 = cpu_reg_mod (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s1)))) Y (snd (fst (get_curr_win () s1)))" have f2: "low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2)))" using a1 by (meson get_curr_win2_low_equal) then have f3: "\w wa. user_reg_val w wa (snd (fst (get_curr_win () s2))) = user_reg_val w wa (snd (fst (get_curr_win () s1)))" by (simp add: user_reg_val_low_equal) have "\is. get_operand2 is (snd (fst (get_curr_win () s2))) = get_operand2 is (snd (fst (get_curr_win () s1)))" using f2 by (simp add: get_operand2_low_equal) then show "low_equal (cpu_reg_mod (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s1)))) Y (snd (fst (get_curr_win () s1)))) (cpu_reg_mod (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s2)))) Y (snd (fst (get_curr_win () s2))))" using f3 f2 by (metis cpu_reg_mod_low_equal) qed next case False then have f1: "\(fst instr = sreg_type WRY)" by auto then show ?thesis proof (cases "fst instr = sreg_type WRASR") case True then have f1_1: "fst instr = sreg_type WRASR" by auto then show ?thesis proof (cases "privileged_ASR (get_operand_w5 (snd instr ! 3)) \ get_S (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))) = 0") case True then show ?thesis using a1 f1 f1_1 apply write_state_reg_instr_privilege_proof apply (clarsimp simp add: get_s_get_curr_win) apply (simp add: raise_trap_def add_trap_set_def simpler_modify_def) apply (clarsimp simp add: get_curr_win3_low_equal) using traps_low_equal mod_trap_low_equal get_curr_win2_low_equal by fastforce next case False then have f1_2: "\ (privileged_ASR (get_operand_w5 (snd instr ! 3)) \ get_S (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))) = 0)" by auto then show ?thesis proof (cases "illegal_instruction_ASR (get_operand_w5 (snd instr ! 3))") case True then show ?thesis using a1 f1 f1_1 f1_2 apply write_state_reg_instr_privilege_proof apply (clarsimp simp add: get_s_get_curr_win) apply auto apply (simp add: raise_trap_def add_trap_set_def simpler_modify_def) apply (clarsimp simp add: get_curr_win3_low_equal) using traps_low_equal mod_trap_low_equal get_curr_win2_low_equal apply fastforce apply (simp add: raise_trap_def add_trap_set_def simpler_modify_def) apply (clarsimp simp add: get_curr_win3_low_equal) using traps_low_equal mod_trap_low_equal get_curr_win2_low_equal by fastforce next case False then show ?thesis using a1 f1 f1_1 f1_2 apply write_state_reg_instr_privilege_proof apply (clarsimp simp add: get_s_get_curr_win) apply auto apply (simp add: simpler_modify_def) apply (simp add: delayed_pool_add_def DELAYNUM_def) apply (auto simp add: get_curr_win_low_equal) using get_curr_win2_low_equal cpu_reg_mod_low_equal user_reg_val_low_equal get_operand2_low_equal proof - assume a1: "low_equal s1 s2" assume "t2 = cpu_reg_mod (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s2)))) (ASR (get_operand_w5 (snd instr ! 3))) (snd (fst (get_curr_win () s2)))" assume "t1 = cpu_reg_mod (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s1)))) (ASR (get_operand_w5 (snd instr ! 3))) (snd (fst (get_curr_win () s1)))" have "low_equal (snd (fst (get_curr_win () s1))) (snd (fst (get_curr_win () s2)))" using a1 by (meson get_curr_win2_low_equal) then show "low_equal (cpu_reg_mod (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s1)))) (ASR (get_operand_w5 (snd instr ! 3))) (snd (fst (get_curr_win () s1)))) (cpu_reg_mod (user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s2)))) (ASR (get_operand_w5 (snd instr ! 3))) (snd (fst (get_curr_win () s2))))" using cpu_reg_mod_low_equal get_operand2_low_equal user_reg_val_low_equal by fastforce next assume f1: "\ illegal_instruction_ASR (get_operand_w5 (snd instr ! 3))" assume f2: "fst instr = sreg_type WRASR" assume f3: "snd (fst (write_state_reg_instr instr s1)) = snd (fst (modify (delayed_pool_add (DELAYNUM, user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s1))), ASR (get_operand_w5 (snd instr ! 3)))) (snd (fst (get_curr_win () s1))))) " assume f4: "snd (fst (write_state_reg_instr instr s2)) = snd (fst (modify (delayed_pool_add (DELAYNUM, user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s2))), ASR (get_operand_w5 (snd instr ! 3)))) (snd (fst (get_curr_win () s2)))))" assume f5: "low_equal s1 s2" assume f6: "(get_S (cpu_reg_val PSR s1)) = 0" assume f7: "(get_S (cpu_reg_val PSR s2)) = 0" assume f8: "t1 = snd (fst (modify (delayed_pool_add (DELAYNUM, user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s1))), ASR (get_operand_w5 (snd instr ! 3)))) (snd (fst (get_curr_win () s1)))))" assume f9: "t2 = snd (fst (modify (delayed_pool_add (DELAYNUM, user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s2))), ASR (get_operand_w5 (snd instr ! 3)))) (snd (fst (get_curr_win () s2)))))" assume f10: "get_S (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))) \ 0" assume f11: "(\s1 s2 t1 t2. low_equal s1 s2 \ t1 = snd (fst (get_curr_win () s1)) \ t2 = snd (fst (get_curr_win () s2)) \ low_equal t1 t2)" assume f12: "(\s1 s2 t1 w cr t2. low_equal s1 s2 \ t1 = cpu_reg_mod w cr s1 \ t2 = cpu_reg_mod w cr s2 \ low_equal t1 t2)" assume f13: "(\s1 s2 win ur. low_equal s1 s2 \ user_reg_val win ur s1 = user_reg_val win ur s2)" assume f14: "(\s1 s2 op_list. low_equal s1 s2 \ get_operand2 op_list s1 = get_operand2 op_list s2)" show "low_equal (snd (fst (modify (delayed_pool_add (DELAYNUM, user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s1))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s1))), ASR (get_operand_w5 (snd instr ! 3)))) (snd (fst (get_curr_win () s1)))))) (snd (fst (modify (delayed_pool_add (DELAYNUM, user_reg_val (fst (fst (get_curr_win () s2))) (get_operand_w5 (snd instr ! Suc 0)) (snd (fst (get_curr_win () s2))) XOR get_operand2 (snd instr) (snd (fst (get_curr_win () s2))), ASR (get_operand_w5 (snd instr ! 3)))) (snd (fst (get_curr_win () s2))))))" using f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 using Sparc_Properties.ucast_0 assms get_curr_win_privilege by blast qed qed qed next case False then have f2: "\(fst instr = sreg_type WRASR)" by auto have f3: "get_S (cpu_reg_val PSR (snd (fst (get_curr_win () s1)))) = 0 \ get_S (cpu_reg_val PSR (snd (fst (get_curr_win () s2)))) = 0" using get_curr_win_privilege a1 by (metis ucast_id) then show ?thesis proof (cases "fst instr = sreg_type WRPSR") case True then show ?thesis using a1 f1 f2 f3 apply write_state_reg_instr_privilege_proof apply (simp add: raise_trap_def add_trap_set_def simpler_modify_def) apply (clarsimp simp add: get_curr_win3_low_equal) using traps_low_equal mod_trap_low_equal get_curr_win2_low_equal by fastforce next case False then have f4: "\(fst instr = sreg_type WRPSR)" by auto then show ?thesis proof (cases "fst instr = sreg_type WRWIM") case True then show ?thesis using a1 f1 f2 f3 f4 apply write_state_reg_instr_privilege_proof apply (simp add: raise_trap_def add_trap_set_def simpler_modify_def) apply (clarsimp simp add: get_curr_win3_low_equal) using traps_low_equal mod_trap_low_equal get_curr_win2_low_equal by fastforce next case False then have f5: "\(fst instr = sreg_type WRWIM)" by auto then show ?thesis using a1 f1 f2 f3 f4 f5 apply write_state_reg_instr_privilege_proof apply (simp add: raise_trap_def add_trap_set_def simpler_modify_def) apply (clarsimp simp add: get_curr_win3_low_equal) using traps_low_equal mod_trap_low_equal get_curr_win2_low_equal by fastforce qed qed qed qed lemma flush_instr_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (flush_instr instr s1)) \ t2 = snd (fst (flush_instr instr s2))" shows "low_equal t1 t2" using a1 apply (simp add: flush_instr_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def simpler_modify_def) apply (simp add: flush_cache_all_def) apply (simp add: low_equal_def) apply (simp add: user_accessible_def) apply (simp add: mem_equal_def) by auto lemma branch_instr_sub1_low_equal: assumes a1: "low_equal s1 s2" shows "branch_instr_sub1 instr_name s1 = branch_instr_sub1 instr_name s2" using a1 apply (simp add: branch_instr_sub1_def) by (simp add: low_equal_def) lemma set_annul_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (set_annul True s1)) \ t2 = snd (fst (set_annul True s2))" shows "low_equal t1 t2" using a1 apply (simp add: set_annul_def) apply (simp add: simpler_modify_def annul_mod_def) using state_var2_low_equal state_var_low_equal by fastforce lemma branch_instr_low_equal_sub0: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (write_cpu (cpu_reg_val PC s2 + sign_ext24 (ucast (get_operand_w22 (snd instr ! Suc 0)) << 2)) nPC (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s1))))) \ t2 = snd (fst (write_cpu (cpu_reg_val PC s2 + sign_ext24 (ucast (get_operand_w22 (snd instr ! Suc 0)) << 2)) nPC (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s2)))))" shows "low_equal t1 t2" proof - from a1 have "low_equal (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s1))) (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s2)))" using write_cpu_low_equal by blast then show ?thesis using a1 write_cpu_low_equal by blast qed lemma branch_instr_low_equal_sub1: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (set_annul True (snd (fst (write_cpu (cpu_reg_val PC s2 + sign_ext24 (ucast (get_operand_w22 (snd instr ! Suc 0)) << 2)) nPC (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s1)))))))) \ t2 = snd (fst (set_annul True (snd (fst (write_cpu (cpu_reg_val PC s2 + sign_ext24 (ucast (get_operand_w22 (snd instr ! Suc 0)) << 2)) nPC (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s2))))))))" shows "low_equal t1 t2" proof - from a1 have "low_equal (snd (fst (write_cpu (cpu_reg_val PC s2 + sign_ext24 (ucast (get_operand_w22 (snd instr ! Suc 0)) << 2)) nPC (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s1)))))) (snd (fst (write_cpu (cpu_reg_val PC s2 + sign_ext24 (ucast (get_operand_w22 (snd instr ! Suc 0)) << 2)) nPC (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s2))))))" using branch_instr_low_equal_sub0 by blast then show ?thesis using a1 using set_annul_low_equal by blast qed lemma branch_instr_low_equal_sub2: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (set_annul True (snd (fst (write_cpu (cpu_reg_val nPC s2 + 4) nPC (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s1)))))))) \ t2 = snd (fst (set_annul True (snd (fst (write_cpu (cpu_reg_val nPC s2 + 4) nPC (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s2))))))))" shows "low_equal t1 t2" proof - from a1 have "low_equal (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s1))) (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s2)))" using write_cpu_low_equal by blast then have "low_equal (snd (fst (write_cpu (cpu_reg_val nPC s2 + 4) nPC (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s1)))))) (snd (fst (write_cpu (cpu_reg_val nPC s2 + 4) nPC (snd (fst (write_cpu (cpu_reg_val nPC s2) PC s2))))))" using write_cpu_low_equal by blast then show ?thesis using a1 using set_annul_low_equal by blast qed lemma branch_instr_low_equal: assumes a1: "low_equal s1 s2 \ t1 = snd (fst (branch_instr instr s1)) \ t2 = snd (fst (branch_instr instr s2))" shows "low_equal t1 t2" using a1 apply (simp add: branch_instr_def) apply (simp add: Let_def simpler_gets_def bind_def h1_def h2_def) apply (simp add: case_prod_unfold return_def) apply clarsimp apply (simp add: branch_instr_sub1_low_equal) apply (simp_all add: cpu_reg_val_low_equal) apply (cases "branch_instr_sub1 (fst instr) s2 = 1") apply clarsimp apply (simp add: bind_def h1_def h2_def Let_def) apply (simp_all add: cpu_reg_val_low_equal) apply (simp add: case_prod_unfold) apply (cases "fst instr = bicc_type BA \ get_operand_flag (snd instr ! 0) = 1") apply clarsimp using branch_instr_low_equal_sub1 apply blast apply clarsimp apply (simp add: return_def) using branch_instr_low_equal_sub0 apply fastforce apply (simp add: bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (cases "get_operand_flag (snd instr ! 0) = 1") apply clarsimp apply (simp_all add: cpu_reg_val_low_equal) using branch_instr_low_equal_sub2 apply metis apply (simp add: return_def) using write_cpu_low_equal by metis lemma dispath_instr_low_equal: assumes a1: "low_equal s1 s2 \ (((get_S (cpu_reg_val PSR s1)))::word1) = 0 \ (((get_S (cpu_reg_val PSR s2)))::word1) = 0 \ \ snd (dispatch_instruction instr s1) \ \ snd (dispatch_instruction instr s2) \ t1 = (snd (fst (dispatch_instruction instr s1))) \ t2 = (snd (fst (dispatch_instruction instr s2)))" shows "low_equal t1 t2" proof (cases "get_trap_set s1 = {}") case True then have f_no_traps: "get_trap_set s1 = {} \ get_trap_set s2 = {}" using a1 by (simp add: low_equal_def get_trap_set_def) then show ?thesis proof (cases "fst instr \ {load_store_type LDSB,load_store_type LDUB, load_store_type LDUBA,load_store_type LDUH,load_store_type LD, load_store_type LDA,load_store_type LDD}") case True then show ?thesis using a1 f_no_traps apply dispath_instr_privilege_proof by (blast intro: load_instr_low_equal) next case False then have f1: "fst instr \ {load_store_type LDSB, load_store_type LDUB, load_store_type LDUBA, load_store_type LDUH, load_store_type LD, load_store_type LDA, load_store_type LDD}" by auto then show ?thesis proof (cases "fst instr \ {load_store_type STB,load_store_type STH, load_store_type ST,load_store_type STA,load_store_type STD}") case True then show ?thesis using a1 f_no_traps f1 apply dispath_instr_privilege_proof using store_instr_low_equal by blast next case False then have f2: "\(fst instr \ {load_store_type STB,load_store_type STH, load_store_type ST,load_store_type STA,load_store_type STD})" by auto then show ?thesis proof (cases "fst instr \ {sethi_type SETHI}") case True then show ?thesis using a1 f_no_traps f1 f2 apply dispath_instr_privilege_proof by (auto intro: sethi_low_equal) next case False then have f3: "\(fst instr \ {sethi_type SETHI})" by auto then show ?thesis proof (cases "fst instr \ {nop_type NOP}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 apply dispath_instr_privilege_proof by (auto intro: nop_low_equal) next case False then have f4: "\(fst instr \ {nop_type NOP})" by auto then show ?thesis proof (cases "fst instr \ {logic_type ANDs,logic_type ANDcc,logic_type ANDN, logic_type ANDNcc,logic_type ORs,logic_type ORcc,logic_type ORN, logic_type XORs,logic_type XNOR}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 apply dispath_instr_privilege_proof using logical_instr_low_equal by blast next case False then have f5: "\(fst instr \ {logic_type ANDs,logic_type ANDcc,logic_type ANDN, logic_type ANDNcc,logic_type ORs,logic_type ORcc,logic_type ORN, logic_type XORs,logic_type XNOR})" by auto then show ?thesis proof (cases "fst instr \ {shift_type SLL,shift_type SRL,shift_type SRA}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 apply dispath_instr_privilege_proof using shift_instr_low_equal by blast next case False then have f6: "\(fst instr \ {shift_type SLL,shift_type SRL,shift_type SRA})" by auto then show ?thesis proof (cases "fst instr \ {arith_type ADD,arith_type ADDcc,arith_type ADDX}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 f6 apply dispath_instr_privilege_proof using add_instr_low_equal by blast next case False then have f7: "\(fst instr \ {arith_type ADD,arith_type ADDcc,arith_type ADDX})" by auto then show ?thesis proof (cases "fst instr \ {arith_type SUB,arith_type SUBcc,arith_type SUBX}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 f6 f7 apply dispath_instr_privilege_proof using sub_instr_low_equal by blast next case False then have f8: "\(fst instr \ {arith_type SUB,arith_type SUBcc,arith_type SUBX})" by auto then show ?thesis proof (cases "fst instr \ {arith_type UMUL,arith_type SMUL,arith_type SMULcc}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 f6 f7 f8 apply dispath_instr_privilege_proof using mul_instr_low_equal by blast next case False then have f9: "\(fst instr \ {arith_type UMUL,arith_type SMUL, arith_type SMULcc})" by auto then show ?thesis proof (cases "fst instr \ {arith_type UDIV,arith_type UDIVcc,arith_type SDIV}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 f6 f7 f8 f9 apply dispath_instr_privilege_proof using div_instr_low_equal by blast next case False then have f10: "\(fst instr \ {arith_type UDIV, arith_type UDIVcc,arith_type SDIV})" by auto then show ?thesis proof (cases "fst instr \ {ctrl_type SAVE,ctrl_type RESTORE}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 apply dispath_instr_privilege_proof using save_restore_instr_low_equal by blast next case False then have f11: "\(fst instr \ {ctrl_type SAVE,ctrl_type RESTORE})" by auto then show ?thesis proof (cases "fst instr \ {call_type CALL}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 apply dispath_instr_privilege_proof using call_instr_low_equal by blast next case False then have f12: "\(fst instr \ {call_type CALL})" by auto then show ?thesis proof (cases "fst instr \ {ctrl_type JMPL}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 apply dispath_instr_privilege_proof using jmpl_instr_low_equal by blast next case False then have f13: "\(fst instr \ {ctrl_type JMPL})" by auto then show ?thesis proof (cases "fst instr \ {ctrl_type RETT}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 apply dispath_instr_privilege_proof using rett_instr_low_equal by blast next case False then have f14: "\(fst instr \ {ctrl_type RETT})" by auto then show ?thesis proof (cases "fst instr \ {sreg_type RDY,sreg_type RDPSR, sreg_type RDWIM, sreg_type RDTBR}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 apply dispath_instr_privilege_proof using read_state_reg_low_equal by blast next case False then have f15: "\(fst instr \ {sreg_type RDY,sreg_type RDPSR, sreg_type RDWIM, sreg_type RDTBR})" by auto then show ?thesis proof (cases "fst instr \ {sreg_type WRY,sreg_type WRPSR, sreg_type WRWIM, sreg_type WRTBR}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 apply dispath_instr_privilege_proof using write_state_reg_low_equal by blast next case False then have f16: "\(fst instr \ {sreg_type WRY,sreg_type WRPSR, sreg_type WRWIM, sreg_type WRTBR})" by auto then show ?thesis proof (cases "fst instr \ {load_store_type FLUSH}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 apply dispath_instr_privilege_proof using flush_instr_low_equal by blast next case False then have f17: "\(fst instr \ {load_store_type FLUSH})" by auto then show ?thesis proof (cases "fst instr \ {bicc_type BE,bicc_type BNE, bicc_type BGU,bicc_type BLE,bicc_type BL,bicc_type BGE, bicc_type BNEG,bicc_type BG,bicc_type BCS,bicc_type BLEU, bicc_type BCC,bicc_type BA,bicc_type BN}") case True then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 apply dispath_instr_privilege_proof using branch_instr_low_equal by blast next case False then show ?thesis using a1 f_no_traps f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 apply dispath_instr_privilege_proof by (simp add: fail_def) qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed qed next case False then have "get_trap_set s1 \ {} \ get_trap_set s2 \ {}" using a1 by (simp add: low_equal_def get_trap_set_def) then show ?thesis using a1 apply (simp add: dispatch_instruction_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def) apply (simp add: Let_def) by (simp add: return_def) qed lemma execute_instr_sub1_low_equal: assumes a1: "low_equal s1 s2 \ \ snd (execute_instr_sub1 instr s1) \ \ snd (execute_instr_sub1 instr s2) \ t1 = (snd (fst (execute_instr_sub1 instr s1))) \ t2 = (snd (fst (execute_instr_sub1 instr s2)))" shows "low_equal t1 t2" proof (cases "get_trap_set s1 = {}") case True then have "get_trap_set s1 = {} \ get_trap_set s2 = {}" using a1 by (simp add: low_equal_def get_trap_set_def) then show ?thesis using a1 apply (simp add: execute_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (case_tac "fst instr \ call_type CALL \ fst instr \ ctrl_type RETT \ fst instr \ ctrl_type JMPL \ fst instr \ bicc_type BE \ fst instr \ bicc_type BNE \ fst instr \ bicc_type BGU \ fst instr \ bicc_type BLE \ fst instr \ bicc_type BL \ fst instr \ bicc_type BGE \ fst instr \ bicc_type BNEG \ fst instr \ bicc_type BG \ fst instr \ bicc_type BCS \ fst instr \ bicc_type BLEU \ fst instr \ bicc_type BCC \ fst instr \ bicc_type BA \ fst instr \ bicc_type BN") apply clarsimp apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: low_equal_def) apply (simp add: cpu_reg_val_def write_cpu_def cpu_reg_mod_def) apply (simp add: simpler_modify_def return_def) apply (simp add: user_accessible_mod_cpu_reg mem_equal_mod_cpu_reg) apply clarsimp by (auto simp add: return_def) next case False then have "get_trap_set s1 \ {} \ get_trap_set s2 \ {}" using a1 by (simp add: low_equal_def get_trap_set_def) then show ?thesis using a1 apply (simp add: execute_instr_sub1_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (simp add: return_def) qed theorem non_interference_step: assumes a1: "(((get_S (cpu_reg_val PSR s1)))::word1) = 0 \ good_context s1 \ get_delayed_pool s1 = [] \ get_trap_set s1 = {} \ (((get_S (cpu_reg_val PSR s2)))::word1) = 0 \ get_delayed_pool s2 = [] \ get_trap_set s2 = {} \ good_context s2 \ low_equal s1 s2" shows "\t1 t2. Some t1 = NEXT s1 \ Some t2 = NEXT s2 \ (((get_S (cpu_reg_val PSR t1)))::word1) = 0 \ (((get_S (cpu_reg_val PSR t2)))::word1) = 0 \ low_equal t1 t2" proof - from a1 have "good_context s1 \ good_context s2" by auto then have "NEXT s1 = Some (snd (fst (execute_instruction () s1))) \ NEXT s2 = Some (snd (fst (execute_instruction () s2)))" by (simp add: single_step) then have "\t1 t2. Some t1 = NEXT s1 \ Some t2 = NEXT s2" by auto then have f0: "snd (execute_instruction() s1) = False \ snd (execute_instruction() s2) = False" by (auto simp add: NEXT_def case_prod_unfold) then have f1: "\t1 t2. Some t1 = NEXT s1 \ Some t2 = NEXT s2 \ (((get_S (cpu_reg_val PSR t1)))::word1) = 0 \ (((get_S (cpu_reg_val PSR t2)))::word1) = 0" using a1 apply (auto simp add: NEXT_def case_prod_unfold) by (auto simp add: safe_privilege) then show ?thesis proof (cases "exe_mode_val s1") case True then have f_exe0: "exe_mode_val s1" by auto then have f_exe: "exe_mode_val s1 \ exe_mode_val s2" proof - have "low_equal s1 s2" using a1 by auto then have "state_var s1 = state_var s2" by (simp add: low_equal_def) then have "exe_mode_val s1 = exe_mode_val s2" by (simp add: exe_mode_val_def) then show ?thesis using f_exe0 by auto qed then show ?thesis proof (cases "\e. fetch_instruction (delayed_pool_write s1) = Inl e") case True then have f_fetch_error: "\e. fetch_instruction (delayed_pool_write s1) = Inl e" by auto then have f_fetch_error2: "(\e. fetch_instruction (delayed_pool_write s1) = Inl e) \ (\e. fetch_instruction (delayed_pool_write s2) = Inl e)" proof - have "cpu_reg s1 = cpu_reg s2" using a1 by (simp add: low_equal_def) then have "cpu_reg_val PC s1 = cpu_reg_val PC s2" by (simp add: cpu_reg_val_def) then have "cpu_reg_val PC s1 = cpu_reg_val PC s2 \ (((get_S (cpu_reg_val PSR (delayed_pool_write s1))))::word1) = 0 \ (((get_S (cpu_reg_val PSR (delayed_pool_write s2))))::word1) = 0" using a1 by (auto simp add: empty_delayed_pool_write_privilege) then show ?thesis using a1 f_fetch_error apply (simp add: fetch_instruction_def) apply (simp add: Let_def) apply clarsimp apply (case_tac "uint (3 AND cpu_reg_val PC (delayed_pool_write s1)) = 0") apply auto apply (case_tac "fst (memory_read 8 (cpu_reg_val PC (delayed_pool_write s1)) (delayed_pool_write s1)) = None") apply auto apply (simp add: case_prod_unfold) using a1 apply (auto simp add: mem_read_delayed_write_low_equal) apply (simp add: case_prod_unfold) using a1 apply (auto simp add: mem_read_delayed_write_low_equal) apply (simp add: delayed_pool_write_def) by (simp add: Let_def get_delayed_write_def) qed then show ?thesis proof (cases "exe_mode_val s1") case True then have "exe_mode_val s1 \ exe_mode_val s2" using exe_mode_low_equal a1 by auto then show ?thesis using f1 apply (simp add: NEXT_def execute_instruction_def) apply (simp add: bind_def h1_def h2_def Let_def simpler_gets_def) using a1 apply clarsimp apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: simpler_modify_def) using f_fetch_error2 apply clarsimp apply (simp add: raise_trap_def simpler_modify_def return_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: return_def simpler_modify_def) apply (simp add: raise_trap_def simpler_modify_def return_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: return_def) apply (simp add: delayed_pool_write_def get_delayed_write_def Let_def) apply (simp add: low_equal_def) apply (simp add: add_trap_set_def) apply (simp add: cpu_reg_val_def) apply clarsimp by (simp add: mem_equal_mod_trap user_accessible_mod_trap) next case False then have "\ (exe_mode_val s1) \ \ (exe_mode_val s2)" using exe_mode_low_equal a1 by auto then show ?thesis using f1 apply (simp add: NEXT_def execute_instruction_def) apply (simp add: bind_def h1_def h2_def Let_def simpler_gets_def) using a1 apply clarsimp apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) by (simp add: return_def) qed next case False then have f_fetch_suc: "(\v. fetch_instruction (delayed_pool_write s1) = Inr v)" using fetch_instr_result_1 by auto then have "(\v. fetch_instruction (delayed_pool_write s1) = Inr v \ fetch_instruction (delayed_pool_write s2) = Inr v)" proof - have "cpu_reg s1 = cpu_reg s2" using a1 by (simp add: low_equal_def) then have "cpu_reg_val PC s1 = cpu_reg_val PC s2" by (simp add: cpu_reg_val_def) then have "cpu_reg_val PC s1 = cpu_reg_val PC s2 \ (((get_S (cpu_reg_val PSR (delayed_pool_write s1))))::word1) = 0 \ (((get_S (cpu_reg_val PSR (delayed_pool_write s2))))::word1) = 0" using a1 by (auto simp add: empty_delayed_pool_write_privilege) then show ?thesis using a1 f_fetch_suc apply (simp add: fetch_instruction_def) apply (simp add: Let_def) apply clarsimp apply (case_tac "uint (3 AND cpu_reg_val PC (delayed_pool_write s1)) = 0") apply auto apply (case_tac "fst (memory_read 8 (cpu_reg_val PC (delayed_pool_write s1)) (delayed_pool_write s1)) = None") apply auto apply (simp add: case_prod_unfold) using a1 apply (auto simp add: mem_read_delayed_write_low_equal) apply (simp add: case_prod_unfold) using a1 apply (auto simp add: mem_read_delayed_write_low_equal) apply (simp add: delayed_pool_write_def) by (simp add: Let_def get_delayed_write_def) qed then have "(\v. fetch_instruction (delayed_pool_write s1) = Inr v \ fetch_instruction (delayed_pool_write s2) = Inr v \ \ (\e. (decode_instruction v) = Inl e))" using dispatch_fail f0 a1 f_exe by auto then have f_fetch_dec: "(\v. fetch_instruction (delayed_pool_write s1) = Inr v \ fetch_instruction (delayed_pool_write s2) = Inr v \ (\v1. (decode_instruction v) = Inr v1))" using decode_instr_result_4 by auto then show ?thesis proof (cases "annul_val (delayed_pool_write s1)") case True then have "annul_val (delayed_pool_write s1) \ annul_val (delayed_pool_write s2)" using a1 apply (simp add: low_equal_def) by (simp add: delayed_pool_write_def get_delayed_write_def annul_val_def) then show ?thesis using a1 f1 f_exe f_fetch_dec apply (simp add: NEXT_def execute_instruction_def) apply (simp add: exec_gets return_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: simpler_modify_def) apply clarsimp apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: write_cpu_def cpu_reg_val_def set_annul_def) apply (simp add: simpler_modify_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: write_cpu_def cpu_reg_val_def set_annul_def) apply (simp add: simpler_modify_def) apply (simp add: cpu_reg_mod_def annul_mod_def) apply (simp add: delayed_pool_write_def get_delayed_write_def) apply (simp add: write_annul_def) apply clarsimp apply (simp add: low_equal_def) apply (simp add: user_accessible_annul mem_equal_annul) by (metis) next case False then have "\ annul_val (delayed_pool_write s1) \ \ annul_val (delayed_pool_write s2)" using a1 apply (simp add: low_equal_def) apply (simp add: delayed_pool_write_def get_delayed_write_def) by (simp add: annul_val_def) then show ?thesis using a1 f1 f_exe f_fetch_dec apply (simp add: NEXT_def execute_instruction_def) apply (simp add: exec_gets return_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: simpler_modify_def) apply clarsimp apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (case_tac "snd (execute_instr_sub1 (a, b) (snd (fst (dispatch_instruction (a, b) (delayed_pool_write s1))))) \ snd (dispatch_instruction (a, b) (delayed_pool_write s1))") apply auto apply (case_tac "snd (execute_instr_sub1 (a, b) (snd (fst (dispatch_instruction (a, b) (delayed_pool_write s2))))) \ snd (dispatch_instruction (a, b) (delayed_pool_write s2))") apply auto apply (simp add: simpler_modify_def) apply (simp add: simpler_gets_def bind_def h1_def h2_def Let_def) apply (simp add: case_prod_unfold) apply (simp add: delayed_pool_write_def get_delayed_write_def) by (meson dispath_instr_low_equal dispath_instr_privilege execute_instr_sub1_low_equal) qed qed next case False then have f_non_exe: "exe_mode_val s1 = False" by auto then have "exe_mode_val s1 = False \ exe_mode_val s2 = False" proof - have "low_equal s1 s2" using a1 by auto then have "state_var s1 = state_var s2" by (simp add: low_equal_def) then have "exe_mode_val s1 = exe_mode_val s2" by (simp add: exe_mode_val_def) then show ?thesis using f_non_exe by auto qed then show ?thesis using f1 a1 apply (simp add: NEXT_def execute_instruction_def) by (simp add: simpler_gets_def bind_def h1_def h2_def Let_def return_def) qed qed function (sequential) SEQ:: "nat \ ('a::len) sparc_state \ ('a) sparc_state option" where "SEQ 0 s = Some s" |"SEQ n s = ( case SEQ (n-1) s of None \ None | Some t \ NEXT t )" by pat_completeness auto termination by lexicographic_order lemma SEQ_suc: "SEQ n s = Some t \ SEQ (Suc n) s = NEXT t" apply (induction n) apply clarsimp by (simp add: option.case_eq_if) definition user_seq_exe:: "nat \ ('a::len) sparc_state \ bool" where "user_seq_exe n s \ \i t. (i \ n \ SEQ i s = Some t) \ (good_context t \ get_delayed_pool t = [] \ get_trap_set t = {})" text \NIA is short for non-interference assumption.\ definition "NIA t1 t2 \ (((get_S (cpu_reg_val PSR t1)))::word1) = 0 \ (((get_S (cpu_reg_val PSR t2)))::word1) = 0 \ good_context t1 \ get_delayed_pool t1 = [] \ get_trap_set t1 = {} \ good_context t2 \ get_delayed_pool t2 = [] \ get_trap_set t2 = {} \ low_equal t1 t2" text \NIC is short for non-interference conclusion.\ definition "NIC t1 t2 \ (\u1 u2. Some u1 = NEXT t1 \ Some u2 = NEXT t2 \ (((get_S (cpu_reg_val PSR u1)))::word1) = 0 \ (((get_S (cpu_reg_val PSR u2)))::word1) = 0 \ low_equal u1 u2)" lemma NIS_short: "\t1 t2. NIA t1 t2 \ NIC t1 t2" apply (simp add: NIA_def NIC_def) using non_interference_step by auto lemma non_interference_induct_case_sub1: assumes a1: "(\t1. Some t1 = SEQ n s1 \ (\t2. Some t2 = SEQ n s2 \ NIA t1 t2))" shows "(\t1. Some t1 = SEQ n s1 \ (\t2. Some t2 = SEQ n s2 \ NIA t1 t2 \ NIC t1 t2))" using NIS_short using assms by auto lemma non_interference_induct_case: assumes a1: "((\i t. i \ n \ SEQ i s1 = Some t \ good_context t \ get_delayed_pool t = [] \ get_trap_set t = {}) \ (\i t. i \ n \ SEQ i s2 = Some t \ good_context t \ get_delayed_pool t = [] \ get_trap_set t = {}) \ (\t1. Some t1 = SEQ n s1 \ (\t2. Some t2 = SEQ n s2 \ (((get_S (cpu_reg_val PSR t1)))::word1) = 0 \ (((get_S (cpu_reg_val PSR t2)))::word1) = 0 \ low_equal t1 t2))) \ (\i t. i \ Suc n \ SEQ i s1 = Some t \ good_context t \ get_delayed_pool t = [] \ get_trap_set t = {}) \ (\i t. i \ Suc n \ SEQ i s2 = Some t \ good_context t \ get_delayed_pool t = [] \ get_trap_set t = {})" shows "\t1. Some t1 = (case SEQ n s1 of None \ None | Some x \ NEXT x) \ (\t2. Some t2 = (case SEQ n s2 of None \ None | Some x \ NEXT x) \ (((get_S (cpu_reg_val PSR t1)))::word1) = 0 \ (((get_S (cpu_reg_val PSR t2)))::word1) = 0 \ low_equal t1 t2)" proof - from a1 have f1: "((\i t. i \ n \ SEQ i s1 = Some t \ good_context t \ get_delayed_pool t = [] \ get_trap_set t = {}) \ (\i t. i \ n \ SEQ i s2 = Some t \ good_context t \ get_delayed_pool t = [] \ get_trap_set t = {}))" by (metis le_SucI) then have f2: "(\t1. Some t1 = SEQ n s1 \ (\t2. Some t2 = SEQ n s2 \ (((get_S (cpu_reg_val PSR t1)))::word1) = 0 \ (((get_S (cpu_reg_val PSR t2)))::word1) = 0 \ low_equal t1 t2))" using a1 by auto then have f3: "(\t1. Some t1 = SEQ n s1 \ (\t2. Some t2 = SEQ n s2 \ NIA t1 t2))" using f1 NIA_def by (metis (full_types) dual_order.refl) then have "(\t1. Some t1 = SEQ n s1 \ (\t2. Some t2 = SEQ n s2 \ NIA t1 t2 \ NIC t1 t2))" using non_interference_induct_case_sub1 by blast then have "(\t1. Some t1 = SEQ n s1 \ (\t2. Some t2 = SEQ n s2 \ ((((get_S (cpu_reg_val PSR t1)))::word1) = 0 \ (((get_S (cpu_reg_val PSR t2)))::word1) = 0 \ good_context t1 \ get_delayed_pool t1 = [] \ get_trap_set t1 = {} \ good_context t2 \ get_delayed_pool t2 = [] \ get_trap_set t2 = {} \ low_equal t1 t2) \ (\u1 u2. Some u1 = NEXT t1 \ Some u2 = NEXT t2 \ (((get_S (cpu_reg_val PSR u1)))::word1) = 0 \ (((get_S (cpu_reg_val PSR u2)))::word1) = 0 \ low_equal u1 u2)))" using NIA_def NIC_def by fastforce then show ?thesis by (metis option.simps(5)) qed lemma non_interference_induct_case_sub2: assumes a1: "(user_seq_exe n s1 \ user_seq_exe n s2 \ (\t1. Some t1 = SEQ n s1 \ (\t2. Some t2 = SEQ n s2 \ (((get_S (cpu_reg_val PSR t1)))::word1) = 0 \ (((get_S (cpu_reg_val PSR t2)))::word1) = 0 \ low_equal t1 t2))) \ user_seq_exe (Suc n) s1 \ user_seq_exe (Suc n) s2" shows "\t1. Some t1 = (case SEQ n s1 of None \ None | Some x \ NEXT x) \ (\t2. Some t2 = (case SEQ n s2 of None \ None | Some x \ NEXT x) \ (((get_S (cpu_reg_val PSR t1)))::word1) = 0 \ (((get_S (cpu_reg_val PSR t2)))::word1) = 0 \ low_equal t1 t2)" using a1 by (simp add: non_interference_induct_case user_seq_exe_def) theorem non_interference: assumes a1: "(((get_S (cpu_reg_val PSR s1)))::word1) = 0 \ good_context s1 \ get_delayed_pool s1 = [] \ get_trap_set s1 = {} \ (((get_S (cpu_reg_val PSR s2)))::word1) = 0 \ get_delayed_pool s2 = [] \ get_trap_set s2 = {} \ good_context s2 \ user_seq_exe n s1 \ user_seq_exe n s2 \ low_equal s1 s2" shows "(\t1 t2. Some t1 = SEQ n s1 \ Some t2 = SEQ n s2 \ (((get_S (cpu_reg_val PSR t1)))::word1) = 0 \ (((get_S (cpu_reg_val PSR t2)))::word1) = 0 \ low_equal t1 t2)" using a1 apply (induction n) apply (simp add: user_seq_exe_def) apply clarsimp by (simp add: non_interference_induct_case_sub2) end end diff --git a/thys/Security_Protocol_Refinement/Key_establish/m1_kerberos.thy b/thys/Security_Protocol_Refinement/Key_establish/m1_kerberos.thy --- a/thys/Security_Protocol_Refinement/Key_establish/m1_kerberos.thy +++ b/thys/Security_Protocol_Refinement/Key_establish/m1_kerberos.thy @@ -1,1229 +1,1229 @@ (******************************************************************************* Project: Development of Security Protocols by Refinement Module: Key_establish/m1_kerberos.thy (Isabelle/HOL 2016) ID: $Id: m1_kerberos.thy 133856 2017-03-20 18:05:54Z csprenge $ Author: Ivano Somaini, ETH Zurich Christoph Sprenger, ETH Zurich Key distribution protocols Abstract version of Kerberos protocol with server-authentication and mutual initiator and responder authentication. Copyright (c) 2009-2016 Ivano Somaini, Christoph Sprenger Licence: LGPL *******************************************************************************) section \Abstract Kerberos core protocol (L1)\ theory m1_kerberos imports m1_keydist_iirn begin text \We augment the basic abstract key distribution model such that the server sends a timestamp along with the session key. We use a cache to guard against replay attacks and timestamp validity checks to ensure recentness of the session key. We establish three refinements for this model, namely that this model refines \begin{enumerate} \item the authenticated key distribution model \m1_keydist_iirn\, \item the injective agreement model \a0i\, instantiated such that the responder agrees with the initiator on the session key, its timestamp and the initiator's authenticator timestamp. \item the injective agreement model \a0i\, instantiated such that the initiator agrees with the responder on the session key, its timestamp and the initiator's authenticator timestamp. \end{enumerate} \ (******************************************************************************) subsection \State\ (******************************************************************************) text \We extend the basic key distribution by adding timestamps. We add a clock variable modeling the current time and an authenticator replay cache recording triples @{term "(A, Kab, Ta)"} of agents, session keys, and authenticator timestamps. The inclusion of the session key avoids false replay rejections for different keys with identical authenticator timestamps. The frames, runs, and observations remain the same as in the previous model, but we will use the @{typ "nat list"}'s to store timestamps. \ type_synonym time = "nat" \ \for clock and timestamps\ consts Ls :: "time" \ \life time for session keys\ La :: "time" \ \life time for authenticators\ text \State and observations\ record m1_state = "m1r_state" + leak :: "(key \ agent \ agent \ nonce \ time) set" \ \key leaked plus context\ clk :: "time" cache :: "(agent \ key \ time) set" type_synonym m1_obs = "m1_state" type_synonym 'x m1_pred = "'x m1_state_scheme set" type_synonym 'x m1_trans = "('x m1_state_scheme \ 'x m1_state_scheme) set" consts END :: "atom" \ \run end marker (for initiator)\ (******************************************************************************) subsection \Events\ (******************************************************************************) definition \ \by @{term "A"}, refines @{term "m1x_step1"}\ m1_step1 :: "[rid_t, agent, agent, nonce] \ 'x m1_trans" where "m1_step1 \ m1a_step1" definition \ \by @{term "B"}, refines @{term "m1x_step2"}\ m1_step2 :: "[rid_t, agent, agent] \ 'x m1_trans" where "m1_step2 \ m1a_step2" definition \ \by @{term "Server"}, refines @{term m1x_step3}\ m1_step3 :: "[rid_t, agent, agent, key, nonce, time] \ 'x m1_trans" where "m1_step3 Rs A B Kab Na Ts \ {(s, s'). \ \new guards:\ Ts = clk s \ \ \fresh timestamp\ \ \rest as before:\ (s, s') \ m1a_step3 Rs A B Kab Na [aNum Ts] }" definition \ \by @{text "A"}, refines @{term m1x_step5}\ m1_step4 :: "[rid_t, agent, agent, nonce, key, time, time] \ 'x m1_trans" where "m1_step4 Ra A B Na Kab Ts Ta \ {(s, s'). \ \previous guards:\ runs s Ra = Some (Init, [A, B], []) \ (Kab \ Domain (leak s) \ (Kab, A) \ azC (runs s)) \ \ \authorization guard\ Na = Ra$na \ \ \fix parameter\ \ \guard for agreement with server on \(Kab, B, Na, isl)\,\ \ \where \isl = take is_len nla\; injectiveness by including \Na\\ (A \ bad \ (\Rs. Kab = sesK (Rs$sk) \ runs s Rs = Some (Serv, [A, B], [aNon Na, aNum Ts]))) \ \ \new guards:\ Ta = clk s \ \ \fresh timestamp\ clk s < Ts + Ls \ \ \ensure session key recentness\ \ \actions:\ s' = s\ runs := (runs s)(Ra \ (Init, [A, B], [aKey Kab, aNum Ts, aNum Ta])) \ }" definition \ \by @{term "B"}, refines @{term m1x_step4}\ m1_step5 :: "[rid_t, agent, agent, key, time, time] \ 'x m1_trans" where "m1_step5 Rb A B Kab Ts Ta \ {(s, s'). \ \previous guards:\ runs s Rb = Some (Resp, [A, B], []) \ (Kab \ Domain (leak s) \ (Kab, B) \ azC (runs s)) \ \ \authorization guard\ \ \guard for showing agreement with server on \(Kab, A, rsl)\,\ \ \where \rsl = take rs_len nlb\; this agreement is non-injective\ (B \ bad \ (\Rs Na. Kab = sesK (Rs$sk) \ runs s Rs = Some (Serv, [A, B], [aNon Na, aNum Ts]))) \ \ \new guards:\ \ \guard for showing agreement with initiator \A\ on \(Kab, Ts, Ta)\\ (A \ bad \ B \ bad \ (\Ra nl. runs s Ra = Some (Init, [A, B], aKey Kab # aNum Ts # aNum Ta # nl))) \ \ \ensure recentness of session key\ clk s < Ts + Ls \ \ \check validity of authenticator and prevent its replay\ \ \'replays' with fresh authenticator ok!\ clk s < Ta + La \ (B, Kab, Ta) \ cache s \ \ \actions:\ s' = s\ runs := (runs s)(Rb \ (Resp, [A, B], [aKey Kab, aNum Ts, aNum Ta])), cache := insert (B, Kab, Ta) (cache s) \ }" definition \ \by @{term "A"}, refines @{term skip}\ m1_step6 :: "[rid_t, agent, agent, nonce, key, time, time] \ 'x m1_trans" where "m1_step6 Ra A B Na Kab Ts Ta \ {(s, s'). runs s Ra = Some (Init, [A, B], [aKey Kab, aNum Ts, aNum Ta]) \ \ \key recv'd before\ Na = Ra$na \ \ \fix parameter\ \ \check key's freshness [NEW]\ \ \\clk s < Ts + Ls \\\ \ \guard for showing agreement with \B\ on \Kab\, \Ts\, and \Ta\\ (A \ bad \ B \ bad \ (\Rb. runs s Rb = Some (Resp, [A, B], [aKey Kab, aNum Ts, aNum Ta]))) \ \ \actions: (redundant) update local state marks successful termination\ s' = s\ runs := (runs s)(Ra \ (Init, [A, B], [aKey Kab, aNum Ts, aNum Ta, END])) \ }" definition \ \by attacker, refines @{term m1a_leak}\ m1_leak :: "[rid_t, agent, agent, nonce, time] \ 'x m1_trans" where "m1_leak Rs A B Na Ts \ {(s, s1). \ \guards:\ runs s Rs = Some (Serv, [A, B], [aNon Na, aNum Ts]) \ (clk s \ Ts + Ls) \ \ \only compromise 'old' session keys\ \ \actions:\ \ \record session key as leaked;\ s1 = s\ leak := insert (sesK (Rs$sk), A, B, Na, Ts) (leak s) \ }" text \Clock tick event\ definition \ \refines @{term skip}\ m1_tick :: "time \ 'x m1_trans" where "m1_tick T \ {(s, s'). s' = s\ clk := clk s + T \ }" text \Purge event: purge cache of expired timestamps\ definition \ \refines @{term skip}\ m1_purge :: "agent \ 'x m1_trans" where "m1_purge A \ {(s, s'). s' = s\ cache := cache s - {(A, K, T) | A K T. (A, K, T) \ cache s \ T + La \ clk s } \ }" (******************************************************************************) subsection \Specification\ (******************************************************************************) definition m1_init :: "m1_state set" where "m1_init \ { \ runs = Map.empty, leak = corrKey \ {undefined}, clk = 0, cache = {} \ }" definition m1_trans :: "'x m1_trans" where "m1_trans \ (\A B Ra Rb Rs Na Kab Ts Ta T. m1_step1 Ra A B Na \ m1_step2 Rb A B \ m1_step3 Rs A B Kab Na Ts \ m1_step4 Ra A B Na Kab Ts Ta \ m1_step5 Rb A B Kab Ts Ta \ m1_step6 Ra A B Na Kab Ts Ta \ m1_leak Rs A B Na Ts \ m1_tick T \ m1_purge A \ Id )" definition m1 :: "(m1_state, m1_obs) spec" where "m1 \ \ init = m1_init, trans = m1_trans, obs = id \" lemmas m1_loc_defs = m1_def m1_init_def m1_trans_def m1_step1_def m1_step2_def m1_step3_def m1_step4_def m1_step5_def m1_step6_def m1_leak_def m1_purge_def m1_tick_def lemmas m1_defs = m1_loc_defs m1a_defs lemma m1_obs_id [simp]: "obs m1 = id" by (simp add: m1_def) (******************************************************************************) subsection \Invariants\ (******************************************************************************) subsubsection \inv0: Finite domain\ (*inv**************************************************************************) text \There are only finitely many runs. This is needed to establish the responder/initiator agreement.\ definition m1_inv0_fin :: "'x m1_pred" where "m1_inv0_fin \ {s. finite (dom (runs s))}" lemmas m1_inv0_finI = m1_inv0_fin_def [THEN setc_def_to_intro, rule_format] lemmas m1_inv0_finE [elim] = m1_inv0_fin_def [THEN setc_def_to_elim, rule_format] lemmas m1_inv0_finD = m1_inv0_fin_def [THEN setc_def_to_dest, rule_format] text \Invariance proofs.\ lemma PO_m1_inv0_fin_init [iff]: "init m1 \ m1_inv0_fin" by (auto simp add: m1_defs intro!: m1_inv0_finI) lemma PO_m1_inv0_fin_trans [iff]: "{m1_inv0_fin} trans m1 {> m1_inv0_fin}" by (auto simp add: PO_hoare_defs m1_defs intro!: m1_inv0_finI) lemma PO_m1_inv0_fin [iff]: "reach m1 \ m1_inv0_fin" by (rule inv_rule_incr, auto del: subsetI) subsubsection \inv1: Caching invariant for responder\ (*inv**************************************************************************) definition m1_inv1r_cache :: "'x m1_pred" where "m1_inv1r_cache \ {s. \Rb A B Kab Ts Ta nl. runs s Rb = Some (Resp, [A, B], aKey Kab # aNum Ts # aNum Ta # nl) \ clk s < Ta + La \ (B, Kab, Ta) \ cache s }" lemmas m1_inv1r_cacheI = m1_inv1r_cache_def [THEN setc_def_to_intro, rule_format] lemmas m1_inv1r_cacheE [elim] = m1_inv1r_cache_def [THEN setc_def_to_elim, rule_format] lemmas m1_inv1r_cacheD = m1_inv1r_cache_def [THEN setc_def_to_dest, rule_format, rotated 1] text \Invariance proof\ lemma PO_m1_inv1r_cache_init [iff]: "init m1 \ m1_inv1r_cache" by (auto simp add: m1_defs intro!: m1_inv1r_cacheI) lemma PO_m1_inv1r_cache_trans [iff]: "{m1_inv1r_cache} trans m1 {> m1_inv1r_cache}" apply (auto simp add: PO_hoare_defs m1_defs intro!: m1_inv1r_cacheI dest: m1_inv1r_cacheD) apply (auto dest: m1_inv1r_cacheD) done lemma PO_m1_inv1r_cache [iff]: "reach m1 \ m1_inv1r_cache" by (rule inv_rule_basic) (auto del: subsetI) (******************************************************************************) subsection \Refinement of \m1a\\ (******************************************************************************) subsubsection \Simulation relation\ (******************************************************************************) text \The abstraction removes all but the first freshness identifiers (corresponding to \Kab\ and \Ts\) from the initiator and responder frames and leaves the server's freshness ids untouched. \ overloading is_len' \ "is_len" rs_len' \ "rs_len" begin definition is_len_def [simp]: "is_len' \ 1::nat" definition rs_len_def [simp]: "rs_len' \ 1::nat" end fun rm1a1 :: "role_t \ atom list \ atom list" where "rm1a1 Init = take (Suc is_len)" \ \take \Kab\, \Ts\; drop \Ta\\ | "rm1a1 Resp = take (Suc rs_len)" \ \take \Kab\, \Ts\; drop \Ta\\ | "rm1a1 Serv = id" \ \take \Na\, \Ts\\ abbreviation runs1a1 :: "runs_t \ runs_t" where "runs1a1 \ map_runs rm1a1" lemma knC_runs1a1 [simp]: "knC (runs1a1 runz) = knC runz" apply (auto simp add: map_runs_def elim!: knC.cases) apply (rename_tac b, case_tac b, auto) apply (rename_tac b, case_tac b, auto) apply (rule knC_init, auto simp add: map_runs_def) apply (rule knC_resp, auto simp add: map_runs_def) apply (rule_tac knC_serv, auto simp add: map_runs_def) done text \med1a1: The mediator function maps a concrete observation (i.e., run) to an abstract one.\ text \R1a1: The simulation relation is defined in terms of the mediator function.\ definition med1a1 :: "m1_obs \ m1a_obs" where "med1a1 s \ \ runs = runs1a1 (runs s), m1x_state.leak = Domain (leak s) \" definition R1a1 :: "(m1a_state \ m1_state) set" where "R1a1 \ {(s, t). s = med1a1 t}" lemmas R1a1_defs = R1a1_def med1a1_def subsubsection \Refinement proof\ (******************************************************************************) lemma PO_m1_step1_refines_m1a_step1: "{R1a1} (m1a_step1 Ra A B Na), (m1_step1 Ra A B Na) {> R1a1}" by (auto simp add: PO_rhoare_defs R1a1_defs m1_defs) lemma PO_m1_step2_refines_m1a_step2: "{R1a1} (m1a_step2 Rb A B), (m1_step2 Rb A B) {> R1a1}" by (auto simp add: PO_rhoare_defs R1a1_defs m1_defs) lemma PO_m1_step3_refines_m1a_step3: "{R1a1} (m1a_step3 Rs A B Kab Na [aNum Ts]), (m1_step3 Rs A B Kab Na Ts) {> R1a1}" by (auto simp add: PO_rhoare_defs R1a1_defs m1_defs) lemma PO_m1_step4_refines_m1a_step4: "{R1a1} (m1a_step4 Ra A B Na Kab [aNum Ts]), (m1_step4 Ra A B Na Kab Ts Ta) {> R1a1}" by (auto simp add: PO_rhoare_defs R1a1_defs m1_defs map_runs_def) lemma PO_m1_step5_refines_m1a_step5: "{R1a1} (m1a_step5 Rb A B Kab [aNum Ts]), (m1_step5 Rb A B Kab Ts Ta) {> R1a1}" by (auto simp add: PO_rhoare_defs R1a1_defs m1_defs map_runs_def) lemma PO_m1_step6_refines_m1a_skip: "{R1a1} Id, (m1_step6 Ra A B Na Kab Ts Ta) {> R1a1}" by (auto simp add: PO_rhoare_defs R1a1_defs m1_defs map_runs_def) lemma PO_m1_leak_refines_m1a_leak: "{R1a1} (m1a_leak Rs), (m1_leak Rs A B Na Ts) {> R1a1}" by (auto simp add: PO_rhoare_defs R1a1_defs m1_defs map_runs_def dest: dom_lemmas) lemma PO_m1_tick_refines_m1a_skip: "{R1a1} Id, (m1_tick T) {> R1a1}" by (auto simp add: PO_rhoare_defs R1a1_defs m1_defs map_runs_def) lemma PO_m1_purge_refines_m1a_skip: "{R1a1} Id, (m1_purge A) {> R1a1}" by (auto simp add: PO_rhoare_defs R1a1_defs m1_defs map_runs_def) text \All together now...\ lemmas PO_m1_trans_refines_m1a_trans = PO_m1_step1_refines_m1a_step1 PO_m1_step2_refines_m1a_step2 PO_m1_step3_refines_m1a_step3 PO_m1_step4_refines_m1a_step4 PO_m1_step5_refines_m1a_step5 PO_m1_step6_refines_m1a_skip PO_m1_leak_refines_m1a_leak PO_m1_tick_refines_m1a_skip PO_m1_purge_refines_m1a_skip lemma PO_m1_refines_init_m1a [iff]: "init m1 \ R1a1``(init m1a)" by (auto simp add: R1a1_defs m1_defs intro!: s0g_secrecyI) lemma PO_m1_refines_trans_m1a [iff]: "{R1a1} (trans m1a), (trans m1) {> R1a1}" apply (auto simp add: m1_def m1_trans_def m1a_def m1a_trans_def intro!: PO_m1_trans_refines_m1a_trans) apply (force intro!: PO_m1_trans_refines_m1a_trans)+ done text \Observation consistency.\ lemma obs_consistent_med1a1 [iff]: "obs_consistent R1a1 med1a1 m1a m1" by (auto simp add: obs_consistent_def R1a1_def m1a_def m1_def) text \Refinement result.\ lemma PO_m1_refines_m1a [iff]: "refines R1a1 med1a1 m1a m1" by (rule Refinement_basic) (auto del: subsetI) lemma m1_implements_m1a [iff]: "implements med1a1 m1a m1" by (rule refinement_soundness) (fast) subsubsection \inv (inherited): Secrecy\ (*invh*************************************************************************) text \Secrecy, as external and internal invariant\ definition m1_secrecy :: "'x m1_pred" where "m1_secrecy \ {s. knC (runs s) \ azC (runs s) \ Domain (leak s) \ UNIV}" lemmas m1_secrecyI = m1_secrecy_def [THEN setc_def_to_intro, rule_format] lemmas m1_secrecyE [elim] = m1_secrecy_def [THEN setc_def_to_elim, rule_format] lemma PO_m1_obs_secrecy [iff]: "oreach m1 \ m1_secrecy" apply (rule_tac Q=m1x_secrecy in external_invariant_translation) apply (auto del: subsetI) apply (fastforce simp add: med1a1_def intro!: m1_secrecyI) done lemma PO_m1_secrecy [iff]: "reach m1 \ m1_secrecy" by (rule external_to_internal_invariant) (auto del: subsetI) subsubsection \inv (inherited): Responder auth server.\ (*invh*************************************************************************) definition m1_inv2r_serv :: "'x m1r_pred" where "m1_inv2r_serv \ {s. \A B Rb Kab Ts nlb. B \ bad \ runs s Rb = Some (Resp, [A, B], aKey Kab # aNum Ts # nlb) \ (\Rs Na. Kab = sesK (Rs$sk) \ runs s Rs = Some (Serv, [A, B], [aNon Na, aNum Ts])) }" lemmas m1_inv2r_servI = m1_inv2r_serv_def [THEN setc_def_to_intro, rule_format] lemmas m1_inv2r_servE [elim] = m1_inv2r_serv_def [THEN setc_def_to_elim, rule_format] lemmas m1_inv2r_servD = m1_inv2r_serv_def [THEN setc_def_to_dest, rule_format, rotated -1] text \Proof of invariance.\ lemma PO_m1_inv2r_serv [iff]: "reach m1 \ m1_inv2r_serv" apply (rule_tac Sa=m1a and Pa=m1a_inv2r_serv and Qa=m1a_inv2r_serv and Q=m1_inv2r_serv in internal_invariant_translation) apply (auto del: subsetI) \ \1 subgoal\ apply (auto simp add: vimage_def intro!: m1_inv2r_servI) apply (simp add: m1a_inv2r_serv_def med1a1_def) apply (rename_tac x A B Rb Kab Ts nlb) apply (drule_tac x=A in spec) apply (drule_tac x=B in spec, clarsimp) apply (drule_tac x=Rb in spec) apply (drule_tac x=Kab in spec) apply (drule_tac x="[aNum Ts]" in spec) apply (auto simp add: map_runs_def) done subsubsection \inv (inherited): Initiator auth server.\ (*invh*************************************************************************) text \Simplified version of invariant \m1a_inv2i_serv\.\ definition m1_inv2i_serv :: "'x m1r_pred" where "m1_inv2i_serv \ {s. \A B Ra Kab Ts nla. A \ bad \ runs s Ra = Some (Init, [A, B], aKey Kab # aNum Ts # nla) \ (\Rs. Kab = sesK (Rs$sk) \ runs s Rs = Some (Serv, [A, B], [aNon (Ra$na), aNum Ts])) }" lemmas m1_inv2i_servI = m1_inv2i_serv_def [THEN setc_def_to_intro, rule_format] lemmas m1_inv2i_servE [elim] = m1_inv2i_serv_def [THEN setc_def_to_elim, rule_format] lemmas m1_inv2i_servD = m1_inv2i_serv_def [THEN setc_def_to_dest, rule_format, rotated -1] text \Proof of invariance.\ lemma PO_m1_inv2i_serv [iff]: "reach m1 \ m1_inv2i_serv" apply (rule_tac Pa=m1a_inv2i_serv and Qa=m1a_inv2i_serv and Q=m1_inv2i_serv in internal_invariant_translation) apply (auto del: subsetI) \ \1 subgoal\ apply (auto simp add: m1a_inv2i_serv_def med1a1_def vimage_def intro!: m1_inv2i_servI) apply (rename_tac x A B Ra Kab Ts nla) apply (drule_tac x=A in spec, clarsimp) apply (drule_tac x=B in spec) apply (drule_tac x=Ra in spec) apply (drule_tac x=Kab in spec) apply (drule_tac x="[aNum Ts]" in spec) apply (auto simp add: map_runs_def) done declare PO_m1_inv2i_serv [THEN subsetD, intro] subsubsection \inv (inherited): Initiator key freshness\ (*invh*************************************************************************) definition m1_inv1_ifresh :: "'x m1_pred" where "m1_inv1_ifresh \ {s. \A A' B B' Ra Ra' Kab nl nl'. runs s Ra = Some (Init, [A, B], aKey Kab # nl) \ runs s Ra' = Some (Init, [A', B'], aKey Kab # nl') \ A \ bad \ B \ bad \ Kab \ Domain (leak s) \ Ra = Ra' }" lemmas m1_inv1_ifreshI = m1_inv1_ifresh_def [THEN setc_def_to_intro, rule_format] lemmas m1_inv1_ifreshE [elim] = m1_inv1_ifresh_def [THEN setc_def_to_elim, rule_format] lemmas m1_inv1_ifreshD = m1_inv1_ifresh_def [THEN setc_def_to_dest, rule_format, rotated 1] lemma PO_m1_ifresh [iff]: "reach m1 \ m1_inv1_ifresh" apply (rule_tac Pa=m1a_inv1_ifresh and Qa=m1a_inv1_ifresh and Q=m1_inv1_ifresh in internal_invariant_translation) apply (auto del: subsetI) apply (auto simp add: med1a1_def map_runs_def vimage_def m1_inv1_ifresh_def) done (******************************************************************************) subsection \Refinement of \a0i\ for responder/initiator\ (******************************************************************************) text \The responder injectively agrees with the initiator on @{term "Kab"}, @{term "Ts"}, and @{term "Ta"}.\ subsubsection \Simulation relation\ (******************************************************************************) text \We define two auxiliary functions to reconstruct the signals of the initial model from completed initiator and responder runs.\ type_synonym risig = "key \ time \ time" abbreviation ri_running :: "[runs_t, agent, agent, key, time, time] \ rid_t set" where "ri_running runz A B Kab Ts Ta \ {Ra. \nl. runz Ra = Some (Init, [A, B], aKey Kab # aNum Ts # aNum Ta # nl) }" abbreviation ri_commit :: "[runs_t, agent, agent, key, time, time] \ rid_t set" where "ri_commit runz A B Kab Ts Ta \ {Rb. \nl. runz Rb = Some (Resp, [A, B], aKey Kab # aNum Ts # aNum Ta # nl) }" fun ri_runs2sigs :: "runs_t \ risig signal \ nat" where "ri_runs2sigs runz (Running [B, A] (Kab, Ts, Ta)) = card (ri_running runz A B Kab Ts Ta)" | "ri_runs2sigs runz (Commit [B, A] (Kab, Ts, Ta)) = card (ri_commit runz A B Kab Ts Ta)" | "ri_runs2sigs runz _ = 0" text \Simulation relation and mediator function. We map completed initiator and responder runs to commit and running signals, respectively.\ definition med_a0iim1_ri :: "m1_obs \ risig a0i_obs" where "med_a0iim1_ri o1 \ \ signals = ri_runs2sigs (runs o1), corrupted = {} \" definition R_a0iim1_ri :: "(risig a0i_state \ m1_state) set" where "R_a0iim1_ri \ {(s, t). signals s = ri_runs2sigs (runs t) \ corrupted s = {} }" lemmas R_a0iim1_ri_defs = R_a0iim1_ri_def med_a0iim1_ri_def subsubsection \Lemmas about the auxiliary functions\ (******************************************************************************) text \Other lemmas\ lemma ri_runs2sigs_empty [simp]: "runz = Map.empty \ ri_runs2sigs runz = (\s. 0)" by (rule ext, erule rev_mp) (rule ri_runs2sigs.induct, auto) lemma finite_ri_running [simp, intro]: "finite (dom runz) \ finite (ri_running runz A B Kab Ts Ta)" by (auto intro: finite_subset dest: dom_lemmas) lemma finite_ri_commit [simp, intro]: "finite (dom runz) \ finite (ri_commit runz A B Kab Ts Ta)" by (auto intro: finite_subset dest: dom_lemmas) text \Update lemmas\ lemma ri_runs2sigs_upd_init_none [simp]: "\ Na \ dom runz \ \ ri_runs2sigs (runz(Na \ (Init, [A, B], []))) = ri_runs2sigs runz" by (rule ext, erule rev_mp, rule ri_runs2sigs.induct) (auto dest: dom_lemmas) lemma ri_runs2sigs_upd_resp_none [simp]: "\ Rb \ dom runz \ \ ri_runs2sigs (runz(Rb \ (Resp, [A, B], []))) = ri_runs2sigs runz" by (rule ext, erule rev_mp, rule ri_runs2sigs.induct) (auto dest: dom_lemmas) lemma ri_runs2sigs_upd_serv [simp]: "\ Rs \ dom runz \ \ ri_runs2sigs (runz(Rs \ (Serv, [A, B], [aNon Na, aNum Ts]))) = ri_runs2sigs runz" by (rule ext, erule rev_mp, rule ri_runs2sigs.induct) (auto dest: dom_lemmas) lemma ri_runs2sigs_upd_init_some [simp]: "\ runz Ra = Some (Init, [A, B], []); finite (dom runz) \ \ ri_runs2sigs (runz(Ra \ (Init, [A, B], [aKey Kab, aNum Ts, aNum Ta]))) = (ri_runs2sigs runz)( Running [B, A] (Kab, Ts, Ta) := Suc (card (ri_running runz A B Kab Ts Ta)))" apply (rule ext, (erule rev_mp)+) apply (rule ri_runs2sigs.induct, auto) \ \1 subgoal\ apply (rename_tac runz) apply (rule_tac s="card (insert Ra (ri_running runz A B Kab Ts Ta))" in trans, fast, auto) done lemma ri_runs2sigs_upd_resp_some [simp]: "\ runz Rb = Some (Resp, [A, B], []); finite (dom runz) \ \ ri_runs2sigs (runz(Rb \ (Resp, [A, B], [aKey Kab, aNum Ts, aNum Ta]))) = (ri_runs2sigs runz)( Commit [B, A] (Kab, Ts, Ta) := Suc (card (ri_commit runz A B Kab Ts Ta)))" apply (rule ext, (erule rev_mp)+) apply (rule ri_runs2sigs.induct, auto) \ \1 subgoal\ apply (rename_tac runz) apply (rule_tac s="card (insert Rb (ri_commit runz A B Kab Ts Ta))" in trans, fast, auto) done lemma ri_runs2sigs_upd_init_some2 [simp]: "\ runz Ra = Some (Init, [A, B], [aKey Kab, aNum Ts, aNum Ta]) \ \ ri_runs2sigs (runz(Ra \ (Init, [A, B], [aKey Kab, aNum Ts, aNum Ta, END]))) = ri_runs2sigs runz" by (rule ext, erule rev_mp, rule ri_runs2sigs.induct) (auto dest: dom_lemmas) subsubsection \Refinement proof\ (******************************************************************************) lemma PO_m1_step1_refines_a0_ri_skip: "{R_a0iim1_ri} Id, (m1_step1 Ra A B Na) {> R_a0iim1_ri}" by (auto simp add: PO_rhoare_defs R_a0iim1_ri_defs m1_defs) lemma PO_m1_step2_refines_a0_ri_skip: "{R_a0iim1_ri} Id, (m1_step2 Rb A B) {> R_a0iim1_ri}" by (auto simp add: PO_rhoare_defs R_a0iim1_ri_defs m1_defs) lemma PO_m1_step3_refines_a0_ri_skip: "{R_a0iim1_ri} Id, (m1_step3 Rs A B Kab Na Ts) {> R_a0iim1_ri}" by (auto simp add: PO_rhoare_defs R_a0iim1_ri_defs m1_defs) lemma PO_m1_step4_refines_a0_ri_running: "{R_a0iim1_ri \ UNIV \ m1_inv0_fin} (a0i_running [B, A] (Kab, Ts, Ta)), (m1_step4 Ra A B Na Kab Ts Ta) {> R_a0iim1_ri}" by (auto simp add: PO_rhoare_defs R_a0iim1_ri_defs a0i_defs m1_defs) lemma PO_m1_step5_refines_a0_ri_commit: "{R_a0iim1_ri \ UNIV \ (m1_inv1r_cache \ m1_inv0_fin)} (a0i_commit [B, A] (Kab, Ts, Ta)), (m1_step5 Rb A B Kab Ts Ta) {> R_a0iim1_ri}" apply (auto simp add: PO_rhoare_defs R_a0iim1_ri_defs a0i_defs m1_defs) \ \2 subgoals\ apply (rename_tac s t aa ab ac ba Rs Na Ra nl, subgoal_tac "card (ri_commit (runs t) A B (sesK (Rs$sk)) Ts Ta) = 0 \ card (ri_running (runs t) A B (sesK (Rs$sk)) Ts Ta) > 0", auto) apply (rename_tac s t Rs Na Ra nl, subgoal_tac "card (ri_commit (runs t) A B (sesK (Rs$sk)) Ts Ta) = 0 \ card (ri_running (runs t) A B (sesK (Rs$sk)) Ts Ta) > 0", auto) done lemma PO_m1_step6_refines_a0_ri_skip: "{R_a0iim1_ri} Id, (m1_step6 Ra A B Na Kab Ts Ta) {> R_a0iim1_ri}" by (auto simp add: PO_rhoare_defs R_a0iim1_ri_defs m1_defs) lemma PO_m1_leak_refines_a0_ri_skip: "{R_a0iim1_ri} Id, (m1_leak Rs A B Na Ts) {> R_a0iim1_ri}" by (auto simp add: PO_rhoare_defs R_a0iim1_ri_defs a0i_defs m1_defs) lemma PO_m1_tick_refines_a0_ri_skip: "{R_a0iim1_ri} Id, (m1_tick T) {> R_a0iim1_ri}" by (auto simp add: PO_rhoare_defs R_a0iim1_ri_defs m1_defs) lemma PO_m1_purge_refines_a0_ri_skip: "{R_a0iim1_ri} Id, (m1_purge A) {> R_a0iim1_ri}" by (auto simp add: PO_rhoare_defs R_a0iim1_ri_defs m1_defs) text \All together now...\ lemmas PO_m1_trans_refines_a0_ri_trans = PO_m1_step1_refines_a0_ri_skip PO_m1_step2_refines_a0_ri_skip PO_m1_step3_refines_a0_ri_skip PO_m1_step4_refines_a0_ri_running PO_m1_step5_refines_a0_ri_commit PO_m1_step6_refines_a0_ri_skip PO_m1_leak_refines_a0_ri_skip PO_m1_tick_refines_a0_ri_skip PO_m1_purge_refines_a0_ri_skip lemma PO_m1_refines_init_a0_ri [iff]: "init m1 \ R_a0iim1_ri``(init a0i)" by (auto simp add: R_a0iim1_ri_defs a0i_defs m1_defs intro!: exI [where x="\signals = \s. 0, corrupted = {} \"]) lemma PO_m1_refines_trans_a0_ri [iff]: "{R_a0iim1_ri \ a0i_inv1_iagree \ (m1_inv1r_cache \ m1_inv0_fin)} (trans a0i), (trans m1) {> R_a0iim1_ri}" by (force simp add: m1_def m1_trans_def a0i_def a0i_trans_def intro!: PO_m1_trans_refines_a0_ri_trans) lemma obs_consistent_med_a0iim1_ri [iff]: "obs_consistent (R_a0iim1_ri \ a0i_inv1_iagree \ (m1_inv1r_cache \ m1_inv0_fin)) med_a0iim1_ri a0i m1" by (auto simp add: obs_consistent_def R_a0iim1_ri_def med_a0iim1_ri_def a0i_def m1_def) text \Refinement result.\ lemma PO_m1_refines_a0ii_ri [iff]: "refines (R_a0iim1_ri \ a0i_inv1_iagree \ (m1_inv1r_cache \ m1_inv0_fin)) med_a0iim1_ri a0i m1" by (rule Refinement_using_invariants) (auto) lemma m1_implements_a0ii_ri: "implements med_a0iim1_ri a0i m1" by (rule refinement_soundness) (fast) subsubsection \inv3 (inherited): Responder and initiator\ (*invh*************************************************************************) text \This is a translation of the agreement property to Level 1. It follows from the refinement and is needed to prove inv4 below.\ definition m1_inv3r_init :: "'x m1_pred" where "m1_inv3r_init \ {s. \A B Rb Kab Ts Ta nlb. B \ bad \ A \ bad \ Kab \ Domain (leak s) \ runs s Rb = Some (Resp, [A, B], aKey Kab # aNum Ts # aNum Ta # nlb) \ (\Ra nla. runs s Ra = Some (Init, [A, B], aKey Kab # aNum Ts # aNum Ta # nla)) }" lemmas m1_inv3r_initI = m1_inv3r_init_def [THEN setc_def_to_intro, rule_format] lemmas m1_inv3r_initE [elim] = m1_inv3r_init_def [THEN setc_def_to_elim, rule_format] lemmas m1_inv3r_initD = m1_inv3r_init_def [THEN setc_def_to_dest, rule_format, rotated -1] text \Invariance proof.\ lemma PO_m1_inv3r_init [iff]: "reach m1 \ m1_inv3r_init" apply (rule INV_from_Refinement_basic [OF PO_m1_refines_a0ii_ri]) apply (auto simp add: R_a0iim1_ri_def a0i_inv1_iagree_def intro!: m1_inv3r_initI) \ \1 subgoal\ apply (rename_tac s A B Rb Kab Ts Ta nlb a) apply (drule_tac x="[B, A]" in spec, clarsimp) apply (drule_tac x="Kab" in spec) apply (drule_tac x="Ts" in spec) apply (drule_tac x="Ta" in spec) apply (subgoal_tac "card (ri_commit (runs s) A B Kab Ts Ta) > 0", auto) done subsubsection \inv4: Key freshness for responder\ (*inv*************************************************************************) definition m1_inv4_rfresh :: "'x m1_pred" where "m1_inv4_rfresh \ {s. \Rb1 Rb2 A1 A2 B1 B2 Kab Ts1 Ts2 Ta1 Ta2. runs s Rb1 = Some (Resp, [A1, B1], [aKey Kab, aNum Ts1, aNum Ta1]) \ runs s Rb2 = Some (Resp, [A2, B2], [aKey Kab, aNum Ts2, aNum Ta2]) \ B1 \ bad \ A1 \ bad \ Kab \ Domain (leak s) \ Rb1 = Rb2 }" lemmas m1_inv4_rfreshI = m1_inv4_rfresh_def [THEN setc_def_to_intro, rule_format] lemmas m1_inv4_rfreshE [elim] = m1_inv4_rfresh_def [THEN setc_def_to_elim, rule_format] lemmas m1_inv4_rfreshD = m1_inv4_rfresh_def [THEN setc_def_to_dest, rule_format, rotated 1] text \Proof of key freshness for responder. All cases except step5 are straightforward.\ lemma PO_m1_inv4_rfresh_step5: "{m1_inv4_rfresh \ m1_inv3r_init \ m1_inv2r_serv \ m1_inv1r_cache \ m1_secrecy \ m1_inv1_ifresh} (m1_step5 Rb A B Kab Ts Ta) {> m1_inv4_rfresh}" apply (auto simp add: PO_hoare_defs m1_defs intro!: m1_inv4_rfreshI) apply (auto dest: m1_inv4_rfreshD) apply (auto dest: m1_inv2r_servD) \ \5 subgoals\ apply (drule m1_inv2r_servD, auto) apply (elim azC.cases, auto) apply (drule m1_inv2r_servD, auto) apply (elim azC.cases, auto) apply (drule m1_inv2r_servD, auto) apply (elim azC.cases, auto) apply (rename_tac Rb2 A2 B2 Ts2 Ta2 s Rs Na Ra nl) apply (case_tac "B2 \ bad") apply (thin_tac "(sesK (Rs$sk), B) \ azC (runs s)") apply (subgoal_tac "(sesK (Rs$sk), B2) \ azC (runs s)") apply (erule azC.cases, auto) apply (erule m1_secrecyE, auto) apply (case_tac "A2 \ bad", auto dest: m1_inv2r_servD) apply (frule m1_inv3r_initD, auto) apply (rename_tac Raa nla, subgoal_tac "Raa = Ra", auto) \ \uses cache invariant\ apply (frule m1_inv3r_initD, auto) apply (rename_tac Raa nla, subgoal_tac "Raa = Ra", auto) \ \uses cache invariant\ done lemmas PO_m1_inv4_rfresh_step5_lemmas = PO_m1_inv4_rfresh_step5 lemma PO_m1_inv4_rfresh_init [iff]: "init m1 \ m1_inv4_rfresh" by (auto simp add: m1_defs intro!: m1_inv4_rfreshI) lemma PO_m1_inv4_rfresh_trans [iff]: "{m1_inv4_rfresh \ m1_inv3r_init \ m1_inv2r_serv \ m1_inv1r_cache \ m1_secrecy \ m1_inv1_ifresh} trans m1 {> m1_inv4_rfresh}" by (auto simp add: m1_def m1_trans_def intro!: PO_m1_inv4_rfresh_step5_lemmas) (auto simp add: PO_hoare_defs m1_defs intro!: m1_inv4_rfreshI dest: m1_inv4_rfreshD) lemma PO_m1_inv4_rfresh [iff]: "reach m1 \ m1_inv4_rfresh" apply (rule_tac J="m1_inv3r_init \ m1_inv2r_serv \ m1_inv1r_cache \ m1_secrecy \ m1_inv1_ifresh" in inv_rule_incr) apply (auto simp add: Int_assoc del: subsetI) done lemma PO_m1_obs_inv4_rfresh [iff]: "oreach m1 \ m1_inv4_rfresh" by (rule external_from_internal_invariant) (auto del: subsetI) (******************************************************************************) subsection \Refinement of \a0i\ for initiator/responder\ (******************************************************************************) text \The initiator injectively agrees with the responder on \Kab\, \Ts\, and \Ta\.\ subsubsection \Simulation relation\ (******************************************************************************) text \We define two auxiliary functions to reconstruct the signals of the initial model from completed initiator and responder runs.\ type_synonym irsig = "key \ time \ time" abbreviation ir_running :: "[runs_t, agent, agent, key, time, time] \ rid_t set" where "ir_running runz A B Kab Ts Ta \ {Rb. \nl. runz Rb = Some (Resp, [A, B], aKey Kab # aNum Ts # aNum Ta # nl) }" abbreviation ir_commit :: "[runs_t, agent, agent, key, time, time] \ rid_t set" where "ir_commit runz A B Kab Ts Ta \ {Ra. \nl. runz Ra = Some (Init, [A, B], aKey Kab # aNum Ts # aNum Ta # END # nl) }" fun ir_runs2sigs :: "runs_t \ risig signal \ nat" where "ir_runs2sigs runz (Running [A, B] (Kab, Ts, Ta)) = card (ir_running runz A B Kab Ts Ta)" | "ir_runs2sigs runz (Commit [A, B] (Kab, Ts, Ta)) = card (ir_commit runz A B Kab Ts Ta)" | "ir_runs2sigs runz _ = 0" text \Simulation relation and mediator function. We map completed initiator and responder runs to commit and running signals, respectively.\ definition med_a0iim1_ir :: "m1_obs \ irsig a0i_obs" where "med_a0iim1_ir o1 \ \ signals = ir_runs2sigs (runs o1), corrupted = {} \" definition R_a0iim1_ir :: "(irsig a0i_state \ m1_state) set" where "R_a0iim1_ir \ {(s, t). signals s = ir_runs2sigs (runs t) \ corrupted s = {} }" lemmas R_a0iim1_ir_defs = R_a0iim1_ir_def med_a0iim1_ir_def subsubsection \Lemmas about the auxiliary functions\ (******************************************************************************) lemma ir_runs2sigs_empty [simp]: "runz = Map.empty \ ir_runs2sigs runz = (\s. 0)" by (rule ext, erule rev_mp) (rule ir_runs2sigs.induct, auto) (* already proven higher up: lemma ir_running_finite [simp, intro]: "finite (dom runz) \ finite (ir_running runz A B Kab Ts Ta)" by (auto intro: finite_subset dest: dom_lemmas) *) lemma ir_commit_finite [simp, intro]: "finite (dom runz) \ finite (ir_commit runz A B Kab Ts Ta)" by (auto intro: finite_subset dest: dom_lemmas) text \Update lemmas\ lemma ir_runs2sigs_upd_init_none [simp]: "\ Ra \ dom runz \ \ ir_runs2sigs (runz(Ra \ (Init, [A, B], []))) = ir_runs2sigs runz" by (rule ext, erule rev_mp) (rule ir_runs2sigs.induct, auto dest: dom_lemmas) lemma ir_runs2sigs_upd_resp_none [simp]: "\ Rb \ dom runz \ \ ir_runs2sigs (runz(Rb \ (Resp, [A, B], []))) = ir_runs2sigs runz" by (rule ext, erule rev_mp) (rule ir_runs2sigs.induct, auto dest: dom_lemmas) lemma ir_runs2sigs_upd_serv [simp]: "\ Rs \ dom (runs y) \ - \ ir_runs2sigs (runs y(Rs \ (Serv, [A, B], [aNon Na, aNum Ts]))) + \ ir_runs2sigs ((runs y)(Rs \ (Serv, [A, B], [aNon Na, aNum Ts]))) = ir_runs2sigs (runs y)" by (rule ext, erule rev_mp) (rule ir_runs2sigs.induct, auto dest: dom_lemmas) lemma ir_runs2sigs_upd_init_some [simp]: "\ runz Ra = Some (Init, [A, B], []) \ \ ir_runs2sigs (runz(Ra \ (Init, [A, B], [aKey Kab, aNum Ts, aNum Ta]))) = ir_runs2sigs runz" by (rule ext, erule rev_mp) (rule ir_runs2sigs.induct, auto dest: dom_lemmas) lemma ir_runs2sigs_upd_resp_some_raw: assumes "runz Rb = Some (Resp, [A, B], [])" "finite (dom runz)" shows "ir_runs2sigs (runz(Rb \ (Resp, [A, B], [aKey Kab, aNum Ts, aNum Ta]))) s = ((ir_runs2sigs runz)( Running [A, B] (Kab, Ts, Ta) := Suc (card (ir_running runz A B Kab Ts Ta)))) s" using assms proof (induct rule: ir_runs2sigs.induct) case (1 runz A B Kab Ts Ta) note H = this hence "Rb \ ir_running runz A B Kab Ts Ta" by auto moreover with H have "card (insert Rb (ir_running runz A B Kab Ts Ta)) = Suc (card (ir_running runz A B Kab Ts Ta))" by auto ultimately show ?case by (auto elim: subst) qed (auto) lemma ir_runs2sigs_upd_resp_some [simp]: "\ runz Rb = Some (Resp, [A, B], []); finite (dom runz) \ \ ir_runs2sigs (runz(Rb \ (Resp, [A, B], [aKey Kab, aNum Ts, aNum Ta]))) = (ir_runs2sigs runz)( Running [A, B] (Kab, Ts, Ta) := Suc (card (ir_running runz A B Kab Ts Ta)))" by (intro ext ir_runs2sigs_upd_resp_some_raw) lemma ir_runs2sigs_upd_init_some2_raw: assumes "runz Ra = Some (Init, [A, B], [aKey Kab, aNum Ts, aNum Ta])" "finite (dom runz)" shows "ir_runs2sigs (runz(Ra \ (Init, [A, B], [aKey Kab, aNum Ts, aNum Ta, END]))) s = ((ir_runs2sigs runz)( Commit [A, B] (Kab, Ts, Ta) := Suc (card (ir_commit runz A B Kab Ts Ta)))) s" using assms proof (induct runz s rule: ir_runs2sigs.induct) case (2 runz A B Kab Ts Ta) note H = this from H have "Ra \ ir_commit runz A B Kab Ts Ta" by auto moreover with H have "card (insert Ra (ir_commit runz A B Kab Ts Ta)) = Suc (card (ir_commit runz A B Kab Ts Ta))" by (auto) ultimately show ?case by (auto elim: subst) qed (auto) lemma ir_runs2sigs_upd_init_some2 [simp]: "\ runz Na = Some (Init, [A, B], [aKey Kab, aNum Ts, aNum Ta]); finite (dom runz) \ \ ir_runs2sigs (runz(Na \ (Init, [A, B], [aKey Kab, aNum Ts, aNum Ta, END]))) = (ir_runs2sigs runz)( Commit [A, B] (Kab, Ts, Ta) := Suc (card (ir_commit runz A B Kab Ts Ta)))" by (intro ir_runs2sigs_upd_init_some2_raw ext) subsubsection \Refinement proof\ (******************************************************************************) lemma PO_m1_step1_refines_ir_a0ii_skip: "{R_a0iim1_ir} Id, (m1_step1 Ra A B Na) {> R_a0iim1_ir}" by (simp add: PO_rhoare_defs R_a0iim1_ir_defs m1_defs, safe, auto) lemma PO_m1_step2_refines_ir_a0ii_skip: "{R_a0iim1_ir} Id, (m1_step2 Rb A B) {> R_a0iim1_ir}" by (simp add: PO_rhoare_defs R_a0iim1_ir_defs m1_defs, safe, auto) lemma PO_m1_step3_refines_ir_a0ii_skip: "{R_a0iim1_ir} Id, (m1_step3 Rs A B Kab Na Ts) {> R_a0iim1_ir}" by (simp add: PO_rhoare_defs R_a0iim1_ir_defs a0i_defs m1_defs, safe, auto) lemma PO_m1_step4_refines_ir_a0ii_skip: "{R_a0iim1_ir} Id, (m1_step4 Ra A B Na Kab Ts Ta) {> R_a0iim1_ir}" by (simp add: PO_rhoare_defs R_a0iim1_ir_defs m1_defs, safe, auto) lemma PO_m1_step5_refines_ir_a0ii_running: "{R_a0iim1_ir \ UNIV \ m1_inv0_fin} (a0i_running [A, B] (Kab, Ts, Ta)), (m1_step5 Rb A B Kab Ts Ta) {> R_a0iim1_ir}" by (simp add: PO_rhoare_defs R_a0iim1_ir_defs a0i_defs m1_defs, safe, auto) lemma PO_m1_step6_refines_ir_a0ii_commit: "{R_a0iim1_ir \ UNIV \ m1_inv0_fin} (a0n_commit [A, B] (Kab, Ts, Ta)), (m1_step6 Ra A B Na Kab Ts Ta) {> R_a0iim1_ir}" by (simp add: PO_rhoare_defs R_a0iim1_ir_defs a0n_defs m1_defs, safe, auto) lemma PO_m1_leak_refines_ir_a0ii_skip: "{R_a0iim1_ir} Id, (m1_leak Rs A B Na Ts) {> R_a0iim1_ir}" by (simp add: PO_rhoare_defs R_a0iim1_ir_defs a0n_defs m1_defs, safe, auto) lemma PO_m1_tick_refines_ir_a0ii_skip: "{R_a0iim1_ir} Id, (m1_tick T) {> R_a0iim1_ir}" by (simp add: PO_rhoare_defs R_a0iim1_ir_defs m1_defs, safe, auto) lemma PO_m1_purge_refines_ir_a0ii_skip: "{R_a0iim1_ir} Id, (m1_purge A) {> R_a0iim1_ir}" by (simp add: PO_rhoare_defs R_a0iim1_ir_defs m1_defs, safe, auto) text \All together now...\ lemmas PO_m1_trans_refines_ir_a0ii_trans = PO_m1_step1_refines_ir_a0ii_skip PO_m1_step2_refines_ir_a0ii_skip PO_m1_step3_refines_ir_a0ii_skip PO_m1_step4_refines_ir_a0ii_skip PO_m1_step5_refines_ir_a0ii_running PO_m1_step6_refines_ir_a0ii_commit PO_m1_leak_refines_ir_a0ii_skip PO_m1_tick_refines_ir_a0ii_skip PO_m1_purge_refines_ir_a0ii_skip lemma PO_m1_refines_init_ir_a0ii [iff]: "init m1 \ R_a0iim1_ir``(init a0n)" by (auto simp add: R_a0iim1_ir_defs a0n_defs m1_defs intro!: exI [where x="\signals = \s. 0, corrupted = {}\"]) lemma PO_m1_refines_trans_ir_a0ii [iff]: "{R_a0iim1_ir \ UNIV \ m1_inv0_fin} (trans a0n), (trans m1) {> R_a0iim1_ir}" by (auto simp add: m1_def m1_trans_def a0n_def a0n_trans_def intro!: PO_m1_trans_refines_ir_a0ii_trans) text \Observation consistency.\ lemma obs_consistent_med_a0iim1_ir [iff]: "obs_consistent (R_a0iim1_ir \ UNIV \ m1_inv0_fin) med_a0iim1_ir a0n m1" by (auto simp add: obs_consistent_def R_a0iim1_ir_def med_a0iim1_ir_def a0n_def m1_def) text \Refinement result.\ lemma PO_m1_refines_a0ii_ir [iff]: "refines (R_a0iim1_ir \ UNIV \ m1_inv0_fin) med_a0iim1_ir a0n m1" by (rule Refinement_using_invariants) (auto) lemma m1_implements_a0ii_ir: "implements med_a0iim1_ir a0n m1" by (rule refinement_soundness) (fast) end diff --git a/thys/Security_Protocol_Refinement/Key_establish/m3_ds.thy b/thys/Security_Protocol_Refinement/Key_establish/m3_ds.thy --- a/thys/Security_Protocol_Refinement/Key_establish/m3_ds.thy +++ b/thys/Security_Protocol_Refinement/Key_establish/m3_ds.thy @@ -1,756 +1,756 @@ (******************************************************************************* Project: Development of Security Protocols by Refinement Module: Key_establish/m3_ds.thy (Isabelle/HOL 2016-1) ID: $Id: m3_ds.thy 132890 2016-12-24 10:25:57Z csprenge $ Authors: Christoph Sprenger, ETH Zurich Ivano Somaini, ETH Zurich Key distribution protocols Level 3: Denning-Sacco protocol Copyright (c) 2009-2016 Christoph Sprenger, Ivano Somaini Licence: LGPL *******************************************************************************) section \Denning-Sacco protocol (L3)\ theory m3_ds imports m2_ds "../Refinement/Message" begin text \ We model the Denning-Sacco protocol: \[ \begin{array}{lll} \mathrm{M1.} & A \rightarrow S: & A, B \\ \mathrm{M2.} & S \rightarrow A: & \{Kab, B, Ts, Na, \{Kab, A, Ts\}_{Kbs} \}_{Kas} \\ \mathrm{M3.} & A \rightarrow B: & \{Kab, A, Ts\}_{Kbs} \end{array} \] \ text \Proof tool configuration. Avoid annoying automatic unfolding of \dom\.\ declare domIff [simp, iff del] (******************************************************************************) subsection \Setup\ (******************************************************************************) text \Now we can define the initial key knowledge.\ overloading ltkeySetup' \ ltkeySetup begin definition ltkeySetup_def: "ltkeySetup' \ {(sharK C, A) | C A. A = C \ A = Sv}" end lemma corrKey_shrK_bad [simp]: "corrKey = shrK`bad" by (auto simp add: keySetup_def ltkeySetup_def corrKey_def) (******************************************************************************) subsection \State\ (******************************************************************************) text \The secure channels are star-shaped to/from the server. Therefore, we have only one agent in the relation.\ record m3_state = "m1_state" + IK :: "msg set" \ \intruder knowledge\ text \Observable state: @{term "runs"}, @{term "leak"}, @{term "clk"}, and @{term "cache"}.\ type_synonym m3_obs = "m2_obs" definition m3_obs :: "m3_state \ m3_obs" where "m3_obs s \ \ runs = runs s, leak = leak s, clk = clk s \" type_synonym m3_pred = "m3_state set" type_synonym m3_trans = "(m3_state \ m3_state) set" (******************************************************************************) subsection \Events\ (******************************************************************************) text \Protocol events.\ definition \ \by @{term "A"}, refines @{term "m2_step1"}\ m3_step1 :: "[rid_t, agent, agent] \ m3_trans" where "m3_step1 Ra A B \ {(s, s1). \ \guards:\ Ra \ dom (runs s) \ \ \\Ra\ is fresh\ \ \actions:\ s1 = s\ runs := (runs s)(Ra \ (Init, [A, B], [])), IK := insert \Agent A, Agent B\ (IK s) \ \send \M1\\ \ }" definition \ \by @{term "B"}, refines @{term "m2_step2"}\ m3_step2 :: "[rid_t, agent, agent] \ m3_trans" where "m3_step2 \ m1_step2" definition \ \by @{text "Server"}, refines @{term m2_step3}\ m3_step3 :: "[rid_t, agent, agent, key, time] \ m3_trans" where "m3_step3 Rs A B Kab Ts \ {(s, s1). \ \guards:\ Rs \ dom (runs s) \ \ \fresh server run\ Kab = sesK (Rs$sk) \ \ \fresh session key\ \Agent A, Agent B\ \ IK s \ \ \recv \M1\\ Ts = clk s \ \ \fresh timestamp\ \ \actions:\ \ \record session key and send \M2\\ s1 = s\ runs := (runs s)(Rs \ (Serv, [A, B], [aNum Ts])), IK := insert (Crypt (shrK A) \ \send \M2\\ \Key Kab, Agent B, Number Ts, Crypt (shrK B) \Key Kab, Agent A, Number Ts\\) (IK s) \ }" definition \ \by @{term "A"}, refines @{term m2_step4}\ m3_step4 :: "[rid_t, agent, agent, key, time, msg] \ m3_trans" where "m3_step4 Ra A B Kab Ts X \ {(s, s1). \ \guards:\ runs s Ra = Some (Init, [A, B], []) \ \ \key not yet recv'd\ Crypt (shrK A) \ \recv \M2\\ \Key Kab, Agent B, Number Ts, X\ \ IK s \ \ \check freshness of session key\ clk s < Ts + Ls \ \ \actions:\ \ \record session key and send \M3\\ s1 = s\ runs := (runs s)(Ra \ (Init, [A, B], [aKey Kab, aNum Ts])), IK := insert X (IK s) \ \send \M3\\ \ }" definition \ \by @{term "B"}, refines @{term m2_step5}\ m3_step5 :: "[rid_t, agent, agent, key, time] \ m3_trans" where "m3_step5 Rb A B Kab Ts \ {(s, s1). \ \guards:\ runs s Rb = Some (Resp, [A, B], []) \ \ \key not yet recv'd\ Crypt (shrK B) \Key Kab, Agent A, Number Ts\ \ IK s \ \ \recv \M3\\ \ \ensure freshness of session key; replays with fresh authenticator ok!\ clk s < Ts + Ls \ \ \actions:\ \ \record session key\ s1 = s\ runs := (runs s)(Rb \ (Resp, [A, B], [aKey Kab, aNum Ts])) \ }" text \Clock tick event\ definition \ \refines @{term "m2_tick"}\ m3_tick :: "time \ m3_trans" where "m3_tick \ m1_tick" text \Session key compromise.\ definition \ \refines @{term m2_leak}\ m3_leak :: "rid_t \ m3_trans" where "m3_leak Rs \ {(s, s1). \ \guards:\ Rs \ dom (runs s) \ fst (the (runs s Rs)) = Serv \ \ \compromise server run \Rs\\ \ \actions:\ \ \record session key as leaked and add it to intruder knowledge\ s1 = s\ leak := insert (sesK (Rs$sk)) (leak s), IK := insert (Key (sesK (Rs$sk))) (IK s) \ }" text \Intruder fake event. The following "Dolev-Yao" event generates all intruder-derivable messages.\ definition \ \refines @{term "m2_fake"}\ m3_DY_fake :: "m3_trans" where "m3_DY_fake \ {(s, s1). \ \actions:\ s1 = s\ IK := synth (analz (IK s)) \ \ \take DY closure\ }" (******************************************************************************) subsection \Transition system\ (******************************************************************************) definition m3_init :: "m3_pred" where "m3_init \ { \ runs = Map.empty, leak = shrK`bad, clk = 0, IK = Key`shrK`bad \ }" definition m3_trans :: "m3_trans" where "m3_trans \ (\A B Ra Rb Rs Kab Ts T X. m3_step1 Ra A B \ m3_step2 Rb A B \ m3_step3 Rs A B Kab Ts \ m3_step4 Ra A B Kab Ts X \ m3_step5 Rb A B Kab Ts \ m3_tick T \ m3_leak Rs \ m3_DY_fake \ Id )" definition m3 :: "(m3_state, m3_obs) spec" where "m3 \ \ init = m3_init, trans = m3_trans, obs = m3_obs \" lemmas m3_loc_defs = m3_def m3_init_def m3_trans_def m3_obs_def m3_step1_def m3_step2_def m3_step3_def m3_step4_def m3_step5_def m3_tick_def m3_leak_def m3_DY_fake_def lemmas m3_defs = m3_loc_defs m2_defs (******************************************************************************) subsection \Invariants\ (******************************************************************************) text \Specialized injection that we can apply more aggressively.\ lemmas analz_Inj_IK = analz.Inj [where H="IK s" for s] lemmas parts_Inj_IK = parts.Inj [where H="IK s" for s] declare parts_Inj_IK [dest!] declare analz_into_parts [dest] subsubsection \inv1: Secrecy of pre-distributed shared keys\ (******************************************************************************) definition m3_inv1_lkeysec :: "m3_pred" where "m3_inv1_lkeysec \ {s. \C. (Key (shrK C) \ parts (IK s) \ C \ bad) \ (C \ bad \ Key (shrK C) \ IK s) }" lemmas m3_inv1_lkeysecI = m3_inv1_lkeysec_def [THEN setc_def_to_intro, rule_format] lemmas m3_inv1_lkeysecE [elim] = m3_inv1_lkeysec_def [THEN setc_def_to_elim, rule_format] lemmas m3_inv1_lkeysecD = m3_inv1_lkeysec_def [THEN setc_def_to_dest, rule_format] text \Invariance proof.\ lemma PO_m3_inv1_lkeysec_init [iff]: "init m3 \ m3_inv1_lkeysec" by (auto simp add: m3_defs intro!: m3_inv1_lkeysecI) lemma PO_m3_inv1_lkeysec_trans [iff]: "{m3_inv1_lkeysec} trans m3 {> m3_inv1_lkeysec}" by (auto simp add: PO_hoare_defs m3_defs intro!: m3_inv1_lkeysecI) (auto dest!: Body) lemma PO_m3_inv1_lkeysec [iff]: "reach m3 \ m3_inv1_lkeysec" by (rule inv_rule_incr) (fast+) text \Useful simplifier lemmas\ lemma m3_inv1_lkeysec_for_parts [simp]: "\ s \ m3_inv1_lkeysec \ \ Key (shrK C) \ parts (IK s) \ C \ bad" by auto lemma m3_inv1_lkeysec_for_analz [simp]: "\ s \ m3_inv1_lkeysec \ \ Key (shrK C) \ analz (IK s) \ C \ bad" by auto subsubsection \inv2: Ticket shape for honestly encrypted M2\ (******************************************************************************) definition m3_inv2_ticket :: "m3_pred" where "m3_inv2_ticket \ {s. \A B T K X. A \ bad \ Crypt (shrK A) \Key K, Agent B, Number T, X\ \ parts (IK s) \ X = Crypt (shrK B) \Key K, Agent A, Number T\ \ K \ range sesK }" lemmas m3_inv2_ticketI = m3_inv2_ticket_def [THEN setc_def_to_intro, rule_format] lemmas m3_inv2_ticketE [elim] = m3_inv2_ticket_def [THEN setc_def_to_elim, rule_format] lemmas m3_inv2_ticketD = m3_inv2_ticket_def [THEN setc_def_to_dest, rule_format, rotated -1] text \Invariance proof.\ lemma PO_m3_inv2_ticket_init [iff]: "init m3 \ m3_inv2_ticket" by (auto simp add: m3_defs intro!: m3_inv2_ticketI) lemma PO_m3_inv2_ticket_trans [iff]: "{m3_inv2_ticket \ m3_inv1_lkeysec} trans m3 {> m3_inv2_ticket}" apply (auto simp add: PO_hoare_defs m3_defs intro!: m3_inv2_ticketI) apply (auto dest: m3_inv2_ticketD) \ \2 subgoals, from step4\ apply (drule parts_cut, drule Body, auto dest: m3_inv2_ticketD)+ done lemma PO_m3_inv2_ticket [iff]: "reach m3 \ m3_inv2_ticket" by (rule inv_rule_incr) (auto del: subsetI) subsubsection \inv3: Session keys not used to encrypt other session keys\ (******************************************************************************) text \Session keys are not used to encrypt other keys. Proof requires generalization to sets of session keys.\ definition m3_inv3_sesK_compr :: "m3_pred" where "m3_inv3_sesK_compr \ {s. \K KK. KK \ range sesK \ (Key K \ analz (Key`KK \ (IK s))) = (K \ KK \ Key K \ analz (IK s)) }" lemmas m3_inv3_sesK_comprI = m3_inv3_sesK_compr_def [THEN setc_def_to_intro, rule_format] lemmas m3_inv3_sesK_comprE = m3_inv3_sesK_compr_def [THEN setc_def_to_elim, rule_format] lemmas m3_inv3_sesK_comprD = m3_inv3_sesK_compr_def [THEN setc_def_to_dest, rule_format] text \Additional lemma\ lemmas insert_commute_Key = insert_commute [where x="Key K" for K] lemmas m3_inv3_sesK_compr_simps = m3_inv3_sesK_comprD m3_inv3_sesK_comprD [where KK="{Kab}" for Kab, simplified] m3_inv3_sesK_comprD [where KK="insert Kab KK" for Kab KK, simplified] insert_commute_Key \ \to get the keys to the front\ text \Invariance proof.\ lemma PO_m3_inv3_sesK_compr_step4: "{m3_inv3_sesK_compr \ m3_inv2_ticket \ m3_inv1_lkeysec} m3_step4 Ra A B Kab Ts X {> m3_inv3_sesK_compr}" proof - { fix K KK s assume H: "s \ m3_inv1_lkeysec" "s \ m3_inv3_sesK_compr" "s \ m3_inv2_ticket" "runs s Ra = Some (Init, [A, B], [])" "KK \ range sesK" "Crypt (shrK A) \Key Kab, Agent B, Number Ts, X\ \ analz (IK s)" have "(Key K \ analz (insert X (Key ` KK \ IK s))) = (K \ KK \ Key K \ analz (insert X (IK s)))" proof (cases "A \ bad") case True show ?thesis proof - note H moreover with \A \ bad\ have "X \ analz (IK s)" by (auto dest!: Decrypt) moreover hence "X \ analz (Key ` KK \ IK s)" by (auto intro: analz_mono [THEN [2] rev_subsetD]) ultimately show ?thesis by (auto simp add: m3_inv3_sesK_compr_simps analz_insert_eq) qed next case False thus ?thesis using H by (fastforce simp add: m3_inv3_sesK_compr_simps dest!: m3_inv2_ticketD [OF analz_into_parts]) qed } thus ?thesis by (auto simp add: PO_hoare_defs m3_defs intro!: m3_inv3_sesK_comprI dest!: analz_Inj_IK) qed text \All together now.\ lemmas PO_m3_inv3_sesK_compr_trans_lemmas = PO_m3_inv3_sesK_compr_step4 lemma PO_m3_inv3_sesK_compr_init [iff]: "init m3 \ m3_inv3_sesK_compr" by (auto simp add: m3_defs intro!: m3_inv3_sesK_comprI) lemma PO_m3_inv3_sesK_compr_trans [iff]: "{m3_inv3_sesK_compr \ m3_inv2_ticket \ m3_inv1_lkeysec} trans m3 {> m3_inv3_sesK_compr}" by (auto simp add: m3_def m3_trans_def intro!: PO_m3_inv3_sesK_compr_trans_lemmas) (auto simp add: PO_hoare_defs m3_defs m3_inv3_sesK_compr_simps intro!: m3_inv3_sesK_comprI) lemma PO_m3_inv3_sesK_compr [iff]: "reach m3 \ m3_inv3_sesK_compr" by (rule_tac J="m3_inv2_ticket \ m3_inv1_lkeysec" in inv_rule_incr) (auto) (******************************************************************************) subsection \Refinement\ (******************************************************************************) subsubsection \Message abstraction and simulation relation\ (******************************************************************************) text \Abstraction function on sets of messages.\ inductive_set abs_msg :: "msg set \ chmsg set" for H :: "msg set" where am_M1: "\Agent A, Agent B\ \ H \ Insec A B (Msg []) \ abs_msg H" | am_M2a: "Crypt (shrK C) \Key K, Agent B, Number T, X\ \ H \ Secure Sv C (Msg [aAgt B, aKey K, aNum T]) \ abs_msg H" | am_M2b: "Crypt (shrK C) \Key K, Agent A, Number T\ \ H \ Secure Sv C (Msg [aKey K, aAgt A, aNum T]) \ abs_msg H" text \R23: The simulation relation. This is a data refinement of the insecure and secure channels of refinement 2.\ definition R23_msgs :: "(m2_state \ m3_state) set" where "R23_msgs \ {(s, t). abs_msg (parts (IK t)) \ chan s }" definition R23_keys :: "(m2_state \ m3_state) set" where "R23_keys \ {(s, t). \KK K. KK \ range sesK \ Key K \ analz (Key`KK \ (IK t)) \ aKey K \ extr (aKey`KK \ ik0) (chan s) }" definition R23_pres :: "(m2_state \ m3_state) set" where "R23_pres \ {(s, t). runs s = runs t \ clk s = clk t \ leak s = leak t}" definition R23 :: "(m2_state \ m3_state) set" where "R23 \ R23_msgs \ R23_keys \ R23_pres" lemmas R23_defs = R23_def R23_msgs_def R23_keys_def R23_pres_def text \The mediator function is the identity here.\ definition med32 :: "m3_obs \ m2_obs" where "med32 \ id" lemmas R23_msgsI = R23_msgs_def [THEN rel_def_to_intro, simplified, rule_format] lemmas R23_msgsE [elim] = R23_msgs_def [THEN rel_def_to_elim, simplified, rule_format] lemmas R23_keysI = R23_keys_def [THEN rel_def_to_intro, simplified, rule_format] lemmas R23_keysE [elim] = R23_keys_def [THEN rel_def_to_elim, simplified, rule_format] lemmas R23_presI = R23_pres_def [THEN rel_def_to_intro, simplified, rule_format] lemmas R23_presE [elim] = R23_pres_def [THEN rel_def_to_elim, simplified, rule_format] lemmas R23_intros = R23_msgsI R23_keysI R23_presI text \Lemmas for various instantiations (for keys).\ lemmas R23_keys_dest = R23_keys_def [THEN rel_def_to_dest, simplified, rule_format, rotated 2] lemmas R23_keys_dests = R23_keys_dest R23_keys_dest [where KK="{}", simplified] R23_keys_dest [where KK="{K'}" for K', simplified] R23_keys_dest [where KK="insert K' KK" for K' KK, simplified, OF _ _ conjI] subsubsection \General lemmas\ (******************************************************************************) text \General facts about @{term "abs_msg"}\ declare abs_msg.intros [intro!] declare abs_msg.cases [elim!] lemma abs_msg_empty: "abs_msg {} = {}" by (auto) lemma abs_msg_Un [simp]: "abs_msg (G \ H) = abs_msg G \ abs_msg H" by (auto) lemma abs_msg_mono [elim]: "\ m \ abs_msg G; G \ H \ \ m \ abs_msg H" by (auto) lemma abs_msg_insert_mono [intro]: "\ m \ abs_msg H \ \ m \ abs_msg (insert m' H)" by (auto) text \Facts about @{term "abs_msg"} concerning abstraction of fakeable messages. This is crucial for proving the refinement of the intruder event.\ lemma abs_msg_DY_subset_fakeable: "\ (s, t) \ R23_msgs; (s, t) \ R23_keys; (s, t) \ R23_non; t \ m3_inv1_lkeysec \ \ abs_msg (synth (analz (IK t))) \ fake ik0 (dom (runs s)) (chan s)" apply (auto) \ \4 subgoals, deal with replays first\ apply (blast) defer 1 apply (blast) \ \remaining 2 subgoals are real fakes\ apply (rule fake_StatCh, auto dest: R23_keys_dests)+ done subsubsection \Refinement proof\ (******************************************************************************) text \Pair decomposition. These were set to \texttt{elim!}, which is too agressive here.\ declare MPair_analz [rule del, elim] declare MPair_parts [rule del, elim] text \Protocol events.\ lemma PO_m3_step1_refines_m2_step1: "{R23} (m2_step1 Ra A B), (m3_step1 Ra A B) {> R23}" by (auto simp add: PO_rhoare_defs R23_def m3_defs intro!: R23_intros) (auto) lemma PO_m3_step2_refines_m2_step2: "{R23} (m2_step2 Rb A B), (m3_step2 Rb A B) {> R23}" by (auto simp add: PO_rhoare_defs R23_def m3_defs intro!: R23_intros) (auto) lemma PO_m3_step3_refines_m2_step3: "{R23 \ (m2_inv3a_sesK_compr) \ (m3_inv3_sesK_compr \ m3_inv1_lkeysec)} (m2_step3 Rs A B Kab Ts), (m3_step3 Rs A B Kab Ts) {> R23}" proof - { fix s t assume H: "(s, t) \ R23_msgs" "(s, t) \ R23_keys" "(s, t) \ R23_pres" "s \ m2_inv3a_sesK_compr" "t \ m3_inv3_sesK_compr" "t \ m3_inv1_lkeysec" "Kab = sesK (Rs$sk)" "Rs \ dom (runs t)" "\Agent A, Agent B\ \ parts (IK t)" let ?s'= - "s\ runs := runs s(Rs \ (Serv, [A, B], [aNum (clk t)])), + "s\ runs := (runs s)(Rs \ (Serv, [A, B], [aNum (clk t)])), chan := insert (Secure Sv A (Msg [aAgt B, aKey Kab, aNum (clk t)])) (insert (Secure Sv B (Msg [aKey Kab, aAgt A, aNum (clk t)])) (chan s)) \" let ?t'= - "t\ runs := runs t(Rs \ (Serv, [A, B], [aNum (clk t)])), + "t\ runs := (runs t)(Rs \ (Serv, [A, B], [aNum (clk t)])), IK := insert (Crypt (shrK A) \ Key Kab, Agent B, Number (clk t), Crypt (shrK B) \ Key Kab, Agent A, Number (clk t) \\) (IK t) \" have "(?s', ?t') \ R23_msgs" using H by (-) (rule R23_intros, auto) moreover have "(?s', ?t') \ R23_keys" using H by (-) (rule R23_intros, auto simp add: m2_inv3a_sesK_compr_simps m3_inv3_sesK_compr_simps dest: R23_keys_dests) moreover have "(?s', ?t') \ R23_pres" using H by (-) (rule R23_intros, auto) moreover note calculation } thus ?thesis by (auto simp add: PO_rhoare_defs R23_def m3_defs) qed lemma PO_m3_step4_refines_m2_step4: "{R23 \ UNIV \ (m3_inv3_sesK_compr \ m3_inv2_ticket \ m3_inv1_lkeysec)} (m2_step4 Ra A B Kab Ts), (m3_step4 Ra A B Kab Ts X) {> R23}" proof - { fix s t assume H: "(s, t) \ R23_msgs" "(s, t) \ R23_keys" "(s, t) \ R23_pres" "t \ m3_inv3_sesK_compr" "t \ m3_inv2_ticket" "t \ m3_inv1_lkeysec" "runs t Ra = Some (Init, [A, B], [])" "Crypt (shrK A) \Key Kab, Agent B, Number Ts, X\ \ analz (IK t)" - let ?s' = "s\ runs := runs s(Ra \ (Init, [A, B], [aKey Kab, aNum Ts])) \" - and ?t' = "t\ runs := runs t(Ra \ (Init, [A, B], [aKey Kab, aNum Ts])), + let ?s' = "s\ runs := (runs s)(Ra \ (Init, [A, B], [aKey Kab, aNum Ts])) \" + and ?t' = "t\ runs := (runs t)(Ra \ (Init, [A, B], [aKey Kab, aNum Ts])), IK := insert X (IK t) \" from H have "Secure Sv A (Msg [aAgt B, aKey Kab, aNum Ts]) \ chan s" by (auto) moreover have "X \ parts (IK t)" using H by (auto dest!: Body MPair_parts) hence "(?s', ?t') \ R23_msgs" using H by (auto intro!: R23_intros) (auto) moreover have "(?s', ?t') \ R23_keys" proof (cases) assume "A \ bad" show ?thesis proof - note H moreover hence "X \ analz (IK t)" using \A \ bad\ by (-) (drule Decrypt, auto) ultimately show ?thesis by (-) (rule R23_intros, auto dest!: analz_cut intro: analz_monotonic) qed next assume "A \ bad" show ?thesis proof - note H moreover with \A \ bad\ have "X = Crypt (shrK B) \Key Kab, Agent A, Number Ts \ \ Kab \ range sesK" by (auto dest!: m3_inv2_ticketD) moreover { assume H1: "Key (shrK B) \ analz (IK t)" have "aKey Kab \ extr ik0 (chan s)" proof - note calculation moreover hence "Secure Sv B (Msg [aKey Kab, aAgt A, aNum Ts]) \ chan s" by (-) (drule analz_into_parts, drule Body, elim MPair_parts, auto) ultimately show ?thesis using H1 by auto qed } ultimately show ?thesis by (-) (rule R23_intros, auto simp add: m3_inv3_sesK_compr_simps) qed qed moreover have "(?s', ?t') \ R23_pres" using H by (auto intro!: R23_intros) moreover note calculation } thus ?thesis by (auto simp add: PO_rhoare_defs R23_def m3_defs dest!: analz_Inj_IK) qed lemma PO_m3_step5_refines_m2_step5: "{R23} (m2_step5 Rb A B Kab Ts), (m3_step5 Rb A B Kab Ts) {> R23}" by (auto simp add: PO_rhoare_defs R23_def m3_defs intro!: R23_intros) (auto) lemma PO_m3_tick_refines_m2_tick: "{R23} (m2_tick T), (m3_tick T) {>R23}" by (auto simp add: PO_rhoare_defs R23_def m3_defs intro!: R23_intros) (auto) text \Intruder events.\ lemma PO_m3_leak_refines_m2_leak: "{R23} (m2_leak Rs), (m3_leak Rs) {>R23}" by (auto simp add: PO_rhoare_defs R23_def m3_defs intro!: R23_intros) (auto dest: R23_keys_dests) lemma PO_m3_DY_fake_refines_m2_fake: "{R23 \ UNIV \ (m3_inv1_lkeysec)} m2_fake, m3_DY_fake {> R23}" apply (auto simp add: PO_rhoare_defs R23_def m3_defs intro!: R23_intros del: abs_msg.cases) apply (auto intro: abs_msg_DY_subset_fakeable [THEN subsetD] del: abs_msg.cases dest: R23_keys_dests) done text \All together now...\ lemmas PO_m3_trans_refines_m2_trans = PO_m3_step1_refines_m2_step1 PO_m3_step2_refines_m2_step2 PO_m3_step3_refines_m2_step3 PO_m3_step4_refines_m2_step4 PO_m3_step5_refines_m2_step5 PO_m3_tick_refines_m2_tick PO_m3_leak_refines_m2_leak PO_m3_DY_fake_refines_m2_fake lemma PO_m3_refines_init_m2 [iff]: "init m3 \ R23``(init m2)" by (auto simp add: R23_def m3_defs ik0_def intro!: R23_intros) lemma PO_m3_refines_trans_m2 [iff]: "{R23 \ (m2_inv3a_sesK_compr) \ (m3_inv3_sesK_compr \ m3_inv2_ticket \ m3_inv1_lkeysec)} (trans m2), (trans m3) {> R23}" by (auto simp add: m3_def m3_trans_def m2_def m2_trans_def) (blast intro!: PO_m3_trans_refines_m2_trans)+ lemma PO_m3_observation_consistent [iff]: "obs_consistent R23 med32 m2 m3" by (auto simp add: obs_consistent_def R23_def med32_def m3_defs) text \Refinement result.\ lemma m3_refines_m2 [iff]: "refines (R23 \ (m2_inv3a_sesK_compr) \ (m3_inv3_sesK_compr \ m3_inv2_ticket \ m3_inv1_lkeysec)) med32 m2 m3" by (rule Refinement_using_invariants) (auto) lemma m3_implements_m2 [iff]: "implements med32 m2 m3" by (rule refinement_soundness) (auto) end diff --git a/thys/Security_Protocol_Refinement/Key_establish/m3_ds_par.thy b/thys/Security_Protocol_Refinement/Key_establish/m3_ds_par.thy --- a/thys/Security_Protocol_Refinement/Key_establish/m3_ds_par.thy +++ b/thys/Security_Protocol_Refinement/Key_establish/m3_ds_par.thy @@ -1,605 +1,605 @@ (******************************************************************************* Project: Development of Security Protocols by Refinement Module: Key_establish/m3_ds_par.thy (Isabelle/HOL 2016-1) ID: $Id: m3_ds_par.thy 132890 2016-12-24 10:25:57Z csprenge $ Authors: Christoph Sprenger, ETH Zurich Ivano Somaini, ETH Zurich Key distribution protocols Level 3: parallel version of Denning-Sacco protocol Copyright (c) 2009-2016 Christoph Sprenger, Ivano Somaini Licence: LGPL *******************************************************************************) section \Denning-Sacco, direct variant (L3)\ theory m3_ds_par imports m2_ds "../Refinement/Message" begin text \ We model a direct implementation of the channel-based Denning-Sacco protocol at Level 2. In this version, there is no ticket forwarding. \[ \begin{array}{lll} \mathrm{M1.} & A \rightarrow S: & A, B \\ \mathrm{M2a.} & S \rightarrow A: & \{Kab, B, Ts\}_{Kas} \\ \mathrm{M2b.} & S \rightarrow B: & \{Kab, A, Ts\}_{Kbs} \end{array} \] \ text \Proof tool configuration. Avoid annoying automatic unfolding of \dom\.\ declare domIff [simp, iff del] (******************************************************************************) subsection \Setup\ (******************************************************************************) text \Now we can define the initial key knowledge.\ overloading ltkeySetup' \ ltkeySetup begin definition ltkeySetup_def: "ltkeySetup' \ {(sharK C, A) | C A. A = C \ A = Sv}" end lemma corrKey_shrK_bad [simp]: "corrKey = shrK`bad" by (auto simp add: keySetup_def ltkeySetup_def corrKey_def) (******************************************************************************) subsection \State\ (******************************************************************************) text \The secure channels are star-shaped to/from the server. Therefore, we have only one agent in the relation.\ record m3_state = "m1_state" + IK :: "msg set" \ \intruder knowledge\ text \Observable state: @{term "runs"}, @{term "leak"}, @{term "clk"}, and @{term "cache"}.\ type_synonym m3_obs = "m2_obs" definition m3_obs :: "m3_state \ m3_obs" where "m3_obs s \ \ runs = runs s, leak = leak s, clk = clk s \" type_synonym m3_pred = "m3_state set" type_synonym m3_trans = "(m3_state \ m3_state) set" (******************************************************************************) subsection \Events\ (******************************************************************************) text \Protocol events.\ definition \ \by @{term "A"}, refines @{term "m2_step1"}\ m3_step1 :: "[rid_t, agent, agent] \ m3_trans" where "m3_step1 Ra A B \ {(s, s1). \ \guards:\ Ra \ dom (runs s) \ \ \\Ra\ is fresh\ \ \actions:\ s1 = s\ runs := (runs s)(Ra \ (Init, [A, B], [])), IK := insert \Agent A, Agent B\ (IK s) \ \send \M1\\ \ }" definition \ \by @{term "B"}, refines @{term "m2_step2"}\ m3_step2 :: "[rid_t, agent, agent] \ m3_trans" where "m3_step2 \ m1_step2" definition \ \by @{text "Server"}, refines @{term m2_step3}\ m3_step3 :: "[rid_t, agent, agent, key, time] \ m3_trans" where "m3_step3 Rs A B Kab Ts \ {(s, s1). \ \guards:\ Rs \ dom (runs s) \ \ \fresh server run\ Kab = sesK (Rs$sk) \ \ \fresh session key\ \Agent A, Agent B\ \ IK s \ \ \recv \M1\\ Ts = clk s \ \ \fresh timestamp\ \ \actions:\ \ \record session key and send \M2\\ s1 = s\ runs := (runs s)(Rs \ (Serv, [A, B], [aNum Ts])), \ \send \M2a\, \M2b\\ IK := insert (Crypt (shrK A) \Agent B, Key Kab, Number Ts\) (insert (Crypt (shrK B) \Key Kab, Agent A, Number Ts\) (IK s)) \ }" definition \ \by @{term "A"}, refines @{term m2_step4}\ m3_step4 :: "[rid_t, agent, agent, key, time] \ m3_trans" where "m3_step4 Ra A B Kab Ts \ {(s, s1). \ \guards:\ runs s Ra = Some (Init, [A, B], []) \ \ \key not yet recv'd\ Crypt (shrK A) \ \recv \M2\\ \Agent B, Key Kab, Number Ts\ \ IK s \ \ \check freshness of session key\ clk s < Ts + Ls \ \ \actions:\ \ \record session key\ s1 = s\ runs := (runs s)(Ra \ (Init, [A, B], [aKey Kab, aNum Ts])) \ }" definition \ \by @{term "B"}, refines @{term m2_step5}\ m3_step5 :: "[rid_t, agent, agent, key, time] \ m3_trans" where "m3_step5 Rb A B Kab Ts \ {(s, s1). \ \guards:\ runs s Rb = Some (Resp, [A, B], []) \ \ \key not yet recv'd\ Crypt (shrK B) \Key Kab, Agent A, Number Ts\ \ IK s \ \ \recv \M3\\ \ \ensure freshness of session key; replays with fresh authenticator ok!\ clk s < Ts + Ls \ \ \actions:\ \ \record session key\ s1 = s\ runs := (runs s)(Rb \ (Resp, [A, B], [aKey Kab, aNum Ts])) \ }" text \Clock tick event\ definition \ \refines @{term "m2_tick"}\ m3_tick :: "time \ m3_trans" where "m3_tick \ m1_tick" text \Session key compromise.\ definition \ \refines @{term m2_leak}\ m3_leak :: "rid_t \ m3_trans" where "m3_leak Rs \ {(s, s1). \ \guards:\ Rs \ dom (runs s) \ fst (the (runs s Rs)) = Serv \ \ \compromise server run \Rs\\ \ \actions:\ \ \record session key as leaked and add it to intruder knowledge\ s1 = s\ leak := insert (sesK (Rs$sk)) (leak s), IK := insert (Key (sesK (Rs$sk))) (IK s) \ }" text \Intruder fake event. The following "Dolev-Yao" event generates all intruder-derivable messages.\ definition \ \refines @{term "m2_fake"}\ m3_DY_fake :: "m3_trans" where "m3_DY_fake \ {(s, s1). \ \actions:\ s1 = s\ IK := synth (analz (IK s)) \ \ \take DY closure\ }" (******************************************************************************) subsection \Transition system\ (******************************************************************************) definition m3_init :: "m3_pred" where "m3_init \ { \ runs = Map.empty, leak = shrK`bad, clk = 0, IK = Key`shrK`bad \ }" definition m3_trans :: "m3_trans" where "m3_trans \ (\A B Ra Rb Rs Kab Ts T. m3_step1 Ra A B \ m3_step2 Rb A B \ m3_step3 Rs A B Kab Ts \ m3_step4 Ra A B Kab Ts \ m3_step5 Rb A B Kab Ts \ m3_tick T \ m3_leak Rs \ m3_DY_fake \ Id )" definition m3 :: "(m3_state, m3_obs) spec" where "m3 \ \ init = m3_init, trans = m3_trans, obs = m3_obs \" lemmas m3_loc_defs = m3_def m3_init_def m3_trans_def m3_obs_def m3_step1_def m3_step2_def m3_step3_def m3_step4_def m3_step5_def m3_tick_def m3_leak_def m3_DY_fake_def lemmas m3_defs = m3_loc_defs m2_defs (******************************************************************************) subsection \Invariants\ (******************************************************************************) text \Specialized injection that we can apply more aggressively.\ lemmas analz_Inj_IK = analz.Inj [where H="IK s" for s] lemmas parts_Inj_IK = parts.Inj [where H="IK s" for s] declare parts_Inj_IK [dest!] declare analz_into_parts [dest] subsubsection \inv1: Secrecy of pre-distributed shared keys\ (******************************************************************************) definition m3_inv1_lkeysec :: "m3_pred" where "m3_inv1_lkeysec \ {s. \C. (Key (shrK C) \ parts (IK s) \ C \ bad) \ (C \ bad \ Key (shrK C) \ IK s) }" lemmas m3_inv1_lkeysecI = m3_inv1_lkeysec_def [THEN setc_def_to_intro, rule_format] lemmas m3_inv1_lkeysecE [elim] = m3_inv1_lkeysec_def [THEN setc_def_to_elim, rule_format] lemmas m3_inv1_lkeysecD = m3_inv1_lkeysec_def [THEN setc_def_to_dest, rule_format] text \Invariance proof.\ lemma PO_m3_inv1_lkeysec_init [iff]: "init m3 \ m3_inv1_lkeysec" by (auto simp add: m3_defs intro!: m3_inv1_lkeysecI) lemma PO_m3_inv1_lkeysec_trans [iff]: "{m3_inv1_lkeysec} trans m3 {> m3_inv1_lkeysec}" by (fastforce simp add: PO_hoare_defs m3_defs intro!: m3_inv1_lkeysecI) lemma PO_m3_inv1_lkeysec [iff]: "reach m3 \ m3_inv1_lkeysec" by (rule inv_rule_incr) (fast+) text \Useful simplifier lemmas\ lemma m3_inv1_lkeysec_for_parts [simp]: "\ s \ m3_inv1_lkeysec \ \ Key (shrK C) \ parts (IK s) \ C \ bad" by auto lemma m3_inv1_lkeysec_for_analz [simp]: "\ s \ m3_inv1_lkeysec \ \ Key (shrK C) \ analz (IK s) \ C \ bad" by auto subsubsection \inv3: Session keys not used to encrypt other session keys\ (******************************************************************************) text \Session keys are not used to encrypt other keys. Proof requires generalization to sets of session keys. NOTE: This invariant will be derived from the corresponding L2 invariant using the simulation relation. \ definition m3_inv3_sesK_compr :: "m3_pred" where "m3_inv3_sesK_compr \ {s. \K KK. KK \ range sesK \ (Key K \ analz (Key`KK \ (IK s))) = (K \ KK \ Key K \ analz (IK s)) }" lemmas m3_inv3_sesK_comprI = m3_inv3_sesK_compr_def [THEN setc_def_to_intro, rule_format] lemmas m3_inv3_sesK_comprE = m3_inv3_sesK_compr_def [THEN setc_def_to_elim, rule_format] lemmas m3_inv3_sesK_comprD = m3_inv3_sesK_compr_def [THEN setc_def_to_dest, rule_format] text \Additional lemma\ lemmas insert_commute_Key = insert_commute [where x="Key K" for K] lemmas m3_inv3_sesK_compr_simps = m3_inv3_sesK_comprD m3_inv3_sesK_comprD [where KK="insert Kab KK" for Kab KK, simplified] m3_inv3_sesK_comprD [where KK="{Kab}" for Kab, simplified] insert_commute_Key (******************************************************************************) subsection \Refinement\ (******************************************************************************) subsubsection \Message abstraction and simulation relation\ (******************************************************************************) text \Abstraction function on sets of messages.\ inductive_set abs_msg :: "msg set \ chmsg set" for H :: "msg set" where am_M1: "\Agent A, Agent B\ \ H \ Insec A B (Msg []) \ abs_msg H" | am_M2a: "Crypt (shrK C) \Agent B, Key K, Number T\ \ H \ Secure Sv C (Msg [aAgt B, aKey K, aNum T]) \ abs_msg H" | am_M2b: "Crypt (shrK C) \Key K, Agent A, Number T\ \ H \ Secure Sv C (Msg [aKey K, aAgt A, aNum T]) \ abs_msg H" text \R23: The simulation relation. This is a data refinement of the insecure and secure channels of refinement 2.\ definition R23_msgs :: "(m2_state \ m3_state) set" where "R23_msgs \ {(s, t). abs_msg (parts (IK t)) \ chan s }" definition R23_keys :: "(m2_state \ m3_state) set" where "R23_keys \ {(s, t). \KK K. KK \ range sesK \ Key K \ analz (Key`KK \ (IK t)) \ aKey K \ extr (aKey`KK \ ik0) (chan s) }" definition R23_pres :: "(m2_state \ m3_state) set" where "R23_pres \ {(s, t). runs s = runs t \ leak s = leak t \ clk s = clk t}" definition R23 :: "(m2_state \ m3_state) set" where "R23 \ R23_msgs \ R23_keys \ R23_pres" lemmas R23_defs = R23_def R23_msgs_def R23_keys_def R23_pres_def text \The mediator function is the identity here.\ definition med32 :: "m3_obs \ m2_obs" where "med32 \ id" lemmas R23_msgsI = R23_msgs_def [THEN rel_def_to_intro, simplified, rule_format] lemmas R23_msgsE [elim] = R23_msgs_def [THEN rel_def_to_elim, simplified, rule_format] lemmas R23_keysI = R23_keys_def [THEN rel_def_to_intro, simplified, rule_format] lemmas R23_keysE [elim] = R23_keys_def [THEN rel_def_to_elim, simplified, rule_format] lemmas R23_presI = R23_pres_def [THEN rel_def_to_intro, simplified, rule_format] lemmas R23_presE [elim] = R23_pres_def [THEN rel_def_to_elim, simplified, rule_format] lemmas R23_intros = R23_msgsI R23_keysI R23_presI text \Simplifier lemmas for various instantiations (for keys).\ lemmas R23_keys_simp = R23_keys_def [THEN rel_def_to_dest, simplified, rule_format] lemmas R23_keys_simps = R23_keys_simp R23_keys_simp [where KK="{}", simplified] R23_keys_simp [where KK="{K'}" for K', simplified] R23_keys_simp [where KK="insert K' KK" for K' KK, simplified, OF _ conjI] subsubsection \General lemmas\ (******************************************************************************) text \General facts about @{term "abs_msg"}\ declare abs_msg.intros [intro!] declare abs_msg.cases [elim!] lemma abs_msg_empty: "abs_msg {} = {}" by (auto) lemma abs_msg_Un [simp]: "abs_msg (G \ H) = abs_msg G \ abs_msg H" by (auto) lemma abs_msg_mono [elim]: "\ m \ abs_msg G; G \ H \ \ m \ abs_msg H" by (auto) lemma abs_msg_insert_mono [intro]: "\ m \ abs_msg H \ \ m \ abs_msg (insert m' H)" by (auto) text \Facts about @{term "abs_msg"} concerning abstraction of fakeable messages. This is crucial for proving the refinement of the intruder event.\ lemma abs_msg_DY_subset_fakeable: "\ (s, t) \ R23_msgs; (s, t) \ R23_keys; t \ m3_inv1_lkeysec \ \ abs_msg (synth (analz (IK t))) \ fake ik0 (dom (runs s)) (chan s)" apply (auto) \ \4 subgoals, deal with replays first\ apply (blast) defer 1 apply (blast) \ \remaining 2 subgoals are real fakes\ apply (rule fake_StatCh, auto simp add: R23_keys_simps)+ done subsubsection \Refinement proof\ (******************************************************************************) text \Pair decomposition. These were set to \texttt{elim!}, which is too agressive here.\ declare MPair_analz [rule del, elim] declare MPair_parts [rule del, elim] text \Protocol events.\ lemma PO_m3_step1_refines_m2_step1: "{R23} (m2_step1 Ra A B), (m3_step1 Ra A B) {> R23}" by (auto simp add: PO_rhoare_defs R23_def m3_defs intro!: R23_intros) (auto) lemma PO_m3_step2_refines_m2_step2: "{R23} (m2_step2 Rb A B), (m3_step2 Rb A B) {> R23}" by (auto simp add: PO_rhoare_defs R23_def m3_defs intro!: R23_intros) (auto) lemma PO_m3_step3_refines_m2_step3: "{R23 \ (m2_inv3a_sesK_compr) \ (m3_inv3_sesK_compr \ m3_inv1_lkeysec)} (m2_step3 Rs A B Kab Ts), (m3_step3 Rs A B Kab Ts) {> R23}" proof - { fix s t assume H: "(s, t) \ R23_msgs" "(s, t) \ R23_keys" "(s, t) \ R23_pres" "s \ m2_inv3a_sesK_compr" "t \ m3_inv3_sesK_compr" "t \ m3_inv1_lkeysec" "Kab = sesK (Rs$sk)" "Rs \ dom (runs t)" "\ Agent A, Agent B \ \ parts (IK t)" let ?s'= - "s\ runs := runs s(Rs \ (Serv, [A, B], [aNum (clk t)])), + "s\ runs := (runs s)(Rs \ (Serv, [A, B], [aNum (clk t)])), chan := insert (Secure Sv A (Msg [aAgt B, aKey Kab, aNum (clk t)])) (insert (Secure Sv B (Msg [aKey Kab, aAgt A, aNum (clk t)])) (chan s)) \" let ?t'= - "t\ runs := runs t(Rs \ (Serv, [A, B], [aNum (clk t)])), + "t\ runs := (runs t)(Rs \ (Serv, [A, B], [aNum (clk t)])), IK := insert (Crypt (shrK A) \ Agent B, Key Kab, Number (clk t) \) (insert (Crypt (shrK B) \ Key Kab, Agent A, Number (clk t) \) (IK t)) \" have "(?s', ?t') \ R23_msgs" using H by (-) (rule R23_intros, auto) moreover have "(?s', ?t') \ R23_keys" using H by (-) (rule R23_intros, auto simp add: m2_inv3a_sesK_compr_simps m3_inv3_sesK_compr_simps, auto simp add: R23_keys_simps) moreover have "(?s', ?t') \ R23_pres" using H by (-) (rule R23_intros, auto) moreover note calculation } thus ?thesis by (auto simp add: PO_rhoare_defs R23_def m3_defs) qed lemma PO_m3_step4_refines_m2_step4: "{R23 \ UNIV \ (m3_inv1_lkeysec)} (m2_step4 Ra A B Kab Ts), (m3_step4 Ra A B Kab Ts) {> R23}" by (auto simp add: PO_rhoare_defs R23_def m3_defs intro!: R23_intros) (auto) lemma PO_m3_step5_refines_m2_step5: "{R23} (m2_step5 Rb A B Kab Ts), (m3_step5 Rb A B Kab Ts) {> R23}" by (auto simp add: PO_rhoare_defs R23_def m3_defs intro!: R23_intros) (auto) lemma PO_m3_tick_refines_m2_tick: "{R23} (m2_tick T), (m3_tick T) {>R23}" by (auto simp add: PO_rhoare_defs R23_def m3_defs intro!: R23_intros) (auto) text \Intruder events.\ lemma PO_m3_leak_refines_m2_leak: "{R23} (m2_leak Rs), (m3_leak Rs) {>R23}" by (auto simp add: PO_rhoare_defs R23_def m3_defs intro!: R23_intros) (auto simp add: R23_keys_simps) lemma PO_m3_DY_fake_refines_m2_fake: "{R23 \ UNIV \ (m3_inv1_lkeysec)} m2_fake, m3_DY_fake {> R23}" apply (auto simp add: PO_rhoare_defs R23_def m3_defs intro!: R23_intros del: abs_msg.cases) apply (auto intro: abs_msg_DY_subset_fakeable [THEN subsetD] del: abs_msg.cases) apply (auto simp add: R23_keys_simps) done text \All together now...\ lemmas PO_m3_trans_refines_m2_trans = PO_m3_step1_refines_m2_step1 PO_m3_step2_refines_m2_step2 PO_m3_step3_refines_m2_step3 PO_m3_step4_refines_m2_step4 PO_m3_step5_refines_m2_step5 PO_m3_tick_refines_m2_tick PO_m3_leak_refines_m2_leak PO_m3_DY_fake_refines_m2_fake lemma PO_m3_refines_init_m2 [iff]: "init m3 \ R23``(init m2)" by (auto simp add: R23_def m3_defs intro!: R23_intros) lemma PO_m3_refines_trans_m2 [iff]: "{R23 \ (m2_inv3a_sesK_compr) \ (m3_inv3_sesK_compr \ m3_inv1_lkeysec)} (trans m2), (trans m3) {> R23}" by (auto simp add: m3_def m3_trans_def m2_def m2_trans_def) (blast intro!: PO_m3_trans_refines_m2_trans)+ lemma PO_m3_observation_consistent [iff]: "obs_consistent R23 med32 m2 m3" by (auto simp add: obs_consistent_def R23_def med32_def m3_defs) text \Refinement result.\ lemma m3_refines_m2 [iff]: "refines (R23 \ (m2_inv3a_sesK_compr) \ (m3_inv1_lkeysec)) med32 m2 m3" proof - have "R23 \ m2_inv3a_sesK_compr \ UNIV \ UNIV \ m3_inv3_sesK_compr" by (auto simp add: R23_def R23_keys_simps intro!: m3_inv3_sesK_comprI) thus ?thesis by (-) (rule Refinement_using_invariants, auto) qed lemma m3_implements_m2 [iff]: "implements med32 m2 m3" by (rule refinement_soundness) (auto) end diff --git a/thys/Security_Protocol_Refinement/Key_establish/m3_kerberos4.thy b/thys/Security_Protocol_Refinement/Key_establish/m3_kerberos4.thy --- a/thys/Security_Protocol_Refinement/Key_establish/m3_kerberos4.thy +++ b/thys/Security_Protocol_Refinement/Key_establish/m3_kerberos4.thy @@ -1,1005 +1,1005 @@ (******************************************************************************* Project: Development of Security Protocols by Refinement Module: Key_establish/m3_kerberos4.thy (Isabelle/HOL 2016-1) ID: $Id: m3_kerberos4.thy 132890 2016-12-24 10:25:57Z csprenge $ Authors Ivano Somaini, ETH Zurich Christoph Sprenger, ETH Zurich Key distribution protocols Third refinement: core Kerberos IV Copyright (c) 2009-2016 Ivano Somaini, Christoph Sprenger Licence: LGPL *******************************************************************************) section \Core Kerberos 4 (L3)\ theory m3_kerberos4 imports m2_kerberos "../Refinement/Message" begin text \ We model the core Kerberos 4 protocol: \[ \begin{array}{lll} \mathrm{M1.} & A \rightarrow S: & A, B \\ \mathrm{M2.} & S \rightarrow A: & \{Kab, B, Ts, Na, \{Kab, A, Ts\}_{Kbs} \}_{Kas} \\ \mathrm{M3.} & A \rightarrow B: & \{A, Ta\}_{Kab}, \{Kab, A, Ts\}_{Kbs} \\ \mathrm{M4.} & B \rightarrow A: & \{Ta\}_{Kab} \\ \end{array} \] \ text \Proof tool configuration. Avoid annoying automatic unfolding of \dom\.\ declare domIff [simp, iff del] (******************************************************************************) subsection \Setup\ (******************************************************************************) text \Now we can define the initial key knowledge.\ overloading ltkeySetup' \ ltkeySetup begin definition ltkeySetup_def: "ltkeySetup' \ {(sharK C, A) | C A. A = C \ A = Sv}" end lemma corrKey_shrK_bad [simp]: "corrKey = shrK`bad" by (auto simp add: keySetup_def ltkeySetup_def corrKey_def) (******************************************************************************) subsection \State\ (******************************************************************************) text \The secure channels are star-shaped to/from the server. Therefore, we have only one agent in the relation.\ record m3_state = "m1_state" + IK :: "msg set" \ \intruder knowledge\ text \Observable state: @{term "runs"}, @{term "clk"}, and @{term "cache"}.\ type_synonym m3_obs = "m2_obs" definition m3_obs :: "m3_state \ m3_obs" where "m3_obs s \ \ runs = runs s, leak = leak s, clk = clk s, cache = cache s \" type_synonym m3_pred = "m3_state set" type_synonym m3_trans = "(m3_state \ m3_state) set" (******************************************************************************) subsection \Events\ (******************************************************************************) text \Protocol events.\ definition \ \by @{term "A"}, refines @{term "m2_step1"}\ m3_step1 :: "[rid_t, agent, agent, nonce] \ m3_trans" where "m3_step1 Ra A B Na \ {(s, s1). \ \guards:\ Ra \ dom (runs s) \ \ \\Ra\ is fresh\ Na = Ra$na \ \ \generated nonce\ \ \actions:\ s1 = s\ runs := (runs s)(Ra \ (Init, [A, B], [])), IK := insert \Agent A, Agent B, Nonce Na\ (IK s) \ \send \M1\\ \ }" definition \ \by @{term "B"}, refines @{term "m2_step2"}\ m3_step2 :: "[rid_t, agent, agent] \ m3_trans" where "m3_step2 \ m1_step2" definition \ \by @{text "Server"}, refines @{term m2_step3}\ m3_step3 :: "[rid_t, agent, agent, key, nonce, time] \ m3_trans" where "m3_step3 Rs A B Kab Na Ts \ {(s, s1). \ \guards:\ Rs \ dom (runs s) \ \ \fresh server run\ Kab = sesK (Rs$sk) \ \ \fresh session key\ \Agent A, Agent B, Nonce Na\ \ IK s \ \ \recv \M1\\ Ts = clk s \ \ \fresh timestamp\ \ \actions:\ \ \record session key and send \M2\\ s1 = s\ runs := (runs s)(Rs \ (Serv, [A, B], [aNon Na, aNum Ts])), IK := insert (Crypt (shrK A) \ \send \M2\\ \Key Kab, Agent B, Number Ts, Nonce Na, Crypt (shrK B) \Key Kab, Agent A, Number Ts\\) (IK s) \ }" definition \ \by @{term "A"}, refines @{term m2_step4}\ m3_step4 :: "[rid_t, agent, agent, nonce, key, time, time, msg] \ m3_trans" where "m3_step4 Ra A B Na Kab Ts Ta X \ {(s, s1). \ \guards:\ runs s Ra = Some (Init, [A, B], []) \ \ \key not yet recv'd\ Na = Ra$na \ \ \generated nonce\ Crypt (shrK A) \ \recv \M2\\ \Key Kab, Agent B, Number Ts, Nonce Na, X\ \ IK s \ \ \read current time\ Ta = clk s \ \ \check freshness of session key\ clk s < Ts + Ls \ \ \actions:\ \ \record session key and send \M3\\ s1 = s\ runs := (runs s)(Ra \ (Init, [A, B], [aKey Kab, aNum Ts, aNum Ta])), IK := insert \Crypt Kab \Agent A, Number Ta\, X\ (IK s) \ \\M3\\ \ }" definition \ \by @{term "B"}, refines @{term m2_step5}\ m3_step5 :: "[rid_t, agent, agent, key, time, time] \ m3_trans" where "m3_step5 Rb A B Kab Ts Ta \ {(s, s1). \ \guards:\ runs s Rb = Some (Resp, [A, B], []) \ \ \key not yet recv'd\ \Crypt Kab \Agent A, Number Ta\, \ \recv \M3\\ Crypt (shrK B) \Key Kab, Agent A, Number Ts\\ \ IK s \ \ \ensure freshness of session key\ clk s < Ts + Ls \ \ \check authenticator's validity and replay; 'replays' with fresh authenticator ok!\ clk s < Ta + La \ (B, Kab, Ta) \ cache s \ \ \actions:\ \ \record session key\ s1 = s\ runs := (runs s)(Rb \ (Resp, [A, B], [aKey Kab, aNum Ts, aNum Ta])), cache := insert (B, Kab, Ta) (cache s), IK := insert (Crypt Kab (Number Ta)) (IK s) \ \send \M4\\ \ }" definition \ \by @{term "A"}, refines @{term m2_step6}\ m3_step6 :: "[rid_t, agent, agent, nonce, key, time, time] \ m3_trans" where "m3_step6 Ra A B Na Kab Ts Ta \ {(s, s'). \ \guards:\ runs s Ra = Some (Init, [A, B], [aKey Kab, aNum Ts, aNum Ta]) \ \ \knows key\ Na = Ra$na \ \ \generated nonce\ clk s < Ts + Ls \ \ \check session key's recentness\ Crypt Kab (Number Ta) \ IK s \ \ \recv \M4\\ \ \actions:\ s' = s\ runs := (runs s)(Ra \ (Init, [A, B], [aKey Kab, aNum Ts, aNum Ta, END])) \ }" text \Clock tick event\ definition \ \refines @{term "m2_tick"}\ m3_tick :: "time \ m3_trans" where "m3_tick \ m1_tick" text \Purge event: purge cache of expired timestamps\ definition \ \refines @{term "m2_purge"}\ m3_purge :: "agent \ m3_trans" where "m3_purge \ m1_purge" text \Session key compromise.\ definition \ \refines @{term m2_leak}\ m3_leak :: "[rid_t, agent, agent, nonce, time] \ m3_trans" where "m3_leak Rs A B Na Ts \ {(s, s1). \ \guards:\ runs s Rs = Some (Serv, [A, B], [aNon Na, aNum Ts]) \ (clk s \ Ts + Ls) \ \ \only compromise 'old' session keys!\ \ \actions:\ \ \record session key as leaked and add it to intruder knowledge\ s1 = s\ leak := insert (sesK (Rs$sk), A, B, Na, Ts) (leak s), IK := insert (Key (sesK (Rs$sk))) (IK s) \ }" text \Intruder fake event. The following "Dolev-Yao" event generates all intruder-derivable messages.\ definition \ \refines @{term "m2_fake"}\ m3_DY_fake :: "m3_trans" where "m3_DY_fake \ {(s, s1). \ \actions:\ s1 = s\ IK := synth (analz (IK s)) \ \ \take DY closure\ }" (******************************************************************************) subsection \Transition system\ (******************************************************************************) definition m3_init :: "m3_pred" where "m3_init \ { \ runs = Map.empty, leak = shrK`bad \ {undefined}, clk = 0, cache = {}, IK = Key`shrK`bad \ }" definition m3_trans :: "m3_trans" where "m3_trans \ (\A B Ra Rb Rs Na Kab Ts Ta T X. m3_step1 Ra A B Na \ m3_step2 Rb A B \ m3_step3 Rs A B Kab Na Ts \ m3_step4 Ra A B Na Kab Ts Ta X \ m3_step5 Rb A B Kab Ts Ta \ m3_step6 Ra A B Na Kab Ts Ta \ m3_tick T \ m3_purge A \ m3_leak Rs A B Na Ts \ m3_DY_fake \ Id )" definition m3 :: "(m3_state, m3_obs) spec" where "m3 \ \ init = m3_init, trans = m3_trans, obs = m3_obs \" lemmas m3_loc_defs = m3_def m3_init_def m3_trans_def m3_obs_def m3_step1_def m3_step2_def m3_step3_def m3_step4_def m3_step5_def m3_step6_def m3_tick_def m3_purge_def m3_leak_def m3_DY_fake_def lemmas m3_defs = m3_loc_defs m2_defs (******************************************************************************) subsection \Invariants\ (******************************************************************************) text \Specialized injection that we can apply more aggressively.\ lemmas analz_Inj_IK = analz.Inj [where H="IK s" for s] lemmas parts_Inj_IK = parts.Inj [where H="IK s" for s] declare parts_Inj_IK [dest!] declare analz_into_parts [dest] subsubsection \inv4: Secrecy of pre-distributed shared keys\ (*inv**************************************************************************) definition m3_inv4_lkeysec :: "m3_pred" where "m3_inv4_lkeysec \ {s. \C. (Key (shrK C) \ parts (IK s) \ C \ bad) \ (C \ bad \ Key (shrK C) \ IK s) }" lemmas m3_inv4_lkeysecI = m3_inv4_lkeysec_def [THEN setc_def_to_intro, rule_format] lemmas m3_inv4_lkeysecE [elim] = m3_inv4_lkeysec_def [THEN setc_def_to_elim, rule_format] lemmas m3_inv4_lkeysecD = m3_inv4_lkeysec_def [THEN setc_def_to_dest, rule_format] text \Invariance proof.\ lemma PO_m3_inv4_lkeysec_init [iff]: "init m3 \ m3_inv4_lkeysec" by (auto simp add: m3_defs intro!: m3_inv4_lkeysecI) lemma PO_m3_inv4_lkeysec_trans [iff]: "{m3_inv4_lkeysec} trans m3 {> m3_inv4_lkeysec}" by (auto simp add: PO_hoare_defs m3_defs intro!: m3_inv4_lkeysecI) (auto dest!: Body) lemma PO_m3_inv4_lkeysec [iff]: "reach m3 \ m3_inv4_lkeysec" by (rule inv_rule_incr) (fast+) text \Useful simplifier lemmas\ lemma m3_inv4_lkeysec_for_parts [simp]: "\ s \ m3_inv4_lkeysec \ \ Key (shrK C) \ parts (IK s) \ C \ bad" by auto lemma m3_inv4_lkeysec_for_analz [simp]: "\ s \ m3_inv4_lkeysec \ \ Key (shrK C) \ analz (IK s) \ C \ bad" by auto subsubsection \inv6: Ticket shape for honestly encrypted M2\ (*inv**************************************************************************) definition m3_inv6_ticket :: "m3_pred" where "m3_inv6_ticket \ {s. \A B T K N X. A \ bad \ Crypt (shrK A) \Key K, Agent B, Number T, Nonce N, X\ \ parts (IK s) \ X = Crypt (shrK B) \Key K, Agent A, Number T\ \ K \ range sesK }" lemmas m3_inv6_ticketI = m3_inv6_ticket_def [THEN setc_def_to_intro, rule_format] lemmas m3_inv6_ticketE [elim] = m3_inv6_ticket_def [THEN setc_def_to_elim, rule_format] lemmas m3_inv6_ticketD = m3_inv6_ticket_def [THEN setc_def_to_dest, rule_format, rotated -1] text \Invariance proof.\ lemma PO_m3_inv6_ticket_init [iff]: "init m3 \ m3_inv6_ticket" by (auto simp add: m3_defs intro!: m3_inv6_ticketI) lemma PO_m3_inv6_ticket_trans [iff]: "{m3_inv6_ticket \ m3_inv4_lkeysec} trans m3 {> m3_inv6_ticket}" apply (auto simp add: PO_hoare_defs m3_defs intro!: m3_inv6_ticketI) apply (auto dest: m3_inv6_ticketD) \ \2 subgoals\ apply (drule parts_cut, drule Body, auto dest: m3_inv6_ticketD)+ done lemma PO_m3_inv6_ticket [iff]: "reach m3 \ m3_inv6_ticket" by (rule inv_rule_incr) (auto del: subsetI) subsubsection \inv7: Session keys not used to encrypt other session keys\ (*inv**************************************************************************) text \Session keys are not used to encrypt other keys. Proof requires generalization to sets of session keys. NOTE: For Kerberos 4, this invariant cannot be inherited from the corresponding L2 invariant. The simulation relation is only an implication not an equivalence. \ definition m3_inv7a_sesK_compr :: "m3_pred" where "m3_inv7a_sesK_compr \ {s. \K KK. KK \ range sesK \ (Key K \ analz (Key`KK \ (IK s))) = (K \ KK \ Key K \ analz (IK s)) }" lemmas m3_inv7a_sesK_comprI = m3_inv7a_sesK_compr_def [THEN setc_def_to_intro, rule_format] lemmas m3_inv7a_sesK_comprE = m3_inv7a_sesK_compr_def [THEN setc_def_to_elim, rule_format] lemmas m3_inv7a_sesK_comprD = m3_inv7a_sesK_compr_def [THEN setc_def_to_dest, rule_format] text \Additional lemma\ lemmas insert_commute_Key = insert_commute [where x="Key K" for K] lemmas m3_inv7a_sesK_compr_simps = m3_inv7a_sesK_comprD m3_inv7a_sesK_comprD [where KK="insert Kab KK" for Kab KK, simplified] m3_inv7a_sesK_comprD [where KK="{Kab}" for Kab, simplified] insert_commute_Key text \Invariance proof.\ lemma PO_m3_inv7a_sesK_compr_step4: "{m3_inv7a_sesK_compr \ m3_inv6_ticket \ m3_inv4_lkeysec} m3_step4 Ra A B Na Kab Ts Ta X {> m3_inv7a_sesK_compr}" proof - { fix K KK s assume H: "s \ m3_inv4_lkeysec" "s \ m3_inv7a_sesK_compr" "s \ m3_inv6_ticket" "runs s Ra = Some (Init, [A, B], [])" "Na = Ra$na" "KK \ range sesK" "Crypt (shrK A) \Key Kab, Agent B, Number Ts, Nonce Na, X\ \ analz (IK s)" have "(Key K \ analz (insert X (Key ` KK \ IK s))) = (K \ KK \ Key K \ analz (insert X (IK s)))" proof (cases "A \ bad") case True with H have "X \ analz (IK s)" by (auto dest!: Decrypt) moreover with H have "X \ analz (Key ` KK \ IK s)" by (auto intro: analz_monotonic) ultimately show ?thesis using H by (auto simp add: m3_inv7a_sesK_compr_simps analz_insert_eq) next case False thus ?thesis using H by (fastforce simp add: m3_inv7a_sesK_compr_simps dest!: m3_inv6_ticketD [OF analz_into_parts]) qed } thus ?thesis by (auto simp add: PO_hoare_defs m3_defs intro!: m3_inv7a_sesK_comprI dest!: analz_Inj_IK) qed text \All together now.\ lemmas PO_m3_inv7a_sesK_compr_trans_lemmas = PO_m3_inv7a_sesK_compr_step4 lemma PO_m3_inv7a_sesK_compr_init [iff]: "init m3 \ m3_inv7a_sesK_compr" by (auto simp add: m3_defs intro!: m3_inv7a_sesK_comprI) lemma PO_m3_inv7a_sesK_compr_trans [iff]: "{m3_inv7a_sesK_compr \ m3_inv6_ticket \ m3_inv4_lkeysec} trans m3 {> m3_inv7a_sesK_compr}" by (auto simp add: m3_def m3_trans_def intro!: PO_m3_inv7a_sesK_compr_trans_lemmas) (auto simp add: PO_hoare_defs m3_defs m3_inv7a_sesK_compr_simps intro!: m3_inv7a_sesK_comprI) lemma PO_m3_inv7a_sesK_compr [iff]: "reach m3 \ m3_inv7a_sesK_compr" by (rule_tac J="m3_inv6_ticket \ m3_inv4_lkeysec" in inv_rule_incr) (auto) subsubsection \inv7b: Session keys not used to encrypt nonces\ (*inv**************************************************************************) text \Session keys are not used to encrypt nonces. The proof requires a generalization to sets of session keys.\ definition m3_inv7b_sesK_compr_non :: "m3_pred" where "m3_inv7b_sesK_compr_non \ {s. \N KK. KK \ range sesK \ (Nonce N \ analz (Key`KK \ (IK s))) = (Nonce N \ analz (IK s)) }" lemmas m3_inv7b_sesK_compr_nonI = m3_inv7b_sesK_compr_non_def [THEN setc_def_to_intro, rule_format] lemmas m3_inv7b_sesK_compr_nonE = m3_inv7b_sesK_compr_non_def [THEN setc_def_to_elim, rule_format] lemmas m3_inv7b_sesK_compr_nonD = m3_inv7b_sesK_compr_non_def [THEN setc_def_to_dest, rule_format] lemmas m3_inv7b_sesK_compr_non_simps = m3_inv7b_sesK_compr_nonD m3_inv7b_sesK_compr_nonD [where KK="insert Kab KK" for Kab KK, simplified] m3_inv7b_sesK_compr_nonD [where KK="{Kab}" for Kab, simplified] insert_commute_Key text \Invariance proof.\ lemma PO_m3_inv7b_sesK_compr_non_step3: "{m3_inv7b_sesK_compr_non} m3_step3 Rs A B Kab Na Ts {> m3_inv7b_sesK_compr_non}" by (auto simp add: PO_hoare_defs m3_defs m3_inv7b_sesK_compr_non_simps intro!: m3_inv7b_sesK_compr_nonI dest!: analz_Inj_IK) lemma PO_m3_inv7b_sesK_compr_non_step4: "{m3_inv7b_sesK_compr_non \ m3_inv6_ticket \ m3_inv4_lkeysec} m3_step4 Ra A B Na Kab Ts Ta X {> m3_inv7b_sesK_compr_non}" proof - { fix N KK s assume H: "s \ m3_inv4_lkeysec""s \ m3_inv7b_sesK_compr_non" "s \ m3_inv6_ticket" "runs s Ra = Some (Init, [A, B], [])" "Na = Ra$na" "KK \ range sesK" "Crypt (shrK A) \Key Kab, Agent B, Number Ts, Nonce Na, X\ \ analz (IK s)" have "(Nonce N \ analz (insert X (Key ` KK \ IK s))) = (Nonce N \ analz (insert X (IK s)))" proof (cases) assume "A \ bad" show ?thesis proof - from \A \ bad\ have "X \ analz (IK s)" using H by (auto dest!: Decrypt) moreover hence "X \ analz (Key ` KK \ IK s)" by (auto intro: analz_monotonic) ultimately show ?thesis using H by (auto simp add: m3_inv7b_sesK_compr_non_simps analz_insert_eq) qed next assume "A \ bad" thus ?thesis using H by (fastforce simp add: m3_inv7b_sesK_compr_non_simps dest!: m3_inv6_ticketD [OF analz_into_parts]) qed } thus ?thesis by (auto simp add: PO_hoare_defs m3_defs intro!: m3_inv7b_sesK_compr_nonI dest!: analz_Inj_IK) qed text \All together now.\ lemmas PO_m3_inv7b_sesK_compr_non_trans_lemmas = PO_m3_inv7b_sesK_compr_non_step3 PO_m3_inv7b_sesK_compr_non_step4 lemma PO_m3_inv7b_sesK_compr_non_init [iff]: "init m3 \ m3_inv7b_sesK_compr_non" by (auto simp add: m3_defs intro!: m3_inv7b_sesK_compr_nonI) lemma PO_m3_inv7b_sesK_compr_non_trans [iff]: "{m3_inv7b_sesK_compr_non \ m3_inv6_ticket \ m3_inv4_lkeysec} trans m3 {> m3_inv7b_sesK_compr_non}" by (auto simp add: m3_def m3_trans_def intro!: PO_m3_inv7b_sesK_compr_non_trans_lemmas) (auto simp add: PO_hoare_defs m3_defs m3_inv7b_sesK_compr_non_simps intro!: m3_inv7b_sesK_compr_nonI) lemma PO_m3_inv7b_sesK_compr_non [iff]: "reach m3 \ m3_inv7b_sesK_compr_non" by (rule_tac J="m3_inv6_ticket \ m3_inv4_lkeysec" in inv_rule_incr) (auto) (******************************************************************************) subsection \Refinement\ (******************************************************************************) subsubsection \Message abstraction and simulation relation\ (******************************************************************************) text \Abstraction function on sets of messages.\ inductive_set abs_msg :: "msg set \ chmsg set" for H :: "msg set" where am_M1: "\Agent A, Agent B, Nonce N\ \ H \ Insec A B (Msg [aNon N]) \ abs_msg H" | am_M2a: "Crypt (shrK C) \Key K, Agent B, Number T, Nonce N, X\ \ H \ Secure Sv C (Msg [aKey K, aAgt B, aNum T, aNon N]) \ abs_msg H" | am_M2b: "Crypt (shrK C) \Key K, Agent A, Number T\ \ H \ Secure Sv C (Msg [aKey K, aAgt A, aNum T]) \ abs_msg H" | am_M3: "Crypt K \Agent A, Number T\ \ H \ dAuth K (Msg [aAgt A, aNum T]) \ abs_msg H" | am_M4: "Crypt K (Number T) \ H \ dAuth K (Msg [aNum T]) \ abs_msg H" text \R23: The simulation relation. This is a data refinement of the insecure and secure channels of refinement 2.\ definition R23_msgs :: "(m2_state \ m3_state) set" where "R23_msgs \ {(s, t). abs_msg (parts (IK t)) \ chan s }" definition R23_keys :: "(m2_state \ m3_state) set" where "R23_keys \ {(s, t). \KK K. KK \ range sesK \ Key K \ analz (Key`KK \ (IK t)) \ aKey K \ extr (aKey`KK \ ik0) (chan s) }" definition R23_non :: "(m2_state \ m3_state) set" where "R23_non \ {(s, t). \KK N. KK \ range sesK \ Nonce N \ analz (Key`KK \ (IK t)) \ aNon N \ extr (aKey`KK \ ik0) (chan s) }" definition R23_pres :: "(m2_state \ m3_state) set" where "R23_pres \ {(s, t). runs s = runs t \ leak s = leak t \ clk s = clk t \ cache s = cache t}" definition R23 :: "(m2_state \ m3_state) set" where "R23 \ R23_msgs \ R23_keys \ R23_non \ R23_pres" lemmas R23_defs = R23_def R23_msgs_def R23_keys_def R23_non_def R23_pres_def text \The mediator function is the identity here.\ definition med32 :: "m3_obs \ m2_obs" where "med32 \ id" lemmas R23_msgsI = R23_msgs_def [THEN rel_def_to_intro, simplified, rule_format] lemmas R23_msgsE [elim] = R23_msgs_def [THEN rel_def_to_elim, simplified, rule_format] lemmas R23_msgsE' [elim] = R23_msgs_def [THEN rel_def_to_dest, simplified, rule_format, THEN subsetD] lemmas R23_keysI = R23_keys_def [THEN rel_def_to_intro, simplified, rule_format] lemmas R23_keysE [elim] = R23_keys_def [THEN rel_def_to_elim, simplified, rule_format] lemmas R23_keysD = R23_keys_def [THEN rel_def_to_dest, simplified, rule_format, rotated 2] lemmas R23_nonI = R23_non_def [THEN rel_def_to_intro, simplified, rule_format] lemmas R23_nonE [elim] = R23_non_def [THEN rel_def_to_elim, simplified, rule_format] lemmas R23_nonD = R23_non_def [THEN rel_def_to_dest, simplified, rule_format, rotated 2] lemmas R23_presI = R23_pres_def [THEN rel_def_to_intro, simplified, rule_format] lemmas R23_presE [elim] = R23_pres_def [THEN rel_def_to_elim, simplified, rule_format] lemmas R23_intros = R23_msgsI R23_keysI R23_nonI R23_presI text \Lemmas for various instantiations (keys and nonces).\ lemmas R23_keys_dests = R23_keysD R23_keysD [where KK="{}", simplified] R23_keysD [where KK="{K}" for K, simplified] R23_keysD [where KK="insert K KK" for K KK, simplified, OF _ _ conjI] lemmas R23_non_dests = R23_nonD R23_nonD [where KK="{}", simplified] R23_nonD [where KK="{K}" for K, simplified] R23_nonD [where KK="insert K KK" for K KK, simplified, OF _ _ conjI] lemmas R23_dests = R23_keys_dests R23_non_dests subsubsection \General lemmas\ (******************************************************************************) text \General facts about @{term "abs_msg"}\ declare abs_msg.intros [intro!] declare abs_msg.cases [elim!] lemma abs_msg_empty: "abs_msg {} = {}" by (auto) lemma abs_msg_Un [simp]: "abs_msg (G \ H) = abs_msg G \ abs_msg H" by (auto) lemma abs_msg_mono [elim]: "\ m \ abs_msg G; G \ H \ \ m \ abs_msg H" by (auto) lemma abs_msg_insert_mono [intro]: "\ m \ abs_msg H \ \ m \ abs_msg (insert m' H)" by (auto) text \Facts about @{term "abs_msg"} concerning abstraction of fakeable messages. This is crucial for proving the refinement of the intruder event.\ lemma abs_msg_DY_subset_fakeable: "\ (s, t) \ R23_msgs; (s, t) \ R23_keys; (s, t) \ R23_non; t \ m3_inv4_lkeysec \ \ abs_msg (synth (analz (IK t))) \ fake ik0 (dom (runs s)) (chan s)" apply (auto) \ \9 subgoals, deal with replays first\ prefer 2 apply (blast) prefer 3 apply (blast) prefer 4 apply (blast) prefer 5 apply (blast) \ \remaining 5 subgoals are real fakes\ apply (intro fake_StatCh fake_DynCh, auto dest: R23_dests)+ done subsubsection \Refinement proof\ (******************************************************************************) text \Pair decomposition. These were set to \texttt{elim!}, which is too agressive here.\ declare MPair_analz [rule del, elim] declare MPair_parts [rule del, elim] text \Protocol events.\ lemma PO_m3_step1_refines_m2_step1: "{R23} (m2_step1 Ra A B Na), (m3_step1 Ra A B Na) {> R23}" by (auto simp add: PO_rhoare_defs R23_def m3_defs intro!: R23_intros) (auto) lemma PO_m3_step2_refines_m2_step2: "{R23} (m2_step2 Rb A B), (m3_step2 Rb A B) {> R23}" by (auto simp add: PO_rhoare_defs R23_def m3_defs intro!: R23_intros) lemma PO_m3_step3_refines_m2_step3: "{R23 \ (m2_inv3a_sesK_compr) \ (m3_inv7a_sesK_compr \ m3_inv4_lkeysec)} (m2_step3 Rs A B Kab Na Ts), (m3_step3 Rs A B Kab Na Ts) {> R23}" proof - { fix s t assume H: "(s, t) \ R23_msgs" "(s, t) \ R23_keys" "(s, t) \ R23_non" "(s, t) \ R23_pres" "s \ m2_inv3a_sesK_compr" "t \ m3_inv7a_sesK_compr" "t \ m3_inv4_lkeysec" "Kab = sesK (Rs$sk)" "Rs \ dom (runs t)" "\Agent A, Agent B, Nonce Na\ \ parts (IK t)" let ?s'= - "s\ runs := runs s(Rs \ (Serv, [A, B], [aNon Na, aNum (clk t)])), + "s\ runs := (runs s)(Rs \ (Serv, [A, B], [aNon Na, aNum (clk t)])), chan := insert (Secure Sv A (Msg [aKey Kab, aAgt B, aNum (clk t), aNon Na])) (insert (Secure Sv B (Msg [aKey Kab, aAgt A, aNum (clk t)])) (chan s)) \" let ?t'= - "t\ runs := runs t(Rs \ (Serv, [A, B], [aNon Na, aNum (clk t)])), + "t\ runs := (runs t)(Rs \ (Serv, [A, B], [aNon Na, aNum (clk t)])), IK := insert (Crypt (shrK A) \ Key Kab, Agent B, Number (clk t), Nonce Na, Crypt (shrK B) \ Key Kab, Agent A, Number (clk t) \\) (IK t) \" \ \here we go\ have "(?s', ?t') \ R23_msgs" using H by (-) (rule R23_intros, auto) moreover have "(?s', ?t') \ R23_keys" using H by (-) (rule R23_intros, auto simp add: m2_inv3a_sesK_compr_simps m3_inv7a_sesK_compr_simps dest: R23_keys_dests) moreover have "(?s', ?t') \ R23_non" using H by (-) (rule R23_intros, auto simp add: m2_inv3a_sesK_compr_simps m3_inv7a_sesK_compr_simps dest: R23_non_dests) moreover have "(?s', ?t') \ R23_pres" using H by (-) (rule R23_intros, auto) moreover note calculation } thus ?thesis by (auto simp add: PO_rhoare_defs R23_def m3_defs) qed lemma PO_m3_step4_refines_m2_step4: "{R23 \ (UNIV) \ (m3_inv7a_sesK_compr \ m3_inv7b_sesK_compr_non \ m3_inv6_ticket \ m3_inv4_lkeysec)} (m2_step4 Ra A B Na Kab Ts Ta), (m3_step4 Ra A B Na Kab Ts Ta X) {> R23}" proof - { fix s t assume H: "(s, t) \ R23_msgs" "(s, t) \ R23_keys" "(s, t) \ R23_non" "(s, t) \ R23_pres" "t \ m3_inv7a_sesK_compr" "t \ m3_inv7b_sesK_compr_non" "t \ m3_inv6_ticket" "t \ m3_inv4_lkeysec" "runs t Ra = Some (Init, [A, B], [])" "Na = Ra$na" "Crypt (shrK A) \Key Kab, Agent B, Number Ts, Nonce Na, X\ \ analz (IK t)" - let ?s' = "s\ runs := runs s(Ra \ (Init, [A, B], [aKey Kab, aNum Ts, aNum (clk t)])), + let ?s' = "s\ runs := (runs s)(Ra \ (Init, [A, B], [aKey Kab, aNum Ts, aNum (clk t)])), chan := insert (dAuth Kab (Msg [aAgt A, aNum (clk t)])) (chan s) \" - and ?t' = "t\ runs := runs t(Ra \ (Init, [A, B], [aKey Kab, aNum Ts, aNum (clk t)])), + and ?t' = "t\ runs := (runs t)(Ra \ (Init, [A, B], [aKey Kab, aNum Ts, aNum (clk t)])), IK := insert \ Crypt Kab \Agent A, Number (clk t)\, X \ (IK t) \" from H have "Secure Sv A (Msg [aKey Kab, aAgt B, aNum Ts, aNon Na]) \ chan s" by (auto) moreover have "X \ parts (IK t)" using H by (auto dest!: Body MPair_parts) hence "(?s', ?t') \ R23_msgs" using H by (auto intro!: R23_intros) (auto) moreover have "(?s', ?t') \ R23_keys" proof (cases) assume "A \ bad" with H have "X \ analz (IK t)" by (-) (drule Decrypt, auto) with H show ?thesis by (-) (rule R23_intros, auto dest!: analz_cut intro: analz_monotonic) next assume "A \ bad" show ?thesis proof - note H moreover with \A \ bad\ have "X = Crypt (shrK B) \Key Kab, Agent A, Number Ts \ \ Kab \ range sesK" by (auto dest!: m3_inv6_ticketD) moreover { assume H1: "Key (shrK B) \ analz (IK t)" have "aKey Kab \ extr ik0 (chan s)" proof - note calculation moreover hence "Secure Sv B (Msg [aKey Kab, aAgt A, aNum Ts]) \ chan s" by (-) (drule analz_into_parts, drule Body, elim MPair_parts, auto) ultimately show ?thesis using H1 by auto qed } ultimately show ?thesis by (-) (rule R23_intros, auto simp add: m3_inv7a_sesK_compr_simps) qed qed moreover have "(?s', ?t') \ R23_non" proof (cases) assume "A \ bad" hence "X \ analz (IK t)" using H by (-) (drule Decrypt, auto) thus ?thesis using H by (-) (rule R23_intros, auto dest!: analz_cut intro: analz_monotonic) next assume "A \ bad" hence "X = Crypt (shrK B) \Key Kab, Agent A, Number Ts \ \ Kab \ range sesK" using H by (auto dest!: m3_inv6_ticketD) thus ?thesis using H by (-) (rule R23_intros, auto simp add: m3_inv7a_sesK_compr_simps m3_inv7b_sesK_compr_non_simps) qed moreover have "(?s', ?t') \ R23_pres" using H by (auto intro!: R23_intros) moreover note calculation } thus ?thesis by (auto simp add: PO_rhoare_defs R23_def m3_defs dest!: analz_Inj_IK) qed lemma PO_m3_step5_refines_m2_step5: "{R23} (m2_step5 Rb A B Kab Ts Ta), (m3_step5 Rb A B Kab Ts Ta) {> R23}" by (auto simp add: PO_rhoare_defs R23_def m3_defs intro!: R23_intros) (auto) lemma PO_m3_step6_refines_m2_step6: "{R23} (m2_step6 Ra A B Na Kab Ts Ta), (m3_step6 Ra A B Na Kab Ts Ta) {> R23}" by (auto simp add: PO_rhoare_defs R23_def m3_defs intro!: R23_intros) lemma PO_m3_tick_refines_m2_tick: "{R23} (m2_tick T), (m3_tick T) {>R23}" by (auto simp add: PO_rhoare_defs R23_def m3_defs intro!: R23_intros) lemma PO_m3_purge_refines_m2_purge: "{R23} (m2_purge A), (m3_purge A) {>R23}" by (auto simp add: PO_rhoare_defs R23_def m3_defs intro!: R23_intros) text \Intruder events.\ lemma PO_m3_leak_refines_m2_leak: "{R23} (m2_leak Rs A B Na Ts), (m3_leak Rs A B Na Ts) {>R23}" by (auto simp add: PO_rhoare_defs R23_def m3_defs intro!: R23_intros) (auto dest: R23_dests) lemma PO_m3_DY_fake_refines_m2_fake: "{R23 \ UNIV \ (m3_inv4_lkeysec)} m2_fake, m3_DY_fake {> R23}" apply (auto simp add: PO_rhoare_defs R23_def m3_defs intro!: R23_intros del: abs_msg.cases) apply (auto intro: abs_msg_DY_subset_fakeable [THEN subsetD] del: abs_msg.cases) apply (auto dest: R23_dests) done text \All together now...\ lemmas PO_m3_trans_refines_m2_trans = PO_m3_step1_refines_m2_step1 PO_m3_step2_refines_m2_step2 PO_m3_step3_refines_m2_step3 PO_m3_step4_refines_m2_step4 PO_m3_step5_refines_m2_step5 PO_m3_step6_refines_m2_step6 PO_m3_tick_refines_m2_tick PO_m3_purge_refines_m2_purge PO_m3_leak_refines_m2_leak PO_m3_DY_fake_refines_m2_fake lemma PO_m3_refines_init_m2 [iff]: "init m3 \ R23``(init m2)" by (auto simp add: R23_def m3_defs intro!: R23_intros) lemma PO_m3_refines_trans_m2 [iff]: "{R23 \ (m2_inv3a_sesK_compr) \ (m3_inv7a_sesK_compr \ m3_inv7b_sesK_compr_non \ m3_inv6_ticket \ m3_inv4_lkeysec)} (trans m2), (trans m3) {> R23}" by (auto simp add: m3_def m3_trans_def m2_def m2_trans_def) (blast intro!: PO_m3_trans_refines_m2_trans)+ lemma PO_m3_observation_consistent [iff]: "obs_consistent R23 med32 m2 m3" by (auto simp add: obs_consistent_def R23_def med32_def m3_defs) text \Refinement result.\ lemma m3_refines_m2 [iff]: "refines (R23 \ (m2_inv3a_sesK_compr) \ (m3_inv7a_sesK_compr \ m3_inv7b_sesK_compr_non \ m3_inv6_ticket \ m3_inv4_lkeysec)) med32 m2 m3" by (rule Refinement_using_invariants) (auto) lemma m3_implements_m2 [iff]: "implements med32 m2 m3" by (rule refinement_soundness) (auto) subsection \Inherited invariants\ (******************************************************************************) subsubsection \inv3 (derived): Key secrecy for initiator\ (*invh*************************************************************************) definition m3_inv3_ikk_init :: "m3_state set" where "m3_inv3_ikk_init \ {s. \A B Ra K Ts nl. runs s Ra = Some (Init, [A, B], aKey K # aNum Ts # nl) \ A \ good \ B \ good \ Key K \ analz (IK s) \ (K, A, B, Ra$na, Ts) \ leak s }" lemmas m3_inv3_ikk_initI = m3_inv3_ikk_init_def [THEN setc_def_to_intro, rule_format] lemmas m3_inv3_ikk_initE [elim] = m3_inv3_ikk_init_def [THEN setc_def_to_elim, rule_format] lemmas m3_inv3_ikk_initD = m3_inv3_ikk_init_def [THEN setc_def_to_dest, rule_format, rotated 1] lemma PO_m3_inv3_ikk_init: "reach m3 \ m3_inv3_ikk_init" proof (rule INV_from_Refinement_using_invariants [OF m3_refines_m2]) show "Range (R23 \ m2_inv3a_sesK_compr \ (m3_inv7a_sesK_compr \ m3_inv7b_sesK_compr_non \ m3_inv6_ticket \ m3_inv4_lkeysec) \ m2_inv6_ikk_init \ UNIV) \ m3_inv3_ikk_init" by (auto simp add: R23_def R23_pres_def intro!: m3_inv3_ikk_initI) (elim m2_inv6_ikk_initE, auto dest: R23_keys_dests) qed auto subsubsection \inv4 (derived): Key secrecy for responder\ (*invh*************************************************************************) definition m3_inv4_ikk_resp :: "m3_state set" where "m3_inv4_ikk_resp \ {s. \A B Rb K Ts nl. runs s Rb = Some (Resp, [A, B], aKey K # aNum Ts # nl) \ A \ good \ B \ good \ Key K \ analz (IK s) \ (\Na. (K, A, B, Na, Ts) \ leak s) }" lemmas m3_inv4_ikk_respI = m3_inv4_ikk_resp_def [THEN setc_def_to_intro, rule_format] lemmas m3_inv4_ikk_respE [elim] = m3_inv4_ikk_resp_def [THEN setc_def_to_elim, rule_format] lemmas m3_inv4_ikk_respD = m3_inv4_ikk_resp_def [THEN setc_def_to_dest, rule_format, rotated 1] lemma PO_m3_inv4_ikk_resp: "reach m3 \ m3_inv4_ikk_resp" proof (rule INV_from_Refinement_using_invariants [OF m3_refines_m2]) show "Range (R23 \ m2_inv3a_sesK_compr \ (m3_inv7a_sesK_compr \ m3_inv7b_sesK_compr_non \ m3_inv6_ticket \ m3_inv4_lkeysec) \ m2_inv7_ikk_resp \ UNIV) \ m3_inv4_ikk_resp" by (auto simp add: R23_def R23_pres_def intro!: m3_inv4_ikk_respI) (elim m2_inv7_ikk_respE, auto dest: R23_keys_dests) qed auto end diff --git a/thys/Security_Protocol_Refinement/Key_establish/m3_kerberos5.thy b/thys/Security_Protocol_Refinement/Key_establish/m3_kerberos5.thy --- a/thys/Security_Protocol_Refinement/Key_establish/m3_kerberos5.thy +++ b/thys/Security_Protocol_Refinement/Key_establish/m3_kerberos5.thy @@ -1,773 +1,773 @@ (******************************************************************************* Project: Development of Security Protocols by Refinement Module: Key_establish/m3_kerberos5.thy (Isabelle/HOL 2016-1) ID: $Id: m3_kerberos5.thy 132890 2016-12-24 10:25:57Z csprenge $ Authors: Ivano Somaini, ETH Zurich Christoph Sprenger, ETH Zurich Key distribution protocols Third refinement: core Kerberos V Copyright (c) 2009-2016 Ivano Somaini, Christoph Sprenger Licence: LGPL *******************************************************************************) section \Core Kerberos 5 (L3)\ theory m3_kerberos5 imports m2_kerberos "../Refinement/Message" begin text \ We model the core Kerberos 5 protocol: \[ \begin{array}{lll} \mathrm{M1.} & A \rightarrow S: & A, B, Na \\ \mathrm{M2.} & S \rightarrow A: & \{Kab, B, Ts, Na\}_{Kas}, \{Kab, A, Ts\}_{Kbs} \\ \mathrm{M3.} & A \rightarrow B: & \{A, Ta\}_{Kab}, \{Kab, A, Ts\}_{Kbs} \\ \mathrm{M4.} & B \rightarrow A: & \{Ta\}_{Kab} \\ \end{array} \] \ text \Proof tool configuration. Avoid annoying automatic unfolding of \dom\.\ declare domIff [simp, iff del] (******************************************************************************) subsection \Setup\ (******************************************************************************) text \Now we can define the initial key knowledge.\ overloading ltkeySetup' \ ltkeySetup begin definition ltkeySetup_def: "ltkeySetup' \ {(sharK C, A) | C A. A = C \ A = Sv}" end lemma corrKey_shrK_bad [simp]: "corrKey = shrK`bad" by (auto simp add: keySetup_def ltkeySetup_def corrKey_def) (******************************************************************************) subsection \State\ (******************************************************************************) text \The secure channels are star-shaped to/from the server. Therefore, we have only one agent in the relation.\ record m3_state = "m1_state" + IK :: "msg set" \ \intruder knowledge\ text \Observable state: @{term "runs"}, @{term "leak"}, @{term "clk"}, and @{term "cache"}.\ type_synonym m3_obs = "m2_obs" definition m3_obs :: "m3_state \ m3_obs" where "m3_obs s \ \ runs = runs s, leak = leak s, clk = clk s, cache = cache s \" type_synonym m3_pred = "m3_state set" type_synonym m3_trans = "(m3_state \ m3_state) set" (******************************************************************************) subsection \Events\ (******************************************************************************) text \Protocol events.\ definition \ \by @{term "A"}, refines @{term "m2_step1"}\ m3_step1 :: "[rid_t, agent, agent, nonce] \ m3_trans" where "m3_step1 Ra A B Na \ {(s, s1). \ \guards:\ Ra \ dom (runs s) \ \ \\Ra\ is fresh\ Na = Ra$na \ \ \generated nonce\ \ \actions:\ s1 = s\ runs := (runs s)(Ra \ (Init, [A, B], [])), IK := insert \Agent A, Agent B, Nonce Na\ (IK s) \ \send M1\ \ }" definition \ \by @{term "B"}, refines @{term "m2_step2"}\ m3_step2 :: "[rid_t, agent, agent] \ m3_trans" where "m3_step2 \ m1_step2" definition \ \by @{text "Server"}, refines @{term m2_step3}\ m3_step3 :: "[rid_t, agent, agent, key, nonce, time] \ m3_trans" where "m3_step3 Rs A B Kab Na Ts \ {(s, s1). \ \guards:\ Rs \ dom (runs s) \ \ \fresh server run\ Kab = sesK (Rs$sk) \ \ \fresh session key\ \Agent A, Agent B, Nonce Na\ \ IK s \ \ \recv \M1\\ Ts = clk s \ \ \fresh timestamp\ \ \actions:\ \ \record session key and send \M2\\ s1 = s\ runs := (runs s)(Rs \ (Serv, [A, B], [aNon Na, aNum Ts])), IK := insert \Crypt (shrK A) \Key Kab, Agent B, Number Ts, Nonce Na\, Crypt (shrK B) \Key Kab, Agent A, Number Ts\\ (IK s) \ }" definition \ \by @{term "A"}, refines @{term m2_step4}\ m3_step4 :: "[rid_t, agent, agent, nonce, key, time, time, msg] \ m3_trans" where "m3_step4 Ra A B Na Kab Ts Ta X \ {(s, s1). \ \guards:\ runs s Ra = Some (Init, [A, B], []) \ \ \key not yet recv'd\ Na = Ra$na \ \ \generated nonce\ \Crypt (shrK A) \ \recv \M2\\ \Key Kab, Agent B, Number Ts, Nonce Na\, X\ \ IK s \ \ \read current time\ Ta = clk s \ \ \check freshness of session key\ clk s < Ts + Ls \ \ \actions:\ \ \record session key and send \M3\\ s1 = s\ runs := (runs s)(Ra \ (Init, [A, B], [aKey Kab, aNum Ts, aNum Ta])), IK := insert \Crypt Kab \Agent A, Number Ta\, X\ (IK s) \ \\M3\\ \ }" definition \ \by @{term "B"}, refines @{term m2_step5}\ m3_step5 :: "[rid_t, agent, agent, key, time, time] \ m3_trans" where "m3_step5 Rb A B Kab Ts Ta \ {(s, s1). \ \guards:\ runs s Rb = Some (Resp, [A, B], []) \ \ \key not yet recv'd\ \Crypt Kab \Agent A, Number Ta\, \ \recv \M3\\ Crypt (shrK B) \Key Kab, Agent A, Number Ts\\ \ IK s \ \ \ensure freshness of session key\ clk s < Ts + Ls \ \ \check authenticator's validity and replay; 'replays' with fresh authenticator ok!\ clk s < Ta + La \ (B, Kab, Ta) \ cache s \ \ \actions:\ \ \record session key\ s1 = s\ runs := (runs s)(Rb \ (Resp, [A, B], [aKey Kab, aNum Ts, aNum Ta])), cache := insert (B, Kab, Ta) (cache s), IK := insert (Crypt Kab (Number Ta)) (IK s) \ \send \M4\\ \ }" definition \ \by @{term "A"}, refines @{term m2_step6}\ m3_step6 :: "[rid_t, agent, agent, nonce, key, time, time] \ m3_trans" where "m3_step6 Ra A B Na Kab Ts Ta \ {(s, s'). \ \guards:\ runs s Ra = Some (Init, [A, B], [aKey Kab, aNum Ts, aNum Ta]) \ \ \knows key\ Na = Ra$na \ \ \generated nonce\ clk s < Ts + Ls \ \ \check session key's recentness\ Crypt Kab (Number Ta) \ IK s \ \ \recv \M4\\ \ \actions:\ s' = s\ runs := (runs s)(Ra \ (Init, [A, B], [aKey Kab, aNum Ts, aNum Ta, END])) \ }" text \Clock tick event\ definition \ \refines @{term "m2_tick"}\ m3_tick :: "time \ m3_trans" where "m3_tick \ m1_tick" text \Purge event: purge cache of expired timestamps\ definition \ \refines @{term "m2_purge"}\ m3_purge :: "agent \ m3_trans" where "m3_purge \ m1_purge" text \Session key compromise.\ definition \ \refines @{term m2_leak}\ m3_leak :: "[rid_t, agent, agent, nonce, time] \ m3_trans" where "m3_leak Rs A B Na Ts \ {(s, s1). \ \guards:\ runs s Rs = Some (Serv, [A, B], [aNon Na, aNum Ts]) \ (clk s \ Ts + Ls) \ \ \only compromise 'old' session keys\ \ \actions:\ \ \record session key as leaked and add it to intruder knowledge\ s1 = s\ leak := insert (sesK (Rs$sk), A, B, Na, Ts) (leak s), IK := insert (Key (sesK (Rs$sk))) (IK s) \ }" text \Intruder fake event. The following "Dolev-Yao" event generates all intruder-derivable messages.\ definition \ \refines @{term "m2_fake"}\ m3_DY_fake :: "m3_trans" where "m3_DY_fake \ {(s, s1). \ \actions:\ s1 = s(| IK := synth (analz (IK s)) |) \ \take DY closure\ }" (******************************************************************************) subsection \Transition system\ (******************************************************************************) definition m3_init :: "m3_pred" where "m3_init \ { \ runs = Map.empty, leak = shrK`bad \ {undefined}, clk = 0, cache = {}, IK = Key`shrK`bad \ }" definition m3_trans :: "m3_trans" where "m3_trans \ (\A B Ra Rb Rs Na Kab Ts Ta T X. m3_step1 Ra A B Na \ m3_step2 Rb A B \ m3_step3 Rs A B Kab Na Ts \ m3_step4 Ra A B Na Kab Ts Ta X \ m3_step5 Rb A B Kab Ts Ta \ m3_step6 Ra A B Na Kab Ts Ta \ m3_tick T \ m3_purge A \ m3_leak Rs A B Na Ts \ m3_DY_fake \ Id )" definition m3 :: "(m3_state, m3_obs) spec" where "m3 \ \ init = m3_init, trans = m3_trans, obs = m3_obs \" lemmas m3_loc_defs = m3_def m3_init_def m3_trans_def m3_obs_def m3_step1_def m3_step2_def m3_step3_def m3_step4_def m3_step5_def m3_step6_def m3_tick_def m3_purge_def m3_leak_def m3_DY_fake_def lemmas m3_defs = m3_loc_defs m2_defs (******************************************************************************) subsection \Invariants\ (******************************************************************************) text \Specialized injection that we can apply more aggressively.\ lemmas analz_Inj_IK = analz.Inj [where H="IK s" for s] lemmas parts_Inj_IK = parts.Inj [where H="IK s" for s] declare parts_Inj_IK [dest!] declare analz_into_parts [dest] subsubsection \inv1: Secrecy of pre-distributed shared keys\ (*inv*************************************************************************) definition m3_inv1_lkeysec :: "m3_pred" where "m3_inv1_lkeysec \ {s. \C. (Key (shrK C) \ parts (IK s) \ C \ bad) \ (C \ bad \ Key (shrK C) \ IK s) }" lemmas m3_inv1_lkeysecI = m3_inv1_lkeysec_def [THEN setc_def_to_intro, rule_format] lemmas m3_inv1_lkeysecE [elim] = m3_inv1_lkeysec_def [THEN setc_def_to_elim, rule_format] lemmas m3_inv1_lkeysecD = m3_inv1_lkeysec_def [THEN setc_def_to_dest, rule_format] text \Invariance proof.\ lemma PO_m3_inv1_lkeysec_init [iff]: "init m3 \ m3_inv1_lkeysec" by (auto simp add: m3_defs intro!: m3_inv1_lkeysecI) lemma PO_m3_inv1_lkeysec_trans [iff]: "{m3_inv1_lkeysec} trans m3 {> m3_inv1_lkeysec}" by (fastforce simp add: PO_hoare_defs m3_defs intro!: m3_inv1_lkeysecI) lemma PO_m3_inv1_lkeysec [iff]: "reach m3 \ m3_inv1_lkeysec" by (rule inv_rule_basic) (fast+) text \Useful simplifier lemmas\ lemma m3_inv1_lkeysec_for_parts [simp]: "\ s \ m3_inv1_lkeysec \ \ Key (shrK C) \ parts (IK s) \ C \ bad" by auto lemma m3_inv1_lkeysec_for_analz [simp]: "\ s \ m3_inv1_lkeysec \ \ Key (shrK C) \ analz (IK s) \ C \ bad" by auto subsubsection \inv2: Session keys not used to encrypt other session keys\ (*invh*************************************************************************) text \Session keys are not used to encrypt other keys. Proof requires generalization to sets of session keys. NOTE: This invariant will be inherted from the corresponding L2 invariant using the simulation relation. \ definition m3_inv2_sesK_compr :: "m3_pred" where "m3_inv2_sesK_compr \ {s. \K KK. KK \ range sesK \ (Key K \ analz (Key`KK \ (IK s))) = (K \ KK \ Key K \ analz (IK s)) }" lemmas m3_inv2_sesK_comprI = m3_inv2_sesK_compr_def [THEN setc_def_to_intro, rule_format] lemmas m3_inv2_sesK_comprE = m3_inv2_sesK_compr_def [THEN setc_def_to_elim, rule_format] lemmas m3_inv2_sesK_comprD = m3_inv2_sesK_compr_def [THEN setc_def_to_dest, rule_format] text \Additional lemma\ lemmas insert_commute_Key = insert_commute [where x="Key K" for K] lemmas m3_inv2_sesK_compr_simps = m3_inv2_sesK_comprD m3_inv2_sesK_comprD [where KK="insert Kab KK" for Kab KK, simplified] m3_inv2_sesK_comprD [where KK="{Kab}" for Kab, simplified] insert_commute_Key (******************************************************************************) subsection \Refinement\ (******************************************************************************) subsubsection \Message abstraction and simulation relation\ (******************************************************************************) text \Abstraction function on sets of messages.\ inductive_set abs_msg :: "msg set \ chmsg set" for H :: "msg set" where am_M1: "\Agent A, Agent B, Nonce N\ \ H \ Insec A B (Msg [aNon N]) \ abs_msg H" | am_M2a: "Crypt (shrK C) \Key K, Agent B, Number T, Nonce N\ \ H \ Secure Sv C (Msg [aKey K, aAgt B, aNum T, aNon N]) \ abs_msg H" | am_M2b: "Crypt (shrK C) \Key K, Agent A, Number T\ \ H \ Secure Sv C (Msg [aKey K, aAgt A, aNum T]) \ abs_msg H" | am_M3: "Crypt K \Agent A, Number T\ \ H \ dAuth K (Msg [aAgt A, aNum T]) \ abs_msg H" | am_M4: "Crypt K (Number T) \ H \ dAuth K (Msg [aNum T]) \ abs_msg H" text \R23: The simulation relation. This is a data refinement of the insecure and secure channels of refinement 2.\ definition R23_msgs :: "(m2_state \ m3_state) set" where "R23_msgs \ {(s, t). abs_msg (parts (IK t)) \ chan s }" definition R23_keys :: "(m2_state \ m3_state) set" where "R23_keys \ {(s, t). \KK K. KK \ range sesK \ Key K \ analz (Key`KK \ (IK t)) \ aKey K \ extr (aKey`KK \ ik0) (chan s) }" definition R23_non :: "(m2_state \ m3_state) set" where "R23_non \ {(s, t). \KK N. KK \ range sesK \ Nonce N \ analz (Key`KK \ (IK t)) \ aNon N \ extr (aKey`KK \ ik0) (chan s) }" definition R23_pres :: "(m2_state \ m3_state) set" where "R23_pres \ {(s, t). runs s = runs t \ leak s = leak t \ clk s = clk t \ cache s = cache t}" definition R23 :: "(m2_state \ m3_state) set" where "R23 \ R23_msgs \ R23_keys \ R23_non \ R23_pres" lemmas R23_defs = R23_def R23_msgs_def R23_keys_def R23_non_def R23_pres_def text \The mediator function is the identity here.\ definition med32 :: "m3_obs \ m2_obs" where "med32 \ id" lemmas R23_msgsI = R23_msgs_def [THEN rel_def_to_intro, simplified, rule_format] lemmas R23_msgsE [elim] = R23_msgs_def [THEN rel_def_to_elim, simplified, rule_format] lemmas R23_msgsE' [elim] = R23_msgs_def [THEN rel_def_to_dest, simplified, rule_format, THEN subsetD] lemmas R23_keysI = R23_keys_def [THEN rel_def_to_intro, simplified, rule_format] lemmas R23_keysE [elim] = R23_keys_def [THEN rel_def_to_elim, simplified, rule_format] lemmas R23_nonI = R23_non_def [THEN rel_def_to_intro, simplified, rule_format] lemmas R23_nonE [elim] = R23_non_def [THEN rel_def_to_elim, simplified, rule_format] lemmas R23_presI = R23_pres_def [THEN rel_def_to_intro, simplified, rule_format] lemmas R23_presE [elim] = R23_pres_def [THEN rel_def_to_elim, simplified, rule_format] lemmas R23_intros = R23_msgsI R23_keysI R23_nonI R23_presI text \Simplifier lemmas for various instantiations (keys and nonces).\ lemmas R23_keys_simp = R23_keys_def [THEN rel_def_to_dest, simplified, rule_format] lemmas R23_keys_simps = R23_keys_simp R23_keys_simp [where KK="{}", simplified] R23_keys_simp [where KK="{K'}" for K', simplified] R23_keys_simp [where KK="insert K' KK" for K' KK, simplified, OF _ conjI] lemmas R23_non_simp = R23_non_def [THEN rel_def_to_dest, simplified, rule_format] lemmas R23_non_simps = R23_non_simp R23_non_simp [where KK="{}", simplified] R23_non_simp [where KK="{K}" for K, simplified] R23_non_simp [where KK="insert K KK" for K KK, simplified, OF _ conjI] lemmas R23_simps = R23_keys_simps R23_non_simps subsubsection \General lemmas\ (******************************************************************************) text \General facts about @{term "abs_msg"}\ declare abs_msg.intros [intro!] declare abs_msg.cases [elim!] lemma abs_msg_empty: "abs_msg {} = {}" by (auto) lemma abs_msg_Un [simp]: "abs_msg (G \ H) = abs_msg G \ abs_msg H" by (auto) lemma abs_msg_mono [elim]: "\ m \ abs_msg G; G \ H \ \ m \ abs_msg H" by (auto) lemma abs_msg_insert_mono [intro]: "\ m \ abs_msg H \ \ m \ abs_msg (insert m' H)" by (auto) text \Facts about @{term "abs_msg"} concerning abstraction of fakeable messages. This is crucial for proving the refinement of the intruder event.\ lemma abs_msg_DY_subset_fakeable: "\ (s, t) \ R23_msgs; (s, t) \ R23_keys; (s, t) \ R23_non; t \ m3_inv1_lkeysec \ \ abs_msg (synth (analz (IK t))) \ fake ik0 (dom (runs s)) (chan s)" apply (auto) \ \9 subgoals, deal with replays first\ prefer 2 apply (blast) prefer 3 apply (blast) prefer 4 apply (blast) prefer 5 apply (blast) \ \remaining 5 subgoals are real fakes\ apply (intro fake_StatCh fake_DynCh, auto simp add: R23_simps)+ done subsubsection \Refinement proof\ (******************************************************************************) text \Pair decomposition. These were set to \texttt{elim!}, which is too agressive here.\ declare MPair_analz [rule del, elim] declare MPair_parts [rule del, elim] text \Protocol events.\ lemma PO_m3_step1_refines_m2_step1: "{R23} (m2_step1 Ra A B Na), (m3_step1 Ra A B Na) {> R23}" by (auto simp add: PO_rhoare_defs R23_def m3_defs intro!: R23_intros) (auto) lemma PO_m3_step2_refines_m2_step2: "{R23} (m2_step2 Rb A B), (m3_step2 Rb A B) {> R23}" by (auto simp add: PO_rhoare_defs R23_def m3_defs intro!: R23_intros) lemma PO_m3_step3_refines_m2_step3: "{R23 \ (m2_inv3a_sesK_compr) \ (m3_inv2_sesK_compr \ m3_inv1_lkeysec)} (m2_step3 Rs A B Kab Na Ts), (m3_step3 Rs A B Kab Na Ts) {> R23}" proof - { fix s t assume H: "(s, t) \ R23_msgs" "(s, t) \ R23_keys" "(s, t) \ R23_non" "(s, t) \ R23_pres" "s \ m2_inv3a_sesK_compr" "t \ m3_inv2_sesK_compr" "t \ m3_inv1_lkeysec" "Kab = sesK (Rs$sk)" "Rs \ dom (runs t)" "\ Agent A, Agent B, Nonce Na \ \ parts (IK t)" let ?s'= - "s\ runs := runs s(Rs \ (Serv, [A, B], [aNon Na, aNum (clk t)])), + "s\ runs := (runs s)(Rs \ (Serv, [A, B], [aNon Na, aNum (clk t)])), chan := insert (Secure Sv A (Msg [aKey Kab, aAgt B, aNum (clk t), aNon Na])) (insert (Secure Sv B (Msg [aKey Kab, aAgt A, aNum (clk t)])) (chan s)) \" let ?t'= - "t\ runs := runs t(Rs \ (Serv, [A, B], [aNon Na, aNum (clk t)])), + "t\ runs := (runs t)(Rs \ (Serv, [A, B], [aNon Na, aNum (clk t)])), IK := insert \ Crypt (shrK A) \ Key Kab, Agent B, Number (clk t), Nonce Na \, Crypt (shrK B) \ Key Kab, Agent A, Number (clk t) \ \ (IK t) \" \ \here we go\ have "(?s', ?t') \ R23_msgs" using H by (-) (rule R23_intros, auto) moreover have "(?s', ?t') \ R23_keys" using H by (-) (rule R23_intros, auto simp add: m2_inv3a_sesK_compr_simps m3_inv2_sesK_compr_simps, auto simp add: R23_keys_simps) moreover have "(?s', ?t') \ R23_non" using H by (-) (rule R23_intros, auto simp add: m2_inv3a_sesK_compr_simps m3_inv2_sesK_compr_simps R23_non_simps) moreover have "(?s', ?t') \ R23_pres" using H by (-) (rule R23_intros, auto) moreover note calculation } thus ?thesis by (auto simp add: PO_rhoare_defs R23_def m3_defs) qed lemma PO_m3_step4_refines_m2_step4: "{R23 \ UNIV \ m3_inv1_lkeysec} (m2_step4 Ra A B Na Kab Ts Ta), (m3_step4 Ra A B Na Kab Ts Ta X) {> R23}" proof - { fix s t assume H: "(s, t) \ R23_msgs" "(s, t) \ R23_keys" "(s, t) \ R23_non" "(s, t) \ R23_pres" "t \ m3_inv1_lkeysec" "runs t Ra = Some (Init, [A, B], [])" "Na = Ra$na" "\ Crypt (shrK A) \Key Kab, Agent B, Number Ts, Nonce Na\, X\ \ analz (IK t)" - let ?s' = "s\ runs := runs s(Ra \ (Init, [A, B], [aKey Kab, aNum Ts, aNum (clk t)])), + let ?s' = "s\ runs := (runs s)(Ra \ (Init, [A, B], [aKey Kab, aNum Ts, aNum (clk t)])), chan := insert (dAuth Kab (Msg [aAgt A, aNum (clk t)])) (chan s) \" - and ?t' = "t\ runs := runs t(Ra \ (Init, [A, B], [aKey Kab, aNum Ts, aNum (clk t)])), + and ?t' = "t\ runs := (runs t)(Ra \ (Init, [A, B], [aKey Kab, aNum Ts, aNum (clk t)])), IK := insert \Crypt Kab \Agent A, Number (clk t)\, X\ (IK t) \" from H have "Secure Sv A (Msg [aKey Kab, aAgt B, aNum Ts, aNon Na]) \ chan s" by (auto dest!: analz_into_parts elim!: MPair_parts) moreover from H have "X \ parts (IK t)" by (auto) with H have "(?s', ?t') \ R23_msgs" by (auto intro!: R23_intros) (auto) moreover from H have "X \ analz (IK t)" by (auto) with H have "(?s', ?t') \ R23_keys" by (auto intro!: R23_intros) (auto dest!: analz_cut intro: analz_monotonic) moreover from H have "X \ analz (IK t)" by (auto) with H have "(?s', ?t') \ R23_non" by (auto intro!: R23_intros) (auto dest!: analz_cut intro: analz_monotonic) moreover have "(?s', ?t') \ R23_pres" using H by (auto intro!: R23_intros) moreover note calculation } thus ?thesis by (auto simp add: PO_rhoare_defs R23_def m3_defs dest!: analz_Inj_IK) qed lemma PO_m3_step5_refines_m2_step5: "{R23} (m2_step5 Rb A B Kab Ts Ta), (m3_step5 Rb A B Kab Ts Ta) {> R23}" by (auto simp add: PO_rhoare_defs R23_def m3_defs intro!: R23_intros) (auto) lemma PO_m3_step6_refines_m2_step6: "{R23} (m2_step6 Ra A B Na Kab Ts Ta), (m3_step6 Ra A B Na Kab Ts Ta) {> R23}" by (auto simp add: PO_rhoare_defs R23_def m3_defs intro!: R23_intros) lemma PO_m3_tick_refines_m2_tick: "{R23} (m2_tick T), (m3_tick T) {>R23}" by (auto simp add: PO_rhoare_defs R23_def m3_defs intro!: R23_intros) lemma PO_m3_purge_refines_m2_purge: "{R23} (m2_purge A), (m3_purge A) {>R23}" by (auto simp add: PO_rhoare_defs R23_def m3_defs intro!: R23_intros) text \Intruder events.\ lemma PO_m3_leak_refines_m2_leak: "{R23} (m2_leak Rs A B Na Ts), (m3_leak Rs A B Na Ts) {>R23}" by (auto simp add: PO_rhoare_defs R23_def m3_defs R23_simps intro!: R23_intros) lemma PO_m3_DY_fake_refines_m2_fake: "{R23 \ m2_inv3a_sesK_compr \ (m3_inv2_sesK_compr \ m3_inv1_lkeysec)} m2_fake, m3_DY_fake {> R23}" apply (auto simp add: PO_rhoare_defs R23_def m3_defs intro!: R23_intros del: abs_msg.cases) apply (auto intro: abs_msg_DY_subset_fakeable [THEN subsetD] del: abs_msg.cases) apply (auto simp add: R23_simps) done text \All together now...\ lemmas PO_m3_trans_refines_m2_trans = PO_m3_step1_refines_m2_step1 PO_m3_step2_refines_m2_step2 PO_m3_step3_refines_m2_step3 PO_m3_step4_refines_m2_step4 PO_m3_step5_refines_m2_step5 PO_m3_step6_refines_m2_step6 PO_m3_tick_refines_m2_tick PO_m3_purge_refines_m2_purge PO_m3_leak_refines_m2_leak PO_m3_DY_fake_refines_m2_fake lemma PO_m3_refines_init_m2 [iff]: "init m3 \ R23``(init m2)" by (auto simp add: R23_def m3_defs intro!: R23_intros) lemma PO_m3_refines_trans_m2 [iff]: "{R23 \ (m2_inv3a_sesK_compr) \ (m3_inv2_sesK_compr \ m3_inv1_lkeysec)} (trans m2), (trans m3) {> R23}" by (auto simp add: m3_def m3_trans_def m2_def m2_trans_def) (blast intro!: PO_m3_trans_refines_m2_trans)+ lemma PO_m3_observation_consistent [iff]: "obs_consistent R23 med32 m2 m3" by (auto simp add: obs_consistent_def R23_def med32_def m3_defs) text \Refinement result.\ lemma m3_refines_m2 [iff]: "refines (R23 \ (m2_inv3a_sesK_compr) \ (m3_inv1_lkeysec)) med32 m2 m3" proof - have "R23 \ m2_inv3a_sesK_compr \ UNIV \ UNIV \ m3_inv2_sesK_compr" by (auto simp add: R23_def R23_keys_simps intro!: m3_inv2_sesK_comprI) thus ?thesis by (-) (rule Refinement_using_invariants, auto) qed lemma m3_implements_m2 [iff]: "implements med32 m2 m3" by (rule refinement_soundness) (auto) subsection \Inherited invariants\ (******************************************************************************) subsubsection \inv3 (derived): Key secrecy for initiator\ (*invh*************************************************************************) definition m3_inv3_ikk_init :: "m3_state set" where "m3_inv3_ikk_init \ {s. \A B Ra K Ts nl. runs s Ra = Some (Init, [A, B], aKey K # aNum Ts # nl) \ A \ good \ B \ good \ Key K \ analz (IK s) \ (K, A, B, Ra$na, Ts) \ leak s }" lemmas m3_inv3_ikk_initI = m3_inv3_ikk_init_def [THEN setc_def_to_intro, rule_format] lemmas m3_inv3_ikk_initE [elim] = m3_inv3_ikk_init_def [THEN setc_def_to_elim, rule_format] lemmas m3_inv3_ikk_initD = m3_inv3_ikk_init_def [THEN setc_def_to_dest, rule_format, rotated 1] lemma PO_m3_inv3_ikk_init: "reach m3 \ m3_inv3_ikk_init" proof (rule INV_from_Refinement_using_invariants [OF m3_refines_m2]) show "Range (R23 \ m2_inv3a_sesK_compr \ m3_inv1_lkeysec \ m2_inv6_ikk_init \ UNIV) \ m3_inv3_ikk_init" by (fastforce simp add: R23_def R23_keys_simps intro!: m3_inv3_ikk_initI) qed auto subsubsection \inv4 (derived): Key secrecy for responder\ (*invh*************************************************************************) definition m3_inv4_ikk_resp :: "m3_state set" where "m3_inv4_ikk_resp \ {s. \A B Rb K Ts nl. runs s Rb = Some (Resp, [A, B], aKey K # aNum Ts # nl) \ A \ good \ B \ good \ Key K \ analz (IK s) \ (\Na. (K, A, B, Na, Ts) \ leak s) }" lemmas m3_inv4_ikk_respI = m3_inv4_ikk_resp_def [THEN setc_def_to_intro, rule_format] lemmas m3_inv4_ikk_respE [elim] = m3_inv4_ikk_resp_def [THEN setc_def_to_elim, rule_format] lemmas m3_inv4_ikk_respD = m3_inv4_ikk_resp_def [THEN setc_def_to_dest, rule_format, rotated 1] lemma PO_m3_inv4_ikk_resp: "reach m3 \ m3_inv4_ikk_resp" proof (rule INV_from_Refinement_using_invariants [OF m3_refines_m2]) show "Range (R23 \ m2_inv3a_sesK_compr \ m3_inv1_lkeysec \ m2_inv7_ikk_resp \ UNIV) \ m3_inv4_ikk_resp" by (auto simp add: R23_def R23_keys_simps intro!: m3_inv4_ikk_respI) (elim m2_inv7_ikk_respE, auto) qed auto end diff --git a/thys/Security_Protocol_Refinement/Key_establish/m3_kerberos_par.thy b/thys/Security_Protocol_Refinement/Key_establish/m3_kerberos_par.thy --- a/thys/Security_Protocol_Refinement/Key_establish/m3_kerberos_par.thy +++ b/thys/Security_Protocol_Refinement/Key_establish/m3_kerberos_par.thy @@ -1,759 +1,759 @@ (******************************************************************************* Project: Development of Security Protocols by Refinement Module: Key_establish/m3_kerberos_par.thy (Isabelle/HOL 2016-1) ID: $Id: m3_kerberos_par.thy 132890 2016-12-24 10:25:57Z csprenge $ Authors: Ivano Somaini, ETH Zurich Christoph Sprenger, ETH Zurich Key distribution protocols Third refinement: parallel version of core Kerberos protocol Copyright (c) 2009-2016 Christoph Sprenger Licence: LGPL *******************************************************************************) section \Core Kerberos, "parallel" variant (L3)\ theory m3_kerberos_par imports m2_kerberos "../Refinement/Message" begin text \ We model a direct implementation of the channel-based core Kerberos protocol at Level 2 without ticket forwarding: \[ \begin{array}{lll} \mathrm{M1.} & A \rightarrow S: & A, B, Na \\ \mathrm{M2a.} & S \rightarrow A: & \{Kab, B, Ts, Na\}_{Kas} \\ \mathrm{M2b.} & S \rightarrow B: & \{Kab, A, Ts\}_{Kbs} \\ \mathrm{M3.} & A \rightarrow B: & \{A, Ta\}_{Kab} \\ \mathrm{M4.} & B \rightarrow A: & \{Ta\}_{Kab} \\ \end{array} \] \ text \Proof tool configuration. Avoid annoying automatic unfolding of \dom\.\ declare domIff [simp, iff del] (******************************************************************************) subsection \Setup\ (******************************************************************************) text \Now we can define the initial key knowledge.\ overloading ltkeySetup' \ ltkeySetup begin definition ltkeySetup_def: "ltkeySetup' \ {(sharK C, A) | C A. A = C \ A = Sv}" end lemma corrKey_shrK_bad [simp]: "corrKey = shrK`bad" by (auto simp add: keySetup_def ltkeySetup_def corrKey_def) (******************************************************************************) subsection \State\ (******************************************************************************) text \The secure channels are star-shaped to/from the server. Therefore, we have only one agent in the relation.\ record m3_state = "m1_state" + IK :: "msg set" \ \intruder knowledge\ text \Observable state: @{term "runs"}, @{term "leak"}, @{term "clk"}, and @{term "cache"}.\ type_synonym m3_obs = "m2_obs" definition m3_obs :: "m3_state \ m3_obs" where "m3_obs s \ \ runs = runs s, leak = leak s, clk = clk s, cache = cache s \" type_synonym m3_pred = "m3_state set" type_synonym m3_trans = "(m3_state \ m3_state) set" (******************************************************************************) subsection \Events\ (******************************************************************************) text \Protocol events.\ definition \ \by @{term "A"}, refines @{term "m2_step1"}\ m3_step1 :: "[rid_t, agent, agent, nonce] \ m3_trans" where "m3_step1 Ra A B Na \ {(s, s1). \ \guards:\ Ra \ dom (runs s) \ \ \\Ra\ is fresh\ Na = Ra$na \ \ \generated nonce\ \ \actions:\ s1 = s\ runs := (runs s)(Ra \ (Init, [A, B], [])), IK := insert \Agent A, Agent B, Nonce Na\ (IK s) \ \send \M1\\ \ }" definition \ \by @{term "B"}, refines @{term "m2_step2"}\ m3_step2 :: "[rid_t, agent, agent] \ m3_trans" where "m3_step2 \ m1_step2" definition \ \by @{text "Server"}, refines @{term m2_step3}\ m3_step3 :: "[rid_t, agent, agent, key, nonce, time] \ m3_trans" where "m3_step3 Rs A B Kab Na Ts \ {(s, s1). \ \guards:\ Rs \ dom (runs s) \ \ \fresh server run\ Kab = sesK (Rs$sk) \ \ \fresh session key\ \Agent A, Agent B, Nonce Na\ \ IK s \ \ \recv \M1\\ Ts = clk s \ \ \fresh timestamp\ \ \actions:\ \ \record session key and send \M2\\ s1 = s\ runs := (runs s)(Rs \ (Serv, [A, B], [aNon Na, aNum Ts])), IK := insert (Crypt (shrK A) \Key Kab, Agent B, Number Ts, Nonce Na\) (insert (Crypt (shrK B) \Key Kab, Agent A, Number Ts\) (IK s)) \ }" definition \ \by @{term "A"}, refines @{term m2_step4}\ m3_step4 :: "[rid_t, agent, agent, nonce, key, time, time] \ m3_trans" where "m3_step4 Ra A B Na Kab Ts Ta \ {(s, s1). \ \guards:\ runs s Ra = Some (Init, [A, B], []) \ \ \key not yet recv'd\ Na = Ra$na \ \ \generated nonce\ Crypt (shrK A) \ \recv \M2a\\ \Key Kab, Agent B, Number Ts, Nonce Na\ \ IK s \ \ \read current time\ Ta = clk s \ \ \check freshness of session key\ clk s < Ts + Ls \ \ \actions:\ \ \record session key and send \M3\\ s1 = s\ runs := (runs s)(Ra \ (Init, [A, B], [aKey Kab, aNum Ts, aNum Ta])), IK := insert (Crypt Kab \Agent A, Number Ta\) (IK s) \ \\M3\\ \ }" definition \ \by @{term "B"}, refines @{term m2_step5}\ m3_step5 :: "[rid_t, agent, agent, key, time, time] \ m3_trans" where "m3_step5 Rb A B Kab Ts Ta \ {(s, s1). \ \guards:\ runs s Rb = Some (Resp, [A, B], []) \ \ \key not yet recv'd\ Crypt (shrK B) \Key Kab, Agent A, Number Ts\ \ IK s \ \ \recv \M2b\\ Crypt Kab \Agent A, Number Ta\ \ IK s \ \ \recv \M3\\ \ \ensure freshness of session key\ clk s < Ts + Ls \ \ \check authenticator's validity and replay; 'replays' with fresh authenticator ok!\ clk s < Ta + La \ (B, Kab, Ta) \ cache s \ \ \actions:\ \ \record session key\ s1 = s\ runs := (runs s)(Rb \ (Resp, [A, B], [aKey Kab, aNum Ts, aNum Ta])), cache := insert (B, Kab, Ta) (cache s), IK := insert (Crypt Kab (Number Ta)) (IK s) \ \send \M4\\ \ }" definition \ \by @{term "A"}, refines @{term m2_step6}\ m3_step6 :: "[rid_t, agent, agent, nonce, key, time, time] \ m3_trans" where "m3_step6 Ra A B Na Kab Ts Ta \ {(s, s'). \ \guards:\ runs s Ra = Some (Init, [A, B], [aKey Kab, aNum Ts, aNum Ta]) \ \ \knows key\ Na = Ra$na \ \ \generated nonce\ clk s < Ts + Ls \ \ \check session key's recentness\ Crypt Kab (Number Ta) \ IK s \ \ \recv \M4\\ \ \actions:\ s' = s\ runs := (runs s)(Ra \ (Init, [A, B], [aKey Kab, aNum Ts, aNum Ta, END])) \ }" text \Clock tick event\ definition \ \refines @{term "m2_tick"}\ m3_tick :: "time \ m3_trans" where "m3_tick \ m1_tick" text \Purge event: purge cache of expired timestamps\ definition \ \refines @{term "m2_purge"}\ m3_purge :: "agent \ m3_trans" where "m3_purge \ m1_purge" text \Session key compromise.\ definition \ \refines @{term m2_leak}\ m3_leak :: "[rid_t, agent, agent, nonce, time] \ m3_trans" where "m3_leak Rs A B Na Ts \ {(s, s1). \ \guards:\ runs s Rs = Some (Serv, [A, B], [aNon Na, aNum Ts]) \ (clk s \ Ts + Ls) \ \ \only compromise 'old' session keys!\ \ \actions:\ \ \record session key as leaked and add it to intruder knowledge\ s1 = s\ leak := insert (sesK (Rs$sk), A, B, Na, Ts) (leak s), IK := insert (Key (sesK (Rs$sk))) (IK s) \ }" text \Intruder fake event. The following "Dolev-Yao" event generates all intruder-derivable messages.\ definition \ \refines @{term "m2_fake"}\ m3_DY_fake :: "m3_trans" where "m3_DY_fake \ {(s, s1). \ \actions:\ s1 = s(| IK := synth (analz (IK s)) |) \ \take DY closure\ }" (******************************************************************************) subsection \Transition system\ (******************************************************************************) definition m3_init :: "m3_pred" where "m3_init \ { \ runs = Map.empty, leak = shrK`bad \ {undefined}, clk = 0, cache = {}, IK = Key`shrK`bad \ }" definition m3_trans :: "m3_trans" where "m3_trans \ (\A B Ra Rb Rs Na Kab Ts Ta T. m3_step1 Ra A B Na \ m3_step2 Rb A B \ m3_step3 Rs A B Kab Na Ts \ m3_step4 Ra A B Na Kab Ts Ta \ m3_step5 Rb A B Kab Ts Ta \ m3_step6 Ra A B Na Kab Ts Ta \ m3_tick T \ m3_purge A \ m3_leak Rs A B Na Ts \ m3_DY_fake \ Id )" definition m3 :: "(m3_state, m3_obs) spec" where "m3 \ \ init = m3_init, trans = m3_trans, obs = m3_obs \" lemmas m3_loc_defs = m3_def m3_init_def m3_trans_def m3_obs_def m3_step1_def m3_step2_def m3_step3_def m3_step4_def m3_step5_def m3_step6_def m3_tick_def m3_purge_def m3_leak_def m3_DY_fake_def lemmas m3_defs = m3_loc_defs m2_defs (******************************************************************************) subsection \Invariants\ (******************************************************************************) text \Specialized injection that we can apply more aggressively.\ lemmas analz_Inj_IK = analz.Inj [where H="IK s" for s] lemmas parts_Inj_IK = parts.Inj [where H="IK s" for s] declare parts_Inj_IK [dest!] declare analz_into_parts [dest] subsubsection \inv1: Secrecy of pre-distributed shared keys\ (******************************************************************************) text \inv1: Secrecy of long-term keys\ definition m3_inv1_lkeysec :: "m3_pred" where "m3_inv1_lkeysec \ {s. \C. (Key (shrK C) \ parts (IK s) \ C \ bad) \ (C \ bad \ Key (shrK C) \ IK s) }" lemmas m3_inv1_lkeysecI = m3_inv1_lkeysec_def [THEN setc_def_to_intro, rule_format] lemmas m3_inv1_lkeysecE [elim] = m3_inv1_lkeysec_def [THEN setc_def_to_elim, rule_format] lemmas m3_inv1_lkeysec_dest = m3_inv1_lkeysec_def [THEN setc_def_to_dest, rule_format] text \Invariance proof.\ lemma PO_m3_inv1_lkeysec_init [iff]: "init m3 \ m3_inv1_lkeysec" by (auto simp add: m3_defs intro!: m3_inv1_lkeysecI) lemma PO_m3_inv1_lkeysec_trans [iff]: "{m3_inv1_lkeysec} trans m3 {> m3_inv1_lkeysec}" by (fastforce simp add: PO_hoare_defs m3_defs intro!: m3_inv1_lkeysecI) lemma PO_m3_inv1_lkeysec [iff]: "reach m3 \ m3_inv1_lkeysec" by (rule inv_rule_basic) (fast+) text \Useful simplifier lemmas\ lemma m3_inv1_lkeysec_for_parts [simp]: "\ s \ m3_inv1_lkeysec \ \ Key (shrK C) \ parts (IK s) \ C \ bad" by auto lemma m3_inv1_lkeysec_for_analz [simp]: "\ s \ m3_inv1_lkeysec \ \ Key (shrK C) \ analz (IK s) \ C \ bad" by auto subsubsection \inv7a: Session keys not used to encrypt other session keys\ (******************************************************************************) text \Session keys are not used to encrypt other keys. Proof requires generalization to sets of session keys. NOTE: This invariant will be derived from the corresponding L2 invariant using the simulation relation. \ definition m3_inv7a_sesK_compr :: "m3_pred" where "m3_inv7a_sesK_compr \ {s. \K KK. KK \ range sesK \ (Key K \ analz (Key`KK \ (IK s))) = (K \ KK \ Key K \ analz (IK s)) }" lemmas m3_inv7a_sesK_comprI = m3_inv7a_sesK_compr_def [THEN setc_def_to_intro, rule_format] lemmas m3_inv7a_sesK_comprE = m3_inv7a_sesK_compr_def [THEN setc_def_to_elim, rule_format] lemmas m3_inv7a_sesK_comprD = m3_inv7a_sesK_compr_def [THEN setc_def_to_dest, rule_format] text \Additional lemma\ lemmas insert_commute_Key = insert_commute [where x="Key K" for K] lemmas m3_inv7a_sesK_compr_simps = m3_inv7a_sesK_comprD m3_inv7a_sesK_comprD [where KK="{Kab}" for Kab, simplified] m3_inv7a_sesK_comprD [where KK="insert Kab KK" for Kab KK, simplified] insert_commute_Key (******************************************************************************) subsection \Refinement\ (******************************************************************************) subsubsection \Message abstraction and simulation relation\ (******************************************************************************) text \Abstraction function on sets of messages.\ inductive_set abs_msg :: "msg set \ chmsg set" for H :: "msg set" where am_M1: "\Agent A, Agent B, Nonce N\ \ H \ Insec A B (Msg [aNon N]) \ abs_msg H" | am_M2a: "Crypt (shrK C) \Key K, Agent B, Number T, Nonce N\ \ H \ Secure Sv C (Msg [aKey K, aAgt B, aNum T, aNon N]) \ abs_msg H" | am_M2b: "Crypt (shrK C) \Key K, Agent A, Number T\ \ H \ Secure Sv C (Msg [aKey K, aAgt A, aNum T]) \ abs_msg H" | am_M3: "Crypt K \Agent A, Number T\ \ H \ dAuth K (Msg [aAgt A, aNum T]) \ abs_msg H" | am_M4: "Crypt K (Number T) \ H \ dAuth K (Msg [aNum T]) \ abs_msg H" text \R23: The simulation relation. This is a data refinement of the insecure and secure channels of refinement 2.\ definition R23_msgs :: "(m2_state \ m3_state) set" where "R23_msgs \ {(s, t). abs_msg (parts (IK t)) \ chan s }" definition R23_keys :: "(m2_state \ m3_state) set" where "R23_keys \ {(s, t). \KK K. KK \ range sesK \ Key K \ analz (Key`KK \ (IK t)) \ aKey K \ extr (aKey`KK \ ik0) (chan s) }" definition R23_non :: "(m2_state \ m3_state) set" where "R23_non \ {(s, t). \KK N. KK \ range sesK \ Nonce N \ analz (Key`KK \ (IK t)) \ aNon N \ extr (aKey`KK \ ik0) (chan s) }" definition R23_pres :: "(m2_state \ m3_state) set" where "R23_pres \ {(s, t). runs s = runs t \ leak s = leak t \ clk s = clk t \ cache s = cache t}" definition R23 :: "(m2_state \ m3_state) set" where "R23 \ R23_msgs \ R23_keys \ R23_non \ R23_pres" lemmas R23_defs = R23_def R23_msgs_def R23_keys_def R23_non_def R23_pres_def text \The mediator function is the identity here.\ definition med32 :: "m3_obs \ m2_obs" where "med32 \ id" lemmas R23_msgsI = R23_msgs_def [THEN rel_def_to_intro, simplified, rule_format] lemmas R23_msgsE [elim] = R23_msgs_def [THEN rel_def_to_elim, simplified, rule_format] lemmas R23_msgsE' [elim] = R23_msgs_def [THEN rel_def_to_dest, simplified, rule_format, THEN subsetD] lemmas R23_keysI = R23_keys_def [THEN rel_def_to_intro, simplified, rule_format] lemmas R23_keysE [elim] = R23_keys_def [THEN rel_def_to_elim, simplified, rule_format] lemmas R23_nonI = R23_non_def [THEN rel_def_to_intro, simplified, rule_format] lemmas R23_nonE [elim] = R23_non_def [THEN rel_def_to_elim, simplified, rule_format] lemmas R23_presI = R23_pres_def [THEN rel_def_to_intro, simplified, rule_format] lemmas R23_presE [elim] = R23_pres_def [THEN rel_def_to_elim, simplified, rule_format] lemmas R23_intros = R23_msgsI R23_keysI R23_nonI R23_presI text \Simplifier lemmas for various instantiations (keys and nonces).\ lemmas R23_keys_simp = R23_keys_def [THEN rel_def_to_dest, simplified, rule_format] lemmas R23_keys_simps = R23_keys_simp R23_keys_simp [where KK="{}", simplified] R23_keys_simp [where KK="{K'}" for K', simplified] R23_keys_simp [where KK="insert K' KK" for K' KK, simplified, OF _ conjI] lemmas R23_non_simp = R23_non_def [THEN rel_def_to_dest, simplified, rule_format] lemmas R23_non_simps = R23_non_simp R23_non_simp [where KK="{}", simplified] R23_non_simp [where KK="{K}" for K, simplified] R23_non_simp [where KK="insert K KK" for K KK, simplified, OF _ conjI] lemmas R23_simps = R23_keys_simps R23_non_simps (* lemmas R23_keys_emptyD = R23_keys_simp [where KK="{}", simplified, THEN iffD1, rotated 1] R23_keys_simp [where KK="{}", simplified, THEN iffD2, rotated 1] lemmas R23_non_emptyD = R23_non_simp [where KK="{}", simplified, THEN iffD1, rotated 1] R23_non_simp [where KK="{}", simplified, THEN iffD2, rotated 1] lemmas R23_emptyD = R23_keys_emptyD R23_non_emptyD *) subsubsection \General lemmas\ (******************************************************************************) text \General facts about @{term "abs_msg"}\ declare abs_msg.intros [intro!] declare abs_msg.cases [elim!] lemma abs_msg_empty: "abs_msg {} = {}" by (auto) lemma abs_msg_Un [simp]: "abs_msg (G \ H) = abs_msg G \ abs_msg H" by (auto) lemma abs_msg_mono [elim]: "\ m \ abs_msg G; G \ H \ \ m \ abs_msg H" by (auto) lemma abs_msg_insert_mono [intro]: "\ m \ abs_msg H \ \ m \ abs_msg (insert m' H)" by (auto) text \Facts about @{term "abs_msg"} concerning abstraction of fakeable messages. This is crucial for proving the refinement of the intruder event.\ lemma abs_msg_DY_subset_fakeable: "\ (s, t) \ R23_msgs; (s, t) \ R23_keys; (s, t) \ R23_non; t \ m3_inv1_lkeysec \ \ abs_msg (synth (analz (IK t))) \ fake ik0 (dom (runs s)) (chan s)" apply (auto) \ \9 subgoals, deal with replays first\ prefer 2 apply (blast) prefer 3 apply (blast) prefer 4 apply (blast) prefer 5 apply (blast) \ \remaining 5 subgoals are real fakes\ apply (intro fake_StatCh fake_DynCh, auto simp add: R23_simps)+ done subsubsection \Refinement proof\ (******************************************************************************) text \Pair decomposition. These were set to \texttt{elim!}, which is too agressive here.\ declare MPair_analz [rule del, elim] declare MPair_parts [rule del, elim] text \Protocol events.\ lemma PO_m3_step1_refines_m2_step1: "{R23} (m2_step1 Ra A B Na), (m3_step1 Ra A B Na) {> R23}" by (auto simp add: PO_rhoare_defs R23_def m3_defs intro!: R23_intros) (auto) lemma PO_m3_step2_refines_m2_step2: "{R23} (m2_step2 Rb A B), (m3_step2 Rb A B) {> R23}" by (auto simp add: PO_rhoare_defs R23_def m3_defs intro!: R23_intros) lemma PO_m3_step3_refines_m2_step3: "{R23 \ (m2_inv3a_sesK_compr) \ (m3_inv7a_sesK_compr \ m3_inv1_lkeysec)} (m2_step3 Rs A B Kab Na Ts), (m3_step3 Rs A B Kab Na Ts) {> R23}" proof - { fix s t assume H: "(s, t) \ R23_msgs" "(s, t) \ R23_keys" "(s, t) \ R23_non" "(s, t) \ R23_pres" "s \ m2_inv3a_sesK_compr" "t \ m3_inv7a_sesK_compr" "t \ m3_inv1_lkeysec" "Kab = sesK (Rs$sk)" "Rs \ dom (runs t)" "\ Agent A, Agent B, Nonce Na \ \ parts (IK t)" let ?s'= - "s\ runs := runs s(Rs \ (Serv, [A, B], [aNon Na, aNum (clk t)])), + "s\ runs := (runs s)(Rs \ (Serv, [A, B], [aNon Na, aNum (clk t)])), chan := insert (Secure Sv A (Msg [aKey Kab, aAgt B, aNum (clk t), aNon Na])) (insert (Secure Sv B (Msg [aKey Kab, aAgt A, aNum (clk t)])) (chan s)) \" let ?t'= - "t\ runs := runs t(Rs \ (Serv, [A, B], [aNon Na, aNum (clk t)])), + "t\ runs := (runs t)(Rs \ (Serv, [A, B], [aNon Na, aNum (clk t)])), IK := insert (Crypt (shrK A) \ Key Kab, Agent B, Number (clk t), Nonce Na \) (insert (Crypt (shrK B) \ Key Kab, Agent A, Number (clk t) \) (IK t)) \" \ \here we go\ have "(?s', ?t') \ R23_msgs" using H by (-) (rule R23_intros, auto) moreover have "(?s', ?t') \ R23_keys" using H by (-) (rule R23_intros, auto simp add: m2_inv3a_sesK_compr_simps m3_inv7a_sesK_compr_simps, auto simp add: R23_keys_simps) moreover have "(?s', ?t') \ R23_non" using H by (-) (rule R23_intros, auto simp add: m2_inv3a_sesK_compr_simps m3_inv7a_sesK_compr_simps R23_non_simps) moreover have "(?s', ?t') \ R23_pres" using H by (-) (rule R23_intros, auto) moreover note calculation } thus ?thesis by (auto simp add: PO_rhoare_defs R23_def m3_defs) qed lemma PO_m3_step4_refines_m2_step4: "{R23 \ UNIV \ (m3_inv1_lkeysec) } (m2_step4 Ra A B Na Kab Ts Ta), (m3_step4 Ra A B Na Kab Ts Ta) {> R23}" by (auto simp add: PO_rhoare_defs R23_def m3_defs intro!: R23_intros) (auto) lemma PO_m3_step5_refines_m2_step5: "{R23} (m2_step5 Rb A B Kab Ts Ta), (m3_step5 Rb A B Kab Ts Ta) {> R23}" by (auto simp add: PO_rhoare_defs R23_def m3_defs intro!: R23_intros) (auto) lemma PO_m3_step6_refines_m2_step6: "{R23} (m2_step6 Ra A B Na Kab Ts Ta), (m3_step6 Ra A B Na Kab Ts Ta) {> R23}" by (auto simp add: PO_rhoare_defs R23_def m3_defs intro!: R23_intros) lemma PO_m3_tick_refines_m2_tick: "{R23} (m2_tick T), (m3_tick T) {>R23}" by (auto simp add: PO_rhoare_defs R23_def m3_defs intro!: R23_intros) lemma PO_m3_purge_refines_m2_purge: "{R23} (m2_purge A), (m3_purge A) {>R23}" by (auto simp add: PO_rhoare_defs R23_def m3_defs intro!: R23_intros) text \Intruder events.\ lemma PO_m3_leak_refines_m2_leak: "{R23} (m2_leak Rs A B Na Ts), (m3_leak Rs A B Na Ts) {>R23}" by (auto simp add: PO_rhoare_defs R23_def m3_defs R23_simps intro!: R23_intros) (* also works, but requires invariants: apply (auto simp add: m2_inv3a_sesK_compr_simps m2_inv3b_sesK_compr_non_simps m3_inv7a_sesK_compr_simps m3_inv7b_sesK_compr_non_simps dest: R23_emptyD) *) lemma PO_m3_DY_fake_refines_m2_fake: "{R23 \ UNIV \ m3_inv1_lkeysec} m2_fake, m3_DY_fake {> R23}" apply (auto simp add: PO_rhoare_defs R23_def m3_defs intro!: R23_intros del: abs_msg.cases) apply (auto intro: abs_msg_DY_subset_fakeable [THEN subsetD] del: abs_msg.cases) apply (auto simp add: R23_simps) done text \All together now...\ lemmas PO_m3_trans_refines_m2_trans = PO_m3_step1_refines_m2_step1 PO_m3_step2_refines_m2_step2 PO_m3_step3_refines_m2_step3 PO_m3_step4_refines_m2_step4 PO_m3_step5_refines_m2_step5 PO_m3_step6_refines_m2_step6 PO_m3_tick_refines_m2_tick PO_m3_purge_refines_m2_purge PO_m3_leak_refines_m2_leak PO_m3_DY_fake_refines_m2_fake lemma PO_m3_refines_init_m2 [iff]: "init m3 \ R23``(init m2)" by (auto simp add: R23_def m3_defs intro!: R23_intros) lemma PO_m3_refines_trans_m2 [iff]: "{R23 \ (m2_inv3a_sesK_compr) \ (m3_inv7a_sesK_compr \ m3_inv1_lkeysec)} (trans m2), (trans m3) {> R23}" apply (auto simp add: m3_def m3_trans_def m2_def m2_trans_def) apply (blast intro!: PO_m3_trans_refines_m2_trans)+ done lemma PO_m3_observation_consistent [iff]: "obs_consistent R23 med32 m2 m3" by (auto simp add: obs_consistent_def R23_def med32_def m3_defs) text \Refinement result.\ lemma m3_refines_m2 [iff]: "refines (R23 \ (m2_inv3a_sesK_compr) \ (m3_inv1_lkeysec)) med32 m2 m3" proof - have "R23 \ m2_inv3a_sesK_compr \ UNIV \ UNIV \ m3_inv7a_sesK_compr" by (auto simp add: R23_def R23_keys_simps intro!: m3_inv7a_sesK_comprI) thus ?thesis by (-) (rule Refinement_using_invariants, auto) qed lemma m3_implements_m2 [iff]: "implements med32 m2 m3" by (rule refinement_soundness) (auto) subsection \Inherited invariants\ (******************************************************************************) subsubsection \inv3 (derived): Key secrecy for initiator\ (*invh*************************************************************************) definition m3_inv3_ikk_init :: "m3_state set" where "m3_inv3_ikk_init \ {s. \A B Ra K Ts nl. runs s Ra = Some (Init, [A, B], aKey K # aNum Ts # nl) \ A \ good \ B \ good \ Key K \ analz (IK s) \ (K, A, B, Ra$na, Ts) \ leak s }" lemmas m3_inv3_ikk_initI = m3_inv3_ikk_init_def [THEN setc_def_to_intro, rule_format] lemmas m3_inv3_ikk_initE [elim] = m3_inv3_ikk_init_def [THEN setc_def_to_elim, rule_format] lemmas m3_inv3_ikk_initD = m3_inv3_ikk_init_def [THEN setc_def_to_dest, rule_format, rotated 1] lemma PO_m3_inv3_ikk_init: "reach m3 \ m3_inv3_ikk_init" proof (rule INV_from_Refinement_using_invariants [OF m3_refines_m2]) show "Range (R23 \ m2_inv3a_sesK_compr \ m3_inv1_lkeysec \ m2_inv6_ikk_init \ UNIV) \ m3_inv3_ikk_init" by (fastforce simp add: R23_def R23_keys_simps intro!: m3_inv3_ikk_initI) qed auto subsubsection \inv4 (derived): Key secrecy for responder\ (*invh*************************************************************************) definition m3_inv4_ikk_resp :: "m3_state set" where "m3_inv4_ikk_resp \ {s. \A B Rb K Ts nl. runs s Rb = Some (Resp, [A, B], aKey K # aNum Ts # nl) \ A \ good \ B \ good \ Key K \ analz (IK s) \ (\Na. (K, A, B, Na, Ts) \ leak s) }" lemmas m3_inv4_ikk_respI = m3_inv4_ikk_resp_def [THEN setc_def_to_intro, rule_format] lemmas m3_inv4_ikk_respE [elim] = m3_inv4_ikk_resp_def [THEN setc_def_to_elim, rule_format] lemmas m3_inv4_ikk_respD = m3_inv4_ikk_resp_def [THEN setc_def_to_dest, rule_format, rotated 1] lemma PO_m3_inv4_ikk_resp: "reach m3 \ m3_inv4_ikk_resp" proof (rule INV_from_Refinement_using_invariants [OF m3_refines_m2]) show "Range (R23 \ m2_inv3a_sesK_compr \ m3_inv1_lkeysec \ m2_inv7_ikk_resp \ UNIV) \ m3_inv4_ikk_resp" by (auto simp add: R23_def R23_keys_simps intro!: m3_inv4_ikk_respI) (elim m2_inv7_ikk_respE, auto) qed auto end diff --git a/thys/Security_Protocol_Refinement/Key_establish/m3_nssk.thy b/thys/Security_Protocol_Refinement/Key_establish/m3_nssk.thy --- a/thys/Security_Protocol_Refinement/Key_establish/m3_nssk.thy +++ b/thys/Security_Protocol_Refinement/Key_establish/m3_nssk.thy @@ -1,925 +1,925 @@ (******************************************************************************* Project: Development of Security Protocols by Refinement Module: Key_establish/m3_nssk.thy (Isabelle/HOL 2016-1) ID: $Id: m3_nssk.thy 132890 2016-12-24 10:25:57Z csprenge $ Author: Christoph Sprenger, ETH Zurich Key distribution protocols Third refinement: Needham-Schroeder Shared Key protocol (NSSK) (Boyd/Mathuria, Protocol 3.19) Copyright (c) 2009-2016 Christoph Sprenger Licence: LGPL *******************************************************************************) section \Needham-Schroeder Shared Key (L3)\ theory m3_nssk imports m2_nssk "../Refinement/Message" begin text \ We model an abstract version of the Needham-Schroeder Shared Key protocol: \[ \begin{array}{lll} \mathrm{M1.} & A \rightarrow S: & A, B, Na \\ \mathrm{M2.} & S \rightarrow A: & \{Na, B, Kab, \{Kab, A\}_{Kbs} \}_{Kas} \\ \mathrm{M3.} & A \rightarrow B: & \{Kab, A\}_{Kbs} \\ \mathrm{M4.} & B \rightarrow A: & \{Nb\}_{Kab} \\ \mathrm{M5.} & A \rightarrow B: & \{Nb - 1 \}_{Kab} \\ \end{array} \] This refinement works with a single insecure channel and introduces the full Dolev-Yao intruder. \ text \Proof tool configuration. Avoid annoying automatic unfolding of \dom\.\ declare domIff [simp, iff del] (******************************************************************************) subsection \Setup\ (******************************************************************************) text \Now we can define the initial key knowledge.\ overloading ltkeySetup' \ ltkeySetup begin definition ltkeySetup_def: "ltkeySetup' \ {(sharK C, A) | C A. A = C \ A = Sv}" end lemma corrKey_shrK_bad [simp]: "corrKey = shrK`bad" by (auto simp add: keySetup_def ltkeySetup_def corrKey_def) (******************************************************************************) subsection \State\ (******************************************************************************) text \The secure channels are star-shaped to/from the server. Therefore, we have only one agent in the relation.\ record m3_state = "m1_state" + IK :: "msg set" \ \intruder knowledge\ text \Observable state: agent's local state.\ type_synonym m3_obs = m2_obs definition m3_obs :: "m3_state \ m3_obs" where "m3_obs s \ \ runs = runs s, leak = leak s \" type_synonym m3_pred = "m3_state set" type_synonym m3_trans = "(m3_state \ m3_state) set" (******************************************************************************) subsection \Events\ (******************************************************************************) text \Protocol events.\ definition \ \by @{term "A"}, refines @{term "m2_step1"}\ m3_step1 :: "[rid_t, agent, agent, nonce] \ m3_trans" where "m3_step1 Ra A B Na \ {(s, s1). \ \guards:\ Ra \ dom (runs s) \ \ \\Ra\ is fresh\ Na = Ra$na \ \ \generate nonce \Na\\ \ \actions:\ s1 = s\ runs := (runs s)(Ra \ (Init, [A, B], [])), IK := insert \Agent A, Agent B, Nonce Na\ (IK s) \ \send msg 1\ \ }" definition \ \by @{term "B"}, refines @{term "m2_step2"}\ m3_step2 :: "[rid_t, agent, agent] \ m3_trans" where "m3_step2 Rb A B \ {(s, s1). \ \guards:\ Rb \ dom (runs s) \ \ \\Rb\ is fresh\ \ \actions:\ \ \create responder thread\ s1 = s\ runs := (runs s)(Rb \ (Resp, [A, B], [])) \ }" definition \ \by @{text "Server"}, refines @{term m2_step3}\ m3_step3 :: "[rid_t, agent, agent, nonce, key] \ m3_trans" where "m3_step3 Rs A B Na Kab \ {(s, s1). \ \guards:\ Rs \ dom (runs s) \ \ \fresh server run\ Kab = sesK (Rs$sk) \ \ \fresh session key\ \Agent A, Agent B, Nonce Na\ \ IK s \ \ \recv msg 1\ \ \actions:\ \ \record session key and send messages 2 and 3\ \ \note that last field in server record is for responder nonce\ s1 = s\ runs := (runs s)(Rs \ (Serv, [A, B], [aNon Na])), IK := insert (Crypt (shrK A) \Nonce Na, Agent B, Key Kab, Crypt (shrK B) \Key Kab, Agent A\\) (IK s) \ }" definition \ \by @{term "A"}, refines @{term m2_step4}\ m3_step4 :: "[rid_t, agent, agent, nonce, key, msg] \ m3_trans" where "m3_step4 Ra A B Na Kab X \ {(s, s1). \ \guards:\ runs s Ra = Some (Init, [A, B], []) \ Na = Ra$na \ Crypt (shrK A) \Nonce Na, Agent B, Key Kab, X\ \ IK s \ \ \recv msg 2\ \ \actions:\ \ \record session key, and forward \X\\ s1 = s\ runs := (runs s)(Ra \ (Init, [A, B], [aKey Kab])), IK := insert X (IK s) \ }" definition \ \by @{term "B"}, refines @{term m2_step5}\ m3_step5 :: "[rid_t, agent, agent, nonce, key] \ m3_trans" where "m3_step5 Rb A B Nb Kab \ {(s, s1). \ \guards:\ runs s Rb = Some (Resp, [A, B], []) \ Nb = Rb$nb \ Crypt (shrK B) \Key Kab, Agent A\ \ IK s \ \ \recv msg 3\ \ \actions:\ \ \record session key\ s1 = s\ runs := (runs s)(Rb \ (Resp, [A, B], [aKey Kab])), IK := insert (Crypt Kab (Nonce Nb)) (IK s) \ }" definition \ \by @{term "A"}, refines @{term m2_step6}\ m3_step6 :: "[rid_t, agent, agent, nonce, nonce, key] \ m3_trans" where "m3_step6 Ra A B Na Nb Kab \ {(s, s'). runs s Ra = Some (Init, [A, B], [aKey Kab]) \ \ \key recv'd before\ Na = Ra$na \ Crypt Kab (Nonce Nb) \ IK s \ \ \receive \M4\\ \ \actions:\ s' = s\ runs := (runs s)(Ra \ (Init, [A, B], [aKey Kab, aNon Nb])), IK := insert (Crypt Kab \Nonce Nb, Nonce Nb\) (IK s) \ }" definition \ \by @{term "B"}, refines @{term m2_step6}\ m3_step7 :: "[rid_t, agent, agent, nonce, key] \ m3_trans" where "m3_step7 Rb A B Nb Kab \ {(s, s'). runs s Rb = Some (Resp, [A, B], [aKey Kab]) \ \ \key recv'd before\ Nb = Rb$nb \ Crypt Kab \Nonce Nb, Nonce Nb\ \ IK s \ \ \receive \M5\\ \ \actions: (redundant) update local state marks successful termination\ s' = s\ runs := (runs s)(Rb \ (Resp, [A, B], [aKey Kab, END])) \ }" text \Session key compromise.\ definition \ \refines @{term m2_leak}\ m3_leak :: "[rid_t, rid_t, rid_t, agent, agent] \ m3_trans" where "m3_leak Rs Ra Rb A B \ {(s, s1). \ \guards:\ runs s Rs = Some (Serv, [A, B], [aNon (Ra$na)]) \ runs s Ra = Some (Init, [A, B], [aKey (sesK (Rs$sk)), aNon (Rb$nb)]) \ runs s Rb = Some (Resp, [A, B], [aKey (sesK (Rs$sk)), END]) \ \ \actions:\ \ \record session key as leaked and add it to intruder knowledge\ s1 = s\ leak := insert (sesK (Rs$sk), Ra$na, Rb$nb) (leak s), IK := insert (Key (sesK (Rs$sk))) (IK s) \ }" text \Intruder fake event.\ definition \ \refines @{term "m2_fake"}\ m3_DY_fake :: "m3_trans" where "m3_DY_fake \ {(s, s1). \ \actions:\ s1 = s(| IK := synth (analz (IK s)) |) }" (******************************************************************************) subsection \Transition system\ (******************************************************************************) definition m3_init :: "m3_state set" where "m3_init \ { \ runs = Map.empty, leak = shrK`bad \ {undefined} \ {undefined}, IK = Key`shrK`bad \ }" definition m3_trans :: "(m3_state \ m3_state) set" where "m3_trans \ (\Ra Rb Rs A B Na Nb Kab X. m3_step1 Ra A B Na \ m3_step2 Rb A B \ m3_step3 Rs A B Na Kab \ m3_step4 Ra A B Na Kab X \ m3_step5 Rb A B Nb Kab \ m3_step6 Ra A B Na Nb Kab \ m3_step7 Rb A B Nb Kab \ m3_leak Rs Ra Rb A B \ m3_DY_fake \ Id )" definition m3 :: "(m3_state, m3_obs) spec" where "m3 \ \ init = m3_init, trans = m3_trans, obs = m3_obs \" lemmas m3_defs = m3_def m3_init_def m3_trans_def m3_obs_def m3_step1_def m3_step2_def m3_step3_def m3_step4_def m3_step5_def m3_step6_def m3_step7_def m3_leak_def m3_DY_fake_def (******************************************************************************) subsection \Invariants\ (******************************************************************************) text \Specialized injection that we can apply more aggressively.\ lemmas analz_Inj_IK = analz.Inj [where H="IK s" for s] lemmas parts_Inj_IK = parts.Inj [where H="IK s" for s] declare parts_Inj_IK [dest!] declare analz_into_parts [dest] subsubsection \inv1: Secrecy of pre-distributed shared keys\ (*inv**************************************************************************) text \inv1: Secrecy of long-term keys\ definition m3_inv1_lkeysec :: "m3_state set" where "m3_inv1_lkeysec \ {s. \C. (Key (shrK C) \ parts (IK s) \ C \ bad) \ (C \ bad \ Key (shrK C) \ IK s) }" lemmas m3_inv1_lkeysecI = m3_inv1_lkeysec_def [THEN setc_def_to_intro, rule_format] lemmas m3_inv1_lkeysecE [elim] = m3_inv1_lkeysec_def [THEN setc_def_to_elim, rule_format] lemmas m3_inv1_lkeysecD = m3_inv1_lkeysec_def [THEN setc_def_to_dest, rule_format] text \Invariance proof.\ lemma PO_m3_inv1_lkeysec_init [iff]: "init m3 \ m3_inv1_lkeysec" by (auto simp add: m3_defs m3_inv1_lkeysec_def) lemma PO_m3_inv1_lkeysec_trans [iff]: "{m3_inv1_lkeysec} trans m3 {> m3_inv1_lkeysec}" by (fastforce simp add: PO_hoare_defs m3_defs intro!: m3_inv1_lkeysecI dest: Body) lemma PO_m3_inv1_lkeysec [iff]: "reach m3 \ m3_inv1_lkeysec" by (rule inv_rule_incr) (fast+) text \Useful simplifier lemmas\ lemma m3_inv1_lkeysec_for_parts [simp]: "\ s \ m3_inv1_lkeysec \ \ Key (shrK C) \ parts (IK s) \ C \ bad" by auto lemma m3_inv1_lkeysec_for_analz [simp]: "\ s \ m3_inv1_lkeysec \ \ Key (shrK C) \ analz (IK s) \ C \ bad" by auto subsubsection \inv2: Ticket shape for honestly encrypted M2\ (*inv**************************************************************************) definition m3_inv2_ticket :: "m3_state set" where "m3_inv2_ticket \ {s. \A B N K X. A \ bad \ Crypt (shrK A) \Nonce N, Agent B, Key K, X\ \ parts (IK s) \ X = Crypt (shrK B) \Key K, Agent A\ \ K \ range sesK }" lemmas m3_inv2_ticketI = m3_inv2_ticket_def [THEN setc_def_to_intro, rule_format] lemmas m3_inv2_ticketE [elim] = m3_inv2_ticket_def [THEN setc_def_to_elim, rule_format] lemmas m3_inv2_ticketD = m3_inv2_ticket_def [THEN setc_def_to_dest, rule_format, rotated -1] text \Invariance proof.\ lemma PO_m3_inv2_ticket_init [iff]: "init m3 \ m3_inv2_ticket" by (auto simp add: m3_defs intro!: m3_inv2_ticketI) lemma PO_m3_inv2_ticket_trans [iff]: "{m3_inv2_ticket \ m3_inv1_lkeysec} trans m3 {> m3_inv2_ticket}" apply (auto simp add: PO_hoare_defs m3_defs intro!: m3_inv2_ticketI) apply (auto dest: m3_inv2_ticketD) \ \2 subgoals, from step4 [?]\ apply (drule Body [where H="IK s" for s], drule parts_cut, auto dest: m3_inv2_ticketD)+ done lemma PO_m3_inv2_ticket [iff]: "reach m3 \ m3_inv2_ticket" by (rule inv_rule_incr) (auto del: subsetI) subsubsection \inv3: Session keys not used to encrypt other session keys\ (*inv**************************************************************************) text \Session keys are not used to encrypt other keys. Proof requires generalization to sets of session keys. NOTE: For NSSK, this invariant cannot be inherited from the corresponding L2 invariant. The simulation relation is only an implication not an equivalence. \ definition m3_inv3_sesK_compr :: "m3_state set" where "m3_inv3_sesK_compr \ {s. \K KK. KK \ range sesK \ (Key K \ analz (Key`KK \ (IK s))) = (K \ KK \ Key K \ analz (IK s)) }" lemmas m3_inv3_sesK_comprI = m3_inv3_sesK_compr_def [THEN setc_def_to_intro, rule_format] lemmas m3_inv3_sesK_comprE = m3_inv3_sesK_compr_def [THEN setc_def_to_elim, rule_format] lemmas m3_inv3_sesK_comprD = m3_inv3_sesK_compr_def [THEN setc_def_to_dest, rule_format] text \Additional lemma\ lemmas insert_commute_Key = insert_commute [where x="Key K" for K] lemmas m3_inv3_sesK_compr_simps = m3_inv3_sesK_comprD m3_inv3_sesK_comprD [where KK="{Kab}" for Kab, simplified] m3_inv3_sesK_comprD [where KK="insert Kab KK" for Kab KK, simplified] insert_commute_Key text \Invariance proof.\ lemma PO_m3_inv3_sesK_compr_step4: "{m3_inv3_sesK_compr \ m3_inv2_ticket \ m3_inv1_lkeysec} m3_step4 Ra A B Na Kab X {> m3_inv3_sesK_compr}" proof - { fix K KK s assume H: "s \ m3_inv1_lkeysec" "s \ m3_inv3_sesK_compr" "s \ m3_inv2_ticket" "runs s Ra = Some (Init, [A, B], [])" "Na = Ra$na" "KK \ range sesK" "Crypt (shrK A) \Nonce Na, Agent B, Key Kab, X\ \ analz (IK s)" have "(Key K \ analz (insert X (Key ` KK \ IK s))) = (K \ KK \ Key K \ analz (insert X (IK s)))" proof (cases "A \ bad") case True with H have "X \ analz (IK s)" by (auto dest!: Decrypt) moreover with H have "X \ analz (Key ` KK \ IK s)" by (auto intro: analz_monotonic) ultimately show ?thesis using H by (auto simp add: m3_inv3_sesK_compr_simps analz_insert_eq) next case False thus ?thesis using H by (fastforce simp add: m3_inv3_sesK_compr_simps dest!: m3_inv2_ticketD [OF analz_into_parts]) qed } thus ?thesis by (auto simp add: PO_hoare_defs m3_defs intro!: m3_inv3_sesK_comprI dest!: analz_Inj_IK) qed text \All together now.\ lemmas PO_m3_inv3_sesK_compr_trans_lemmas = PO_m3_inv3_sesK_compr_step4 lemma PO_m3_inv3_sesK_compr_init [iff]: "init m3 \ m3_inv3_sesK_compr" by (auto simp add: m3_defs intro!: m3_inv3_sesK_comprI) lemma PO_m3_inv3_sesK_compr_trans [iff]: "{m3_inv3_sesK_compr \ m3_inv2_ticket \ m3_inv1_lkeysec} trans m3 {> m3_inv3_sesK_compr}" by (auto simp add: m3_def m3_trans_def intro!: PO_m3_inv3_sesK_compr_trans_lemmas) (auto simp add: PO_hoare_defs m3_defs m3_inv3_sesK_compr_simps intro!: m3_inv3_sesK_comprI) lemma PO_m3_inv3_sesK_compr [iff]: "reach m3 \ m3_inv3_sesK_compr" by (rule_tac J="m3_inv2_ticket \ m3_inv1_lkeysec" in inv_rule_incr) (auto) (******************************************************************************) subsection \Refinement\ (******************************************************************************) subsubsection \Message abstraction and simulation relation\ (******************************************************************************) text \Abstraction function on sets of messages.\ inductive_set abs_msg :: "msg set \ chmsg set" for H :: "msg set" where am_M1: "\Agent A, Agent B, Nonce Na\ \ H \ Insec A B (Msg [aNon Na]) \ abs_msg H" | am_M2: "Crypt (shrK C) \Nonce N, Agent B, Key K, X\ \ H \ Secure Sv C (Msg [aNon N, aAgt B, aKey K]) \ abs_msg H" | am_M3: "Crypt (shrK C) \Key K, Agent A\ \ H \ Secure Sv C (Msg [aKey K, aAgt A]) \ abs_msg H" | am_M4: "Crypt K (Nonce N) \ H \ dAuth K (Msg [aNon N]) \ abs_msg H" | am_M5: "Crypt K \Nonce N, Nonce N'\ \ H \ dAuth K (Msg [aNon N, aNon N']) \ abs_msg H" text \R23: The simulation relation. This is a data refinement of the insecure and secure channels of refinement 2.\ definition R23_msgs :: "(m2_state \ m3_state) set" where "R23_msgs \ {(s, t). abs_msg (parts (IK t)) \ chan s}" definition R23_keys :: "(m2_state \ m3_state) set" where \ \only an implication!\ "R23_keys \ {(s, t). \KK K. KK \ range sesK \ Key K \ analz (Key`KK \ IK t) \ aKey K \ extr (aKey`KK \ ik0) (chan s) }" definition R23_non :: "(m2_state \ m3_state) set" where \ \only an implication!\ "R23_non \ {(s, t). \KK N. KK \ range sesK \ Nonce N \ analz (Key`KK \ IK t) \ aNon N \ extr (aKey`KK \ ik0) (chan s) }" definition R23_pres :: "(m2_state \ m3_state) set" where "R23_pres \ {(s, t). runs s = runs t \ leak s = leak t}" definition R23 :: "(m2_state \ m3_state) set" where "R23 \ R23_msgs \ R23_keys \ R23_non \ R23_pres" lemmas R23_defs = R23_def R23_msgs_def R23_keys_def R23_non_def R23_pres_def text \The mediator function is the identity here.\ definition med32 :: "m3_obs \ m2_obs" where "med32 \ id" lemmas R23_msgsI = R23_msgs_def [THEN rel_def_to_intro, simplified, rule_format] lemmas R23_msgsE [elim] = R23_msgs_def [THEN rel_def_to_elim, simplified, rule_format] lemmas R23_msgsE' [elim] = R23_msgs_def [THEN rel_def_to_dest, simplified, rule_format, THEN subsetD] lemmas R23_keysI = R23_keys_def [THEN rel_def_to_intro, simplified, rule_format] lemmas R23_keysE [elim] = R23_keys_def [THEN rel_def_to_elim, simplified, rule_format] lemmas R23_keysD = R23_keys_def [THEN rel_def_to_dest, simplified, rule_format, rotated 2] lemmas R23_nonI = R23_non_def [THEN rel_def_to_intro, simplified, rule_format] lemmas R23_nonE [elim] = R23_non_def [THEN rel_def_to_elim, simplified, rule_format] lemmas R23_nonD = R23_non_def [THEN rel_def_to_dest, simplified, rule_format, rotated 2] lemmas R23_presI = R23_pres_def [THEN rel_def_to_intro, simplified, rule_format] lemmas R23_presE [elim] = R23_pres_def [THEN rel_def_to_elim, simplified, rule_format] lemmas R23_intros = R23_msgsI R23_keysI R23_nonI R23_presI text \Further lemmas: general lemma for simplifier and different instantiations.\ lemmas R23_keys_dests = R23_keysD R23_keysD [where KK="{}", simplified] R23_keysD [where KK="{K}" for K, simplified] R23_keysD [where KK="insert K KK" for K KK, simplified, OF _ _ conjI] lemmas R23_non_dests = R23_nonD R23_nonD [where KK="{}", simplified] R23_nonD [where KK="{K}" for K, simplified] R23_nonD [where KK="insert K KK" for K KK, simplified, OF _ _ conjI] lemmas R23_dests = R23_keys_dests R23_non_dests (* lemmas insert_commute_Key = insert_commute [where x="Key K" for K] *) (* lemmas R23_keys_emptyD = R23_keysD [where KK="{}", simplified, THEN iffD1, rotated 1] R23_keysD [where KK="{}", simplified, THEN iffD2, rotated 1] *) subsubsection \General lemmas\ (******************************************************************************) text \General facts about @{term "abs_msg"}\ declare abs_msg.intros [intro!] declare abs_msg.cases [elim!] lemma abs_msg_empty: "abs_msg {} = {}" by (auto) lemma abs_msg_Un [simp]: "abs_msg (G \ H) = abs_msg G \ abs_msg H" by (auto) lemma abs_msg_mono [elim]: "\ m \ abs_msg G; G \ H \ \ m \ abs_msg H" by (auto) lemma abs_msg_insert_mono [intro]: "\ m \ abs_msg H \ \ m \ abs_msg (insert m' H)" by (auto) text \Facts about @{term "abs_msg"} concerning abstraction of fakeable messages. This is crucial for proving the refinement of the intruder event.\ lemma abs_msg_DY_subset_fakeable: "\ (s, t) \ R23_msgs; (s, t) \ R23_keys; (s, t) \ R23_non; t \ m3_inv1_lkeysec \ \ abs_msg (synth (analz (IK t))) \ fake ik0 (dom (runs s)) (chan s)" apply (auto) \ \9 subgoals, deal with replays first\ prefer 2 apply (blast) prefer 3 apply (blast) prefer 4 apply (blast) prefer 5 apply (blast) \ \remaining 5 subgoals are real fakes\ apply (intro fake_StatCh fake_DynCh, auto dest: R23_dests)+ done subsubsection \Refinement proof\ (******************************************************************************) text \Pair decomposition. These were set to \texttt{elim!}, which is too agressive here.\ declare MPair_analz [rule del, elim] declare MPair_parts [rule del, elim] text \Protocol events.\ lemma PO_m3_step1_refines_m2_step1: "{R23} (m2_step1 Ra A B Na), (m3_step1 Ra A B Na) {> R23}" by (auto simp add: PO_rhoare_defs R23_def m2_defs m3_defs intro!: R23_intros) (auto) lemma PO_m3_step2_refines_m2_step2: "{R23} (m2_step2 Rb A B), (m3_step2 Rb A B) {> R23}" by (auto simp add: PO_rhoare_defs R23_def m2_defs m3_defs intro!: R23_intros) lemma PO_m3_step3_refines_m2_step3: "{R23 \ (m2_inv3a_sesK_compr) \ (m3_inv3_sesK_compr \ m3_inv1_lkeysec)} (m2_step3 Rs A B Na Kab), (m3_step3 Rs A B Na Kab) {> R23}" proof - { fix s t assume H: "(s, t) \ R23_msgs" "(s, t) \ R23_keys" "(s, t) \ R23_non" "(s, t) \ R23_pres" "s \ m2_inv3a_sesK_compr" "t \ m3_inv3_sesK_compr" "t \ m3_inv1_lkeysec" "Kab = sesK (Rs$sk)" "Rs \ dom (runs t)" "\ Agent A, Agent B, Nonce Na \ \ parts (IK t)" let ?s'= - "s\ runs := runs s(Rs \ (Serv, [A, B], [aNon Na])), + "s\ runs := (runs s)(Rs \ (Serv, [A, B], [aNon Na])), chan := insert (Secure Sv A (Msg [aNon Na, aAgt B, aKey Kab])) (insert (Secure Sv B (Msg [aKey Kab, aAgt A])) (chan s)) \" let ?t'= - "t\ runs := runs t(Rs \ (Serv, [A, B], [aNon Na])), + "t\ runs := (runs t)(Rs \ (Serv, [A, B], [aNon Na])), IK := insert (Crypt (shrK A) \Nonce Na, Agent B, Key Kab, Crypt (shrK B) \Key Kab, Agent A\\) (IK t) \" have "(?s', ?t') \ R23_msgs" using H by (-) (rule R23_intros, auto) moreover have "(?s', ?t') \ R23_keys" using H by (-) (rule R23_intros, auto simp add: m2_inv3a_sesK_compr_simps m3_inv3_sesK_compr_simps dest: R23_keys_dests) moreover have "(?s', ?t') \ R23_non" using H by (-) (rule R23_intros, auto simp add: m2_inv3a_sesK_compr_simps m3_inv3_sesK_compr_simps dest: R23_non_dests) moreover have "(?s', ?t') \ R23_pres" using H by (-) (rule R23_intros, auto) moreover note calculation } thus ?thesis by (auto simp add: PO_rhoare_defs R23_def m2_defs m3_defs) qed (* with equality in R23_keys [OLD, before adding session key compromise]: goal (1 subgoal): 1. \s t. \t \ m3_inv3_keyuse; t \ m3_inv1_lkeysec; t \ m3_inv5_badkeys; Kab \ dom (runs t); Kab \ range sesK; \Agent A, Agent B, Nonce Na\ \ parts (IK t); Kab \ keysFor (analz (IK t)); Key (shrK A) \ analz (IK t); (s, t) \ R23_msgs; (s, t) \ R23_keys; (s, t) \ R23_pres; Kab \ shrK A; Key (shrK B) \ analz (IK t); Kab \ shrK B; A \ bad; B \ bad\ \ Key Kab \ analz (IK t) This does NOT hold, since Kab is revealed in abstract model when A \ bad; B \ bad, but not in concrete one, since here it is protected by A's encryption. *) lemma PO_m3_step4_refines_m2_step4: "{R23 \ (m2_inv3b_sesK_compr_non) \ (m3_inv3_sesK_compr \ m3_inv2_ticket \ m3_inv1_lkeysec)} (m2_step4 Ra A B Na Kab), (m3_step4 Ra A B Na Kab X) {> R23}" proof - { fix s t assume H: "(s, t) \ R23_msgs" "(s, t) \ R23_keys" "(s, t) \ R23_non" "(s, t) \ R23_pres" "s \ m2_inv3b_sesK_compr_non" "t \ m3_inv3_sesK_compr" "t \ m3_inv2_ticket" "t \ m3_inv1_lkeysec" "runs t Ra = Some (Init, [A, B], [])" "Na = Ra$na" "Crypt (shrK A) \Nonce Na, Agent B, Key Kab, X\ \ analz (IK t)" - let ?s' = "s\ runs := runs s(Ra \ (Init, [A, B], [aKey Kab])) \" - and ?t' = "t\ runs := runs t(Ra \ (Init, [A, B], [aKey Kab])), + let ?s' = "s\ runs := (runs s)(Ra \ (Init, [A, B], [aKey Kab])) \" + and ?t' = "t\ runs := (runs t)(Ra \ (Init, [A, B], [aKey Kab])), IK := insert X (IK t) \" from H have "Secure Sv A (Msg [aNon Na, aAgt B, aKey Kab]) \ chan s" by auto moreover have "X \ parts (IK t)" using H by (auto dest!: Body MPair_parts) hence "(?s', ?t') \ R23_msgs" using H by (auto intro!: R23_intros) moreover have "(?s', ?t') \ R23_keys" proof (cases) assume "A \ bad" with H have "X \ analz (IK t)" by (-) (drule Decrypt, auto) with H show ?thesis by (-) (rule R23_intros, auto dest!: analz_cut intro: analz_monotonic) next assume "A \ bad" show ?thesis proof - note H moreover with \A \ bad\ have "X = Crypt (shrK B) \Key Kab, Agent A\ \ Kab \ range sesK" by (auto dest!: m3_inv2_ticketD) moreover { assume H1: "Key (shrK B) \ analz (IK t)" have "aKey Kab \ extr ik0 (chan s)" proof - note calculation moreover hence "Secure Sv B (Msg [aKey Kab, aAgt A]) \ chan s" by (-) (drule analz_into_parts, drule Body, elim MPair_parts, auto) ultimately show ?thesis using H1 by auto qed } ultimately show ?thesis by (-) (rule R23_intros, auto simp add: m3_inv3_sesK_compr_simps) qed qed moreover have "(?s', ?t') \ R23_non" proof (cases) assume "A \ bad" hence "X \ analz (IK t)" using H by (-) (drule Decrypt, auto) thus ?thesis using H by (-) (rule R23_intros, auto dest!: analz_cut intro: analz_monotonic) next assume "A \ bad" hence "X = Crypt (shrK B) \Key Kab, Agent A \ \ Kab \ range sesK" using H by (auto dest!: m3_inv2_ticketD) thus ?thesis using H by (-) (rule R23_intros, auto simp add: m2_inv3b_sesK_compr_non_simps m3_inv3_sesK_compr_simps dest: R23_non_dests) qed moreover have "(?s', ?t') \ R23_pres" using H by (auto intro!: R23_intros) moreover note calculation } thus ?thesis by (auto simp add: PO_rhoare_defs R23_def m2_defs m3_defs dest!: analz_Inj_IK) qed lemma PO_m3_step5_refines_m2_step5: "{R23} (m2_step5 Rb A B Nb Kab), (m3_step5 Rb A B Nb Kab) {> R23}" by (auto simp add: PO_rhoare_defs R23_def m2_defs m3_defs intro!: R23_intros) (auto) lemma PO_m3_step6_refines_m2_step6: "{R23} (m2_step6 Ra A B Na Nb Kab), (m3_step6 Ra A B Na Nb Kab) {> R23}" by (auto simp add: PO_rhoare_defs R23_def m2_defs m3_defs intro!: R23_intros) (auto) lemma PO_m3_step7_refines_m2_step7: "{R23} (m2_step7 Rb A B Nb Kab), (m3_step7 Rb A B Nb Kab) {> R23}" by (auto simp add: PO_rhoare_defs R23_def m2_defs m3_defs intro!: R23_intros) text \Intruder events.\ lemma PO_m3_leak_refines_m2_leak: "{R23} m2_leak Rs Ra Rb A B, m3_leak Rs Ra Rb A B {> R23}" by (auto simp add: PO_rhoare_defs R23_def m2_defs m3_defs intro!: R23_intros) (auto dest: R23_dests) lemma PO_m3_DY_fake_refines_m2_fake: "{R23 \ UNIV \ m3_inv1_lkeysec} m2_fake, m3_DY_fake {> R23}" apply (auto simp add: PO_rhoare_defs R23_def m2_defs m3_defs intro!: R23_intros del: abs_msg.cases) apply (auto intro: abs_msg_DY_subset_fakeable [THEN subsetD] del: abs_msg.cases) apply (auto dest: R23_dests) done text \All together now...\ lemmas PO_m3_trans_refines_m2_trans = PO_m3_step1_refines_m2_step1 PO_m3_step2_refines_m2_step2 PO_m3_step3_refines_m2_step3 PO_m3_step4_refines_m2_step4 PO_m3_step5_refines_m2_step5 PO_m3_step6_refines_m2_step6 PO_m3_step7_refines_m2_step7 PO_m3_leak_refines_m2_leak PO_m3_DY_fake_refines_m2_fake lemma PO_m3_refines_init_m2 [iff]: "init m3 \ R23``(init m2)" by (auto simp add: R23_def m2_defs m3_defs intro!: R23_intros) lemma PO_m3_refines_trans_m2 [iff]: "{R23 \ (m2_inv3a_sesK_compr \ m2_inv3b_sesK_compr_non) \ (m3_inv3_sesK_compr \ m3_inv2_ticket \ m3_inv1_lkeysec)} (trans m2), (trans m3) {> R23}" apply (auto simp add: m3_def m3_trans_def m2_def m2_trans_def) apply (blast intro!: PO_m3_trans_refines_m2_trans)+ done lemma PO_m3_observation_consistent [iff]: "obs_consistent R23 med32 m2 m3" by (auto simp add: obs_consistent_def R23_def med32_def m2_defs m3_defs) text \Refinement result.\ lemma m3_refines_m2 [iff]: "refines (R23 \ (m2_inv3a_sesK_compr \ m2_inv3b_sesK_compr_non) \ (m3_inv3_sesK_compr \ m3_inv2_ticket \ m3_inv1_lkeysec)) med32 m2 m3" by (rule Refinement_using_invariants) (auto) lemma m3_implements_m2 [iff]: "implements med32 m2 m3" by (rule refinement_soundness) (auto) subsection \Inherited invariants\ (******************************************************************************) subsubsection \inv4 (derived): Key secrecy for initiator\ (*invh*************************************************************************) definition m3_inv4_ikk_init :: "m3_state set" where "m3_inv4_ikk_init \ {s. \Ra K A B al. runs s Ra = Some (Init, [A, B], aKey K # al) \ A \ good \ B \ good \ Key K \ analz (IK s) \ (\Nb'. (K, Ra $ na, Nb') \ leak s) }" lemmas m3_inv4_ikk_initI = m3_inv4_ikk_init_def [THEN setc_def_to_intro, rule_format] lemmas m3_inv4_ikk_initE [elim] = m3_inv4_ikk_init_def [THEN setc_def_to_elim, rule_format] lemmas m3_inv4_ikk_initD = m3_inv4_ikk_init_def [THEN setc_def_to_dest, rule_format, rotated 1] lemma PO_m3_inv4_ikk_init: "reach m3 \ m3_inv4_ikk_init" proof (rule INV_from_Refinement_using_invariants [OF m3_refines_m2]) show "Range (R23 \ (m2_inv3a_sesK_compr \ m2_inv3b_sesK_compr_non) \ (m3_inv3_sesK_compr \ m3_inv2_ticket \ m3_inv1_lkeysec) \ m2_inv6_ikk_init \ UNIV) \ m3_inv4_ikk_init" by (auto simp add: R23_def R23_pres_def intro!: m3_inv4_ikk_initI) (elim m2_inv6_ikk_initE, auto dest: R23_keys_dests) qed auto subsubsection \inv5 (derived): Key secrecy for responder\ (*invh*************************************************************************) definition m3_inv5_ikk_resp :: "m3_state set" where "m3_inv5_ikk_resp \ {s. \Rb K A B al. runs s Rb = Some (Resp, [A, B], aKey K # al) \ A \ good \ B \ good \ Key K \ analz (IK s) \ K \ Domain (leak s) }" lemmas m3_inv5_ikk_respI = m3_inv5_ikk_resp_def [THEN setc_def_to_intro, rule_format] lemmas m3_inv5_ikk_respE [elim] = m3_inv5_ikk_resp_def [THEN setc_def_to_elim, rule_format] lemmas m3_inv5_ikk_respD = m3_inv5_ikk_resp_def [THEN setc_def_to_dest, rule_format, rotated 1] lemma PO_m3_inv4_ikk_resp: "reach m3 \ m3_inv5_ikk_resp" proof (rule INV_from_Refinement_using_invariants [OF m3_refines_m2]) show "Range (R23 \ (m2_inv3a_sesK_compr \ m2_inv3b_sesK_compr_non) \ (m3_inv3_sesK_compr \ m3_inv2_ticket \ m3_inv1_lkeysec) \ m2_inv7_ikk_resp \ UNIV) \ m3_inv5_ikk_resp" by (auto simp add: R23_def R23_pres_def intro!: m3_inv5_ikk_respI) (elim m2_inv7_ikk_respE, auto dest: R23_keys_dests) qed auto end diff --git a/thys/Security_Protocol_Refinement/Key_establish/m3_nssk_par.thy b/thys/Security_Protocol_Refinement/Key_establish/m3_nssk_par.thy --- a/thys/Security_Protocol_Refinement/Key_establish/m3_nssk_par.thy +++ b/thys/Security_Protocol_Refinement/Key_establish/m3_nssk_par.thy @@ -1,754 +1,754 @@ (******************************************************************************* Project: Development of Security Protocols by Refinement Module: Key_establish/m3_nssk_par.thy (Isabelle/HOL 2016-1) ID: $Id: m3_nssk_par.thy 132890 2016-12-24 10:25:57Z csprenge $ Author: Christoph Sprenger, ETH Zurich Key distribution protocols Third refinement: parallel version of Needham-Schroeder Shared Key protocol (NSSK, Boyd/Mathuria, Protocol 3.19) Copyright (c) 2009-2016 Christoph Sprenger Licence: LGPL *******************************************************************************) section \Needham-Schroeder Shared Key, "parallel" variant (L3)\ theory m3_nssk_par imports m2_nssk "../Refinement/Message" begin text \ We model an abstract version of the Needham-Schroeder Shared Key protocol: \[ \begin{array}{lll} \mathrm{M1.} & A \rightarrow S: & A, B, Na \\ \mathrm{M2.} & S \rightarrow A: & \{Na, B, Kab, \{Kab, A\}_{Kbs} \}_{Kas} \\ \mathrm{M3.} & A \rightarrow B: & \{Kab, A\}_{Kbs} \\ \mathrm{M4.} & B \rightarrow A: & \{Nb\}_{Kab} \\ \mathrm{M5.} & A \rightarrow B: & \{Nb - 1 \}_{Kab} \\ \end{array} \] We model a "parallel" version of the NSSK protocol: \[ \begin{array}{lll} \mathrm{M1.} & A \rightarrow S: & A, B, Na \\ \mathrm{M2.} & S \rightarrow A: & \{Na, B, Kab\}_{Kas} \\ \mathrm{M3.} & S \rightarrow B: & \{Kab, A\}_{Kbs} \\ \mathrm{M4.} & B \rightarrow A: & \{Nb\}_{Kab} \\ \mathrm{M5.} & A \rightarrow B: & \{Nb - 1 \}_{Kab} \\ \end{array} \] \ text \Proof tool configuration. Avoid annoying automatic unfolding of \dom\.\ declare domIff [simp, iff del] (******************************************************************************) subsection \Setup\ (******************************************************************************) text \Now we can define the initial key knowledge.\ overloading ltkeySetup' \ ltkeySetup begin definition ltkeySetup_def: "ltkeySetup' \ {(sharK C, A) | C A. A = C \ A = Sv}" end lemma corrKey_shrK_bad [simp]: "corrKey = shrK`bad" by (auto simp add: keySetup_def ltkeySetup_def corrKey_def) (******************************************************************************) subsection \State\ (******************************************************************************) text \The secure channels are star-shaped to/from the server. Therefore, we have only one agent in the relation.\ record m3_state = "m1_state" + IK :: "msg set" \ \intruder knowledge\ text \Observable state: agent's local state.\ type_synonym m3_obs = m2_obs definition m3_obs :: "m3_state \ m3_obs" where "m3_obs s \ \ runs = runs s, leak = leak s \" type_synonym m3_pred = "m3_state set" type_synonym m3_trans = "(m3_state \ m3_state) set" (******************************************************************************) subsection \Events\ (******************************************************************************) text \Protocol events.\ definition \ \by @{term "A"}, refines @{term "m2_step1"}\ m3_step1 :: "[rid_t, agent, agent, nonce] \ m3_trans" where "m3_step1 Ra A B Na \ {(s, s1). \ \guards:\ Ra \ dom (runs s) \ \ \\Ra\ is fresh\ Na = Ra$na \ \ \generate nonce \Na\\ \ \actions:\ s1 = s\ runs := (runs s)(Ra \ (Init, [A, B], [])), IK := insert \Agent A, Agent B, Nonce Na\ (IK s) \ \send msg 1\ \ }" definition \ \by @{term "B"}, refines @{term "m2_step2"}\ m3_step2 :: "[rid_t, agent, agent] \ m3_trans" where "m3_step2 Rb A B \ {(s, s1). \ \guards:\ Rb \ dom (runs s) \ \ \\Rb\ is fresh\ \ \actions:\ \ \create responder thread\ s1 = s\ runs := (runs s)(Rb \ (Resp, [A, B], [])) \ }" definition \ \by @{text "Server"}, refines @{term m2_step3}\ m3_step3 :: "[rid_t, agent, agent, nonce, key] \ m3_trans" where "m3_step3 Rs A B Na Kab \ {(s, s1). \ \guards:\ Rs \ dom (runs s) \ \ \fresh server run\ Kab = sesK (Rs$sk) \ \ \fresh session key\ \Agent A, Agent B, Nonce Na\ \ IK s \ \ \recv msg 1\ \ \actions:\ \ \record session key and send messages 2 and 3\ \ \note that last field in server record is for responder nonce\ s1 = s\ runs := (runs s)(Rs \ (Serv, [A, B], [aNon Na])), IK := {Crypt (shrK A) \Nonce Na, Agent B, Key Kab\, Crypt (shrK B) \Key Kab, Agent A\} \ IK s \ }" definition \ \by @{term "A"}, refines @{term m2_step4}\ m3_step4 :: "[rid_t, agent, agent, nonce, key] \ m3_trans" where "m3_step4 Ra A B Na Kab \ {(s, s1). \ \guards:\ runs s Ra = Some (Init, [A, B], []) \ Na = Ra$na \ Crypt (shrK A) \Nonce Na, Agent B, Key Kab\ \ IK s \ \ \recv msg 2\ \ \actions:\ \ \record session key\ s1 = s\ runs := (runs s)(Ra \ (Init, [A, B], [aKey Kab])) \ }" definition \ \by @{term "B"}, refines @{term m2_step5}\ m3_step5 :: "[rid_t, agent, agent, nonce, key] \ m3_trans" where "m3_step5 Rb A B Nb Kab \ {(s, s1). \ \guards:\ runs s Rb = Some (Resp, [A, B], []) \ Nb = Rb$nb \ Crypt (shrK B) \Key Kab, Agent A\ \ IK s \ \ \recv msg 3\ \ \actions:\ \ \record session key\ s1 = s\ runs := (runs s)(Rb \ (Resp, [A, B], [aKey Kab])), IK := insert (Crypt Kab (Nonce Nb)) (IK s) \ }" definition \ \by @{term "A"}, refines @{term m2_step6}\ m3_step6 :: "[rid_t, agent, agent, nonce, nonce, key] \ m3_trans" where "m3_step6 Ra A B Na Nb Kab \ {(s, s'). runs s Ra = Some (Init, [A, B], [aKey Kab]) \ \ \key recv'd before\ Na = Ra$na \ Crypt Kab (Nonce Nb) \ IK s \ \ \receive \M4\\ \ \actions:\ s' = s\ runs := (runs s)(Ra \ (Init, [A, B], [aKey Kab, aNon Nb])), IK := insert (Crypt Kab \Nonce Nb, Nonce Nb\) (IK s) \ }" definition \ \by @{term "B"}, refines @{term m2_step6}\ m3_step7 :: "[rid_t, agent, agent, nonce, key] \ m3_trans" where "m3_step7 Rb A B Nb Kab \ {(s, s'). runs s Rb = Some (Resp, [A, B], [aKey Kab]) \ \ \key recv'd before\ Nb = Rb$nb \ Crypt Kab \Nonce Nb, Nonce Nb\ \ IK s \ \ \receive \M5\\ \ \actions: (redundant) update local state marks successful termination\ s' = s\ runs := (runs s)(Rb \ (Resp, [A, B], [aKey Kab, END])) \ }" text \Session key compromise.\ definition \ \refines @{term m2_leak}\ m3_leak :: "[rid_t, rid_t, rid_t, agent, agent] \ m3_trans" where "m3_leak Rs Ra Rb A B \ {(s, s1). \ \guards:\ runs s Rs = Some (Serv, [A, B], [aNon (Ra$na)]) \ runs s Ra = Some (Init, [A, B], [aKey (sesK (Rs$sk)), aNon (Rb$nb)]) \ runs s Rb = Some (Resp, [A, B], [aKey (sesK (Rs$sk)), END]) \ \ \actions:\ \ \record session key as leaked and add it to intruder knowledge\ s1 = s\ leak := insert (sesK (Rs$sk), Ra$na, Rb$nb) (leak s), IK := insert (Key (sesK (Rs$sk))) (IK s) \ }" text \Intruder fake event.\ definition \ \refines @{term "m2_fake"}\ m3_DY_fake :: "m3_trans" where "m3_DY_fake \ {(s, s1). \ \actions:\ s1 = s(| IK := synth (analz (IK s)) |) }" (******************************************************************************) subsection \Transition system\ (******************************************************************************) definition m3_init :: "m3_state set" where "m3_init \ { \ runs = Map.empty, leak = shrK`bad \ {undefined} \ {undefined}, IK = Key`shrK`bad \ }" definition m3_trans :: "(m3_state \ m3_state) set" where "m3_trans \ (\Ra Rb Rs A B Na Nb Kab. m3_step1 Ra A B Na \ m3_step2 Rb A B \ m3_step3 Rs A B Na Kab \ m3_step4 Ra A B Na Kab \ m3_step5 Rb A B Nb Kab \ m3_step6 Ra A B Na Nb Kab \ m3_step7 Rb A B Nb Kab \ m3_leak Rs Ra Rb A B \ m3_DY_fake \ Id )" definition m3 :: "(m3_state, m3_obs) spec" where "m3 \ \ init = m3_init, trans = m3_trans, obs = m3_obs \" lemmas m3_defs = m3_def m3_init_def m3_trans_def m3_obs_def m3_step1_def m3_step2_def m3_step3_def m3_step4_def m3_step5_def m3_step6_def m3_step7_def m3_leak_def m3_DY_fake_def (******************************************************************************) subsection \Invariants\ (******************************************************************************) text \Specialized injection that we can apply more aggressively.\ lemmas analz_Inj_IK = analz.Inj [where H="IK s" for s] lemmas parts_Inj_IK = parts.Inj [where H="IK s" for s] declare parts_Inj_IK [dest!] declare analz_into_parts [dest] subsubsection \inv1: Secrecy of pre-distributed shared keys\ (******************************************************************************) text \inv1: Secrecy of long-term keys\ definition m3_inv1_lkeysec :: "m3_state set" where "m3_inv1_lkeysec \ {s. \C. (Key (shrK C) \ parts (IK s) \ C \ bad) \ (C \ bad \ Key (shrK C) \ IK s) }" lemmas m3_inv1_lkeysecI = m3_inv1_lkeysec_def [THEN setc_def_to_intro, rule_format] lemmas m3_inv1_lkeysecE [elim] = m3_inv1_lkeysec_def [THEN setc_def_to_elim, rule_format] lemmas m3_inv1_lkeysecD = m3_inv1_lkeysec_def [THEN setc_def_to_dest, rule_format] text \Invariance proof.\ lemma PO_m3_inv1_lkeysec_init [iff]: "init m3 \ m3_inv1_lkeysec" by (auto simp add: m3_defs m3_inv1_lkeysec_def) lemma PO_m3_inv1_lkeysec_trans [iff]: "{m3_inv1_lkeysec} trans m3 {> m3_inv1_lkeysec}" by (fastforce simp add: PO_hoare_defs m3_defs intro!: m3_inv1_lkeysecI) lemma PO_m3_inv1_lkeysec [iff]: "reach m3 \ m3_inv1_lkeysec" by (rule inv_rule_incr) (fast+) text \Useful simplifier lemmas\ lemma m3_inv1_lkeysec_for_parts [simp]: "\ s \ m3_inv1_lkeysec \ \ Key (shrK C) \ parts (IK s) \ C \ bad" by auto lemma m3_inv1_lkeysec_for_analz [simp]: "\ s \ m3_inv1_lkeysec \ \ Key (shrK C) \ analz (IK s) \ C \ bad" by auto subsubsection \inv7a: Session keys not used to encrypt other session keys\ (******************************************************************************) text \Session keys are not used to encrypt other keys. Proof requires generalization to sets of session keys. NOTE: This invariant will be derived from the corresponding L2 invariant using the simulation relation. \ definition m3_inv7a_sesK_compr :: "m3_pred" where "m3_inv7a_sesK_compr \ {s. \K KK. KK \ range sesK \ (Key K \ analz (Key`KK \ (IK s))) = (K \ KK \ Key K \ analz (IK s)) }" lemmas m3_inv7a_sesK_comprI = m3_inv7a_sesK_compr_def [THEN setc_def_to_intro, rule_format] lemmas m3_inv7a_sesK_comprE = m3_inv7a_sesK_compr_def [THEN setc_def_to_elim, rule_format] lemmas m3_inv7a_sesK_comprD = m3_inv7a_sesK_compr_def [THEN setc_def_to_dest, rule_format] text \Additional lemma\ lemmas insert_commute_Key = insert_commute [where x="Key K" for K] lemmas m3_inv7a_sesK_compr_simps = m3_inv7a_sesK_comprD m3_inv7a_sesK_comprD [where KK="{Kab}" for Kab, simplified] m3_inv7a_sesK_comprD [where KK="insert Kab KK" for Kab KK, simplified] insert_commute_Key (******************************************************************************) subsection \Refinement\ (******************************************************************************) subsubsection \Message abstraction and simulation relation\ (******************************************************************************) text \Abstraction function on sets of messages.\ inductive_set abs_msg :: "msg set \ chmsg set" for H :: "msg set" where am_M1: "\Agent A, Agent B, Nonce Na\ \ H \ Insec A B (Msg [aNon Na]) \ abs_msg H" | am_M2: "Crypt (shrK C) \Nonce N, Agent B, Key K\ \ H \ Secure Sv C (Msg [aNon N, aAgt B, aKey K]) \ abs_msg H" | am_M3: "Crypt (shrK C) \Key K, Agent A\ \ H \ Secure Sv C (Msg [aKey K, aAgt A]) \ abs_msg H" | am_M4: "Crypt K (Nonce N) \ H \ dAuth K (Msg [aNon N]) \ abs_msg H" | am_M5: "Crypt K \Nonce N, Nonce N'\ \ H \ dAuth K (Msg [aNon N, aNon N']) \ abs_msg H" text \R23: The simulation relation. This is a data refinement of the insecure and secure channels of refinement 2.\ definition R23_msgs :: "(m2_state \ m3_state) set" where "R23_msgs \ {(s, t). abs_msg (parts (IK t)) \ chan s}" definition R23_keys :: "(m2_state \ m3_state) set" where \ \equivalence!\ "R23_keys \ {(s, t). \KK K. KK \ range sesK \ Key K \ analz (Key`KK \ IK t) \ aKey K \ extr (aKey`KK \ ik0) (chan s) }" definition R23_non :: "(m2_state \ m3_state) set" where \ \only an implication!\ "R23_non \ {(s, t). \KK N. KK \ range sesK \ Nonce N \ analz (Key`KK \ IK t) \ aNon N \ extr (aKey`KK \ ik0) (chan s) }" definition R23_pres :: "(m2_state \ m3_state) set" where "R23_pres \ {(s, t). runs s = runs t \ leak s = leak t}" definition R23 :: "(m2_state \ m3_state) set" where "R23 \ R23_msgs \ R23_keys \ R23_non \ R23_pres" lemmas R23_defs = R23_def R23_msgs_def R23_keys_def R23_non_def R23_pres_def text \The mediator function is the identity here.\ definition med32 :: "m3_obs \ m2_obs" where "med32 \ id" lemmas R23_msgsI = R23_msgs_def [THEN rel_def_to_intro, simplified, rule_format] lemmas R23_msgsE [elim] = R23_msgs_def [THEN rel_def_to_elim, simplified, rule_format] lemmas R23_keysI = R23_keys_def [THEN rel_def_to_intro, simplified, rule_format] lemmas R23_keysE [elim] = R23_keys_def [THEN rel_def_to_elim, simplified, rule_format] lemmas R23_keysD = R23_keys_def [THEN rel_def_to_dest, simplified, rule_format] lemmas R23_nonI = R23_non_def [THEN rel_def_to_intro, simplified, rule_format] lemmas R23_nonE [elim] = R23_non_def [THEN rel_def_to_elim, simplified, rule_format] lemmas R23_nonD = R23_non_def [THEN rel_def_to_dest, simplified, rule_format, rotated 2] lemmas R23_presI = R23_pres_def [THEN rel_def_to_intro, simplified, rule_format] lemmas R23_presE [elim] = R23_pres_def [THEN rel_def_to_elim, simplified, rule_format] lemmas R23_intros = R23_msgsI R23_keysI R23_nonI R23_presI text \Further lemmas: general lemma for simplifier and different instantiations.\ lemmas R23_keys_simps = R23_keysD R23_keysD [where KK="{}", simplified] R23_keysD [where KK="{K'}" for K', simplified] R23_keysD [where KK="insert K' KK" for K' KK, simplified, OF _ conjI] lemmas R23_non_dests = R23_nonD R23_nonD [where KK="{}", simplified] R23_nonD [where KK="{K}" for K, simplified] R23_nonD [where KK="insert K KK" for K KK, simplified, OF _ _ conjI] subsubsection \General lemmas\ (******************************************************************************) text \General facts about @{term "abs_msg"}\ declare abs_msg.intros [intro!] declare abs_msg.cases [elim!] lemma abs_msg_empty: "abs_msg {} = {}" by (auto) lemma abs_msg_Un [simp]: "abs_msg (G \ H) = abs_msg G \ abs_msg H" by (auto) lemma abs_msg_mono [elim]: "\ m \ abs_msg G; G \ H \ \ m \ abs_msg H" by (auto) lemma abs_msg_insert_mono [intro]: "\ m \ abs_msg H \ \ m \ abs_msg (insert m' H)" by (auto) text \Facts about @{term "abs_msg"} concerning abstraction of fakeable messages. This is crucial for proving the refinement of the intruder event.\ lemma abs_msg_DY_subset_fakeable: "\ (s, t) \ R23_msgs; (s, t) \ R23_keys; (s, t) \ R23_non; t \ m3_inv1_lkeysec \ \ abs_msg (synth (analz (IK t))) \ fake ik0 (dom (runs s)) (chan s)" apply (auto) \ \9 subgoals, deal with replays first\ prefer 2 apply (blast) prefer 3 apply (blast) prefer 4 apply (blast) prefer 5 apply (blast) \ \remaining 5 subgoals are real fakes\ apply (intro fake_StatCh fake_DynCh, auto simp add: R23_keys_simps dest: R23_non_dests)+ done subsubsection \Refinement proof\ (******************************************************************************) text \Pair decomposition. These were set to \texttt{elim!}, which is too agressive here.\ declare MPair_analz [rule del, elim] declare MPair_parts [rule del, elim] text \Protocol events.\ lemma PO_m3_step1_refines_m2_step1: "{R23} (m2_step1 Ra A B Na), (m3_step1 Ra A B Na) {> R23}" by (auto simp add: PO_rhoare_defs R23_def m2_defs m3_defs intro!: R23_intros) (auto) lemma PO_m3_step2_refines_m2_step2: "{R23} (m2_step2 Rb A B), (m3_step2 Rb A B) {> R23}" by (auto simp add: PO_rhoare_defs R23_def m2_defs m3_defs intro!: R23_intros) (auto) lemma PO_m3_step3_refines_m2_step3: "{R23 \ (m2_inv3a_sesK_compr) \ (m3_inv7a_sesK_compr \ m3_inv1_lkeysec)} (m2_step3 Rs A B Na Kab), (m3_step3 Rs A B Na Kab) {> R23}" proof - { fix s t assume H: "(s, t) \ R23_msgs" "(s, t) \ R23_keys" "(s, t) \ R23_non" "(s, t) \ R23_pres" "s \ m2_inv3a_sesK_compr" "t \ m3_inv7a_sesK_compr" "t \ m3_inv1_lkeysec" "Kab = sesK (Rs$sk)" "Rs \ dom (runs t)" "\ Agent A, Agent B, Nonce Na \ \ parts (IK t)" let ?s'= - "s\ runs := runs s(Rs \ (Serv, [A, B], [aNon Na])), + "s\ runs := (runs s)(Rs \ (Serv, [A, B], [aNon Na])), chan := insert (Secure Sv A (Msg [aNon Na, aAgt B, aKey Kab])) (insert (Secure Sv B (Msg [aKey Kab, aAgt A])) (chan s)) \" let ?t'= - "t\ runs := runs t(Rs \ (Serv, [A, B], [aNon Na])), + "t\ runs := (runs t)(Rs \ (Serv, [A, B], [aNon Na])), IK := insert (Crypt (shrK A) \ Nonce Na, Agent B, Key Kab \) (insert (Crypt (shrK B) \ Key Kab, Agent A \) (IK t)) \" have "(?s', ?t') \ R23_msgs" using H by (-) (rule R23_intros, auto) moreover have "(?s', ?t') \ R23_keys" using H by (-) (rule R23_intros, auto simp add: m2_inv3a_sesK_compr_simps m3_inv7a_sesK_compr_simps, auto simp add: R23_keys_simps) moreover have "(?s', ?t') \ R23_non" using H by (-) (rule R23_intros, auto simp add: m2_inv3a_sesK_compr_simps m3_inv7a_sesK_compr_simps dest: R23_non_dests) moreover have "(?s', ?t') \ R23_pres" using H by (-) (rule R23_intros, auto) moreover note calculation } thus ?thesis by (auto simp add: PO_rhoare_defs R23_def m2_defs m3_defs) qed (* with equality in R23_keys [OLD, before adding session key compromise]: goal (1 subgoal): 1. \s t. \t \ m3_inv3_keyuse; t \ m3_inv1_lkeysec; t \ m3_inv5_badkeys; Kab \ dom (runs t); Kab \ range shrK; \Agent A, Agent B, Nonce Na\ \ parts (IK t); Kab \ keysFor (analz (IK t)); Key (shrK A) \ analz (IK t); (s, t) \ R23_msgs; (s, t) \ R23_keys; (s, t) \ R23_pres; Kab \ shrK A; Key (shrK B) \ analz (IK t); Kab \ shrK B; A \ bad; B \ bad\ \ Key Kab \ analz (IK t) This does NOT hold, since Kab is revealed in abstract model when A \ bad; B \ bad, but not in concrete one, since here it is protected by A's encryption. *) lemma PO_m3_step4_refines_m2_step4: "{R23} (m2_step4 Ra A B Na Kab), (m3_step4 Ra A B Na Kab) {> R23}" by (auto simp add: PO_rhoare_defs R23_def m2_defs m3_defs intro!: R23_intros) (auto) lemma PO_m3_step5_refines_m2_step5: "{R23} (m2_step5 Rb A B Nb Kab), (m3_step5 Rb A B Nb Kab) {> R23}" by (auto simp add: PO_rhoare_defs R23_def m2_defs m3_defs intro!: R23_intros) (auto) lemma PO_m3_step6_refines_m2_step6: "{R23} (m2_step6 Ra A B Na Nb Kab), (m3_step6 Ra A B Na Nb Kab) {> R23}" by (auto simp add: PO_rhoare_defs R23_def m2_defs m3_defs intro!: R23_intros) (auto) lemma PO_m3_step7_refines_m2_step7: "{R23} (m2_step7 Rb A B Nb Kab), (m3_step7 Rb A B Nb Kab) {> R23}" by (auto simp add: PO_rhoare_defs R23_def m2_defs m3_defs intro!: R23_intros) (auto) text \Intruder events.\ lemma PO_m3_leak_refines_m2_leak: "{R23} (m2_leak Rs Ra Rb A B), (m3_leak Rs Ra Rb A B) {>R23}" by (auto simp add: PO_rhoare_defs R23_def m2_defs m3_defs intro!: R23_intros) (auto simp add: R23_keys_simps dest: R23_non_dests) lemma PO_m3_DY_fake_refines_m2_fake: "{R23 \ UNIV \ m3_inv1_lkeysec} m2_fake, m3_DY_fake {> R23}" apply (auto simp add: PO_rhoare_defs R23_def m2_defs m3_defs intro!: R23_intros del: abs_msg.cases) apply (auto intro: abs_msg_DY_subset_fakeable [THEN subsetD] del: abs_msg.cases) apply (auto simp add: R23_keys_simps dest: R23_non_dests) done text \All together now...\ lemmas PO_m3_trans_refines_m2_trans = PO_m3_step1_refines_m2_step1 PO_m3_step2_refines_m2_step2 PO_m3_step3_refines_m2_step3 PO_m3_step4_refines_m2_step4 PO_m3_step5_refines_m2_step5 PO_m3_step6_refines_m2_step6 PO_m3_step7_refines_m2_step7 PO_m3_leak_refines_m2_leak PO_m3_DY_fake_refines_m2_fake lemma PO_m3_refines_init_m2 [iff]: "init m3 \ R23``(init m2)" by (auto simp add: R23_def m2_defs m3_defs intro!: R23_intros) lemma PO_m3_refines_trans_m2 [iff]: "{R23 \ (m2_inv3a_sesK_compr) \ (m3_inv7a_sesK_compr \ m3_inv1_lkeysec)} (trans m2), (trans m3) {> R23}" apply (auto simp add: m3_def m3_trans_def m2_def m2_trans_def) apply (blast intro!: PO_m3_trans_refines_m2_trans)+ done lemma PO_m3_observation_consistent [iff]: "obs_consistent R23 med32 m2 m3" by (auto simp add: obs_consistent_def R23_def med32_def m2_defs m3_defs) text \Refinement result.\ lemma m3_refines_m2 [iff]: "refines (R23 \ m2_inv3a_sesK_compr \ m3_inv1_lkeysec) med32 m2 m3" proof - have "R23 \ m2_inv3a_sesK_compr \ UNIV \ UNIV \ m3_inv7a_sesK_compr" by (auto simp add: R23_def R23_keys_simps intro!: m3_inv7a_sesK_comprI) thus ?thesis by (-) (rule Refinement_using_invariants, auto) qed lemma m3_implements_m2 [iff]: "implements med32 m2 m3" by (rule refinement_soundness) (auto) subsection \Inherited invariants\ (******************************************************************************) subsubsection \inv4 (derived): Key secrecy for initiator\ (*invh*************************************************************************) definition m3_inv4_ikk_init :: "m3_state set" where "m3_inv4_ikk_init \ {s. \Ra K A B al. runs s Ra = Some (Init, [A, B], aKey K # al) \ A \ good \ B \ good \ Key K \ analz (IK s) \ (\Nb'. (K, Ra $ na, Nb') \ leak s) }" lemmas m3_inv4_ikk_initI = m3_inv4_ikk_init_def [THEN setc_def_to_intro, rule_format] lemmas m3_inv4_ikk_initE [elim] = m3_inv4_ikk_init_def [THEN setc_def_to_elim, rule_format] lemmas m3_inv4_ikk_initD = m3_inv4_ikk_init_def [THEN setc_def_to_dest, rule_format, rotated 1] lemma PO_m3_inv4_ikk_init: "reach m3 \ m3_inv4_ikk_init" proof (rule INV_from_Refinement_using_invariants [OF m3_refines_m2]) show "Range (R23 \ m2_inv3a_sesK_compr \ m3_inv1_lkeysec \ m2_inv6_ikk_init \ UNIV) \ m3_inv4_ikk_init" by (auto simp add: R23_def R23_pres_def R23_keys_simps intro!: m3_inv4_ikk_initI) qed auto subsubsection \inv5 (derived): Key secrecy for responder\ (*invh*************************************************************************) definition m3_inv5_ikk_resp :: "m3_state set" where "m3_inv5_ikk_resp \ {s. \Rb K A B al. runs s Rb = Some (Resp, [A, B], aKey K # al) \ A \ good \ B \ good \ Key K \ analz (IK s) \ K \ Domain (leak s) }" lemmas m3_inv5_ikk_respI = m3_inv5_ikk_resp_def [THEN setc_def_to_intro, rule_format] lemmas m3_inv5_ikk_respE [elim] = m3_inv5_ikk_resp_def [THEN setc_def_to_elim, rule_format] lemmas m3_inv5_ikk_respD = m3_inv5_ikk_resp_def [THEN setc_def_to_dest, rule_format, rotated 1] lemma PO_m3_inv4_ikk_resp: "reach m3 \ m3_inv5_ikk_resp" proof (rule INV_from_Refinement_using_invariants [OF m3_refines_m2]) show "Range (R23 \ m2_inv3a_sesK_compr \ m3_inv1_lkeysec \ m2_inv7_ikk_resp \ UNIV) \ m3_inv5_ikk_resp" by (auto simp add: R23_def R23_pres_def R23_keys_simps intro!: m3_inv5_ikk_respI) (elim m2_inv7_ikk_respE, auto) qed auto end diff --git a/thys/Separation_Logic_Imperative_HOL/Examples/Array_Map_Impl.thy b/thys/Separation_Logic_Imperative_HOL/Examples/Array_Map_Impl.thy --- a/thys/Separation_Logic_Imperative_HOL/Examples/Array_Map_Impl.thy +++ b/thys/Separation_Logic_Imperative_HOL/Examples/Array_Map_Impl.thy @@ -1,174 +1,174 @@ theory Array_Map_Impl imports "../Sep_Main" Imp_Map_Spec Array_Blit "HOL-Library.Code_Target_Numeral" begin subsection "Array Map" type_synonym 'v array_map = "'v option array" definition "iam_initial_size \ 8::nat" definition "iam_of_list l i \ if i'a) \ ('a::heap) array_map \ assn" where "is_iam m a \ \\<^sub>Al. a\\<^sub>al * \(m=iam_of_list l)" definition iam_new_sz :: "nat \ ('v::heap) array_map Heap" where "iam_new_sz sz \ Array.new sz None" definition iam_new :: "('v::heap) array_map Heap" where "iam_new \ iam_new_sz iam_initial_size" definition iam_lookup :: "nat \ ('v::heap) array_map \ 'v option Heap" where "iam_lookup k a = do { l\Array.len a; if k < l then Array.nth a k else return None }" lemma [code]: "iam_lookup k a \ nth_oo None a k" unfolding nth_oo_def iam_lookup_def . definition iam_delete :: "nat \ ('v::heap) array_map \ ('v::heap) array_map Heap" where "iam_delete k a = do { l\Array.len a; if k < l then Array.upd k None a else return a }" lemma [code]: "iam_delete k a \ upd_oo (return a) k None a" unfolding upd_oo_def iam_delete_def . definition iam_update :: "nat \ 'v::heap \ 'v array_map \ 'v array_map Heap" where "iam_update k v a = do { l\Array.len a; a\if k>=l then do { let newsz = max (k+1) (2 * l + 3); array_grow a newsz None } else return a; Array.upd k (Some v) a }" lemma [code]: "iam_update k v a = upd_oo (do { l\Array.len a; let newsz = max (k+1) (2 * l + 3); a\array_grow a newsz None; Array.upd k (Some v) a }) k (Some v) a" proof - have [simp]: "\x t e. do { l\Array.len a; if x l then t l else do { l'\Array.len a; e l l' } } = do { l\Array.len a; if x l then t l else e l l }" apply (auto simp: bind_def execute_len split: option.split intro!: ext ) done show ?thesis unfolding upd_oo_def iam_update_def apply simp apply (rule cong[OF arg_cong, where f1=bind]) apply simp apply (rule ext) apply auto done qed lemma precise_iam: "precise is_iam" apply rule by (auto simp add: is_iam_def dest: preciseD[OF snga_prec]) lemma iam_new_abs: "iam_of_list (replicate n None) = Map.empty" unfolding iam_of_list_def[abs_def] by auto lemma iam_new_sz_rule: " iam_new_sz n < is_iam Map.empty >" unfolding iam_new_sz_def is_iam_def[abs_def] by (sep_auto simp: iam_new_abs) lemma iam_new_rule: " iam_new < is_iam Map.empty >" unfolding iam_new_def by (sep_auto heap: iam_new_sz_rule) lemma iam_lookup_abs1: "k iam_of_list l k = l!k" by (simp add: iam_of_list_def) lemma iam_lookup_abs2: "\k iam_of_list l k = None" by (simp add: iam_of_list_def) lemma iam_lookup_rule: "< is_iam m p > iam_lookup k p <\r. is_iam m p * \(r=m k) >" unfolding iam_lookup_def is_iam_def by (sep_auto simp: iam_lookup_abs1 iam_lookup_abs2) lemma iam_delete_abs1: "k iam_of_list (l[k := None]) = iam_of_list l |` (- {k})" unfolding iam_of_list_def[abs_def] by (auto intro!: ext simp: restrict_map_def) lemma iam_delete_abs2: "\k iam_of_list l |` (- {k}) = iam_of_list l" unfolding iam_of_list_def[abs_def] by (auto intro!: ext simp: restrict_map_def) lemma iam_delete_rule: "< is_iam m p > iam_delete k p <\r. is_iam (m|`(-{k})) r>" unfolding is_iam_def iam_delete_def by (sep_auto simp: iam_delete_abs1 iam_delete_abs2) lemma iam_update_abs1: "iam_of_list (l@replicate n None) = iam_of_list l" unfolding iam_of_list_def[abs_def] by (auto intro!: ext simp: nth_append) lemma iam_update_abs2: "\ length l \ k - \ iam_of_list (l[k := Some v]) = iam_of_list l(k \ v)" + \ iam_of_list (l[k := Some v]) = (iam_of_list l)(k \ v)" unfolding iam_of_list_def[abs_def] by auto lemma iam_update_rule: "< is_iam m p > iam_update k v p <\r. is_iam (m(k\v)) r>\<^sub>t" unfolding is_iam_def iam_update_def by (sep_auto decon: decon_if_split simp: iam_update_abs1 iam_update_abs2) interpretation iam: imp_map is_iam apply unfold_locales by (rule precise_iam) interpretation iam: imp_map_empty is_iam iam_new apply unfold_locales by (sep_auto heap: iam_new_rule) interpretation iam_sz: imp_map_empty is_iam "iam_new_sz sz" apply unfold_locales by (sep_auto heap: iam_new_sz_rule) interpretation iam: imp_map_lookup is_iam iam_lookup apply unfold_locales by (sep_auto heap: iam_lookup_rule) interpretation iam: imp_map_delete is_iam iam_delete apply unfold_locales by (sep_auto heap: iam_delete_rule) interpretation iam: imp_map_update is_iam iam_update apply unfold_locales by (sep_auto heap: iam_update_rule) end diff --git a/thys/Separation_Logic_Imperative_HOL/Examples/Hash_Map.thy b/thys/Separation_Logic_Imperative_HOL/Examples/Hash_Map.thy --- a/thys/Separation_Logic_Imperative_HOL/Examples/Hash_Map.thy +++ b/thys/Separation_Logic_Imperative_HOL/Examples/Hash_Map.thy @@ -1,1102 +1,1102 @@ section "Hash-Maps" theory Hash_Map imports Hash_Table begin subsection \Auxiliary Lemmas\ lemma map_of_ls_update: "map_of (fst (ls_update k v l)) = (map_of l)(k \ v)" apply (induct l rule: ls_update.induct) by (auto simp add: ext Let_def) lemma map_of_concat: "k \ dom (map_of(concat l)) \ \i. k \ dom (map_of(l!i)) \ i < length l" apply (induct l) apply simp apply auto apply (rule_tac x = 0 in exI) apply auto by (metis Suc_mono domI nth_Cons_Suc) lemma map_of_concat': "k \ dom (map_of(l!i)) \ i < length l \ k \ dom (map_of(concat l))" apply (induct l arbitrary: i) apply simp apply auto apply (case_tac i) apply auto done lemma map_of_concat''': assumes "\i. k \ dom (map_of(l!i)) \ i < length l" shows "k \ dom (map_of(concat l))" proof - from assms obtain i where "k \ dom (map_of (l ! i)) \ i < length l" by blast from map_of_concat'[OF this] show ?thesis . qed lemma map_of_concat'': "(k \ dom (map_of(concat l))) \ (\i. k \ dom (map_of(l!i)) \ i < length l)" apply rule using map_of_concat[of k l] apply simp using map_of_concat'[of k l] by blast lemma abs_update_length: "length (abs_update k v l) = length l" by (simp add: abs_update_def) lemma ls_update_map_of_eq: "map_of (fst (ls_update k v ls)) k = Some v" apply (induct ls rule: ls_update.induct) by (simp_all add: Let_def) lemma ls_update_map_of_neq: "x \ k \ map_of (fst (ls_update k v ls)) x = map_of ls x" apply (induct ls rule: ls_update.induct) by (auto simp add: Let_def) subsection \Main Definitions and Lemmas\ definition is_hashmap' :: "('k, 'v) map \ ('k \ 'v) list list \ ('k::{heap,hashable}, 'v::heap) hashtable \ assn" where "is_hashmap' m l ht = is_hashtable l ht * \ (map_of (concat l) = m)" definition is_hashmap :: "('k, 'v) map \ ('k::{heap,hashable}, 'v::heap) hashtable \ assn" where "is_hashmap m ht = (\\<^sub>Al. is_hashmap' m l ht)" lemma is_hashmap'_prec: "\s s'. h\(is_hashmap' m l ht * F1) \\<^sub>A (is_hashmap' m' l' ht * F2) \ l=l' \ m=m'" unfolding is_hashmap'_def apply (auto simp add: preciseD[OF is_hashtable_prec]) apply (subgoal_tac "l = l'") by (auto simp add: preciseD[OF is_hashtable_prec]) lemma is_hashmap_prec: "precise is_hashmap" unfolding is_hashmap_def[abs_def] apply rule by (auto simp add: is_hashmap'_prec) abbreviation "hm_new \ ht_new" lemma hm_new_rule': " hm_new::('k::{heap,hashable}, 'v::heap) hashtable Heap " apply (rule cons_post_rule) using complete_ht_new apply simp apply (simp add: is_hashmap'_def) done lemma hm_new_rule: " hm_new " apply (rule cons_post_rule) using complete_ht_new apply simp apply (simp add: is_hashmap_def is_hashmap'_def) apply sep_auto done lemma ht_hash_distinct: "ht_hash l \ \i j . i\j \ i < length l \ j < length l \ set (l!i) \ set (l!j) = {}" apply (auto simp add: ht_hash_def) apply metis done lemma ht_hash_in_dom_in_dom_bounded_hashcode_nat: assumes "ht_hash l" assumes "k \ dom (map_of(concat l))" shows "k \ dom (map_of(l!bounded_hashcode_nat (length l) k))" proof - from map_of_concat[OF assms(2)] obtain i where i: "k \ dom (map_of (l ! i)) \ i < length l" by blast thm ht_hash_def hence "\v. (k,v)\set(l!i)" by (auto dest: map_of_SomeD) from this obtain v where v: "(k,v)\set(l!i)" by blast from assms(1)[unfolded ht_hash_def] i v bounded_hashcode_nat_bounds have "bounded_hashcode_nat (length l) k = i" by (metis fst_conv) with i show ?thesis by simp qed lemma ht_hash_in_dom_bounded_hashcode_nat_in_dom: assumes "ht_hash l" assumes "1 < length l" assumes "k \ dom (map_of(l!bounded_hashcode_nat (length l) k))" shows "k \ dom (map_of(concat l))" using map_of_concat'[of k l "bounded_hashcode_nat (length l) k"] assms(2,3) bounded_hashcode_nat_bounds[of "length l" k] by simp lemma ht_hash_in_dom_in_dom_bounded_hashcode_nat_eq: assumes "ht_hash l" assumes "1 < length l" shows "(k \ dom (map_of(concat l))) = (k \ dom (map_of(l!bounded_hashcode_nat (length l) k)))" apply rule using ht_hash_in_dom_in_dom_bounded_hashcode_nat[OF assms(1)] ht_hash_in_dom_bounded_hashcode_nat_in_dom[OF assms] by simp_all lemma ht_hash_in_dom_i_bounded_hashcode_nat_i: assumes "ht_hash l" assumes "1 < length l" assumes "i < length l" assumes "k \ dom (map_of (l!i))" shows "i = bounded_hashcode_nat (length l) k" using assms using bounded_hashcode_nat_bounds by (auto simp add: ht_hash_def ht_distinct_def dom_map_of_conv_image_fst) lemma ht_hash_in_bounded_hashcode_nat_not_i_not_in_dom_i: assumes "ht_hash l" assumes "1 < length l" assumes "i < length l" assumes "i \ bounded_hashcode_nat (length l) k" shows "k \ dom (map_of (l!i))" using assms using bounded_hashcode_nat_bounds by (auto simp add: ht_hash_def ht_distinct_def dom_map_of_conv_image_fst) lemma ht_hash_ht_distinct_in_dom_unique_value: assumes "ht_hash l" assumes "ht_distinct l" assumes "1 < length l" assumes "k \ dom (map_of (concat l))" shows "\!v. (k,v) \ set (concat l)" proof - from assms(4) have ex: "\v. (k,v) \ set (concat l)" by (auto dest!: map_of_SomeD) have "v = w" if kv: "(k,v) \ set (concat l)" and kw: "(k,w) \ set (concat l)" for v w proof - from ht_hash_in_dom_in_dom_bounded_hashcode_nat[OF assms(1,4)] have a: "k \ dom (map_of (l ! bounded_hashcode_nat (length l) k))" . have "k \ dom(map_of(l!i))" if "i < length l" and "i \ bounded_hashcode_nat (length l) k" for i proof - from ht_hash_in_bounded_hashcode_nat_not_i_not_in_dom_i[OF assms(1,3) that] show ?thesis . qed have v: "(k,v) \ set (l ! bounded_hashcode_nat (length l) k)" proof - from kv have "\i. i < length l \ (k, v) \ set (l!i)" by auto (metis in_set_conv_nth) from this obtain i where i: "i < length l \ (k, v) \ set (l!i)" by blast hence "k \ dom (map_of (l!i))" by (metis (no_types) prod.exhaust a assms(1) fst_conv ht_hash_def) from i ht_hash_in_dom_i_bounded_hashcode_nat_i[OF assms(1,3) _ this] have "i = bounded_hashcode_nat (length l) k" by simp with i show ?thesis by simp qed have w: "(k,w) \ set (l ! bounded_hashcode_nat (length l) k)" proof - from kw have "\i. i < length l \ (k, w) \ set (l!i)" by auto (metis in_set_conv_nth) from this obtain i where i: "i < length l \ (k, w) \ set (l!i)" by blast hence "k \ dom (map_of (l!i))" by (metis (no_types) prod.exhaust a assms(1) fst_conv ht_hash_def) from i ht_hash_in_dom_i_bounded_hashcode_nat_i[OF assms(1,3) _ this] have "i = bounded_hashcode_nat (length l) k" by simp with i show ?thesis by simp qed from assms(2,3) have "distinct (map fst (l ! bounded_hashcode_nat (length l) k))" by (simp add: ht_distinct_def bounded_hashcode_nat_bounds) from Map.map_of_is_SomeI[OF this v] Map.map_of_is_SomeI[OF this w] show "v = w" by simp qed with ex show ?thesis by blast qed lemma ht_hash_ht_distinct_map_of: assumes "ht_hash l" assumes "ht_distinct l" assumes "1 < length l" shows "map_of (concat l) k = map_of(l!bounded_hashcode_nat (length l) k) k" proof (cases "k \ dom (map_of(concat l))") case False hence a: "map_of (concat l) k = None" by auto from ht_hash_in_dom_in_dom_bounded_hashcode_nat_eq[OF assms(1,3)] False have "k \ dom (map_of (l ! bounded_hashcode_nat (length l) k))" by simp hence "map_of(l!bounded_hashcode_nat (length l) k) k = None" by auto with a show ?thesis by simp next case True from True obtain y where y: "map_of (concat l) k = Some y" by auto hence a: "(k,y) \ set (concat l)" by (metis map_of_SomeD) from ht_hash_in_dom_in_dom_bounded_hashcode_nat_eq[OF assms(1,3)] True have "k \ dom (map_of (l ! bounded_hashcode_nat (length l) k))" by simp from this obtain z where z: "map_of(l!bounded_hashcode_nat (length l) k) k = Some z" by auto hence "(k,z) \ set (l ! bounded_hashcode_nat (length l) k)" by (metis map_of_SomeD) with bounded_hashcode_nat_bounds[OF assms(3), of k] have b: "(k,z) \ set (concat l)" by auto from ht_hash_ht_distinct_in_dom_unique_value[OF assms True] a b have "y = z" by auto with y z show ?thesis by simp qed lemma ls_lookup_map_of_pre: "distinct (map fst l) \ ls_lookup k l = map_of l k" apply (induct l) apply simp apply (case_tac a) by simp lemma ls_lookup_map_of: assumes "ht_hash l" assumes "ht_distinct l" assumes "1 < length l" shows "ls_lookup k (l ! bounded_hashcode_nat (length l) k) = map_of (concat l) k" proof - from assms(2,3) have "distinct (map fst (l ! bounded_hashcode_nat (length l) k))" by (simp add: ht_distinct_def bounded_hashcode_nat_bounds) from ls_lookup_map_of_pre[OF this] have "ls_lookup k (l ! bounded_hashcode_nat (length l) k) = map_of (l ! bounded_hashcode_nat (length l) k) k" . also from ht_hash_ht_distinct_map_of[OF assms] have "map_of (l ! bounded_hashcode_nat (length l) k) k = map_of (concat l) k" by simp finally show ?thesis . qed abbreviation "hm_lookup \ ht_lookup" lemma hm_lookup_rule': " hm_lookup k ht <\r. is_hashmap' m l ht * \(r = m k)>" unfolding is_hashmap'_def apply sep_auto apply (rule cons_post_rule) using complete_ht_lookup[of l ht k] apply simp apply sep_auto by (simp add: ls_lookup_map_of is_hashtable_def) lemma hm_lookup_rule: " hm_lookup k ht <\r. is_hashmap m ht * \(r = m k)>" unfolding is_hashmap_def apply sep_auto apply (rule cons_post_rule[OF hm_lookup_rule']) by sep_auto lemma abs_update_map_of'': assumes "ht_hash l" assumes "ht_distinct l" assumes "1 < length l" shows "map_of (concat (abs_update k v l)) k = Some v" proof - from ht_hash_ht_distinct_map_of[ OF ht_hash_update[OF assms(1)] ht_distinct_update[OF assms(2)] length_update[OF assms(3)], of k v k] have "map_of (concat (abs_update k v l)) k = map_of ((abs_update k v l) ! bounded_hashcode_nat (length l) k) k" by (simp add: abs_update_length) also have "\ = map_of (fst (ls_update k v (l ! bounded_hashcode_nat (length l) k))) k" by (simp add: abs_update_def bounded_hashcode_nat_bounds[OF assms(3)]) also have "... = Some v" by (simp add: ls_update_map_of_eq) finally show ?thesis . qed lemma abs_update_map_of_hceq: assumes "ht_hash l" assumes "ht_distinct l" assumes "1 < length l" assumes "x \ k" assumes "bounded_hashcode_nat (length l) x = bounded_hashcode_nat (length l) k" shows "map_of (concat (abs_update k v l)) x = map_of (concat l) x" proof - from ht_hash_ht_distinct_map_of[ OF ht_hash_update[OF assms(1)] ht_distinct_update[OF assms(2)] length_update[OF assms(3)], of k v x] have "map_of (concat (abs_update k v l)) x = map_of ((abs_update k v l) ! bounded_hashcode_nat (length l) x) x" by (simp add: abs_update_length) also from assms(5) have "\ = map_of (fst (ls_update k v (l ! bounded_hashcode_nat (length l) k))) x" by (simp add: abs_update_def bounded_hashcode_nat_bounds[OF assms(3)]) also have "\ = map_of (l ! bounded_hashcode_nat (length l) x) x" by (simp add: ls_update_map_of_neq[OF assms(4)] assms(5)) also from ht_hash_ht_distinct_map_of[OF assms(1-3)] have "\ = map_of (concat l) x" by simp finally show ?thesis . qed lemma abs_update_map_of_hcneq: assumes "ht_hash l" assumes "ht_distinct l" assumes "1 < length l" assumes "x \ k" assumes "bounded_hashcode_nat (length l) x \ bounded_hashcode_nat (length l) k" shows "map_of (concat (abs_update k v l)) x = map_of (concat l) x" proof - from ht_hash_ht_distinct_map_of[ OF ht_hash_update[OF assms(1)] ht_distinct_update[OF assms(2)] length_update[OF assms(3)], of k v x] have "map_of (concat (abs_update k v l)) x = map_of ((abs_update k v l) ! bounded_hashcode_nat (length l) x) x" by (simp add: abs_update_length) also from assms(5) have "\ = map_of (l ! bounded_hashcode_nat (length l) x) x" by (simp add: abs_update_def bounded_hashcode_nat_bounds[OF assms(3)]) also from ht_hash_ht_distinct_map_of[OF assms(1-3)] have "\ = map_of (concat l) x" by simp finally show ?thesis . qed lemma abs_update_map_of''': assumes "ht_hash l" assumes "ht_distinct l" assumes "1 < length l" assumes "x \ k" shows "map_of (concat (abs_update k v l)) x = map_of (concat l) x" apply (cases "bounded_hashcode_nat (length l) x = bounded_hashcode_nat (length l) k") by (auto simp add: abs_update_map_of_hceq[OF assms] abs_update_map_of_hcneq[OF assms]) lemma abs_update_map_of': assumes "ht_hash l" assumes "ht_distinct l" assumes "1 < length l" shows "map_of (concat (abs_update k v l)) x - = (map_of (concat l)(k \ v)) x" + = ((map_of (concat l))(k \ v)) x" apply (cases "x = k") apply (simp add: abs_update_map_of''[OF assms]) by (simp add: abs_update_map_of'''[OF assms]) lemma abs_update_map_of: assumes "ht_hash l" assumes "ht_distinct l" assumes "1 < length l" shows "map_of (concat (abs_update k v l)) - = map_of (concat l)(k \ v) " + = (map_of (concat l))(k \ v) " apply (rule ext) by (simp add: abs_update_map_of'[OF assms]) lemma ls_insls_map_of: assumes "ht_hash ld" assumes "ht_distinct ld" assumes "1 < length ld" assumes "distinct (map fst xs)" shows "map_of (concat (ls_insls xs ld)) = map_of (concat ld) ++ map_of xs" using assms apply (induct xs arbitrary: ld) apply simp apply (case_tac a) apply (simp only: ls_insls.simps) subgoal premises prems proof - from prems(5) prems(1)[OF ht_hash_update[OF prems(2)] ht_distinct_update[OF prems(3)] length_update[OF prems(4)]] abs_update_map_of[OF prems(2-4)] show ?thesis apply simp apply (rule map_add_upd_left) apply (metis dom_map_of_conv_image_fst) done qed done lemma ls_insls_map_of': assumes "ht_hash ls" assumes "ht_distinct ls" assumes "ht_hash ld" assumes "ht_distinct ld" assumes "1 < length ld" assumes "n < length ls" shows "map_of (concat (ls_insls (ls ! n) ld)) ++ map_of (concat (take n ls)) = map_of (concat ld) ++ map_of (concat (take (Suc n) ls))" proof - from assms(2,6) have "distinct (map fst (ls ! n))" by (simp add: ht_distinct_def) from ls_insls_map_of[OF assms(3-5) this] assms(6) show ?thesis by (simp add: List.take_Suc_conv_app_nth) qed lemma ls_copy_map_of: assumes "ht_hash ls" assumes "ht_distinct ls" assumes "ht_hash ld" assumes "ht_distinct ld" assumes "1 < length ld" assumes "n \ length ls" shows "map_of (concat (ls_copy n ls ld)) = map_of (concat ld) ++ map_of (concat (take n ls))" using assms apply (induct n arbitrary: ld) apply simp subgoal premises prems for n ld proof - note a = ht_hash_ls_insls[OF prems(4), of "ls ! n"] note b = ht_distinct_ls_insls[OF prems(5), of "ls ! n"] note c = length_ls_insls[OF prems(6), of "ls ! n"] from prems have "n < length ls" by simp with ls_insls_map_of'[OF prems(2-6) this] prems(1)[OF assms(1,2) a b c] show ?thesis by simp qed done lemma ls_rehash_map_of: assumes "ht_hash l" assumes "ht_distinct l" assumes "1 < length l" shows "map_of (concat (ls_rehash l)) = map_of (concat l)" using assms(3) ls_copy_map_of[OF assms(1,2) ht_hash_replicate ht_distinct_replicate] by (simp add: ls_rehash_def) lemma abs_update_rehash_map_of: assumes "ht_hash l" assumes "ht_distinct l" assumes "1 < length l" shows "map_of (concat (abs_update k v (ls_rehash l))) - = map_of (concat l)(k \ v)" + = (map_of (concat l))(k \ v)" proof - note a = ht_hash_ls_rehash[of l] note b = ht_distinct_ls_rehash[of l] note c = length_ls_rehash[OF assms(3)] from abs_update_map_of[OF a b c] ls_rehash_map_of[OF assms] show ?thesis by simp qed abbreviation "hm_update \ ht_update" lemma hm_update_rule': " hm_update k v ht <\r. is_hashmap (m(k \ v)) r * true>" proof (cases "length l * load_factor \ the_size ht * 100") case True show ?thesis unfolding is_hashmap'_def apply sep_auto apply (rule cons_post_rule[OF complete_ht_update_rehash[OF True]]) unfolding is_hashmap_def is_hashmap'_def apply sep_auto apply (simp add: abs_update_rehash_map_of is_hashtable_def) done next case False show ?thesis unfolding is_hashmap'_def is_hashtable_def apply sep_auto apply (rule cons_post_rule) using complete_ht_update_normal[OF False, simplified is_hashtable_def, simplified, of k v] apply auto unfolding is_hashmap_def is_hashmap'_def apply sep_auto by (simp add: abs_update_map_of is_hashtable_def) qed lemma hm_update_rule: " hm_update k v ht <\r. is_hashmap (m(k \ v)) r * true>" unfolding is_hashmap_def[of m] by (sep_auto heap add: hm_update_rule') lemma ls_delete_map_of: assumes "distinct (map fst l)" shows "map_of (fst (ls_delete k l)) x = ((map_of l) |` (- {k})) x" using assms apply (induct l rule: ls_delete.induct) apply simp apply (auto simp add: map_of_eq_None_iff Let_def) by (metis ComplD ComplI Compl_insert option.set(2) insertE insertI2 map_upd_eq_restrict restrict_map_def) lemma update_ls_delete_map_of: assumes "ht_hash l" assumes "ht_distinct l" assumes "ht_hash (l[bounded_hashcode_nat (length l) k := fst (ls_delete k (l ! bounded_hashcode_nat (length l) k))])" assumes "ht_distinct (l[bounded_hashcode_nat (length l) k := fst (ls_delete k (l ! bounded_hashcode_nat (length l) k))])" assumes "1 < length l" shows "map_of (concat (l[bounded_hashcode_nat (length l) k := fst (ls_delete k (l ! bounded_hashcode_nat (length l) k))])) x = ((map_of (concat l)) |` (- {k})) x" proof - from assms(2) bounded_hashcode_nat_bounds[OF assms(5)] have distinct: "distinct (map fst (l ! bounded_hashcode_nat (length l) k))" by (auto simp add: ht_distinct_def) note id1 = ht_hash_ht_distinct_map_of[OF assms(3,4), simplified, OF assms(5)[simplified], of x] note id2 = ht_hash_ht_distinct_map_of[OF assms(1,2,5), of x] show ?thesis proof (cases "bounded_hashcode_nat (length l) x = bounded_hashcode_nat (length l) k") case True with id1 have "map_of (concat (l[bounded_hashcode_nat (length l) k := fst (ls_delete k (l ! bounded_hashcode_nat (length l) k))])) x = map_of (l[bounded_hashcode_nat (length l) k := fst (ls_delete k (l ! bounded_hashcode_nat (length l) k))] ! bounded_hashcode_nat (length l) k) x" by simp also have "\ = map_of (fst (ls_delete k (l ! bounded_hashcode_nat (length l) k))) x" by (simp add: bounded_hashcode_nat_bounds[OF assms(5)]) also from ls_delete_map_of[OF distinct] have "\ = (map_of (l ! bounded_hashcode_nat (length l) k) |` (- {k})) x" by simp finally show ?thesis by (cases "x = k") (simp_all add: id2 True) next case False with bounded_hashcode_nat_bounds[OF assms(5)] id1 id2[symmetric] show ?thesis by (cases "x = k") simp_all qed qed abbreviation "hm_delete \ ht_delete" lemma hm_delete_rule': " hm_delete k ht " apply (simp only: is_hashmap'_def[of m] is_hashtable_def) apply sep_auto apply (rule cons_post_rule) using complete_ht_delete[simplified is_hashtable_def] apply sep_auto apply (simp add: is_hashmap_def is_hashmap'_def) apply (sep_auto) apply (simp add: is_hashtable_def) apply (sep_auto) by (auto simp add: update_ls_delete_map_of) lemma hm_delete_rule: " hm_delete k ht " unfolding is_hashmap_def[of m] by (sep_auto heap add: hm_delete_rule') definition hm_isEmpty :: "('k, 'v) hashtable \ bool Heap" where "hm_isEmpty ht \ return (the_size ht = 0)" lemma hm_isEmpty_rule': " hm_isEmpty ht <\r. is_hashmap' m l ht * \(r \ m=Map.empty)>" unfolding hm_isEmpty_def unfolding is_hashmap_def is_hashmap'_def is_hashtable_def ht_size_def apply (cases ht, simp) apply sep_auto done lemma hm_isEmpty_rule: " hm_isEmpty ht <\r. is_hashmap m ht * \(r \ m=Map.empty)>" unfolding is_hashmap_def apply (sep_auto heap: hm_isEmpty_rule') done definition hm_size :: "('k, 'v) hashtable \ nat Heap" where "hm_size ht \ return (the_size ht)" lemma length_card_dom_map_of: assumes "distinct (map fst l)" shows "length l = card (dom (map_of l))" using assms apply (induct l) apply simp apply simp apply (case_tac a) apply (auto intro!: fst_conv map_of_SomeD) apply (subgoal_tac "aa \ dom (map_of l)") apply simp by (metis dom_map_of_conv_image_fst) lemma ht_hash_dom_map_of_disj: assumes "ht_hash l" assumes "i < length l" assumes "j < length l" assumes "i \ j" shows "dom (map_of (l!i)) \ dom (map_of(l!j)) = {}" using assms unfolding ht_hash_def apply auto by (metis fst_conv map_of_SomeD) lemma ht_hash_dom_map_of_disj_drop: assumes "ht_hash l" assumes "i < length l" shows "dom (map_of (l!i)) \ dom (map_of (concat (drop (Suc i) l))) = {}" apply auto subgoal premises prems for x y z proof - from prems(2) have "x \ dom (map_of (concat (drop (Suc i) l)))" by auto hence "\j. j < length (drop (Suc i) l) \ x \ dom (map_of ((drop (Suc i) l)!j))" by (metis Hash_Map.map_of_concat \x \ dom (map_of (concat (drop (Suc i) l)))\ length_drop) from this obtain j where j: "j < length (drop (Suc i) l) \ x \ dom (map_of ((drop (Suc i) l)!j))" by blast hence length: "(Suc i + j) < length l" by auto from j have neq: "i \ (Suc i + j)" by simp from j have in_dom: "x \ dom (map_of (l!(Suc i + j)))" by auto from prems(1) have in_dom2: "x \ dom (map_of (l ! i))" by auto from ht_hash_dom_map_of_disj[OF assms length neq] in_dom in_dom2 show ?thesis by auto qed done lemma sum_list_length_card_dom_map_of_concat: assumes "ht_hash l" assumes "ht_distinct l" shows "sum_list (map length l) = card (dom (map_of (concat l)))" using assms proof - from ht_hash_dom_map_of_disj_drop[OF assms(1)] have "\i. i < length l \ dom (map_of (l ! i)) \ dom (map_of (concat (drop (Suc i) l))) = {}" by auto with assms(2) show ?thesis proof (induct l) case Nil thus ?case by simp next case (Cons l ls) from Cons(2) have a: "ht_distinct ls" by (auto simp add: ht_distinct_def) from Cons(3) have b: "\i < length ls. dom (map_of (ls ! i)) \ dom (map_of (concat (drop (Suc i) ls))) = {}" apply simp apply (rule allI) apply (rule_tac x="Suc i" and P="(\i. i dom (map_of ((l # ls) ! i)) \ dom (map_of (concat (drop i ls))) = {})" in allE) by simp_all from Cons(2) have "distinct (map fst l)" by (auto simp add: ht_distinct_def) note l = length_card_dom_map_of[OF this] from Cons(3) have c: "dom (map_of l) \ dom (map_of (concat ls)) = {}" apply (rule_tac x="0" and P="(\i. i dom (map_of ((l # ls) ! i)) \ dom (map_of (concat (drop i ls))) = {})" in allE) by simp_all from Cons(1)[OF a b] l c show ?case by (simp add: card_Un_disjoint) qed qed lemma hm_size_rule': " hm_size ht <\r. is_hashmap' m l ht * \(r = card (dom m))>" unfolding hm_size_def is_hashmap_def is_hashmap'_def is_hashtable_def apply sep_auto apply (cases ht) apply (simp add: ht_size_def) apply sep_auto by (simp add: sum_list_length_card_dom_map_of_concat) lemma hm_size_rule: " hm_size ht <\r. is_hashmap m ht * \(r = card (dom m))>" unfolding is_hashmap_def by (sep_auto heap: hm_size_rule') subsection \Iterators\ subsubsection \Definitions\ type_synonym ('k,'v) hm_it = "(nat \ ('k\'v) list \ ('k,'v) hashtable)" fun hm_it_adjust :: "nat \ ('k::{heap,hashable},'v::heap) hashtable \ nat Heap" where "hm_it_adjust 0 ht = return 0" | "hm_it_adjust n ht = do { l \ Array.nth (the_array ht) n; case l of [] \ hm_it_adjust (n - 1) ht | _ \ return n }" definition hm_it_init :: "('k::{heap,hashable},'v::heap) hashtable \ ('k,'v) hm_it Heap" where "hm_it_init ht \ do { n\Array.len (the_array ht); if n = 0 then return (0,[],ht) else do { i\hm_it_adjust (n - 1) ht; l\Array.nth (the_array ht) i; return (i,l,ht) } }" definition hm_it_has_next :: "('k::{heap,hashable},'v::heap) hm_it \ bool Heap" where "hm_it_has_next it \ return (case it of (0,[],_) \ False | _ \ True)" definition hm_it_next :: "('k::{heap,hashable},'v::heap) hm_it \ (('k\'v)\('k,'v) hm_it) Heap" where "hm_it_next it \ case it of (i,a#b#l,ht) \ return (a,(i,b#l,ht)) | (0,[a],ht) \ return (a,(0,[],ht)) | (Suc i,[a],ht) \ do { i \ hm_it_adjust i ht; l \ Array.nth (the_array ht) i; return (a,(i,rev l,ht)) } " definition "hm_is_it' l ht l' it \ is_hashtable l ht * \(let (i,r,ht')=it in ht = ht' \ l' = (concat (take i l) @ rev r) \ distinct (map fst (l')) \ i \ length l \ (r=[] \ i=0) )" definition "hm_is_it m ht m' it \ \\<^sub>Al l'. hm_is_it' l ht l' it * \(map_of (concat l) = m \ map_of l' = m') " subsubsection \Auxiliary Lemmas\ lemma concat_take_Suc_empty: "\ n < length l; l!n=[] \ \ concat (take (Suc n) l) = concat (take n l)" apply (induct n arbitrary: l) apply (case_tac l) apply auto [2] apply (case_tac l) apply auto [2] done lemma nth_concat_splitE: assumes "i(i < length l)" hence 1: "concat (l#ls)!i = concat ls ! (i - length l)" by (auto simp: nth_append) obtain j k where "j < length ls" and "k < length (ls!j)" and "concat ls ! (i - length l) = ls!j!k" and "i - length l = length (concat (take j ls)) + k" apply (rule Cons.hyps[of "i - length l"]) using Cons.prems L by auto thus ?case using L apply (rule_tac Cons.prems(1)[of "Suc j" k]) apply (auto simp: nth_append) done qed qed lemma is_hashmap'_distinct: "is_hashtable l ht \\<^sub>A is_hashtable l ht * \(distinct (map fst (concat l)))" apply (simp add: distinct_conv_nth) proof (intro allI impI, elim exE) fix i j a b assume 1: "i < length (concat l)" assume 2: "j < length (concat l)" assume 3: "i\j" assume HM: "(a,b) \ is_hashtable l ht" from 1 obtain ji ki where IFMT: "i = length (concat (take ji l)) + ki" and JI_LEN: "ji < length l" and KI_LEN: "ki < length (l!ji)" and [simp]: "concat l ! i = l!ji!ki" by (blast elim: nth_concat_splitE) from 2 obtain jj kj where JFMT: "j = length (concat (take jj l)) + kj" and JJ_LEN: "jj < length l" and KJ_LEN: "kj < length (l!jj)" and [simp]: "concat l ! j = l!jj!kj" by (blast elim: nth_concat_splitE) show "fst (concat l ! i) \ fst (concat l ! j)" proof cases assume [simp]: "ji=jj" with IFMT JFMT \i\j\ have "ki\kj" by auto moreover from HM JJ_LEN have "distinct (map fst (l!jj))" unfolding is_hashmap'_def is_hashtable_def ht_distinct_def by auto ultimately show ?thesis using KI_LEN KJ_LEN by (simp add: distinct_conv_nth) next assume NE: "ji\jj" from HM have "\x\set (l!ji). bounded_hashcode_nat (length l) (fst x) = ji" "\x\set (l!jj). bounded_hashcode_nat (length l) (fst x) = jj" unfolding is_hashmap'_def is_hashtable_def ht_hash_def using JI_LEN JJ_LEN by auto with KI_LEN KJ_LEN NE show ?thesis apply (auto) by (metis nth_mem) qed qed lemma take_set: "set (take n l) = { l!i | i. i i length l" "x \ n" shows "\i. x i\n \ l!i=[]" proof - have "take (Suc n) l = take (Suc x + (n - x)) l" by simp also have "\ = take (Suc x) l @ take (n - x) (drop (Suc x) l)" by (simp only: take_add) finally have "concat (take (Suc x) l) = concat (take (Suc x) l) @ concat (take (n - x) (drop (Suc x) l))" using A by simp hence 1: "\l\set (take (n - x) (drop (Suc x) l)). l=[]" by simp show ?thesis proof safe fix i assume "xn" hence "l!i \ set (take (n - x) (drop (Suc x) l))" using L[simp del] apply (auto simp: take_set) apply (rule_tac x="i - Suc x" in exI) apply auto done with 1 show "l!i=[]" by blast qed qed lemma take_Suc0: "l\[] \ take (Suc 0) l = [l!0]" "0 < length l \ take (Suc 0) l = [l!0]" "Suc n \ length l \ take (Suc 0) l = [l!0]" by (cases l, auto)+ lemma concat_take_Suc_app_nth: assumes "x < length l" shows "concat (take (Suc x) l) = concat (take x l) @ l ! x" using assms by (auto simp: take_Suc_conv_app_nth) lemma hm_hashcode_eq: assumes "j < length (l!i)" assumes "i < length l" assumes "h \ is_hashtable l ht" shows "bounded_hashcode_nat (length l) (fst (l!i!j)) = i" using assms unfolding is_hashtable_def ht_hash_def apply (cases "l!i!j") apply (force simp: set_conv_nth) done lemma distinct_imp_distinct_take: "distinct (map fst (concat l)) \ distinct (map fst (concat (take x l)))" apply (subst (asm) append_take_drop_id[of x l,symmetric]) apply (simp del: append_take_drop_id) done lemma hm_it_adjust_rule: "i hm_it_adjust i ht <\j. is_hashtable l ht * \( j\i \ (concat (take (Suc i) l) = concat (take (Suc j) l)) \ (j=0 \ l!j \ []) ) >" proof (induct i) case 0 thus ?case by sep_auto next case (Suc n) show ?case using Suc.prems by (sep_auto heap add: Suc.hyps simp: concat_take_Suc_empty split: list.split) qed lemma hm_it_next_rule': "l'\[] \ hm_it_next it <\((k,v),it'). hm_is_it' l ht (butlast l') it' * \(last l' = (k,v) \ distinct (map fst l') )>" unfolding hm_it_next_def hm_is_it'_def is_hashmap'_def using [[hypsubst_thin = true]] apply (sep_auto (plain) split: nat.split list.split heap: hm_it_adjust_rule simp: take_Suc0) apply (simp split: prod.split nat.split list.split) apply (intro allI impI conjI) apply auto [] apply auto [] apply sep_auto [] apply (sep_auto (plain) heap: hm_it_adjust_rule) apply auto [] apply sep_auto apply (cases l, auto) [] apply (metis SUP_upper fst_image_mp image_mono set_concat) apply (drule skip_empty_aux, simp_all) [] defer apply (auto simp: concat_take_Suc_app_nth) [] apply auto [] apply sep_auto apply (auto simp: butlast_append) [] apply (auto simp: butlast_append) [] apply sep_auto apply (auto simp: butlast_append) [] apply (auto simp: butlast_append) [] by (metis Ex_list_of_length Suc_leD concat_take_Suc_app_nth le_neq_implies_less le_trans nat.inject not_less_eq_eq) subsubsection \Main Lemmas\ lemma hm_it_next_rule: "m'\Map.empty \ hm_it_next it <\((k,v),it'). hm_is_it m ht (m' |` (-{k})) it' * \(m' k = Some v)>" proof - { fix ys a have aux3: " \distinct (map fst ys); a \ fst ` set ys\ \ map_of ys a = None" by (induct ys) auto } note aux3 = this assume "m'\Map.empty" thus ?thesis unfolding hm_is_it_def apply (sep_auto heap: hm_it_next_rule') apply (case_tac l' rule: rev_cases, auto simp: restrict_map_def aux3 intro!: ext) [] apply (case_tac l' rule: rev_cases, auto) done qed lemma hm_it_init_rule: fixes ht :: "('k::{heap,hashable},'v::heap) hashtable" shows " hm_it_init ht \<^sub>t" unfolding hm_it_init_def is_hashmap_def is_hashmap'_def hm_is_it_def hm_is_it'_def apply (sep_auto simp del: map_of_append heap add: hm_it_adjust_rule) apply (case_tac l, auto) [] apply (sep_auto simp del: concat_eq_Nil_conv map_of_append) apply (auto simp: distinct_imp_distinct_take dest: ent_fwd[OF _ is_hashmap'_distinct]) [] apply (drule sym) apply (auto simp: is_hashtable_def ht_distinct_def rev_map[symmetric]) [] apply (auto simp: set_conv_nth) [] apply (hypsubst_thin) apply (drule_tac j=ia in hm_hashcode_eq, simp_all) [] apply (drule_tac j=ib in hm_hashcode_eq, simp_all) [] apply (auto simp: is_hashmap'_def is_hashtable_def ht_distinct_def) [] apply (clarsimp) apply (drule ent_fwd[OF _ is_hashmap'_distinct]) apply clarsimp apply (subst concat_take_Suc_app_nth) apply (case_tac l,auto) [] apply (simp) apply (hypsubst_thin) apply (subst (asm) (2) concat_take_Suc_app_nth) apply (case_tac l,auto) [] apply (subst map_of_rev_distinct) apply auto done lemma hm_it_has_next_rule: " hm_it_has_next it <\r. hm_is_it m ht m' it * \(r\m'\Map.empty)>" unfolding is_hashmap'_def hm_is_it_def hm_is_it'_def hm_it_has_next_def by (sep_auto split: nat.split list.split) lemma hm_it_finish: "hm_is_it m p m' it \\<^sub>A is_hashmap m p" unfolding hm_is_it_def hm_is_it'_def is_hashmap_def is_hashmap'_def by sep_auto end diff --git a/thys/Separation_Logic_Imperative_HOL/Examples/Hash_Set_Impl.thy b/thys/Separation_Logic_Imperative_HOL/Examples/Hash_Set_Impl.thy --- a/thys/Separation_Logic_Imperative_HOL/Examples/Hash_Set_Impl.thy +++ b/thys/Separation_Logic_Imperative_HOL/Examples/Hash_Set_Impl.thy @@ -1,150 +1,150 @@ section "Hash-Sets" theory Hash_Set_Impl imports Imp_Set_Spec Hash_Map_Impl begin subsection \Auxiliary Definitions\ definition map_of_set:: "'a set \ 'a\unit" where "map_of_set S x \ if x\S then Some () else None" lemma ne_some_unit_eq: "x\Some () \ x=None" by (cases x) auto lemma map_of_set_simps[simp]: "dom (map_of_set s) = s" "map_of_set (dom m) = m" "map_of_set {} = Map.empty" "map_of_set s x = None \ x\s" "map_of_set s x = Some u \ x\s" - "map_of_set s (x\()) = map_of_set (insert x s)" + "(map_of_set s) (x\()) = map_of_set (insert x s)" "(map_of_set s) |` (-{x}) = map_of_set (s -{x})" apply (auto simp: map_of_set_def dom_def ne_some_unit_eq restrict_map_def intro!: ext) done lemma map_of_set_eq': "map_of_set a = map_of_set b \ a=b" apply (auto simp: map_of_set_def[abs_def]) apply (metis option.simps(3))+ done lemma map_of_set_eq[simp]: "map_of_set s = m \ dom m=s" apply (auto simp: dom_def map_of_set_def[abs_def] ne_some_unit_eq intro!: ext) apply (metis option.simps(3)) done subsection \Main Definitions\ type_synonym 'a hashset = "('a,unit) hashtable" definition "is_hashset s ht \ is_hashmap (map_of_set s) ht" lemma hs_set_impl: "imp_set is_hashset" apply unfold_locales apply rule unfolding is_hashset_def apply (subst map_of_set_eq'[symmetric]) by (metis preciseD[OF is_hashmap_prec]) interpretation hs: imp_set is_hashset by (rule hs_set_impl) definition hs_new :: "'a::{heap,hashable} hashset Heap" where "hs_new = hm_new" lemma hs_new_impl: "imp_set_empty is_hashset hs_new" apply unfold_locales apply (sep_auto heap: hm_new_rule simp: is_hashset_def hs_new_def) done interpretation hs: imp_set_empty is_hashset hs_new by (rule hs_new_impl) definition hs_memb:: "'a::{heap,hashable} \ 'a hashset \ bool Heap" where "hs_memb x s \ do { r\hm_lookup x s; return (case r of Some _ \ True | None \ False) }" lemma hs_memb_impl: "imp_set_memb is_hashset hs_memb" apply unfold_locales unfolding hs_memb_def apply (sep_auto heap: hm_lookup_rule simp: is_hashset_def split: option.split) done interpretation hs: imp_set_memb is_hashset hs_memb by (rule hs_memb_impl) definition hs_ins:: "'a::{heap,hashable} \ 'a hashset \ 'a hashset Heap" where "hs_ins x ht \ hm_update x () ht" lemma hs_ins_impl: "imp_set_ins is_hashset hs_ins" apply unfold_locales apply (sep_auto heap: hm_update_rule simp: hs_ins_def is_hashset_def) done interpretation hs: imp_set_ins is_hashset hs_ins by (rule hs_ins_impl) definition hs_delete :: "'a::{heap,hashable} \ 'a hashset \ 'a hashset Heap" where "hs_delete x ht \ hm_delete x ht" lemma hs_delete_impl: "imp_set_delete is_hashset hs_delete" apply unfold_locales apply (sep_auto heap: hm_delete_rule simp: is_hashset_def hs_delete_def) done interpretation hs: imp_set_delete is_hashset hs_delete by (rule hs_delete_impl) definition "hs_isEmpty == hm_isEmpty" lemma hs_is_empty_impl: "imp_set_is_empty is_hashset hs_isEmpty" apply unfold_locales apply (sep_auto heap: hm_isEmpty_rule simp: is_hashset_def hs_isEmpty_def) done interpretation hs: imp_set_is_empty is_hashset hs_isEmpty by (rule hs_is_empty_impl) definition "hs_size == hm_size" lemma hs_size_impl: "imp_set_size is_hashset hs_size" apply unfold_locales apply (sep_auto heap: hm_size_rule simp: is_hashset_def hs_size_def) done interpretation hs: imp_set_size is_hashset hs_size by (rule hs_size_impl) type_synonym ('a) hs_it = "('a,unit) hm_it" definition "hs_is_it s hs its it \ hm_is_it (map_of_set s) hs (map_of_set its) it" definition hs_it_init :: "('a::{heap,hashable}) hashset \ 'a hs_it Heap" where "hs_it_init \ hm_it_init" definition hs_it_has_next :: "('a::{heap,hashable}) hs_it \ bool Heap" where "hs_it_has_next \ hm_it_has_next" definition hs_it_next :: "('a::{heap,hashable}) hs_it \ ('a\'a hs_it) Heap" where "hs_it_next it \ do { ((x,_),it) \ hm_it_next it; return (x,it) }" lemma hs_iterate_impl: "imp_set_iterate is_hashset hs_is_it hs_it_init hs_it_has_next hs_it_next" apply unfold_locales unfolding hs_it_init_def hs_it_next_def hs_it_has_next_def hs_is_it_def is_hashset_def apply sep_auto apply sep_auto apply sep_auto apply (sep_auto eintros: hm.quit_iteration) done interpretation hs: imp_set_iterate is_hashset hs_is_it hs_it_init hs_it_has_next hs_it_next by (rule hs_iterate_impl) export_code hs_new hs_memb hs_ins hs_delete hs_isEmpty hs_size hs_it_init hs_it_has_next hs_it_next checking SML_imp end diff --git a/thys/UPF/ServiceExample.thy b/thys/UPF/ServiceExample.thy --- a/thys/UPF/ServiceExample.thy +++ b/thys/UPF/ServiceExample.thy @@ -1,137 +1,137 @@ (***************************************************************************** * HOL-TestGen --- theorem-prover based test case generation * http://www.brucker.ch/projects/hol-testgen/ * * This file is part of HOL-TestGen. * * Copyright (c) 2010-2012 ETH Zurich, Switzerland * 2010-2014 Achim D. Brucker, Germany * 2010-2014 Université Paris-Sud, 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. ******************************************************************************) section \Instantiating Our Secure Service Example\ theory ServiceExample imports Service begin text \ In the following, we briefly present an instantiations of our secure service example from the last section. We assume three different members of the health care staff and two patients: \ subsection \Access Control Configuration\ definition alice :: user where "alice = 1" definition bob :: user where "bob = 2" definition charlie :: user where "charlie = 3" definition patient1 :: patient where "patient1 = 5" definition patient2 :: patient where "patient2 = 6" definition UC0 :: \ where - "UC0 = Map.empty(alice\Nurse)(bob\ClinicalPractitioner)(charlie\Clerical)" + "UC0 = Map.empty(alice\Nurse, bob\ClinicalPractitioner, charlie\Clerical)" definition entry1 :: entry where "entry1 = (Open,alice, dummyContent)" definition entry2 :: entry where "entry2 = (Closed,bob, dummyContent)" definition entry3 :: entry where "entry3 = (Closed,alice, dummyContent)" definition SCR1 :: SCR where "SCR1 = (Map.empty(1\entry1))" definition SCR2 :: SCR where "SCR2 = (Map.empty)" definition Spine0 :: DB where - "Spine0 = Map.empty(patient1\SCR1)(patient2\SCR2)" + "Spine0 = Map.empty(patient1\SCR1, patient2\SCR2)" definition LR1 :: LR where "LR1 =(Map.empty(1\{alice}))" definition \0 :: \ where "\0 = (Map.empty(patient1\LR1))" subsection \The Initial System State\ definition \0 :: "DB \ \\\" where "\0 = (Spine0,\0,UC0)" subsection\Basic Properties\ lemma [simp]: "(case a of allow d \ \X\ | deny d2 \ \Y\) = \ \ False" by (case_tac a,simp_all) lemma [cong,simp]: "((if hasLR urp1_alice 1 \0 then \allow ()\ else \deny ()\) = \) = False" by (simp) lemmas MonSimps = valid_SE_def unit_SE_def bind_SE_def lemmas Psplits = option.splits unit.splits prod.splits decision.splits lemmas PolSimps = valid_SE_def unit_SE_def bind_SE_def if_splits policy2MON_def SE_LR_RBAC_ST_Policy_def map_add_def id_def LRsimps prod_2_def RBACPolicy_def SE_LR_Policy_def SEPolicy_def RBAC_def deleteEntrySE_def editEntrySE_def readEntrySE_def \0_def \0_def UC0_def patient1_def patient2_def LR1_def alice_def bob_def charlie_def get_entry_def SE_LR_RBAC_Policy_def Allow_def Deny_def dom_restrict_def policy_range_comp_def prod_orA_def prod_orD_def ST_Allow_def ST_Deny_def Spine0_def SCR1_def SCR2_def entry1_def entry2_def entry3_def FunPolicy_def SE_LR_FUN_Policy_def o_def image_def UPFDefs lemma "SE_LR_RBAC_Policy ((createSCR alice Clerical patient1),\0)= Some (deny ())" by (simp add: PolSimps) lemma exBool[simp]: "\a::bool. a" by auto lemma deny_allow[simp]: " \deny ()\ \ Some ` range allow" by auto lemma allow_deny[simp]: " \allow ()\ \ Some ` range deny" by auto text\Policy as monad. Alice using her first urp can read the SCR of patient1.\ lemma "(\0 \ (os \ mbind [(createSCR alice Clerical patient1)] (PolMon); (return (os = [(deny (Out) )]))))" by (simp add: PolMon_def MonSimps PolSimps) text\Presenting her other urp, she is not allowed to read it.\ lemma "SE_LR_RBAC_Policy ((appendEntry alice Clerical patient1 ei d),\0)= \deny ()\" by (simp add: PolSimps) end