diff --git a/thys/Interpreter_Optimizations/AList_Extra.thy b/thys/Interpreter_Optimizations/AList_Extra.thy new file mode 100644 --- /dev/null +++ b/thys/Interpreter_Optimizations/AList_Extra.thy @@ -0,0 +1,61 @@ +theory AList_Extra + imports "HOL-Library.AList" List_util +begin + +lemma list_all2_rel_prod_updateI: + assumes "list_all2 (rel_prod (=) R) xs ys" and "R xval yval" + shows "list_all2 (rel_prod (=) R) (AList.update k xval xs) (AList.update k yval ys)" + using assms(1,1,2) + by (induction xs ys rule: list.rel_induct) auto + +lemma length_map_entry[simp]: "length (AList.map_entry k f al) = length al" + by (induction al) simp_all + +lemma map_entry_id0[simp]: "AList.map_entry k id = id" +proof (rule ext) + fix xs + show "AList.map_entry k id xs = id xs" + by (induction xs) auto +qed + +lemma map_entry_id: "AList.map_entry k id xs = xs" + by simp + +lemma map_entry_map_of_Some_conv: + "map_of xs k = Some v \ AList.map_entry k f xs = AList.update k (f v) xs" + by (induction xs) auto + +lemma map_entry_map_of_None_conv: + "map_of xs k = None \ AList.map_entry k f xs = xs" + by (induction xs) auto + +lemma list_all2_rel_prod_map_entry: + assumes + "list_all2 (rel_prod (=) R) xs ys" and + "\xval yval. map_of xs k = Some xval \ map_of ys k = Some yval \ R (f xval) (g yval)" + shows "list_all2 (rel_prod (=) R) (AList.map_entry k f xs) (AList.map_entry k g ys)" + using assms(1)[THEN rel_option_map_of, of k] +proof (cases rule: option.rel_cases) + case None + thus ?thesis + using assms(1) by (simp add: map_entry_map_of_None_conv) +next + case (Some xval yval) + then show ?thesis + using assms(1,2) + by (auto simp add: map_entry_map_of_Some_conv intro!: list_all2_rel_prod_updateI) +qed + +lemmas list_all2_rel_prod_map_entry1 = list_all2_rel_prod_map_entry[where g = id, simplified] +lemmas list_all2_rel_prod_map_entry2 = list_all2_rel_prod_map_entry[where f = id, simplified] + +lemma list_all_updateI: + assumes "list_all P xs" and "P (k, v)" + shows "list_all P (AList.update k v xs)" + using assms + by (induction xs) auto + +lemma set_update: "set (AList.update k v xs) \ {(k, v)} \ set xs" + by (induction xs) auto + +end \ No newline at end of file diff --git a/thys/Interpreter_Optimizations/Dynamic.thy b/thys/Interpreter_Optimizations/Dynamic.thy --- a/thys/Interpreter_Optimizations/Dynamic.thy +++ b/thys/Interpreter_Optimizations/Dynamic.thy @@ -1,24 +1,25 @@ theory Dynamic imports Main begin locale dynval = fixes + uninitialized :: 'dyn and is_true :: "'dyn \ bool" and is_false :: "'dyn \ bool" assumes not_true_and_false: "\ (is_true d \ is_false d)" begin lemma false_not_true: "is_false d \ \ is_true d" using not_true_and_false by blast lemma true_not_false: "is_true d \ \ is_false d" using not_true_and_false by blast lemma is_true_and_is_false_implies_False: "is_true d \ is_false d \ False" using true_not_false false_not_true by simp end end \ No newline at end of file diff --git a/thys/Interpreter_Optimizations/Env.thy b/thys/Interpreter_Optimizations/Env.thy --- a/thys/Interpreter_Optimizations/Env.thy +++ b/thys/Interpreter_Optimizations/Env.thy @@ -1,60 +1,164 @@ theory Env - imports Main + imports Main "HOL-Library.Library" begin +section \Generic lemmas\ + +lemma map_of_list_allI: + assumes "\k v. f k = Some v \ P (k, v)" and + "\k v. map_of kvs k = Some v \ f k = Some v" and + "distinct (map fst kvs)" + shows "list_all P kvs" + using assms(2-) +proof (induction kvs) + case Nil + then show ?case by simp +next + case (Cons kv kvs) + from Cons.prems(1) have "f (fst kv) = Some (snd kv)" + by simp + hence"P kv" + using assms(1) + by (cases kv; simp) + moreover have "list_all P kvs" + using Cons.IH Cons.prems + by (metis Some_eq_map_of_iff distinct.simps(2) list.set_intros(2) list.simps(9)) + ultimately show ?case by simp +qed + section \Environment\ locale env = fixes empty :: 'env and get :: "'env \ 'key \ 'val option" and add :: "'env \ 'key \ 'val \ 'env" and to_list :: "'env \ ('key \ 'val) list" assumes get_empty: "get empty x = None" and get_add_eq: "get (add e x v) x = Some v" and get_add_neq: "x \ y \ get (add e x v) y = get e y" and - to_list_correct: "map_of (to_list e) = get e" + to_list_correct: "map_of (to_list e) = get e" and + to_list_distinct: "distinct (map fst (to_list e))" begin declare get_empty[simp] declare get_add_eq[simp] declare get_add_neq[simp] definition singleton where "singleton \ add empty" fun add_list :: "'env \ ('key \ 'val) list \ 'env" where "add_list e [] = e" | "add_list e (x # xs) = add (add_list e xs) (fst x) (snd x)" definition from_list :: "('key \ 'val) list \ 'env" where "from_list \ add_list empty" -lemma from_list_correct: "map_of xs k = get (from_list xs) k" +lemma from_list_correct: "get (from_list xs) = map_of xs" proof (induction xs) case Nil then show ?case using get_empty by (simp add: from_list_def) next case (Cons x xs) - then show ?case - using get_add_eq get_add_neq by (simp add: from_list_def) + show ?case + using get_add_eq get_add_neq Cons.IH by (auto simp: from_list_def) qed lemma from_list_Nil[simp]: "from_list [] = empty" by (simp add: from_list_def) lemma get_from_list_to_list: "get (from_list (to_list e)) = get e" proof fix x show "get (from_list (to_list e)) x = get e x" - unfolding from_list_correct[symmetric] - unfolding to_list_correct[symmetric] - by simp + unfolding from_list_correct + unfolding to_list_correct + by (rule refl) +qed + +lemma to_list_list_allI: + assumes "\k v. get e k = Some v \ P (k, v)" + shows "list_all P (to_list e)" +proof (rule map_of_list_allI[of "get e"]) + fix k v + show "get e k = Some v \ P (k, v)" + using assms by simp +next + fix k v + show "map_of (to_list e) k = Some v \ get e k = Some v" + unfolding to_list_correct by assumption +next + show "distinct (map fst (to_list e))" + by (rule to_list_distinct) +qed + +definition map_entry where + "map_entry e k f \ case get e k of None \ e | Some v \ add e k (f v)" + +lemma get_map_entry_eq[simp]: "get (map_entry e k f) k = map_option f (get e k)" + unfolding map_entry_def + by (cases "get e k") simp_all + +lemma get_map_entry_neq[simp]: "x \ y \ get (map_entry e x f) y = get e y" + unfolding map_entry_def + by (cases "get e x") simp_all + +lemma dom_map_entry[simp]: "dom (get (map_entry e k f)) = dom (get e)" + unfolding dom_def + apply (rule Collect_cong) + by (metis None_eq_map_option_iff get_map_entry_eq get_map_entry_neq) + +lemma get_map_entry_conv: + "get (map_entry e x f) y = map_option (\v. if x = y then f v else v) (get e y)" + unfolding map_entry_def + by (cases "get e x"; cases "x = y"; simp add: option.map_ident) + +lemma map_option_comp_map_entry: + assumes "\x \ ran (get e). f (g x) = f x" + shows "map_option f \ get (map_entry e k g) = map_option f \ get e" +proof (intro ext) + fix x + show "(map_option f \ get (map_entry e k g)) x = (map_option f \ get e) x" + proof (cases "k = x") + case True + thus ?thesis + using assms(1) + by (auto simp: get_map_entry_eq option.map_comp intro!: map_option_cong ranI) + next + case False + then show ?thesis + by (simp add: get_map_entry_neq) + qed +qed + +lemma map_option_comp_get_add: + assumes "k \ dom (get e)" and "\x \ ran (get e). f v = f x" + shows "map_option f \ get (add e k v) = map_option f \ get e" +proof (intro ext) + fix x + show "(map_option f \ get (add e k v)) x = (map_option f \ get e) x" + proof (cases "x = k") + case True + show ?thesis + proof (cases "get e x") + case None + thus ?thesis + using True assms(1) by auto + next + case (Some v') + thus ?thesis + using True assms(2) by (auto intro: ranI) + qed + next + case False + thus ?thesis by simp + qed qed end end \ No newline at end of file diff --git a/thys/Interpreter_Optimizations/Env_list.thy b/thys/Interpreter_Optimizations/Env_list.thy --- a/thys/Interpreter_Optimizations/Env_list.thy +++ b/thys/Interpreter_Optimizations/Env_list.thy @@ -1,60 +1,93 @@ theory Env_list - imports Env + imports Env "HOL-Library.Library" begin + +subsection \Generic lemmas\ + +lemma map_of_filter: + "x \ y \ map_of (filter (\z. fst z \ y) zs) x = map_of zs x" + by (induction zs) auto + + subsection \List-based implementation of environment\ context begin -qualified -datatype ('key, 'val) t = T "('key \ 'val) list" +qualified type_synonym ('key, 'val) t = "('key \ 'val) list" qualified definition empty :: "('key, 'val) t" where - "empty \ T []" + "empty \ []" -qualified fun get :: "('key, 'val) t \ 'key \ 'val option" where - "get (T xs) = Map.map_of xs" +qualified definition get :: "('key, 'val) t \ 'key \ 'val option" where + "get \ map_of" -qualified fun add :: "('key, 'val) t \ 'key \ 'val \ ('key, 'val) t" where - "add (T xs) k v = T ((k, v) # xs)" +qualified definition add :: "('key, 'val) t \ 'key \ 'val \ ('key, 'val) t" where + "add e k v \ AList.update k v e" + +term filter qualified fun to_list :: "('key, 'val) t \ ('key \ 'val) list" where - "to_list (T xs) = xs" + "to_list [] = []" | + "to_list (x # xs) = x # to_list (filter (\(k, v). k \ fst x) xs)" lemma get_empty: "get empty x = None" - by (simp add: empty_def) + unfolding get_def empty_def + by simp lemma get_add_eq: "get (add e x v) x = Some v" - by (cases e; simp) + unfolding get_def add_def + by (simp add: update_Some_unfold) lemma get_add_neq: "x \ y \ get (add e x v) y = get e y" - by (cases e; simp) + unfolding get_def add_def + by (simp add: update_conv') lemma to_list_correct: "map_of (to_list e) = get e" -proof (rule ext) - fix k - obtain xs where "e = T xs" - by (cases e; simp) - show "map_of (to_list e) k = get e k" - unfolding \e = T xs\ - by (induction xs) simp_all + unfolding get_def +proof (induction e rule: to_list.induct) + case 1 + then show ?case by simp +next + case (2 x xs) + show ?case + proof (rule ext) + fix k' + show "map_of (to_list (x # xs)) k' = map_of (x # xs) k'" + using "2.IH" map_of_filter + by (auto simp add: case_prod_beta') + qed +qed + +lemma set_to_list: "set (to_list e) \ set e" + by (induction e rule: to_list.induct) auto + +lemma to_list_distinct: "distinct (map fst (to_list e))" +proof (induction e rule: to_list.induct) + case 1 + then show ?case by simp +next + case (2 x xs) + then show ?case + using set_to_list + by fastforce qed end global_interpretation env_list: env Env_list.empty Env_list.get Env_list.add Env_list.to_list defines singleton = env_list.singleton and add_list = env_list.add_list and from_list = env_list.from_list apply (unfold_locales) - by (simp_all add: get_empty get_add_eq get_add_neq to_list_correct) + by (simp_all add: get_empty get_add_eq get_add_neq to_list_correct to_list_distinct) export_code Env_list.empty Env_list.get Env_list.add Env_list.to_list singleton add_list from_list in SML module_name Env end \ No newline at end of file diff --git a/thys/Interpreter_Optimizations/Global.thy b/thys/Interpreter_Optimizations/Global.thy --- a/thys/Interpreter_Optimizations/Global.thy +++ b/thys/Interpreter_Optimizations/Global.thy @@ -1,76 +1,442 @@ theory Global - imports Main Result + imports "HOL-Library.Library" Result Env List_util Option_Extra Map_Extra AList_Extra begin sledgehammer_params [timeout = 30] -sledgehammer_params [provers = "cvc4 e spass vampire z3"] +sledgehammer_params [provers = "cvc4 e spass vampire z3 zipperposition"] + +declare K_record_comp[simp] lemma if_then_Some_else_None_eq[simp]: "(if a then Some b else None) = Some c \ a \ b = c" "(if a then Some b else None) = None \ \ a" by (cases a) simp_all lemma if_then_else_distributive: "(if a then f b else f c) = f (if a then b else c)" by simp -fun traverse :: "('a \ 'b option) \ 'a list \ 'b list option" where - "traverse f [] = Some []" | - "traverse f (x # xs) = do { - x' \ f x; - xs' \ traverse f xs; - Some (x' # xs') - }" -lemma traverse_length: "traverse f xs = Some ys \ length ys = length xs" -proof (induction xs arbitrary: ys) +section \Rest\ + +lemma map_ofD: + fixes xs k opt + assumes "map_of xs k = opt" + shows "opt = None \ (\n < length xs. opt = Some (snd (xs ! n)))" + using assms + by (metis in_set_conv_nth map_of_SomeD not_Some_eq snd_conv) + +lemma list_all2_assoc_map_rel_get: + assumes "list_all2 (=) (map fst xs) (map fst ys)" and "list_all2 R (map snd xs) (map snd ys)" + shows "rel_option R (map_of xs k) (map_of ys k)" + using assms[unfolded list.rel_map] +proof (induction xs ys rule: list.rel_induct) case Nil - then show ?case by simp + thus ?case by simp next - case (Cons x xs) - then show ?case - by (auto simp add: Option.bind_eq_Some_conv) + case (Cons x xs y ys) + thus ?case by (cases "k = fst x") auto qed -datatype 'instr fundef = - Fundef (body: "'instr list") (arity: nat) +subsection \Function definition\ -lemma rel_fundef_arities: "rel_fundef r gd1 gd2 \ arity gd1 = arity gd2" +datatype ('label, 'instr) fundef = + Fundef (body: "('label \ 'instr list) list") (arity: nat) (return: nat) (fundef_locals: nat) + +lemma rel_fundef_arities: "rel_fundef r1 r2 gd1 gd2 \ arity gd1 = arity gd2" + by (simp add: fundef.rel_sel) + +lemma rel_fundef_return: "rel_fundef R1 R2 gd1 gd2 \ return gd1 = return gd2" + by (simp add: fundef.rel_sel) + +lemma rel_fundef_locals: "rel_fundef R1 R2 gd1 gd2 \ fundef_locals gd1 = fundef_locals gd2" by (simp add: fundef.rel_sel) lemma rel_fundef_body_length[simp]: - "rel_fundef r fd1 fd2 \ length (body fd1) = length (body fd2)" + "rel_fundef r1 r2 fd1 fd2 \ length (body fd1) = length (body fd2)" by (auto intro: list_all2_lengthD simp add: fundef.rel_sel) +definition funtype where + "funtype fd \ (arity fd, return fd)" + +lemma rel_fundef_funtype[simp]: "rel_fundef R1 R2 fd1 fd2 \ funtype fd1 = funtype fd2" + by (simp add: fundef.rel_sel funtype_def) + +lemma rel_fundef_rel_fst_hd_bodies: + assumes "rel_fundef R1 R2 fd1 fd2" and "body fd1 \ [] \ body fd2 \ []" + shows "R1 (fst (hd (body fd1))) (fst (hd (body fd2)))" + using assms + unfolding fundef.rel_sel + by (auto intro: list_all2_rel_prod_fst_hd[THEN conjunct1]) + +lemma map_option_comp_conv: + assumes + "\x. rel_option R (f x) (g x)" + "\fd1 fd2. fd1 \ ran f \ fd2 \ ran g \ R fd1 fd2 \ h fd1 = i fd2" + shows "map_option h \ f = map_option i \ g" +proof (intro ext) + fix x + show "(map_option h \ f) x = (map_option i \ g) x" + using assms(1)[of x] + by (cases rule: option.rel_cases) (auto intro: ranI assms(2)) +qed + +lemma map_option_arity_comp_conv: + assumes "(\x. rel_option (rel_fundef R S) (f x) (g x))" + shows "map_option arity \ f = map_option arity \ g" +proof (rule map_option_comp_conv) + fix x show "rel_option (rel_fundef R S) (f x) (g x)" + by (rule assms(1)) +next + fix fd1 fd2 + show "rel_fundef R S fd1 fd2 \ arity fd1 = arity fd2 " + by (rule rel_fundef_arities) +qed + +definition wf_fundef where + "wf_fundef fd \ body fd \ [] \ distinct (map fst (body fd))" + +lemma wf_fundef_non_empty_bodyD[dest,intro]: "wf_fundef fd \ body fd \ []" + by (simp add: wf_fundef_def) + +lemma wf_fundef_distinct_basic_blocksD[dest,intro]: "wf_fundef fd \ distinct (map fst (body fd))" + by (simp add: wf_fundef_def) + +lemma wf_fundef_last_basic_block: + assumes "wf_fundef fd" and "last (body fd) = (l, instrs)" + shows "map_of (body fd) l = Some instrs" + using assms by (cases fd; auto simp: wf_fundef_def dest: last_in_set) + +definition wf_fundefs where + "wf_fundefs F \ (\f fd. F f = Some fd \ wf_fundef fd)" + +lemma wf_fundefsI: + assumes "\f fd. F f = Some fd \ wf_fundef fd" + shows "wf_fundefs F" + using assms by (simp add: wf_fundefs_def) + +lemma wf_fundefsI': + assumes "\f. pred_option wf_fundef (F f)" + shows "wf_fundefs F" + using assms unfolding wf_fundefs_def + by (metis option.pred_inject(2)) + +lemma wf_fundefs_imp_wf_fundef: + assumes "wf_fundefs F" and "F f = Some fd" + shows "wf_fundef fd" + using assms by (simp add: wf_fundefs_def) + +hide_fact wf_fundefs_def + +subsection \Program\ + datatype ('fenv, 'henv, 'fun) prog = - Prog (fun_env: 'fenv) (heap: 'henv) (main_fun: 'fun) + Prog (prog_fundefs: 'fenv) (heap: 'henv) (main_fun: 'fun) -datatype ('fun, 'operand) frame = - Frame 'fun (prog_counter: nat) (operand_stack: "'operand list") +definition wf_prog where + "wf_prog Get p \ wf_fundefs (Get (prog_fundefs p))" + +subsection \Stack frame\ + +datatype ('fun, 'label, 'operand) frame = + Frame 'fun 'label (prog_counter: nat) (regs: "'operand list") (operand_stack: "'operand list") + +definition instr_at where + "instr_at fd label pc = + (case map_of (body fd) label of + Some instrs \ + if pc < length instrs then + Some (instrs ! pc) + else + None + | None \ None)" + +lemma instr_atD: + assumes "instr_at fd l pc = Some instr" + shows "\instrs. map_of (body fd) l = Some instrs \ pc < length instrs \ instrs ! pc = instr" + using assms + by (cases "map_of (body fd) l") (auto simp: instr_at_def) + +lemma rel_fundef_imp_rel_option_instr_at: + assumes "rel_fundef (=) R fd1 fd2" + shows "rel_option R (instr_at fd1 l pc) (instr_at fd2 l pc)" + using assms +proof (cases rule: fundef.rel_cases) + case (Fundef bblocks1 _ _ bblocks2 ) + then show ?thesis + by (auto simp: instr_at_def list_all2_lengthD + intro: list_all2_nthD2 elim!: option.rel_cases dest!: rel_option_map_of[of _ _ _ l]) +qed + +definition next_instr where + "next_instr F f label pc \ do { + fd \ F f; + instr_at fd label pc + }" + +lemma next_instr_eq_Some_conv: + "next_instr F f l pc = Some instr \ (\fd. F f = Some fd \ instr_at fd l pc = Some instr)" + by (simp add: next_instr_def bind_eq_Some_conv) + +lemma next_instrD: + assumes "next_instr F f l pc = Some instr" + shows "\fd. F f = Some fd \ instr_at fd l pc = Some instr" + using assms unfolding next_instr_def + by (cases "F f"; simp) + +lemma next_instr_pc_lt_length_instrsI: + assumes + "next_instr F f l pc = Some instr" and + "F f = Some fd" and + "map_of (body fd) l = Some instrs" + shows "pc < length instrs" + using assms + by (auto dest!: next_instrD instr_atD) + +lemma next_instr_get_map_ofD: + assumes + "next_instr F f l pc = Some instr" and + "F f = Some fd" and + "map_of (body fd) l = Some instrs" + shows "pc < length instrs" and "instrs ! pc = instr" + using assms + by (auto dest!: next_instrD instr_atD) + +lemma next_instr_length_instrs: + assumes + "F f = Some fd" and + "map_of (body fd) label = Some instrs" + shows "next_instr F f label (length instrs) = None" + by (simp add: assms next_instr_def instr_at_def) + +lemma next_instr_take_Suc_conv: + assumes "next_instr F f l pc = Some instr" and + "F f = Some fd" and + "map_of (body fd) l = Some instrs" + shows "take (Suc pc) instrs = take pc instrs @ [instr]" + using assms + by (auto simp: take_Suc_conv_app_nth dest!: next_instrD instr_atD) + + +subsection \Dynamic state\ datatype ('fenv, 'henv, 'frame) state = - State (fun_env: 'fenv) (heap: 'henv) (callstack: "'frame list") - -definition rewrite :: "'instr list \ nat \ 'instr \ 'instr list" where - "rewrite p pc i = list_update p pc i" - -fun rewrite_fundef_body :: "'instr fundef \ nat \ 'instr \ 'instr fundef" where - "rewrite_fundef_body (Fundef xs ar) n x = Fundef (rewrite xs n x) ar" + State (state_fundefs: 'fenv) (heap: 'henv) (callstack: "'frame list") -lemmas length_rewrite[simp] = length_list_update[folded rewrite_def] -lemmas nth_rewrite_eq[simp] = nth_list_update_eq[folded rewrite_def] -lemmas nth_rewrite_neq[simp] = nth_list_update_neq[folded rewrite_def] -lemmas take_rewrite[simp] = take_update_cancel[folded rewrite_def] -lemmas take_rewrite_swap = take_update_swap[folded rewrite_def] -lemmas map_rewrite = map_update[folded rewrite_def] -lemmas list_all2_rewrite_cong[intro] = list_all2_update_cong[folded rewrite_def] +definition state_ok where + "state_ok Get s \ wf_fundefs (Get (state_fundefs s))" -lemma body_rewrite_fundef_body[simp]: "body (rewrite_fundef_body fd n x) = rewrite (body fd) n x" - by (cases fd) simp +inductive final for F_get Return where + finalI: "next_instr (F_get F) f l pc = Some Return \ + final F_get Return (State F H [Frame f l pc R \])" -lemma arity_rewrite_fundef_body[simp]: "arity (rewrite_fundef_body fd n x) = arity fd" - by (cases fd) simp +definition allocate_frame where + "allocate_frame f fd args uninitialized \ + Frame f (fst (hd (body fd))) 0 (args @ replicate (fundef_locals fd) uninitialized) []" -lemma if_eq_const_conv: "(if x then y else z) = w \ x \ y = w \ \x \ z = w" +inductive load for F_get Uninitialized where + "F_get F main = Some fd \ arity fd = 0 \ body fd \ [] \ + load F_get Uninitialized (Prog F H main) (State F H [allocate_frame main fd [] Uninitialized])" + +lemma length_neq_imp_not_list_all2: + assumes "length xs \ length ys" + shows "\ list_all2 R xs ys" +proof (rule notI) + assume "list_all2 R xs ys" + hence "length xs = length ys" + by (rule list_all2_lengthD) + thus False + using assms by contradiction +qed + +lemma list_all2_rel_prod_conv: + "list_all2 (rel_prod R S) xs ys \ + list_all2 R (map fst xs) (map fst ys) \ list_all2 S (map snd xs) (map snd ys)" +proof (cases "length xs = length ys") + case True + then show ?thesis + by (induction xs ys rule: list_induct2) (auto simp: rel_prod_sel) +next + case False + hence neq_length_maps: + "length (map fst xs) \ length (map fst ys)" + "length (map snd xs) \ length (map snd ys)" + by simp_all + show ?thesis + using length_neq_imp_not_list_all2[OF False] + using neq_length_maps[THEN length_neq_imp_not_list_all2] + by simp +qed + +definition rewrite_fundef_body :: + "('label, 'instr) fundef \ 'label \ nat \ 'instr \ ('label, 'instr) fundef" where + "rewrite_fundef_body fd l n instr = + (case fd of Fundef bblocks ar ret locals \ + Fundef (AList.map_entry l (\instrs. instrs[n := instr]) bblocks) ar ret locals)" + +lemma rewrite_fundef_body_cases[case_names invalid_label valid_label]: + assumes + "\bs ar ret locals. fd = Fundef bs ar ret locals \ map_of bs l = None \ P" + "\bs ar ret locals instrs. fd = Fundef bs ar ret locals \ map_of bs l = Some instrs \ P" + shows "P" + using assms + by (cases fd; cases "map_of (body fd) l") auto + +lemma arity_rewrite_fundef_body[simp]: "arity (rewrite_fundef_body fd l pc instr) = arity fd" + by (cases fd; simp add: rewrite_fundef_body_def option.case_eq_if) + +lemma return_rewrite_fundef_body[simp]: "return (rewrite_fundef_body fd l pc instr) = return fd" + by (cases fd) (simp add: rewrite_fundef_body_def option.case_eq_if) + +lemma funtype_rewrite_fundef_body[simp]: "funtype (rewrite_fundef_body fd l pc instr') = funtype fd" + by (simp add: funtype_def) + +lemma length_body_rewrite_fundef_body[simp]: + "length (body (rewrite_fundef_body fd l pc instr)) = length (body fd)" + by (cases fd) (simp add: rewrite_fundef_body_def) + +lemma prod_in_set_fst_image_conv: "(x, y) \ set xys \ x \ fst ` set xys" + by (metis fstI image_eqI) + +lemma map_of_body_rewrite_fundef_body_conv_neq[simp]: + assumes "l \ l'" + shows "map_of (body (rewrite_fundef_body fd l pc instr)) l' = map_of (body fd) l'" + using assms + by (cases fd) (simp add: rewrite_fundef_body_def map_of_map_entry) + +lemma map_of_body_rewrite_fundef_body_conv_eq[simp]: + "map_of (body (rewrite_fundef_body fd l pc instr)) l = + map_option (\xs. xs[pc := instr]) (map_of (body fd) l)" + by (cases fd) (simp add: rewrite_fundef_body_def map_of_map_entry map_option_case) + +lemma instr_at_rewrite_fundef_body_conv: + "instr_at (rewrite_fundef_body fd l' pc' instr') l pc = + map_option (\instr. if l = l' \ pc = pc' then instr' else instr) (instr_at fd l pc)" +proof (cases fd) + case (Fundef bblocks ar ret locals) + then show ?thesis + proof (cases "instr_at fd l pc") + case None + then show ?thesis + by (cases "map_of bblocks l") + (auto simp add: Fundef rewrite_fundef_body_def instr_at_def map_of_map_entry) + next + case (Some instrs) + then show ?thesis + by (cases "map_of bblocks l") + (auto simp add: Fundef rewrite_fundef_body_def instr_at_def map_of_map_entry) + qed +qed + +lemma fundef_locals_rewrite_fundef_body[simp]: + "fundef_locals (rewrite_fundef_body fd l pc instr) = fundef_locals fd" + by (cases fd; simp add: rewrite_fundef_body_def option.case_eq_if) + +lemma rel_fundef_rewrite_body1: + assumes + rel_fd1_fd2: "rel_fundef (=) R fd1 fd2" and + instr_at_l_pc: "instr_at fd1 l pc = Some instr" and + R_iff: "\x. R instr x \ R instr' x" + shows "rel_fundef (=) R (rewrite_fundef_body fd1 l pc instr') fd2" + using assms(1) +proof (cases rule: fundef.rel_cases) + case (Fundef xs ar' ret' locals' ys ar ret locals) + show ?thesis + using Fundef(3,3,1,2,4-) instr_at_l_pc + proof (induction xs ys arbitrary: fd1 fd2 rule: list.rel_induct) + case Nil + then show ?case by (simp add: rewrite_fundef_body_def) + next + case (Cons x xs y ys) + then show ?case + apply (simp add: rewrite_fundef_body_def) + apply safe + using assms(3) + apply (smt (verit, ccfv_SIG) fun_upd_apply fundef.sel(1) instr_atD list_all2_lengthD + list_all2_nthD2 list_all2_update1_cong map_of.simps(2) option.simps(1) prod.sel(1,2) + rel_prod_sel) + by (simp add: instr_at_def) + qed +qed + +lemma rel_fundef_rewrite_body: + assumes rel_fd1_fd2: "rel_fundef (=) R fd1 fd2" and R_i1_i2: "R i1 i2" + shows "rel_fundef (=) R (rewrite_fundef_body fd1 l pc i1) (rewrite_fundef_body fd2 l pc i2)" + using assms(1) +proof (cases rule: fundef.rel_cases) + case (Fundef bblocks1 ar1 ret1 locals1 bblocks2 ar2 ret2 locals2) + then show ?thesis + using R_i1_i2 + by (auto simp: rewrite_fundef_body_def + intro!: list_all2_update_cong + intro!: list_all2_rel_prod_map_entry + dest: rel_option_map_of[of _ _ _ l]) +qed + +lemma rewrite_fundef_body_triv: + "instr_at fd l pc = Some instr \ rewrite_fundef_body fd l pc instr = fd" + by (cases fd) + (auto simp add: rewrite_fundef_body_def map_entry_map_of_Some_conv update_triv dest: instr_atD) + +lemma rel_fundef_rewrite_body2: + assumes + rel_fd1_fd2: "rel_fundef (=) R fd1 fd2" and + instr_at_l_pc: "instr_at fd2 l pc = Some instr" and + R_iff: "\x. R x instr \ R x instr'" + shows "rel_fundef (=) R fd1 (rewrite_fundef_body fd2 l pc instr')" + using assms(1) +proof (cases rule: fundef.rel_cases) + case (Fundef xs ar' ret' locals' ys ar ret locals) + moreover obtain instrs2 where + "map_of (body fd2) l = Some instrs2" and "pc < length instrs2" and "instrs2 ! pc = instr" + using instr_atD[OF instr_at_l_pc] by auto + moreover then obtain instrs1 instr1 where + "map_of (body fd1) l = Some instrs1" and "pc < length instrs1" and "instrs1 ! pc = instr1" + using rel_fundef_imp_rel_option_instr_at[OF rel_fd1_fd2, of l pc] + apply (auto simp add: instr_at_def option_rel_Some2) + by (metis instr_atD instr_at_def) + moreover have "list_all2 R instrs1 instrs2" + using Fundef \map_of (body fd2) l = Some instrs2\ \map_of (body fd1) l = Some instrs1\ + by (auto dest: rel_option_map_of[of _ _ _ l]) + moreover hence "R (instrs1 ! pc) (instrs2 ! pc)" + using \pc < length instrs1\ + by (auto intro: list_all2_nthD) + ultimately show ?thesis + by (auto simp: rewrite_fundef_body_def R_iff + intro!: list_all2_rel_prod_map_entry2 + intro: list_all2_update2_cong) +qed + +lemma rel_fundef_rewrite_body2': + assumes + rel_fd1_fd2: "rel_fundef (=) R fd1 fd2" and + instr_at1: "instr_at fd1 l pc = Some instr1" and + R_instr1_instr2: "R instr1 instr2'" + shows "rel_fundef (=) R fd1 (rewrite_fundef_body fd2 l pc instr2')" + using assms(1) +proof (cases rule: fundef.rel_cases) + case (Fundef bblocks1 ar1 ret1 locals1 bblocks2 ar2 ret2 locals2) + moreover obtain instrs1 where + "map_of bblocks1 l = Some instrs1" and "pc < length instrs1" and "instrs1 ! pc = instr1" + using Fundef instr_at1[THEN instr_atD] by auto + + moreover then obtain instrs2 instr2 where + "map_of (body fd2) l = Some instrs2" and "pc < length instrs2" and "instrs2 ! pc = instr2" + using Fundef rel_fd1_fd2 + apply (auto simp: option_rel_Some1 dest!: rel_option_map_of[where l = l]) + by (metis list_all2_lengthD) + ultimately show ?thesis + using R_instr1_instr2 + by (auto simp: rewrite_fundef_body_def + intro!: list_all2_update2_cong list_all2_rel_prod_map_entry2 + dest: rel_option_map_of[of _ _ _ l]) +qed + +thm rel_fundef_rewrite_body + +lemma if_eq_const_conv: "(if x then y else z) = w \ x \ y = w \ \ x \ z = w" by simp +lemma const_eq_if_conv: "w = (if x then y else z) \ x \ w = y \ \ x \ w = z" + by auto + end diff --git a/thys/Interpreter_Optimizations/Inca.thy b/thys/Interpreter_Optimizations/Inca.thy --- a/thys/Interpreter_Optimizations/Inca.thy +++ b/thys/Interpreter_Optimizations/Inca.thy @@ -1,142 +1,167 @@ theory Inca imports Global OpInl Env Dynamic "VeriComp.Language" begin + section \Inline caching\ subsection \Static representation\ -datatype ('dyn, 'var, 'fun, 'op, 'opinl) instr = +datatype ('dyn, 'var, 'fun, 'label, 'op, 'opinl) instr = IPush 'dyn | IPop | + IGet nat | + ISet nat | ILoad 'var | IStore 'var | IOp 'op | IOpInl 'opinl | - ICJump nat | - ICall 'fun - -subsection \Dynamic representation\ + is_jump: ICJump 'label 'label | + ICall 'fun | + is_return: IReturn locale inca = Fenv: env F_empty F_get F_add F_to_list + Henv: env heap_empty heap_get heap_add heap_to_list + - dynval is_true is_false + + dynval uninitialized is_true is_false + nary_operations_inl \\ \\\\\ \\\\\ \\\ \\\\\ \
\\\\ for + \ \Functions environment\ F_empty and - F_get :: "'fenv \ 'fun \ ('dyn, 'var, 'fun, 'op, 'opinl) instr fundef option" and + F_get :: "'fenv \ 'fun \ ('label, ('dyn, 'var, 'fun, 'label, 'op, 'opinl) instr) fundef option" and F_add and F_to_list and + + \ \Memory heap\ heap_empty and heap_get :: "'henv \ 'var \ 'dyn \ 'dyn option" and heap_add and heap_to_list and - is_true :: "'dyn \ bool" and is_false and + + \ \Dynamic values\ + uninitialized :: 'dyn and is_true and is_false and + + \ \n-ary operations\ \\ :: "'op \ 'dyn list \ 'dyn" and \\\\\ and \\\\\ and \\\ and \\\\\ and \
\\\\ :: "'opinl \ 'op" begin -inductive final :: "('fenv, 'henv, ('fun, 'dyn) frame) state \ bool" where - "F_get F f = Some fd \ pc = length (body fd) \ final (State F H [Frame f pc \])" subsection \Semantics\ -inductive step :: - "('fenv, 'henv, ('fun, 'dyn) frame) state \ ('fenv, 'henv, ('fun, 'dyn) frame) state \ bool" (infix "\" 55) where - +inductive step (infix "\" 55) where step_push: - "F_get F f = Some fd \ pc < length (body fd) \ body fd ! pc = IPush d \ - State F H (Frame f pc \ # st) \ State F H (Frame f (Suc pc) (d # \) # st)" | + "next_instr (F_get F) f l pc = Some (IPush d) \ + State F H (Frame f l pc R \ # st) \ State F H (Frame f l (Suc pc) R (d # \) # st)" | step_pop: - "F_get F f = Some fd \ pc < length (body fd) \ body fd ! pc = IPop \ - State F H (Frame f pc (d # \) # st) \ State F H (Frame f (Suc pc) \ # st)" | + "next_instr (F_get F) f l pc = Some IPop \ + State F H (Frame f l pc R (d # \) # st) \ State F H (Frame f l (Suc pc) R \ # st)" | + + step_get: + "next_instr (F_get F) f l pc = Some (IGet n) \ + n < length R \ d = R ! n \ + State F H (Frame f l pc R \ # st) \ State F H (Frame f l (Suc pc) R (d # \) # st)" | + + step_set: + "next_instr (F_get F) f l pc = Some (ISet n) \ + n < length R \ R' = R[n := d] \ + State F H (Frame f l pc R (d # \) # st) \ State F H (Frame f l (Suc pc) R' \ # st)" | step_load: - "F_get F f = Some fd \ pc < length (body fd) \ body fd ! pc = ILoad x \ + "next_instr (F_get F) f l pc = Some (ILoad x) \ heap_get H (x, y) = Some d \ - State F H (Frame f pc (y # \) # st) \ State F H (Frame f (Suc pc) (d # \) # st)" | + State F H (Frame f l pc R (y # \) # st) \ State F H (Frame f l (Suc pc) R (d # \) # st)" | step_store: - "F_get F f = Some fd \ pc < length (body fd) \ body fd ! pc = IStore x \ + "next_instr (F_get F) f l pc = Some (IStore x) \ heap_add H (x, y) d = H' \ - State F H (Frame f pc (y # d # \) # st) \ State F H' (Frame f (Suc pc) \ # st)" | + State F H (Frame f l pc R (y # d # \) # st) \ State F H' (Frame f l (Suc pc) R \ # st)" | step_op: - "F_get F f = Some fd \ pc < length (body fd) \ body fd ! pc = IOp op \ + "next_instr (F_get F) f l pc = Some (IOp op) \ \\\\\ op = ar \ ar \ length \ \ \\\ op (take ar \) = None \ \\ op (take ar \) = x \ - State F H (Frame f pc \ # st) \ State F H (Frame f (Suc pc) (x # drop ar \) # st)" | + State F H (Frame f l pc R \ # st) \ State F H (Frame f l (Suc pc) R (x # drop ar \) # st)" | step_op_inl: - "F_get F f = Some fd \ pc < length (body fd) \ body fd ! pc = IOp op \ + "next_instr (F_get F) f l pc = Some (IOp op) \ \\\\\ op = ar \ ar \ length \ \ \\\ op (take ar \) = Some opinl \ \\\\\ opinl (take ar \) = x \ - F_add F f (rewrite_fundef_body fd pc (IOpInl opinl)) = F' \ - State F H (Frame f pc \ # st) \ State F' H (Frame f (Suc pc) (x # drop ar \) # st)" | + F' = Fenv.map_entry F f (\fd. rewrite_fundef_body fd l pc (IOpInl opinl)) \ + State F H (Frame f l pc R \ # st) \ State F' H (Frame f l (Suc pc) R (x # drop ar \) # st)" | step_op_inl_hit: - "F_get F f = Some fd \ pc < length (body fd) \ body fd ! pc = IOpInl opinl \ + "next_instr (F_get F) f l pc = Some (IOpInl opinl) \ \\\\\ (\
\\\\ opinl) = ar \ ar \ length \ \ \\\\\ opinl (take ar \) \ \\\\\ opinl (take ar \) = x \ - State F H (Frame f pc \ # st) \ State F H (Frame f (Suc pc) (x # drop ar \) # st)" | + State F H (Frame f l pc R \ # st) \ State F H (Frame f l (Suc pc) R (x # drop ar \) # st)" | step_op_inl_miss: - "F_get F f = Some fd \ pc < length (body fd) \ body fd ! pc = IOpInl opinl \ + "next_instr (F_get F) f l pc = Some (IOpInl opinl) \ \\\\\ (\
\\\\ opinl) = ar \ ar \ length \ \ \ \\\\\ opinl (take ar \) \ \\\\\ opinl (take ar \) = x \ - F_add F f (rewrite_fundef_body fd pc (IOp (\
\\\\ opinl))) = F' \ - State F H (Frame f pc \ # st) \ State F' H (Frame f (Suc pc) (x # drop ar \) # st)" | - - step_cjump_true: - "F_get F f = Some fd \ pc < length (body fd) \ body fd ! pc = ICJump n \ - is_true d \ - State F H (Frame f pc (d # \) # st) \ State F H (Frame f n \ # st)" | + F' = Fenv.map_entry F f (\fd. rewrite_fundef_body fd l pc (IOp (\
\\\\ opinl))) \ + State F H (Frame f l pc R \ # st) \ State F' H (Frame f l (Suc pc) R (x # drop ar \) # st)" | - step_cjump_false: - "F_get F f = Some fd \ pc < length (body fd) \ body fd ! pc = ICJump n \ - is_false d \ - State F H (Frame f pc (d # \) # st) \ State F H (Frame f (Suc pc) \ # st)" | + step_cjump: + "next_instr (F_get F) f l pc = Some (ICJump l\<^sub>t l\<^sub>f) \ + is_true d \ l' = l\<^sub>t \ is_false d \ l' = l\<^sub>f \ + State F H (Frame f l pc R (d # \) # st) \ State F H (Frame f l' 0 R \ # st)" | - step_fun_call: - "F_get F f = Some fd \ pc < length (body fd) \ body fd ! pc = ICall g \ + step_call: + "next_instr (F_get F) f l pc = Some (ICall g) \ F_get F g = Some gd \ arity gd \ length \ \ - frame\<^sub>f = Frame f pc \ \ frame\<^sub>g = Frame g 0 (take (arity gd) \) \ - State F H (frame\<^sub>f # st) \ State F H (frame\<^sub>g # frame\<^sub>f # st)" | + frame\<^sub>g = allocate_frame g gd (take (arity gd) \) uninitialized \ + State F H (Frame f l pc R \ # st) \ State F H (frame\<^sub>g # Frame f l pc R \ # st)" | - step_fun_end: - "F_get F g = Some gd \ arity gd \ length \\<^sub>f \ pc\<^sub>g = length (body gd) \ - frame\<^sub>g = Frame g pc\<^sub>g \\<^sub>g \ frame\<^sub>f = Frame f pc\<^sub>f \\<^sub>f \ - frame\<^sub>f' = Frame f (Suc pc\<^sub>f) (\\<^sub>g @ drop (arity gd) \\<^sub>f) \ - State F H (frame\<^sub>g # frame\<^sub>f # st) \ State F H (frame\<^sub>f' # st)" + step_return: + "next_instr (F_get F) g l\<^sub>g pc\<^sub>g = Some IReturn \ + F_get F g = Some gd \ arity gd \ length \\<^sub>f \ + length \\<^sub>g = return gd \ + frame\<^sub>f' = Frame f l\<^sub>f (Suc pc\<^sub>f) R\<^sub>f (\\<^sub>g @ drop (arity gd) \\<^sub>f) \ + State F H (Frame g l\<^sub>g pc\<^sub>g R\<^sub>g \\<^sub>g # Frame f l\<^sub>f pc\<^sub>f R\<^sub>f \\<^sub>f # st) \ State F H (frame\<^sub>f' # st)" lemma step_deterministic: - "s1 \ s2 \ s1 \ s3 \ s2 = s3" - by (induction rule: step.cases; - auto elim!: step.cases dest: is_true_and_is_false_implies_False) + assumes "s1 \ s2" and "s1 \ s3" + shows "s2 = s3" + using assms + apply (cases s1 s2 rule: step.cases) + apply (auto elim: step.cases dest: is_true_and_is_false_implies_False) [3] + apply (auto elim: step.cases dest: is_true_and_is_false_implies_False) [3] + apply (auto elim: step.cases dest: is_true_and_is_false_implies_False) [3] + apply (auto elim: step.cases dest: is_true_and_is_false_implies_False) [3] + apply (auto elim: step.cases dest: is_true_and_is_false_implies_False) [1] + done -lemma final_finished: "final s \ finished step s" +lemma step_right_unique: "right_unique step" + using step_deterministic + by (rule right_uniqueI) + +lemma final_finished: + assumes "final F_get IReturn s" + shows "finished step s" unfolding finished_def -proof - assume "final s" and "\x. step s x" - then obtain s' where "step s s'" - by auto +proof (rule notI) + assume "\x. step s x" + then obtain s' where "step s s'" by auto thus False - using \final s\ - by (auto elim!: step.cases final.cases) + using \final F_get IReturn s\ + by (auto + simp: state_ok_def + wf_fundefs_imp_wf_fundef[THEN wf_fundef_last_basic_block, THEN next_instr_length_instrs[rotated]] + elim!: step.cases final.cases) qed -sublocale inca_sem: semantics step final +sublocale semantics step "final F_get IReturn" using final_finished step_deterministic by unfold_locales -inductive load :: "('fenv, 'a, 'fun) prog \ ('fenv, 'a, ('fun, 'b) frame) state \ bool" where - "F_get F main = Some fd \ arity fd = 0 \ - load (Prog F H main) (State F H [Frame main 0 []])" +definition load where + "load \ Global.load F_get uninitialized" -sublocale inca_lang: language step final load +sublocale language step "final F_get IReturn" load by unfold_locales end end \ No newline at end of file diff --git a/thys/Interpreter_Optimizations/Inca_to_Ubx_compiler.thy b/thys/Interpreter_Optimizations/Inca_to_Ubx_compiler.thy --- a/thys/Interpreter_Optimizations/Inca_to_Ubx_compiler.thy +++ b/thys/Interpreter_Optimizations/Inca_to_Ubx_compiler.thy @@ -1,558 +1,789 @@ theory Inca_to_Ubx_compiler imports Inca_to_Ubx_simulation Result "VeriComp.Compiler" "HOL-Library.Monad_Syntax" begin section \Generic program rewriting\ -context - fixes rewrite_instr :: "nat \ 'a \ 'stack \ ('err, 'b \ 'stack) result" -begin - -fun rewrite_prog :: "nat \ 'a list \ 'stack \ ('err, 'b list \ 'stack) result" where - "rewrite_prog _ [] \ = Ok ([], \)" | - "rewrite_prog n (x # xs) \ = do { - (x', \') \ rewrite_instr n x \; - (xs', \'') \ rewrite_prog (Suc n) xs \'; - Ok (x' # xs', \'') +primrec monadic_fold_map where + "monadic_fold_map f acc [] = Some (acc, [])" | + "monadic_fold_map f acc (x # xs) = do { + (acc', x') \ f acc x; + (acc'', xs') \ monadic_fold_map f acc' xs; + Some (acc'', x' # xs') }" -lemma rewrite_prog_map_f: - assumes "\x \1 n x' \2. rewrite_instr n x \1 = Ok (x', \2) \ f x' = x" - shows "rewrite_prog n xs \1 = Ok (ys, \2) \ map f ys = xs" - by (induction xs arbitrary: \1 n ys; auto simp: assms) +lemma monadic_fold_map_length: + "monadic_fold_map f acc xs = Some (acc', xs') \ length xs = length xs'" + by (induction xs arbitrary: acc xs') (auto simp: bind_eq_Some_conv) -end +lemma monadic_fold_map_ConsD[dest]: + assumes "monadic_fold_map f a (x # xs) = Some (c, ys)" + shows "\y ys' b. ys = y # ys' \ f a x = Some (b, y) \ monadic_fold_map f b xs = Some (c, ys')" + using assms + by (auto simp add: bind_eq_Some_conv) + +lemma monadic_fold_map_list_all2: + assumes "monadic_fold_map f acc xs = Some (acc', ys)" and + "\acc acc' x y. f acc x = Some (acc', y) \ P x y" + shows "list_all2 P xs ys" + using assms(1) +proof (induction xs arbitrary: acc ys) + case Nil + then show ?case by simp +next + case (Cons x xs) + show ?case + using Cons.prems + by (auto simp: bind_eq_Some_conv intro: assms(2) Cons.IH) +qed + +lemma monadic_fold_map_list_all: + assumes "monadic_fold_map f acc xs = Some (acc', ys)" and + "\acc acc' x y. f acc x = Some (acc', y) \ P y" + shows "list_all P ys" +proof - + have "list_all2 (\_. P) xs ys" + using assms + by (auto elim: monadic_fold_map_list_all2) + thus ?thesis + by (auto elim: list_rel_imp_pred2) +qed fun gen_pop_push where "gen_pop_push instr (domain, codomain) \ = ( let ar = length domain in if ar \ length \ \ take ar \ = domain then - Ok (instr, codomain # drop ar \) + Some (instr, codomain @ drop ar \) else - Error () + None )" context inca_to_ubx_simulation begin -lemma sp_rewrite_prog: - assumes - "rewrite_prog f n p1 \1 = Ok (p2, \2)" and - "\x \1 n x' \2. f n x \1 = Ok (x', \2) \ Subx.sp_instr F x' \1 = Ok \2" - shows "Subx.sp F p2 \1 = Ok \2" - using assms(1) - by (induction p1 arbitrary: \1 n p2; auto simp: assms(2)) - - section \Lifting\ -context - fixes - get_arity :: "'fun \ nat option" and - load_oracle :: "nat \ type option" -begin - fun lift_instr where - "lift_instr (Inca.IPush x) \ = Ok (IPush x, None # \)" | - "lift_instr Inca.IPop (_ # \) = Ok (IPop, \)" | - "lift_instr (Inca.ILoad x) (None # \) = Ok (ILoad x, None # \)" | - "lift_instr (Inca.IStore x) (None # None # \) = Ok (IStore x, \)" | - "lift_instr (Inca.IOp op) \ = gen_pop_push (IOp op) (replicate (\\\\\ op) None, None) \" | - "lift_instr (Inca.IOpInl opinl) \ = - gen_pop_push (IOpInl opinl) (replicate (\\\\\ (\
\\\\ opinl)) None, None) \" | - "lift_instr (Inca.ICall f) \ = do { - ar \ Result.of_option () (get_arity f); - gen_pop_push (ICall f) (replicate ar None, None) \ + "lift_instr F L ret N (Inca.IPush d) \ = Some (IPush d, None # \)" | + "lift_instr F L ret N Inca.IPop (_ # \) = Some (IPop, \)" | + "lift_instr F L ret N (Inca.IGet n) \ = (if n < N then Some (IGet n, None # \) else None)" | + "lift_instr F L ret N (Inca.ISet n) (None # \) = (if n < N then Some (ISet n, \) else None)" | + "lift_instr F L ret N (Inca.ILoad x) (None # \) = Some (ILoad x, None # \)" | + "lift_instr F L ret N (Inca.IStore x) (None # None # \) = Some (IStore x, \)" | + "lift_instr F L ret N (Inca.IOp op) \ = + gen_pop_push (IOp op) (replicate (\\\\\ op) None, [None]) \" | + "lift_instr F L ret N (Inca.IOpInl opinl) \ = + gen_pop_push (IOpInl opinl) (replicate (\\\\\ (\
\\\\ opinl)) None, [None]) \" | + "lift_instr F L ret N (Inca.ICJump l\<^sub>t l\<^sub>f) [None] = + (if List.member L l\<^sub>t \ List.member L l\<^sub>f then Some (ICJump l\<^sub>t l\<^sub>f, []) else None)" | + "lift_instr F L ret N (Inca.ICall f) \ = do { + (ar, ret) \ F f; + gen_pop_push (ICall f) (replicate ar None, replicate ret None) \ }" | - "lift_instr _ _ = Error ()" - -definition lift where - "lift = rewrite_prog (\_. lift_instr)" + "lift_instr F L ret N Inca.IReturn \ = + (if \ = replicate ret None then Some (IReturn, []) else None)" | + "lift_instr _ _ _ _ _ _ = None" -lemma sp_lift_instr: - assumes - "lift_instr instr \1 = Ok (instr', \2)" and - "\x. rel_option (\ar fd. arity fd = ar) (get_arity x) (F x)" - shows "Subx.sp_instr F instr' \1 = Ok \2" - using assms(1) -proof (induction instr \1 rule: lift_instr.induct) - fix f \ n - assume "lift_instr (Inca.ICall f) \ = Ok (instr', \2)" - thus "Subx.sp_instr F instr' \ = Ok \2" - using assms(2)[of f] - by (auto simp: option_rel_Some1) -qed (auto simp add: Let_def) +definition lift_instrs where + "lift_instrs F L ret N \ + monadic_fold_map (\\ instr. map_option prod.swap (lift_instr F L ret N instr \))" -lemma sp_lift: - assumes - "lift n p1 \1 = Ok (p2, \2)" and - "\x. rel_option (\ar fd. arity fd = ar) (get_arity x) (F x)" - shows "Subx.sp F p2 \1 = Ok \2" - using assms - by (auto elim!: sp_rewrite_prog sp_lift_instr simp: lift_def) +lemma lift_instrs_length: + assumes "lift_instrs F L ret N \i xs = Some (\o, ys)" + shows "length xs = length ys" + using assms unfolding lift_instrs_def + by (auto intro: monadic_fold_map_length) -lemma norm_lift_instr: "lift_instr x \1 = Ok (x', \2) \ norm_instr x' = x" - by (induction x \1 rule: lift_instr.induct; - auto simp: Let_def ) +lemma lift_instrs_not_Nil: "lift_instrs F L ret N \i xs = Some (\o, ys) \ xs \ [] \ ys \ []" + using lift_instrs_length by fastforce -lemma norm_lift: - assumes "lift n xs \1 = Ok (ys, \2)" - shows "map norm_instr ys = xs" - by (auto intro!: rewrite_prog_map_f[OF _ assms[unfolded lift_def]] simp: norm_lift_instr) +lemma lift_instrs_NilD[dest]: + assumes "lift_instrs F L ret N \i [] = Some (\o, ys)" + shows "\o = \i \ ys = []" + using assms + by (simp_all add: lift_instrs_def) + +lemmas Some_eq_bind_conv = + bind_eq_Some_conv[unfolded eq_commute[of "Option.bind f g" "Some x" for f g x]] + +lemma lift_instr_is_jump: + assumes "lift_instr F L ret N x \i = Some (y, \o)" + shows "Inca.is_jump x \ Ubx.is_jump y" + using assms + by (rule lift_instr.elims) + (auto simp add: if_split_eq2 Let_def Some_eq_bind_conv) + +lemma lift_instr_is_return: + assumes "lift_instr F L ret N x \i = Some (y, \o)" + shows "Inca.is_return x \ Ubx.is_return y" + using assms + by (rule lift_instr.elims) + (auto simp add: if_split_eq2 Let_def Some_eq_bind_conv) + +lemma lift_instrs_all_not_jump_not_return: + assumes "lift_instrs F L ret N \i xs = Some (\o, ys)" + shows + "list_all (\i. \ Inca.is_jump i \ \ Inca.is_return i) xs \ + list_all (\i. \ Ubx.is_jump i \ \ Ubx.is_return i) ys" + using assms +proof (induction xs arbitrary: \i \o ys) + case Nil + then show ?case by (simp add: lift_instrs_def) +next + case (Cons x xs) + from Cons.prems show ?case + apply (simp add: lift_instrs_def bind_eq_Some_conv) + apply (fold lift_instrs_def) + by (auto simp add: Cons.IH lift_instr_is_jump lift_instr_is_return) +qed + +lemma lift_instrs_all_butlast_not_jump_not_return: + assumes "lift_instrs F L ret N \i xs = Some (\o, ys)" + shows + "list_all (\i. \ Inca.is_jump i \ \ Inca.is_return i) (butlast xs) \ + list_all (\i. \ Ubx.is_jump i \ \ Ubx.is_return i) (butlast ys)" + using lift_instrs_length[OF assms(1)] assms unfolding lift_instrs_def +proof (induction xs ys arbitrary: \i \o rule: list_induct2) + case Nil + then show ?case by simp +next + case (Cons x xs y ys) + thus ?case + by (auto simp add: bind_eq_Some_conv lift_instr_is_jump lift_instr_is_return) +qed + +lemma lift_instr_sp: + assumes "lift_instr F L ret N x \i = Some (y, \o)" + shows "Subx.sp_instr F ret y \i \o" + using assms + apply (induction F L ret N x \i rule: lift_instr.induct; + auto simp: Let_def intro: Subx.sp_instr.intros) + apply (rule Subx.sp_instr.Op, metis append_take_drop_id) + apply (rule Subx.sp_instr.OpInl, metis append_take_drop_id) + apply (auto simp add: bind_eq_Some_conv intro!: Subx.sp_instr.Call, metis append_take_drop_id) + done + +lemma lift_instrs_sp: + assumes "lift_instrs F L ret N \i xs = Some (\o, ys)" + shows "Subx.sp_instrs F ret ys \i \o" + using assms unfolding lift_instrs_def +proof (induction xs arbitrary: \i \o ys) + case Nil + thus ?case by (auto intro: Subx.sp_instrs.Nil) +next + case (Cons x xs) + from Cons.prems show ?case + by (auto simp add: bind_eq_Some_conv intro: Subx.sp_instrs.Cons lift_instr_sp Cons.IH) +qed + +lemma lift_instr_fun_call_in_range: + assumes "lift_instr F L ret N x \i = Some (y, \o)" + shows "Subx.fun_call_in_range F y" + using assms + by (induction F L ret N x \i rule: lift_instr.induct) (auto simp: Let_def bind_eq_Some_conv) + +lemma lift_instrs_all_fun_call_in_range: + assumes "lift_instrs F L ret N \i xs = Some (\o, ys)" + shows "list_all (Subx.fun_call_in_range F) ys" + using assms unfolding lift_instrs_def + by (auto intro!: monadic_fold_map_list_all intro: lift_instr_fun_call_in_range) + +lemma lift_instr_local_var_in_range: + assumes "lift_instr F L ret N x \i = Some (y, \o)" + shows "Subx.local_var_in_range N y" + using assms + by (induction F L ret N x \i rule: lift_instr.induct) (auto simp: Let_def bind_eq_Some_conv) + +lemma lift_instrs_all_local_var_in_range: + assumes "lift_instrs F L ret N \i xs = Some (\o, ys)" + shows "list_all (Subx.local_var_in_range N) ys" + using assms unfolding lift_instrs_def + by (auto intro!: monadic_fold_map_list_all intro: lift_instr_local_var_in_range) + +lemma lift_instr_jump_in_range: + assumes "lift_instr F L ret N x \i = Some (y, \o)" + shows "Subx.jump_in_range (set L) y" + using assms + by (induction F L ret N x \i rule: lift_instr.induct) + (auto simp: Let_def bind_eq_Some_conv in_set_member) + +lemma lift_instrs_all_jump_in_range: + assumes "lift_instrs F L ret N \i xs = Some (\o, ys)" + shows "list_all (Subx.jump_in_range (set L)) ys" + using assms unfolding lift_instrs_def + by (auto intro!: monadic_fold_map_list_all intro: lift_instr_jump_in_range) + +lemma lift_instr_norm: + "lift_instr F L ret N instr1 \1 = Some (instr2, \2) \ norm_eq instr1 instr2" + by (induction instr1 \1 rule: lift_instr.induct) (auto simp: Let_def bind_eq_Some_conv) + +lemma lift_instrs_all_norm: + assumes "lift_instrs F L ret N \1 instrs1 = Some (\2, instrs2)" + shows "list_all2 norm_eq instrs1 instrs2" + using assms unfolding lift_instrs_def + by (auto simp: lift_instr_norm elim!: monadic_fold_map_list_all2) section \Optimization\ -fun result_alternative :: "('a, 'b) result \ ('a, 'b) result \ ('a, 'b) result" (infixr "<|>" 51) where - "result_alternative (Ok x) _ = Ok x" | - "result_alternative _ (Ok x) = Ok x" | - "result_alternative (Error e) _ = Error e" +context + fixes load_oracle :: "nat \ type option" +begin -definition try_unbox where - "try_unbox \ x \ unbox mk_instr \ - (case unbox x of Some n \ Ok (mk_instr n, Some \ # \) | None \ Error ())" +definition orelse :: "'a option \ 'a option \ 'a option" (infixr "orelse" 55) where + "x orelse y = (case x of Some x' \ Some x' | None \ y)" + +lemma orelse_eq_Some_conv: + "x orelse y = Some z \ (x = Some z \ x = None \ y = Some z)" + by (cases x) (simp_all add: orelse_def) + +lemma orelse_eq_SomeE: + assumes + "x orelse y = Some z" and + "x = Some z \ P" and + "x = None \ y = Some z \ P" + shows "P" + using assms(1) + unfolding orelse_def + by (cases x; auto intro: assms(2,3)) + +fun drop_prefix where + "drop_prefix [] ys = Some ys" | + "drop_prefix (x # xs) (y # ys) = (if x = y then drop_prefix xs ys else None)" | + "drop_prefix _ _ = None " + +lemma drop_prefix_eq_Some_conv: "drop_prefix xs ys = Some zs \ ys = xs @ zs" + by (induction xs ys arbitrary: zs rule: drop_prefix.induct) + (auto simp: if_split_eq1) fun optim_instr where - "optim_instr _ (IPush d) \ = - try_unbox Ubx1 d \ unbox_ubx1 IPushUbx1 <|> - try_unbox Ubx2 d \ unbox_ubx2 IPushUbx2 <|> - Ok (IPush d, None # \) + "optim_instr _ _ _ (IPush d) \ = + Some Pair \ (Some IPushUbx1 \ (unbox_ubx1 d)) \ Some (Some Ubx1 # \) orelse + Some Pair \ (Some IPushUbx2 \ (unbox_ubx2 d)) \ Some (Some Ubx2 # \) orelse + Some (IPush d, None # \) " | - "optim_instr _ (IPushUbx1 n) \ = Ok (IPushUbx1 n, Some Ubx1 # \)" | - "optim_instr _ (IPushUbx2 b) \ = Ok (IPushUbx2 b, Some Ubx2 # \)" | - "optim_instr _ IPop (_ # \) = Ok (IPop, \)" | - "optim_instr n (ILoad x) (None # \) = ( - case load_oracle n of - Some \ \ Ok (ILoadUbx \ x, Some \ # \) | - _ \ Ok (ILoad x, None # \) - )" | - "optim_instr _ (ILoadUbx \ x) (None # \) = Ok (ILoadUbx \ x, Some \ # \)" | - "optim_instr _ (IStore x) (None # None # \) = Ok (IStore x, \)" | - "optim_instr _ (IStore x) (None # Some \ # \) = Ok (IStoreUbx \ x, \)" | - "optim_instr _ (IStoreUbx \\<^sub>1 x) (None # Some \\<^sub>2 # \) = (if \\<^sub>1 = \\<^sub>2 then Ok (IStoreUbx \\<^sub>1 x, \) else Error ())" | - "optim_instr _ (IOp op) \ = gen_pop_push (IOp op) (replicate (\\\\\ op) None, None) \" | - "optim_instr _ (IOpInl opinl) \ = ( + "optim_instr _ _ _ (IPushUbx1 n) \ = Some (IPushUbx1 n, Some Ubx1 # \)" | + "optim_instr _ _ _ (IPushUbx2 b) \ = Some (IPushUbx2 b, Some Ubx2 # \)" | + "optim_instr _ _ _ IPop (_ # \) = Some (IPop, \)" | + "optim_instr _ _ pc (IGet n) \ = + map_option (\\. (IGetUbx \ n, Some \ # \)) (load_oracle pc) orelse + Some (IGet n, None # \)" | + "optim_instr _ _ pc (IGetUbx \ n) \ = Some (IGetUbx \ n, Some \ # \)" | + "optim_instr _ _ _ (ISet n) (None # \) = Some (ISet n, \)" | + "optim_instr _ _ _ (ISet n) (Some \ # \) = Some (ISetUbx \ n, \)" | + "optim_instr _ _ _ (ISetUbx _ n) (None # \) = Some (ISet n, \)" | + "optim_instr _ _ _ (ISetUbx _ n) (Some \ # \) = Some (ISetUbx \ n, \)" | + "optim_instr _ _ pc (ILoad x) (None # \) = + map_option (\\. (ILoadUbx \ x, Some \ # \)) (load_oracle pc) orelse + Some (ILoad x, None # \)" | + "optim_instr _ _ _ (ILoadUbx \ x) (None # \) = Some (ILoadUbx \ x, Some \ # \)" | + "optim_instr _ _ _ (IStore x) (None # None # \) = Some (IStore x, \)" | + "optim_instr _ _ _ (IStore x) (None # Some \ # \) = Some (IStoreUbx \ x, \)" | + "optim_instr _ _ _ (IStoreUbx _ x) (None # None # \) = Some (IStore x, \)" | + "optim_instr _ _ _ (IStoreUbx _ x) (None # Some \ # \) = Some (IStoreUbx \ x, \)" | + "optim_instr _ _ _ (IOp op) \ = + map_option (\\o. (IOp op, None # \o)) (drop_prefix (replicate (\\\\\ op) None) \)" | + "optim_instr _ _ _ (IOpInl opinl) \ = ( let ar = \\\\\ (\
\\\\ opinl) in if ar \ length \ then case \\\ opinl (take ar \) of - None \ gen_pop_push (IOpInl opinl) (replicate ar None, None) \ | - Some opubx \ Ok (IOpUbx opubx, snd (\\\\\\\\ opubx) # drop ar \) + None \ map_option (\\o. (IOpInl opinl, None # \o)) (drop_prefix (replicate ar None) \) | + Some opubx \ map_option (\\o. (IOpUbx opubx, snd (\\\\\\\\ opubx) # \o)) + (drop_prefix (fst (\\\\\\\\ opubx)) \) else - Error () + None )" | - "optim_instr _ (IOpUbx opubx) \ = gen_pop_push (IOpUbx opubx) (\\\\\\\\ opubx) \" | - "optim_instr _ (ICall f) \ = do { - ar \ Result.of_option () (get_arity f); - gen_pop_push (ICall f) (replicate ar None, None) \ + "optim_instr _ _ _ (IOpUbx opubx) \ = + (let p = \\\\\\\\ opubx in + map_option (\\o. (IOpUbx opubx, snd p # \o)) (drop_prefix (fst p) \))" | + "optim_instr _ _ _ (ICJump l\<^sub>t l\<^sub>f) [None] = Some (ICJump l\<^sub>t l\<^sub>f, []) " | + "optim_instr F _ _ (ICall f) \ = do { + (ar, ret) \ F f; + \o \ drop_prefix (replicate ar None) \; + Some (ICall f, replicate ret None @ \o) }" | - "optim_instr _ _ _ = Error ()" - -definition optim where - "optim \ rewrite_prog optim_instr" - -lemma - assumes - "optim 0 [IPush d\<^sub>1, IPush d\<^sub>2, IStore y] [] = Ok (xs, ys)" - "unbox_ubx1 d\<^sub>1 = Some x" - "unbox_ubx1 d\<^sub>2 = None" "unbox_ubx2 d\<^sub>2 = None" - shows "xs = [IPushUbx1 x, IPush d\<^sub>2, IStoreUbx Ubx1 y] \ ys = []" - using assms(1) - by (simp add: assms optim_def try_unbox_def) - -lemma norm_optim_instr: "optim_instr n x \1 = Ok (x', \2) \ norm_instr x' = norm_instr x" - for x \1 n x' \2 -proof (induction n x \1 rule: optim_instr.induct) - fix d \1 n - assume "optim_instr n (Ubx.IPush d) \1 = Ok (x', \2)" - thus "norm_instr x' = norm_instr (Ubx.instr.IPush d)" - using Subx.box_unbox_inverse - by (auto elim!: result_alternative.elims simp: try_unbox_def option.case_eq_if) -next - fix x \1 n - assume assms: "optim_instr n (Ubx.ILoad x) (None # \1) = Ok (x', \2)" - thus "norm_instr x' = norm_instr (Ubx.ILoad x)" - proof (cases "load_oracle n") - case None - then show ?thesis using assms by simp - next - case (Some x) - then show ?thesis - using assms by (cases x) auto - qed -next - fix opinl \1 n - assume assms: "optim_instr n (IOpInl opinl) \1 = Ok (x', \2)" - then have arity_le_\1: "\\\\\ (\
\\\\ opinl) \ length \1" - using prod.inject by fastforce + "optim_instr _ ret _ IReturn \ = (if \ = replicate ret None then Some (IReturn, []) else None)" | + "optim_instr _ _ _ _ _ = None" - show "norm_instr x' = norm_instr (IOpInl opinl)" - proof (cases "\\\ opinl (take (\\\\\ (\
\\\\ opinl)) \1)") - case None - then show ?thesis - using assms arity_le_\1 - by (simp add: Let_def) - next - case (Some opubx) - then show ?thesis - using assms arity_le_\1 - by (auto simp: case_prod_beta Subx.\\\_invertible) - qed -next - fix opubx \1 n - assume assms: "optim_instr n (IOpUbx opubx) \1 = Ok (x', \2)" - then show "norm_instr x' = norm_instr (IOpUbx opubx)" - by (cases "\\\\\\\\ opubx"; auto simp: Let_def) -qed (auto simp: Let_def) - -lemma norm_optim: - assumes "optim n xs \1 = Ok (ys, \2)" - shows "map norm_instr ys = map norm_instr xs" - using assms - unfolding optim_def - by (induction xs arbitrary: \1 n ys \2; auto simp: norm_optim_instr) +definition optim_instrs where + "optim_instrs F ret \ \pc \i instrs. + map_option (\((_, \o), instrs'). (\o, instrs')) + (monadic_fold_map (\(pc, \) instr. + map_option (\(instr', \o). ((Suc pc, \o), instr')) (optim_instr F ret pc instr \)) + (pc, \i) instrs)" -lemma sp_optim_instr: - assumes - "optim_instr n instr \1 = Ok (instr', \2)" and - "\x. rel_option (\ar fd. arity fd = ar) (get_arity x) (F x) " - shows "Subx.sp_instr F instr' \1 = Ok \2" - using assms(1) - apply (induction n instr \1 rule: optim_instr.induct; - (auto simp: Let_def; fail)?) - subgoal for _ d - by (auto elim!: result_alternative.elims simp: try_unbox_def option.case_eq_if) - subgoal for n - by (cases "load_oracle n"; auto) - subgoal for _ opinl \ - apply (simp add: Let_def) - apply (cases "\\\ opinl (take (\\\\\ (\
\\\\ opinl)) \)") - apply (auto dest: list_all_eq_const_imp_replicate - simp: case_prod_beta Let_def min.absorb2) - subgoal for opubx - using Subx.\\\\\\\\_\\\\\[of opubx] - by (cases "\\\\\\\\ opubx"; auto simp: Subx.\\\_invertible dest: Subx.\\\_opubx_type) - done - subgoal for _ opubx - by (cases "\\\\\\\\ opubx"; auto simp add: Let_def) - subgoal for _ f - using assms(2)[of f] - by (auto simp: option_rel_Some1) - done +lemma optim_instrs_length: + assumes "optim_instrs F ret pc \i xs = Some (\o, ys)" + shows "length xs = length ys" + using assms unfolding optim_instrs_def + by (auto intro: monadic_fold_map_length) -lemma sp_optim: - assumes - "optim n p1 \1 = Ok (p2, \2)" and - "\x. rel_option (\ar fd. arity fd = ar) (get_arity x) (F x)" - shows "Subx.sp F p2 \1 = Ok \2" +lemma optim_instrs_not_Nil: "optim_instrs F ret pc \i xs = Some (\o, ys) \ xs \ [] \ ys \ []" + using optim_instrs_length by fastforce + +lemma optim_instrs_NilD[dest]: + assumes "optim_instrs F ret pc \i [] = Some (\o, ys)" + shows "\o = \i \ ys = []" using assms - by (auto elim!: sp_rewrite_prog sp_optim_instr simp: optim_def) + by (simp_all add: optim_instrs_def) + +lemma optim_instrs_ConsD[dest]: + assumes "optim_instrs F ret pc \i (x # xs) = Some (\o, ys)" + shows "\y ys' \. ys = y # ys' \ + optim_instr F ret pc x \i = Some (y, \) \ + optim_instrs F ret (Suc pc) \ xs = Some (\o, ys')" + using assms + unfolding optim_instrs_def + by (auto simp: bind_eq_Some_conv) + +lemma optim_instr_norm: + assumes "optim_instr F ret pc instr1 \1 = Some (instr2, \2)" + shows "norm_instr instr1 = norm_instr instr2" + using assms + by (cases "(F, ret, pc, instr1, \1)" rule: optim_instr.cases) + (auto simp: ap_option_eq_Some_conv Let_def if_split_eq1 bind_eq_Some_conv option.case_eq_if + orelse_eq_Some_conv + dest!: Subx.box_unbox_inverse dest: Subx.\\\_invertible) + +lemma optim_instrs_all_norm: + assumes "optim_instrs F ret pc \1 instrs1 = Some (\2, instrs2)" + shows "list_all2 (\i1 i2. norm_instr i1 = norm_instr i2) instrs1 instrs2" + using assms unfolding optim_instrs_def + by (auto simp: optim_instr_norm elim!: monadic_fold_map_list_all2) + +lemma optim_instr_is_jump: + assumes "optim_instr F ret pc x \i = Some (y, \o)" + shows "is_jump x \ is_jump y" + using assms + by (cases "(F, ret, pc, x, \i)" rule: optim_instr.cases; + simp add: orelse_eq_Some_conv ap_option_eq_Some_conv bind_eq_Some_conv + Let_def if_split_eq1 option.case_eq_if; + safe; simp) + +lemma optim_instr_is_return: + assumes "optim_instr F ret pc x \i = Some (y, \o)" + shows "is_return x \ is_return y" + using assms + by (cases "(F, ret, pc, x, \i)" rule: optim_instr.cases; + simp add: orelse_eq_Some_conv ap_option_eq_Some_conv bind_eq_Some_conv + Let_def if_split_eq1 option.case_eq_if; + safe; simp) + +lemma optim_instrs_all_butlast_not_jump_not_return: + assumes "optim_instrs F ret pc \i xs = Some (\o, ys)" + shows + "list_all (\i. \ is_jump i \ \ is_return i) (butlast xs) \ + list_all (\i. \ is_jump i \ \ is_return i) (butlast ys)" + using optim_instrs_length[OF assms(1)] assms +proof (induction xs ys arbitrary: pc \i \o rule: list_induct2) + case Nil + thus ?case by simp +next + case (Cons x xs y ys) + from Cons.prems obtain \ where + optim_x: "optim_instr F ret pc x \i = Some (y, \)" and + optim_xs: "optim_instrs F ret (Suc pc) \ xs = Some (\o, ys)" + by auto + show ?case + using Cons.hyps + using optim_x optim_xs + apply (simp add: Cons.IH optim_instr_is_jump optim_instr_is_return) + by fastforce +qed + +lemma optim_instr_jump_in_range: + assumes "optim_instr F ret pc x \i = Some (y, \o)" + shows "Subx.jump_in_range L x \ Subx.jump_in_range L y" + using assms + by (cases "(F, ret, pc, x, \i)" rule: optim_instr.cases) + (auto simp: ap_option_eq_Some_conv Let_def if_split_eq1 option.case_eq_if + bind_eq_Some_conv orelse_eq_Some_conv) + +lemma optim_instrs_all_jump_in_range: + assumes "optim_instrs F ret pc \i xs = Some (\o, ys)" + shows "list_all (Subx.jump_in_range L) xs \ list_all (Subx.jump_in_range L) ys" + using assms + by (induction xs arbitrary: pc \i \o ys) (auto simp: optim_instr_jump_in_range) + +lemma optim_instr_fun_call_in_range: + assumes "optim_instr F ret pc x \i = Some (y, \o)" + shows "Subx.fun_call_in_range F x \ Subx.fun_call_in_range F y" + using assms + by (cases "(F, ret, pc, x, \i)" rule: optim_instr.cases) + (auto simp: ap_option_eq_Some_conv Let_def if_split_eq1 option.case_eq_if + bind_eq_Some_conv orelse_eq_Some_conv) + +lemma optim_instrs_all_fun_call_in_range: + assumes "optim_instrs F ret pc \i xs = Some (\o, ys)" + shows "list_all (Subx.fun_call_in_range F) xs \ list_all (Subx.fun_call_in_range F) ys" + using assms + by (induction xs arbitrary: pc \i \o ys) (auto simp: optim_instr_fun_call_in_range) + +lemma optim_instr_local_var_in_range: + assumes "optim_instr F ret pc x \i = Some (y, \o)" + shows "Subx.local_var_in_range N x \ Subx.local_var_in_range N y" + using assms + by (cases "(F, ret, pc, x, \i)" rule: optim_instr.cases) + (auto simp: ap_option_eq_Some_conv Let_def if_split_eq1 option.case_eq_if + bind_eq_Some_conv orelse_eq_Some_conv) + +lemma optim_instrs_all_local_var_in_range: + assumes "optim_instrs F ret pc \i xs = Some (\o, ys)" + shows "list_all (Subx.local_var_in_range N) xs \ list_all (Subx.local_var_in_range N) ys" + using assms + by (induction xs arbitrary: pc \i \o ys) (auto simp: optim_instr_local_var_in_range) + +lemma optim_instr_sp: + assumes "optim_instr F ret pc x \i = Some (y, \o)" + shows "Subx.sp_instr F ret y \i \o" + using assms + by (cases "(F, ret, pc, x, \i)" rule: optim_instr.cases) + (auto simp add: Let_def if_split_eq1 option.case_eq_if + simp: ap_option_eq_Some_conv orelse_eq_Some_conv drop_prefix_eq_Some_conv bind_eq_Some_conv + intro: Subx.sp_instr.intros) + +lemma optim_instrs_sp: + assumes "optim_instrs F ret pc \i xs = Some (\o, ys)" + shows "Subx.sp_instrs F ret ys \i \o" + using assms + by (induction xs arbitrary: pc \i \o ys) + (auto intro!: Subx.sp_instrs.intros optim_instr_sp) section \Compilation of function definition\ -fun compile_fundef where - "compile_fundef (Fundef instrs ar) = do { - (xs, \\<^sub>1) \ lift 0 instrs (replicate ar None :: type option list); - - \ \Ensure that the function returns a single dynamic result\ - () \ if \\<^sub>1 = [None] then Ok () else Error (); +definition lift_basic_block where + "lift_basic_block F L ret N \ + ap_map_prod Some (\i1. do { + _ \ if i1 \ [] then Some () else None; + _ \ if list_all (\i. \ Inca.is_jump i \ \ Inca.is_return i) (butlast i1) then Some () else None; + (\o, i2) \ lift_instrs F L ret N ([] :: type option list) i1; + (\o', i2') \ optim_instrs F ret 0 ([] :: type option list) i2; + (if \o' = [] then + Some i2' + else (if \o = [] then + Some i2 + else + None)) + })" - (ys, \\<^sub>2) \ optim 0 xs (replicate ar None); +lemma lift_basic_block_rel_prod_all_norm_eq: + assumes "lift_basic_block F L ret N bblock1 = Some bblock2" + shows "rel_prod (=) (list_all2 norm_eq) bblock1 bblock2" + using assms + unfolding lift_basic_block_def + apply (auto simp add: ap_map_prod_eq_Some_conv bind_eq_Some_conv + simp: if_split_eq1 + intro: lift_instrs_all_norm + dest!: optim_instrs_all_norm lift_instrs_all_norm) + subgoal for _ xs zs ys + using list_all2_trans[of norm_eq "\i. norm_eq (norm_instr i)" norm_eq xs ys zs, simplified] + by simp + done - Ok (Fundef ( - if \\<^sub>2 = [None] then - ys \ \use optimization\ - else - xs \ \cancel optimization\ - ) ar) +lemma list_all_iff_butlast_last: + assumes "xs \ []" + shows "list_all P xs \ list_all P (butlast xs) \ P (last xs)" + using assms + by (induction xs) auto + +lemma lift_basic_block_wf: + assumes "lift_basic_block F L ret N x = Some y" + shows "Subx.wf_basic_block F (set L) ret N y" +proof - + obtain f instrs1 \2 instrs2 \3 instrs3 instrs4 where + x_def: "x = (f, instrs1)" and + y_def: "y = (f, instrs4)" and + "instrs1 \ []" and + all_not_jump_not_return_instrs1: + "list_all (\i. \ Inca.is_jump i \ \ Inca.is_return i) (butlast instrs1)" and + lift_instrs1: "lift_instrs F L ret N ([] :: type option list) instrs1 = Some (\2, instrs2)" and + optim_instrs2: "optim_instrs F ret 0 ([] :: type option list) instrs2 = Some (\3, instrs3)" and + instr4_defs: "instrs4 = instrs3 \ \3 = [] \ instrs4 = instrs2 \ \2 = []" + using assms + unfolding lift_basic_block_def ap_map_prod_eq_Some_conv + apply (auto simp: bind_eq_Some_conv if_split_eq1) + by auto + + have "instrs4 \ []" + using instr4_defs \instrs1 \ []\ + using lift_instrs_not_Nil[OF lift_instrs1] + using optim_instrs_not_Nil[OF optim_instrs2] + by meson + moreover have "list_all (Subx.local_var_in_range N) instrs4" + using instr4_defs lift_instrs1 optim_instrs2 + by (auto dest: lift_instrs_all_local_var_in_range simp: optim_instrs_all_local_var_in_range) + moreover have "list_all (Subx.fun_call_in_range F) instrs4" + using instr4_defs lift_instrs1 optim_instrs2 + by (auto dest: lift_instrs_all_fun_call_in_range simp: optim_instrs_all_fun_call_in_range) + moreover have "list_all (Subx.jump_in_range (set L)) instrs4" + using instr4_defs lift_instrs1 optim_instrs2 + by (auto dest: lift_instrs_all_jump_in_range simp: optim_instrs_all_jump_in_range) + moreover have "list_all (\i. \ Ubx.instr.is_jump i \ \ Ubx.instr.is_return i) (butlast instrs4)" + using instr4_defs lift_instrs1 optim_instrs2 all_not_jump_not_return_instrs1 + by (auto simp: + lift_instrs_all_butlast_not_jump_not_return + optim_instrs_all_butlast_not_jump_not_return) + moreover have "Subx.sp_instrs F ret instrs4 [] []" + using instr4_defs lift_instrs1 optim_instrs2 + by (auto intro: lift_instrs_sp optim_instrs_sp) + ultimately show ?thesis + by (auto simp: y_def intro!: Subx.wf_basic_blockI) +qed + +fun compile_fundef where + "compile_fundef F (Fundef bblocks1 ar ret locals) = do { + _ \ if bblocks1 = [] then None else Some (); + bblocks2 \ ap_map_list (lift_basic_block F (map fst bblocks1) ret (ar + locals)) bblocks1; + Some (Fundef bblocks2 ar ret locals) }" -lemma rel_compile_fundef: - assumes "compile_fundef fd1 = Ok fd2" - shows "rel_fundef norm_eq fd1 fd2" -proof (cases fd1) - case (Fundef xs ar) - obtain ys zs \2 where - lift_xs: "lift 0 xs (replicate ar None :: type option list) = Ok (ys, [None])" and - optim_ys: "optim 0 ys (replicate ar None :: type option list) = Ok (zs, \2)" and - check: "fd2 = Fundef (if \2 = [None] then zs else ys) ar" - using assms unfolding Fundef - unfolding compile_fundef.simps - by (auto simp: Nitpick.case_unit_unfold if_then_else_distributive) +lemma compile_fundef_arities: "compile_fundef F fd1 = Some fd2 \ arity fd1 = arity fd2" + by (cases fd1) (auto simp: bind_eq_Some_conv) - then have "map norm_instr ys = xs" - by (auto intro: norm_lift) +lemma compile_fundef_returns: "compile_fundef F fd1 = Some fd2 \ return fd1 = return fd2" + by (cases fd1) (auto simp: bind_eq_Some_conv) - show ?thesis - proof (cases "\2 = [None]") - case True - then show ?thesis - using True Fundef - using check norm_lift[OF lift_xs, symmetric] - using norm_optim[OF optim_ys, symmetric] - by (simp add: list.rel_map list.rel_refl) +lemma compile_fundef_locals: + "compile_fundef F fd1 = Some fd2 \ fundef_locals fd1 = fundef_locals fd2" + by (cases fd1) (auto simp: bind_eq_Some_conv) + +lemma if_then_None_else_Some_eq[simp]: + "(if a then None else Some b) = Some c \ \ a \ b = c" + "(if a then None else Some b) = None \ a" + by (cases a) simp_all + +lemma + assumes "compile_fundef F fd1 = Some fd2" + shows + rel_compile_fundef: "rel_fundef (=) norm_eq fd1 fd2" (is ?REL) and + wf_compile_fundef: "Subx.wf_fundef F fd2" (is ?WF) + unfolding atomize_conj +proof (cases fd1) + case (Fundef bblocks1 ar ret locals) + with assms obtain bblocks2 where + "bblocks1 \ []" and + lift_bblocks1: + "ap_map_list (lift_basic_block F (map fst bblocks1) ret (ar + locals)) bblocks1 = Some bblocks2" and + fd2_def: "fd2 = Fundef bblocks2 ar ret locals" + by (auto simp add: bind_eq_Some_conv) + + show "?REL \ ?WF" + proof (rule conjI) + show ?REL + unfolding Fundef fd2_def + proof (rule fundef.rel_intros) + show "list_all2 (rel_prod (=) (list_all2 norm_eq)) bblocks1 bblocks2" + using lift_bblocks1 + unfolding ap_map_list_iff_list_all2 + by (auto elim: list.rel_mono_strong intro: lift_basic_block_rel_prod_all_norm_eq) + qed simp_all next - case False - then show ?thesis - using Fundef check norm_lift[OF lift_xs, symmetric] - by (simp add: list.rel_map list.rel_refl) + have "bblocks2 \ []" + using \bblocks1 \ []\ length_ap_map_list[OF lift_bblocks1] by force + moreover have "list_all (Subx.wf_basic_block F (fst ` set bblocks1) ret (ar + locals)) bblocks2" + using lift_bblocks1 + unfolding ap_map_list_iff_list_all2 + by (auto elim!: list_rel_imp_pred2 dest: lift_basic_block_wf) + moreover have "fst ` set bblocks1 = fst ` set bblocks2" + using lift_bblocks1 + unfolding ap_map_list_iff_list_all2 + by (induction bblocks1 bblocks2 rule: list.rel_induct) + (auto simp add: lift_basic_block_def ap_map_prod_eq_Some_conv) + ultimately show ?WF + unfolding fd2_def + by (auto intro: Subx.wf_fundefI) qed qed -lemma sp_compile_fundef: - assumes - "compile_fundef fd1 = Ok fd2" and - "\x. rel_option (\ar fd. arity fd = ar) (get_arity x) (F x)" - shows "sp_fundef F fd2 (body fd2) = Ok [None]" -proof (cases fd1) - case (Fundef xs ar) - with assms show ?thesis - by (auto elim!: sp_optim sp_lift simp: sp_fundef_def Nitpick.case_unit_unfold if_eq_const_conv) -qed +end end -end locale inca_ubx_compiler = inca_to_ubx_simulation Finca_empty Finca_get for Finca_empty and Finca_get :: "_ \ 'fun \ _ option" + fixes load_oracle :: "'fun \ nat \ type option" begin + section \Compilation of function environment\ definition compile_env_entry where - "compile_env_entry \ \ x = map_result id (Pair (fst x)) (compile_fundef \ (\ (fst x)) (snd x))" + "compile_env_entry F \ \p. ap_map_prod Some (compile_fundef (load_oracle (fst p)) F) p" lemma rel_compile_env_entry: - assumes "compile_env_entry \ \ (f, fd1) = Ok (f, fd2)" - shows "rel_fundef norm_eq fd1 fd2" + assumes "compile_env_entry F (f, fd1) = Some (f, fd2)" + shows "rel_fundef (=) norm_eq fd1 fd2" using assms unfolding compile_env_entry_def - using rel_compile_fundef - by auto - -lemma sp_compile_env_entry: - assumes - "compile_env_entry \ \ (f, fd1) = Ok (f, fd2)" and - "\x. rel_option (\ar fd. arity fd = ar) (\ x) (F x)" - shows "sp_fundef F fd2 (body fd2) = Ok [None]" - using assms unfolding compile_env_entry_def - by (auto elim: sp_compile_fundef) + by (auto simp: ap_map_prod_eq_Some_conv intro!: rel_compile_fundef) definition compile_env where - "compile_env \ \ e \ - map_result id Subx.Fenv.from_list - (Result.those (map (compile_env_entry \ \) (Finca_to_list e))) " + "compile_env e \ do { + let fundefs1 = Finca_to_list e; + fundefs2 \ ap_map_list (compile_env_entry (map_option funtype \ Finca_get e)) fundefs1; + Some (Subx.Fenv.from_list fundefs2) + }" -lemma list_assoc_those_compile_env_entries: - "Result.those (map (compile_env_entry \ \) xs) = Ok ys \ - rel_option (\fd1 fd2. compile_fundef \ (\ f) fd1 = Ok fd2) (map_of xs f) (map_of ys f)" +lemma rel_ap_map_list_ap_map_list_compile_env_entries: + assumes "ap_map_list (compile_env_entry F) xs = Some ys" + shows "rel_fundefs (Finca_get (Sinca.Fenv.from_list xs)) (Fubx_get (Subx.Fenv.from_list ys))" + using assms proof (induction xs arbitrary: ys) case Nil - then show ?case by simp -next - case (Cons x xs) - then obtain y ys' where - xs_def: "ys = y # ys'" and - comp_x: "compile_env_entry \ \ x = Ok y" and - comp_xs: "Result.those (map (compile_env_entry \ \) xs) = Ok ys'" - by auto - - obtain g fd1 fd2 where prods: "x = (g, fd1)" "y = (g, fd2)" - using comp_x - by (cases x) (auto simp: compile_env_entry_def) - - have "compile_fundef \ (\ g) fd1 = Ok fd2" - using comp_x unfolding prods compile_env_entry_def - by auto - - then show ?case - unfolding prods xs_def - using Cons.IH[OF comp_xs] - by simp -qed - -lemma compile_env_rel_compile_fundef: - assumes "compile_env \ \ F1 = Ok F2" - shows "rel_option (\fd1 fd2. compile_fundef \ (\ f) fd1 = Ok fd2) (Finca_get F1 f) (Fubx_get F2 f)" -proof - - obtain xs where those_xs: "Result.those (map (compile_env_entry \ \) (Finca_to_list F1)) = Ok xs" - using assms[unfolded compile_env_def] - by auto - hence F2_def: "F2 = Subx.Fenv.from_list xs" - using assms[unfolded compile_env_def] - by simp - show ?thesis - proof (cases "map_of (Finca_to_list F1) f") - case None - then show ?thesis - unfolding F2_def Subx.Fenv.from_list_correct[symmetric] - unfolding Sinca.Fenv.to_list_correct[symmetric] - using list_assoc_those_compile_env_entries[OF those_xs, of f] - by simp - next - case (Some fd1) - then show ?thesis - unfolding F2_def Subx.Fenv.from_list_correct[symmetric] - unfolding Sinca.Fenv.to_list_correct[symmetric] - using list_assoc_those_compile_env_entries[OF those_xs, of f] - by simp - qed -qed - -lemma rel_those_compile_env_entries: - "Result.those (map (compile_env_entry \ \) xs) = Ok ys \ - rel_fundefs (Finca_get (Sinca.Fenv.from_list xs)) (Fubx_get (Subx.Fenv.from_list ys))" -proof (induction xs arbitrary: ys) - case Nil - then show ?case + thus ?case using rel_fundefs_empty by simp next case (Cons x xs) - then obtain y ys' where - xs_def: "ys = y # ys'" and - comp_x: "compile_env_entry \ \ x = Ok y" and - comp_xs: "Result.those (map (compile_env_entry \ \) xs) = Ok ys'" - by auto - - obtain f fd1 fd2 where prods: "x = (f, fd1)" "y = (f, fd2)" - using comp_x - by (cases x) (auto simp: compile_env_entry_def) + from Cons.prems obtain y ys' where + ys_def: "ys = y # ys'" and + compile_env_x: "compile_env_entry F x = Some y" and + compile_env_xs: "ap_map_list (compile_env_entry F) xs = Some ys'" + by (auto simp add: ap_option_eq_Some_conv) - have "rel_fundef norm_eq fd1 fd2" - using rel_compile_env_entry[OF comp_x[unfolded prods]] . + obtain f fd1 fd2 where + prods: "x = (f, fd1)" "y = (f, fd2)" and "compile_fundef (load_oracle f) F fd1 = Some fd2" + using compile_env_x + by (cases x) (auto simp: compile_env_entry_def eq_fst_iff ap_map_prod_eq_Some_conv) - then show ?case - unfolding prods xs_def - unfolding Sinca.Fenv.from_list_correct[symmetric] Subx.Fenv.from_list_correct[symmetric] - unfolding rel_fundefs_def - apply auto - using Cons.IH[OF comp_xs] - unfolding Sinca.Fenv.from_list_correct[symmetric] Subx.Fenv.from_list_correct[symmetric] - by (simp add: rel_fundefs_def) + have "rel_fundef (=) norm_eq fd1 fd2" + using compile_env_x[unfolded prods] + by (auto intro: rel_compile_env_entry) + thus ?case + using Cons.IH[OF compile_env_xs, THEN rel_fundefsD] + unfolding prods ys_def + unfolding Sinca.Fenv.from_list_correct Subx.Fenv.from_list_correct + by (auto intro: rel_fundefsI) qed lemma rel_fundefs_compile_env: - assumes "compile_env \ \ e = Ok e'" - shows "rel_fundefs (Finca_get e) (Fubx_get e')" + assumes "compile_env F1 = Some F2" + shows "rel_fundefs (Finca_get F1) (Fubx_get F2)" proof - - obtain xs where - FOO: "Result.those (map (compile_env_entry \ \) (Finca_to_list e)) = Ok xs" and - BAR: "e' = Subx.Fenv.from_list xs" - using assms - by (auto simp: compile_env_def) + from assms obtain xs where + ap_map_list_F1: "ap_map_list (compile_env_entry (map_option funtype \ Finca_get F1)) (Finca_to_list F1) = Some xs" and + F2_def: "F2 = Subx.Fenv.from_list xs" + by (auto simp: compile_env_def bind_eq_Some_conv) show ?thesis - using rel_those_compile_env_entries[OF FOO] - unfolding BAR - unfolding Sinca.Fenv.get_from_list_to_list - . + using rel_ap_map_list_ap_map_list_compile_env_entries[OF ap_map_list_F1] + unfolding F2_def Sinca.Fenv.get_from_list_to_list + by assumption qed -lemma sp_fundefs_compile_env: - assumes - "compile_env \ \ F1 = Ok F2" and - "\x. rel_option (\ar fd. arity fd = ar) (\ x) (Finca_get F1 x)" - shows "sp_fundefs (Fubx_get F2)" - unfolding sp_fundefs_def - proof (intro allI impI) - fix f fd2 - assume F2_f: "Fubx_get F2 f = Some fd2" - then obtain fd1 where - "Finca_get F1 f = Some fd1" and - compile_fd1: "compile_fundef \ (\ f) fd1 = Ok fd2" - using compile_env_rel_compile_fundef[OF assms(1), of f] - by (auto simp: option_rel_Some2) - - note rel_F1_F2 = rel_fundefs_compile_env[OF assms(1)] - - have "rel_option (\ar fd. arity fd = ar) (\ f) (Fubx_get F2 f)" for f - using assms(2)[of f] - proof (cases rule: option.rel_cases) - case None - then show ?thesis - using rel_fundefs_None1[OF rel_F1_F2] - by simp - next - case (Some ar fd1) - then obtain fd2 where - "Fubx_get F2 f = Some fd2" and - rel_fd1_fd2: "rel_fundef (\x y. x = norm_instr y) fd1 fd2" - using rel_fundefs_Some1[OF rel_F1_F2] - by (meson rel_fundefs_Some1) - with Some show ?thesis - by (simp add: rel_fundef_arities) - qed - thus "sp_fundef (Fubx_get F2) fd2 (body fd2) = Ok [None]" - by (auto intro!: sp_compile_fundef[OF compile_fd1]) - qed - section \Compilation of program\ fun compile where - "compile (Prog F H f) = Result.to_option (do { - F' \ compile_env (map_option arity \ Finca_get F) load_oracle F; - Ok (Prog F' H f) - })" + "compile (Prog F1 H f) = Some Prog \ compile_env F1 \ Some H \ Some f" + +lemma ap_map_list_cong: + assumes "\x. x \ set ys \ f x = g x" and "xs = ys" + shows "ap_map_list f xs = ap_map_list g ys" + using assms + by (induction xs arbitrary: ys) auto + +lemma compile_env_wf_fundefs: + assumes "compile_env F1 = Some F2" + shows "Subx.wf_fundefs (Fubx_get F2)" +proof (intro Subx.wf_fundefsI allI) + fix f + obtain xs where + ap_map_list_F1: "ap_map_list (compile_env_entry (map_option funtype \ Finca_get F1)) (Finca_to_list F1) = Some xs" and + F2_def: "F2 = Subx.Fenv.from_list xs" + using assms by (auto simp: compile_env_def bind_eq_Some_conv) + + have rel_map_of_F1_xs: + "\f. rel_option (\x y. compile_fundef (load_oracle f) (map_option funtype \ Finca_get F1) x = Some y) + (map_of (Finca_to_list F1) f) (map_of xs f)" + using ap_map_list_F1 + by (auto simp: compile_env_entry_def ap_map_prod_eq_Some_conv + dest: ap_map_list_imp_rel_option_map_of) + + have funtype_F1_eq_funtype_F2: + "map_option funtype \ Finca_get F1 = map_option funtype \ Fubx_get F2" + proof (rule ext, simp) + fix x + show "map_option funtype (Finca_get F1 x) = map_option funtype (Fubx_get F2 x)" + unfolding F2_def Subx.Fenv.from_list_correct Sinca.Fenv.to_list_correct[symmetric] + using rel_map_of_F1_xs[of x] + by (cases rule: option.rel_cases) + (simp_all add: funtype_def compile_fundef_arities compile_fundef_returns) + qed + + show "pred_option (Subx.wf_fundef (map_option funtype \ Fubx_get F2)) (Fubx_get F2 f) " + proof (cases "map_of (Finca_to_list F1) f") + case None + thus ?thesis + using rel_map_of_F1_xs[of f, unfolded None] + by (simp add: F2_def Subx.Fenv.from_list_correct) + next + case (Some fd1) + show ?thesis + using rel_map_of_F1_xs[of f, unfolded Some option_rel_Some1] + unfolding funtype_F1_eq_funtype_F2 F2_def Subx.Fenv.from_list_correct + by (auto intro: wf_compile_fundef) + qed +qed lemma compile_load: assumes compile_p1: "compile p1 = Some p2" and load: "Subx.load p2 s2" shows "\s1. Sinca.load p1 s1 \ match s1 s2" proof - obtain F1 H main where p1_def: "p1 = Prog F1 H main" by (cases p1) simp then obtain F2 where - compile_F1: "compile_env (map_option arity \ Finca_get F1) load_oracle F1 = Ok F2" and + compile_F1: "compile_env F1 = Some F2" and p2_def: "p2 = Prog F2 H main" using compile_p1 - by auto + by (auto simp: ap_option_eq_Some_conv) note rel_F1_F2 = rel_fundefs_compile_env[OF compile_F1] show ?thesis - using assms(2) unfolding p2_def - proof (cases _ s2 rule: Subx.load.cases) + using assms(2) unfolding p2_def Subx.load_def + proof (cases _ _ _ s2 rule: Global.load.cases) case (1 fd2) - let ?s1 = "State F1 H [Frame main 0 []]" + then obtain fd1 where + F1_main: "Finca_get F1 main = Some fd1" and rel_fd1_fd2: "rel_fundef (=) norm_eq fd1 fd2" + using rel_fundefs_Some2[OF rel_F1_F2] + by auto + + let ?s1 = "State F1 H [allocate_frame main fd1 [] uninitialized]" - from 1 obtain fd1 where - Fstd_main: "Finca_get F1 main = Some fd1" and - "compile_fundef (map_option arity \ Finca_get F1) (load_oracle main) fd1 = Ok fd2" - using compile_env_rel_compile_fundef[OF compile_F1, of main] - by (auto simp: option_rel_Some2) - with 1 have "arity fd1 = 0" - using fundef.rel_sel rel_compile_fundef by metis - hence "Sinca.load p1 ?s1" - unfolding p1_def - by (auto intro!: Sinca.load.intros[OF Fstd_main]) - moreover have "?s1 \ s2" - proof - - have "sp_fundefs (Fubx_get F2)" - by (auto intro: sp_fundefs_compile_env[OF compile_F1] simp: option.rel_map option.rel_refl) - moreover have "rel_stacktraces (Fubx_get F2) [Frame main 0 []] [Frame main 0 []] None" - using 1 by (auto intro!: rel_stacktraces.intros simp: sp_fundef_def) - ultimately show ?thesis + show ?thesis + proof (intro exI conjI) + show "Sinca.load p1 ?s1" + unfolding Sinca.load_def p1_def + using 1 F1_main rel_fd1_fd2 + by (auto simp: rel_fundef_arities intro!: Global.load.intros dest: rel_fundef_body_length) + next + have "Subx.wf_state s2" unfolding 1 - using rel_F1_F2 - by (auto intro!: match.intros) + using compile_F1 + by (auto intro!: Subx.wf_stateI intro: compile_env_wf_fundefs) + then show "match ?s1 s2" + using 1 rel_F1_F2 rel_fd1_fd2 + by (auto simp: allocate_frame_def rel_fundef_locals + simp: rel_fundef_rel_fst_hd_bodies[OF rel_fd1_fd2 disjI2] + intro!: match.intros rel_stacktraces.intros intro: Subx.sp_instrs.Nil) qed - ultimately show ?thesis by auto qed qed interpretation std_to_inca_compiler: - compiler Sinca.step Subx.step Sinca.final Subx.final Sinca.load Subx.load + compiler Sinca.step Subx.step "final Finca_get Inca.IReturn" "final Fubx_get Ubx.IReturn" + Sinca.load Subx.load "\_ _. False" "\_. match" compile using compile_load by unfold_locales auto end -end +end \ No newline at end of file diff --git a/thys/Interpreter_Optimizations/Inca_to_Ubx_simulation.thy b/thys/Interpreter_Optimizations/Inca_to_Ubx_simulation.thy --- a/thys/Interpreter_Optimizations/Inca_to_Ubx_simulation.thy +++ b/thys/Interpreter_Optimizations/Inca_to_Ubx_simulation.thy @@ -1,1934 +1,2377 @@ theory Inca_to_Ubx_simulation - imports List_util Option_applicative Result + imports List_util Result "VeriComp.Simulation" - Inca Ubx Ubx_type_inference Unboxed_lemmas + Inca Ubx Ubx_Verification Unboxed_lemmas begin +lemma take_:"Suc n = length xs \ take n xs = butlast xs" + using butlast_conv_take diff_Suc_1 append_butlast_last_id + by (metis butlast_conv_take diff_Suc_1) + +lemma append_take_singleton_conv:"Suc n = length xs \ xs = take n xs @ [xs ! n]" +proof (induction xs arbitrary: n) + case Nil + then show ?case by simp +next + case (Cons x xs) + then show ?case + proof (cases n) + case 0 + then show ?thesis + using Cons + by simp + next + case (Suc n') + have "Suc n' = length xs" + by (rule Cons.prems[unfolded Suc, simplified]) + from Suc show ?thesis + by (auto intro: Cons.IH[OF \Suc n' = length xs\]) + qed +qed + section \Locale imports\ locale inca_to_ubx_simulation = Sinca: inca Finca_empty Finca_get Finca_add Finca_to_list heap_empty heap_get heap_add heap_to_list - is_true is_false + uninitialized is_true is_false \\ \\\\\ \\\\\ \\\ \\\\\ \
\\\\ + - Subx: ubx_sp + Subx: ubx Fubx_empty Fubx_get Fubx_add Fubx_to_list heap_empty heap_get heap_add heap_to_list - is_true is_false box_ubx1 unbox_ubx1 box_ubx2 unbox_ubx2 + uninitialized is_true is_false + box_ubx1 unbox_ubx1 box_ubx2 unbox_ubx2 \\ \\\\\ \\\\\ \\\ \\\\\ \
\\\\ \\\\\ \\\ \\\ \\\\\\\\ for \ \Functions environments\ Finca_empty and - Finca_get :: "'fenv_inca \ 'fun \ ('dyn, 'var, 'fun, 'op, 'opinl) Inca.instr fundef option" and + Finca_get :: "'fenv_inca \ 'fun \ ('label, ('dyn, 'var, 'fun, 'label, 'op, 'opinl) Inca.instr) fundef option" and Finca_add and Finca_to_list and Fubx_empty and - Fubx_get :: "'fenv_ubx \ 'fun \ ('dyn, 'var, 'fun, 'op, 'opinl, 'opubx, 'ubx1, 'ubx2) Ubx.instr fundef option" and + Fubx_get :: "'fenv_ubx \ 'fun \ ('label, ('dyn, 'var, 'fun, 'label, 'op, 'opinl, 'opubx, 'ubx1, 'ubx2) Ubx.instr) fundef option" and Fubx_add and Fubx_to_list and \ \Memory heap\ heap_empty and heap_get :: "'henv \ 'var \ 'dyn \ 'dyn option" and heap_add and heap_to_list and + \ \Dynamic values\ + uninitialized :: 'dyn and is_true and is_false and + \ \Unboxed values\ - is_true and is_false and box_ubx1 and unbox_ubx1 and box_ubx2 and unbox_ubx2 and \ \n-ary operations\ \\ and \\\\\ and \\\\\ and \\\ and \\\\\ and \
\\\\ and \\\\\ and \\\ and \\\ and \\\\\\\\ begin -declare Subx.sp_append[simp] - - -section \Strongest postcondition\ - -definition sp_fundef where - "sp_fundef F fd xs \ Subx.sp F xs (replicate (arity fd) None)" - -lemmas sp_fundef_generalize = - Subx.sp_generalize[where \ = "replicate (arity fd) None" for fd, simplified, folded sp_fundef_def] - -lemma eq_sp_to_eq_sp_fundef: - assumes "Subx.sp F1 = (Subx.sp F2 :: ('dyn, 'id, _, _, _, _, 'num, 'bool) Ubx.instr list \ _)" - shows "sp_fundef F1 = (sp_fundef F2 :: _ \ ('dyn, 'id, _, _, _, _, 'num, 'bool) Ubx.instr list \ _)" - using assms(1) - by (intro ext; simp add: sp_fundef_def) - -definition sp_fundefs where - "sp_fundefs F \ \f fd. F f = Some fd \ sp_fundef F fd (body fd) = Ok [None]" section \Normalization\ fun norm_instr where "norm_instr (Ubx.IPush d) = Inca.IPush d" | "norm_instr (Ubx.IPushUbx1 n) = Inca.IPush (box_ubx1 n)" | "norm_instr (Ubx.IPushUbx2 b) = Inca.IPush (box_ubx2 b)" | "norm_instr Ubx.IPop = Inca.IPop" | + "norm_instr (Ubx.IGet n) = Inca.IGet n" | + "norm_instr (Ubx.IGetUbx _ n) = Inca.IGet n" | + "norm_instr (Ubx.ISet n) = Inca.ISet n" | + "norm_instr (Ubx.ISetUbx _ n) = Inca.ISet n" | "norm_instr (Ubx.ILoad x) = Inca.ILoad x" | "norm_instr (Ubx.ILoadUbx _ x) = Inca.ILoad x" | "norm_instr (Ubx.IStore x) = Inca.IStore x" | "norm_instr (Ubx.IStoreUbx _ x) = Inca.IStore x" | "norm_instr (Ubx.IOp op) = Inca.IOp op" | "norm_instr (Ubx.IOpInl op) = Inca.IOpInl op" | "norm_instr (Ubx.IOpUbx op) = Inca.IOpInl (\\\ op)" | - "norm_instr (Ubx.ICJump n) = Inca.ICJump n" | - "norm_instr (Ubx.ICall x) = Inca.ICall x" + "norm_instr (Ubx.ICJump l\<^sub>t l\<^sub>f) = Inca.ICJump l\<^sub>t l\<^sub>f" | + "norm_instr (Ubx.ICall x) = Inca.ICall x" | + "norm_instr Ubx.IReturn = Inca.IReturn" -definition rel_fundefs where - "rel_fundefs f g = (\x. rel_option (rel_fundef (\y z. y = norm_instr z)) (f x) (g x))" - -abbreviation (input) norm_eq where - "norm_eq x y \ x = norm_instr y" - -lemma norm_generalize_instr: "norm_instr (Subx.generalize_instr instr) = norm_instr instr" +lemma norm_generalize_instr[simp]: "norm_instr (Subx.generalize_instr instr) = norm_instr instr" by (cases instr) simp_all -lemma rel_fundef_body_nth: - assumes "rel_fundef norm_eq fd1 fd2" and "pc < length (body fd1)" - shows "body fd1 ! pc = norm_instr (body fd2 ! pc)" - using assms - by (auto dest: list_all2_nthD simp: fundef.rel_sel) +abbreviation norm_eq where + "norm_eq x y \ x = norm_instr y" -lemma rel_fundef_rewrite_body: - assumes - "rel_fundef norm_eq fd1 fd2" and - "norm_instr (body fd2 ! pc) = norm_instr instr" - shows "rel_fundef norm_eq fd1 (rewrite_fundef_body fd2 pc instr)" - using assms(1) -proof (cases rule: fundef.rel_cases) - case (Fundef xs ar' ys ar) - hence "length xs = length ys" - by (simp add: list_all2_conv_all_nth) - hence "length xs = length (rewrite ys pc instr)" - by simp - hence "list_all2 norm_eq xs (rewrite ys pc instr)" - proof (elim list_all2_all_nthI) - fix n - assume "n < length xs" - hence "n < length ys" - by (simp add: \length xs = length ys\) - thus "xs ! n = norm_instr (rewrite ys pc instr ! n)" - using list_all2_nthD[OF \list_all2 norm_eq xs ys\ \n < length xs\, symmetric] - using assms(2)[unfolded Fundef(2), simplified] - by (cases "pc = n"; simp) - qed - thus ?thesis - using Fundef by simp +definition rel_fundefs where + "rel_fundefs f g = (\x. rel_option (rel_fundef (=) norm_eq) (f x) (g x))" + +lemma rel_fundefsI: + assumes "\x. rel_option (rel_fundef (=) norm_eq) (F1 x) (F2 x)" + shows "rel_fundefs F1 F2" + using assms + by (simp add: rel_fundefs_def) + +lemma rel_fundefsD: + assumes "rel_fundefs F1 F2" + shows "rel_option (rel_fundef (=) norm_eq) (F1 x) (F2 x)" + using assms + by (simp add: rel_fundefs_def) + +lemma rel_fundefs_next_instr: + assumes rel_F1_F2: "rel_fundefs F1 F2" + shows "rel_option norm_eq (next_instr F1 f l pc) (next_instr F2 f l pc)" + using rel_F1_F2[THEN rel_fundefsD, of f] +proof (cases rule: option.rel_cases) + case None + thus ?thesis by (simp add: next_instr_def) +next + case (Some fd1 fd2) + then show ?thesis + by (auto simp: next_instr_def intro: rel_fundef_imp_rel_option_instr_at) qed -lemma rel_fundef_generalize: - assumes "rel_fundef norm_eq fd1 fd2" - shows "rel_fundef norm_eq fd1 (Subx.generalize_fundef fd2)" - using assms(1) -proof (cases rule: fundef.rel_cases) - case (Fundef xs ar' ys ar) - hence "list_all2 (\x y. x = norm_instr y) xs (map Subx.generalize_instr ys)" - by (auto elim!: list_all2_mono simp: list.rel_map norm_generalize_instr) - thus ?thesis - using Fundef by simp -qed +lemma rel_fundefs_next_instr1: + assumes rel_F1_F2: "rel_fundefs F1 F2" and next_instr1: "next_instr F1 f l pc = Some instr1" + shows "\instr2. next_instr F2 f l pc = Some instr2 \ norm_eq instr1 instr2" + using rel_fundefs_next_instr[OF rel_F1_F2, of f l pc] + unfolding next_instr1 + unfolding option_rel_Some1 + by assumption + +lemma rel_fundefs_next_instr2: + assumes rel_F1_F2: "rel_fundefs F1 F2" and next_instr2: "next_instr F2 f l pc = Some instr2" + shows "\instr1. next_instr F1 f l pc = Some instr1 \ norm_eq instr1 instr2" + using rel_fundefs_next_instr[OF rel_F1_F2, of f l pc] + unfolding next_instr2 + unfolding option_rel_Some2 + by assumption lemma rel_fundefs_empty: "rel_fundefs (\_. None) (\_. None)" by (simp add: rel_fundefs_def) lemma rel_fundefs_None1: assumes "rel_fundefs f g" and "f x = None" shows "g x = None" by (metis assms rel_fundefs_def rel_option_None1) lemma rel_fundefs_None2: assumes "rel_fundefs f g" and "g x = None" shows "f x = None" by (metis assms rel_fundefs_def rel_option_None2) lemma rel_fundefs_Some1: assumes "rel_fundefs f g" and "f x = Some y" - shows "\z. g x = Some z \ rel_fundef norm_eq y z" + shows "\z. g x = Some z \ rel_fundef (=) norm_eq y z" proof - - from assms(1) have "rel_option (rel_fundef norm_eq) (f x) (g x)" + from assms(1) have "rel_option (rel_fundef (=) norm_eq) (f x) (g x)" unfolding rel_fundefs_def by simp with assms(2) show ?thesis by (simp add: option_rel_Some1) qed lemma rel_fundefs_Some2: assumes "rel_fundefs f g" and "g x = Some y" - shows "\z. f x = Some z \ rel_fundef norm_eq z y" + shows "\z. f x = Some z \ rel_fundef (=) norm_eq z y" proof - - from assms(1) have "rel_option (rel_fundef norm_eq) (f x) (g x)" + from assms(1) have "rel_option (rel_fundef (=) norm_eq) (f x) (g x)" unfolding rel_fundefs_def by simp with assms(2) show ?thesis by (simp add: option_rel_Some2) qed lemma rel_fundefs_rel_option: - assumes "rel_fundefs f g" and "\x y. rel_fundef norm_eq x y \ h x y" + assumes "rel_fundefs f g" and "\x y. rel_fundef (=) norm_eq x y \ h x y" shows "rel_option h (f z) (g z)" proof - - have "rel_option (rel_fundef norm_eq) (f z) (g z)" + have "rel_option (rel_fundef (=) norm_eq) (f z) (g z)" using assms(1)[unfolded rel_fundefs_def] by simp then show ?thesis unfolding rel_option_unfold by (auto simp add: assms(2)) qed +lemma rel_fundef_generalizeI: + assumes "rel_fundef (=) norm_eq fd1 fd2" + shows "rel_fundef (=) norm_eq fd1 (Subx.generalize_fundef fd2)" + using assms + by (cases rule: fundef.rel_cases) + (auto simp: map_ran_def list.rel_map elim: list.rel_mono_strong) + +lemma rel_fundefs_generalizeI: + assumes "rel_fundefs (Finca_get F1) (Fubx_get F2)" + shows "rel_fundefs (Finca_get F1) (Fubx_get (Subx.Fenv.map_entry F2 f Subx.generalize_fundef))" +proof (rule rel_fundefsI) + fix x + show "rel_option (rel_fundef (=) norm_eq) + (Finca_get F1 x) (Fubx_get (Subx.Fenv.map_entry F2 f Subx.generalize_fundef) x)" + unfolding Subx.Fenv.get_map_entry_conv + unfolding option.rel_map + using assms(1)[THEN rel_fundefsD, of x] + by (auto intro: rel_fundef_generalizeI elim: option.rel_mono_strong) +qed + +lemma rel_fundefs_rewriteI: + assumes + rel_F1_F2: "rel_fundefs (Finca_get F1) (Fubx_get F2)" and + "norm_eq instr1' instr2'" + shows "rel_fundefs + (Finca_get (Sinca.Fenv.map_entry F1 f (\fd. rewrite_fundef_body fd l pc instr1'))) + (Fubx_get (Subx.Fenv.map_entry F2 f (\fd. rewrite_fundef_body fd l pc instr2')))" + (is "rel_fundefs (Finca_get ?F1') (Fubx_get ?F2')") +proof (rule rel_fundefsI) + fix x + show "rel_option (rel_fundef (=) norm_eq) (Finca_get ?F1' x) (Fubx_get ?F2' x)" + proof (cases "f = x") + case True + show ?thesis + using rel_F1_F2[THEN rel_fundefsD, of x] True assms(2) + by (cases rule: option.rel_cases) (auto intro: rel_fundef_rewrite_body) + next + case False + then show ?thesis + using rel_F1_F2[THEN rel_fundefsD, of x] by simp + qed +qed + section \Equivalence of call stacks\ definition norm_stack :: "('dyn, 'ubx1, 'ubx2) unboxed list \ 'dyn list" where "norm_stack \ \ List.map Subx.norm_unboxed \" lemma norm_stack_Nil[simp]: "norm_stack [] = []" by (simp add: norm_stack_def) lemma norm_stack_Cons[simp]: "norm_stack (d # \) = Subx.norm_unboxed d # norm_stack \" by (simp add: norm_stack_def) lemma norm_stack_append: "norm_stack (xs @ ys) = norm_stack xs @ norm_stack ys" by (simp add: norm_stack_def) lemmas drop_norm_stack = drop_map[where f = Subx.norm_unboxed, folded norm_stack_def] lemmas take_norm_stack = take_map[where f = Subx.norm_unboxed, folded norm_stack_def] lemmas norm_stack_map = map_map[where f = Subx.norm_unboxed, folded norm_stack_def] lemma norm_box_stack[simp]: "norm_stack (map Subx.box_operand \) = norm_stack \" by (induction \) (auto simp: norm_stack_def) lemma length_norm_stack[simp]: "length (norm_stack xs) = length xs" by (simp add: norm_stack_def) definition is_valid_fun_call where - "is_valid_fun_call get fd2 n \2 g \ n < length (body fd2) \ body fd2 ! n = ICall g \ - (\gd. get g = Some gd \ list_all is_dyn_operand (take (arity gd) \2))" - -inductive rel_stacktraces for get where - rel_stacktraces_Nil: - "rel_stacktraces get [] [] opt" | - - rel_stacktraces_Cons: - "rel_stacktraces get st1 st2 (Some f) \ - \1 = norm_stack \2 \ - get f = Some fd2 \ - sp_fundef get fd2 (take n (body fd2)) = Ok (map typeof \2) \ - pred_option (is_valid_fun_call get fd2 n \2) opt \ - rel_stacktraces get (Frame f n \1 # st1) (Frame f n \2 # st2) opt" - -definition all_same_arities where - "all_same_arities F1 F2 \ \f. rel_option (\fd gd. arity fd = arity gd) (F1 f) (F2 f)" - -lemma all_same_arities_commutative: "all_same_arities F1 F2 = all_same_arities F2 F1" -proof - assume "all_same_arities F1 F2" - then show "all_same_arities F2 F1" - unfolding all_same_arities_def - by (simp add: rel_option_unfold) -next - assume "all_same_arities F2 F1" - then show "all_same_arities F1 F2" - unfolding all_same_arities_def - by (simp add: rel_option_unfold) -qed + "is_valid_fun_call F f l pc \ g \ next_instr F f l pc = Some (ICall g) \ + (\gd. F g = Some gd \ arity gd \ length \ \ list_all is_dyn_operand (take (arity gd) \))" -lemma sp_instr_same_arities: - "all_same_arities F1 F2 \ Subx.sp_instr F1 x ys = Subx.sp_instr F2 x ys" -proof (induction F1 x ys rule: Subx.sp_instr.induct) - fix F1 f \ - assume assms: "all_same_arities F1 F2" - show "Subx.sp_instr F1 (Ubx.ICall f) \ = Subx.sp_instr F2 (Ubx.ICall f) \" - proof (cases "F1 f") - case None - then have "F2 f = None" - using HOL.spec[OF assms[unfolded all_same_arities_def], of f] by simp - with None show ?thesis by simp - next - case (Some a) - then obtain b where "F2 f = Some b" and "arity a = arity b" - using HOL.spec[OF assms[unfolded all_same_arities_def], of f] by (auto simp: option_rel_Some1) - with Some show ?thesis by simp - qed -qed simp_all - -lemma sp_same_arities: - assumes "all_same_arities F1 F2" - shows "Subx.sp F1 = Subx.sp F2" -proof (intro ext allI) - fix xs ys - show "Subx.sp F1 xs ys = Subx.sp F2 xs ys" - proof (induction xs arbitrary: ys) - case Nil - then show ?case by simp - next - case (Cons x xs) - show ?case - apply simp - apply (cases "(F1, x, ys)" rule: Subx.sp_instr.cases; - simp add: Cons.IH Let_def) - subgoal for opubx - apply (cases "\\\\\\\\ opubx") - by (auto simp: Cons.IH Let_def) - subgoal for f - apply (rule option.rel_induct[of "(\fd gd. arity fd = arity gd)" "F1 f" "F2 f"]) - using assms(1)[unfolded all_same_arities_def] - by (auto simp add: Cons.IH) - done - qed +lemma is_valid_funcall_map_entry_generalize_fundefI: + assumes "is_valid_fun_call (Fubx_get F2) g l pc \ z" + shows "is_valid_fun_call (Fubx_get (Subx.Fenv.map_entry F2 f Subx.generalize_fundef)) g l pc \ z" +proof (cases "f = z") + case True + then show ?thesis + using assms + by (cases "z = g") + (auto simp: is_valid_fun_call_def next_instr_def Subx.instr_at_generalize_fundef_conv) +next + case False + then show ?thesis + using assms + by (cases "Fubx_get F2 g") + (auto simp: is_valid_fun_call_def next_instr_def + Subx.instr_at_generalize_fundef_conv Subx.Fenv.get_map_entry_conv) qed -lemmas sp_fundef_same_arities = sp_same_arities[THEN eq_sp_to_eq_sp_fundef] - -lemma all_same_arities_add: - assumes "Fubx_get F f = Some fd1" and "arity fd1 = arity fd2" - shows "all_same_arities (Fubx_get F) (Fubx_get (Fubx_add F f fd2))" - unfolding all_same_arities_def -proof (intro allI) - fix g - show "rel_option (\fd gd. arity fd = arity gd) - (Fubx_get F g) - (Fubx_get (Fubx_add F f fd2) g)" - using assms - by (cases "f = g") (simp_all add: option.rel_refl) -qed - -lemma all_same_arities_generalize_fundef: - assumes "Fubx_get F f = Some fd" - shows "all_same_arities (Fubx_get F) (Fubx_get (Fubx_add F f (Subx.generalize_fundef fd)))" - using all_same_arities_add[OF assms(1)] - using ubx.arity_generalize_fundef - by simp +lemma is_valid_fun_call_map_box_operandI: + assumes "is_valid_fun_call (Fubx_get F2) g l pc \ z" + shows "is_valid_fun_call (Fubx_get F2) g l pc (map Subx.box_operand \) z" + using assms + unfolding is_valid_fun_call_def + by (auto simp: take_map list.pred_map list.pred_True) -lemma rel_stacktraces_box: - assumes - "rel_stacktraces F1 xs ys opt" and - "F2 f = map_option Subx.generalize_fundef (F1 f)" and - "\g. f \ g \ F2 g = F1 g" and - "all_same_arities F1 F2" - shows "rel_stacktraces F2 xs (Subx.box_stack f ys) opt" - using assms(1) -proof (induction xs ys opt rule: rel_stacktraces.induct) - case rel_stacktraces_Nil - then show ?case - by (auto intro: rel_stacktraces.intros) -next - case (rel_stacktraces_Cons st1 st2 g \1 \2 gd n opt) - note sp_F1_eq_sp_F2[simp] = sp_same_arities[OF assms(4)] - note sp_fundef_F1_eq_sp_fundef_F2[simp] = eq_sp_to_eq_sp_fundef[OF sp_F1_eq_sp_F2] - show ?case - apply simp - proof (intro conjI impI) - assume "f = g" - then have get2_g: "F2 g = Some (Subx.generalize_fundef gd)" - using assms(2) \F1 g = Some gd\ by simp - - show "rel_stacktraces F2 (Frame g n \1 # st1) - (Frame g n (map Subx.box_operand \2) # Subx.box_stack g st2) opt" - proof (intro rel_stacktraces.intros) - show "rel_stacktraces F2 st1 (Subx.box_stack g st2) (Some g)" - using rel_stacktraces_Cons.IH - unfolding \f = g\ - by simp - next - show "\1 = norm_stack (map Subx.box_operand \2)" - using \\1 = norm_stack \2\ by simp - next - show "F2 g = Some (Subx.generalize_fundef gd)" - using get2_g . - next - show "sp_fundef F2 (Subx.generalize_fundef gd) (take n (body (Subx.generalize_fundef gd))) = - Ok (map typeof (map Subx.box_operand \2))" - using sp_fundef_generalize[OF rel_stacktraces_Cons.hyps(4)] - by (simp add: take_map sp_fundef_def) - next - show "pred_option - (is_valid_fun_call F2 (Subx.generalize_fundef gd) n (map Subx.box_operand \2)) opt" - proof (cases opt) - case (Some h) - then show ?thesis - using rel_stacktraces_Cons.hyps(5) - apply (simp add: take_map list.pred_map list.pred_True is_valid_fun_call_def) - using \f = g\ assms(3) get2_g by fastforce - qed simp - qed +lemma inst_at_rewrite_fundef_body_disj: + "instr_at (rewrite_fundef_body fd l pc instr) l pc = Some instr \ + instr_at (rewrite_fundef_body fd l pc instr) l pc = None" +proof (cases fd) + case (Fundef bblocks ar locals) + show ?thesis + proof (cases "map_of bblocks l") + case None + thus ?thesis + using Fundef + by (simp add: rewrite_fundef_body_def instr_at_def map_entry_map_of_None_conv) next - assume "f \ g" - show "rel_stacktraces F2 (Frame g n \1 # st1) (Frame g n \2 # Subx.box_stack f st2) opt" - using rel_stacktraces_Cons assms(3)[OF \f \ g\] - apply (auto intro!: rel_stacktraces.intros; - cases opt; simp add: sp_fundef_def is_valid_fun_call_def) - subgoal for h - using assms(2,3) by (cases "f = h"; auto) - done + case (Some instr') + moreover hence "l \ fst ` set bblocks" + by (meson domI domIff map_of_eq_None_iff) + ultimately show ?thesis + using Fundef + apply (auto simp add: rewrite_fundef_body_def instr_at_def map_entry_map_of_Some_conv) + by (smt (verit, ccfv_threshold) length_list_update nth_list_update_eq option.case_eq_if option.distinct(1) option.sel update_Some_unfold) qed qed -lemma rel_stacktraces_generalize: - assumes - "rel_stacktraces (Fubx_get F) st1 st2 (Some f)" and - "Fubx_get F f = Some fd " and - sp_prefix: "sp_fundef (Fubx_get F) fd (take pc (body fd)) = Ok (None # map typeof \2)" and - norm_stacks: "\1 = norm_stack \2" and - pc_in_range: "pc < length (body fd)" and - sp_instr: "Subx.sp_instr (Fubx_get F) (Subx.generalize_instr (body fd ! pc)) - (None # map (\_. None) \2) = Ok (None # map (typeof \ Subx.box_operand) \2)" - shows "rel_stacktraces (Fubx_get (Fubx_add F f (Subx.generalize_fundef fd))) - (Frame f (Suc pc) (d # \1) # st1) - (Frame f (Suc pc) (OpDyn d # map Subx.box_operand \2) # Subx.box_stack f st2) None" -proof - - let ?fd' = "Subx.generalize_fundef fd" - let ?F' = "Fubx_add F f ?fd'" - show ?thesis - proof (intro rel_stacktraces_Cons) - show "rel_stacktraces (Fubx_get ?F') st1 (Subx.box_stack f st2) (Some f)" - using assms(2) all_same_arities_generalize_fundef - by (auto intro: rel_stacktraces_box[OF assms(1)]) - next - show "d # \1 = norm_stack (OpDyn d # map Subx.box_operand \2)" - using norm_stacks by simp - next - show "Fubx_get ?F' f = Some ?fd'" - by simp - next - show "sp_fundef (Fubx_get ?F') ?fd' (take (Suc pc) (body ?fd')) = - Ok (map typeof (OpDyn d # map Subx.box_operand \2))" - unfolding all_same_arities_generalize_fundef[THEN sp_fundef_same_arities, - OF assms(2), symmetric] - using sp_fundef_generalize[OF sp_prefix] sp_instr - by (auto simp add: sp_fundef_def take_map take_Suc_conv_app_nth[OF pc_in_range]) - qed simp_all -qed - -lemma rel_fundefs_rewrite: - assumes - rel_F1_F2: "rel_fundefs (Finca_get F1) (Fubx_get F2)" and - F2_get_f: "Fubx_get F2 f = Some fd2" and - F2_add_f: "Fubx_add F2 f (rewrite_fundef_body fd2 pc instr) = F2'" and - eq_norm: "norm_instr (body fd2 ! pc) = norm_instr instr" - shows "rel_fundefs (Finca_get F1) (Fubx_get F2')" - unfolding rel_fundefs_def -proof - fix x - show "rel_option (rel_fundef norm_eq) (Finca_get F1 x) (Fubx_get F2' x)" - proof (cases "x = f") - case True - then have F2'_get_x: "Fubx_get F2' x = Some (rewrite_fundef_body fd2 pc instr)" - using F2_add_f by auto - obtain fd1 where "Finca_get F1 f = Some fd1" and rel_fd1_fd2: "rel_fundef norm_eq fd1 fd2" - using rel_fundefs_Some2[OF rel_F1_F2 F2_get_f] by auto - thus ?thesis - unfolding F2'_get_x option_rel_Some2 - using True rel_fundef_rewrite_body[OF rel_fd1_fd2 eq_norm] - by auto - next - case False - then have "Fubx_get F2' x = Fubx_get F2 x" - using F2_add_f by auto - then show ?thesis - using rel_F1_F2 rel_fundefs_def - by fastforce - qed -qed - -lemma rel_fundef_rewrite_both: - assumes "rel_fundef norm_eq fd1 fd2" and "norm_instr y = x" - shows "rel_fundef norm_eq (rewrite_fundef_body fd1 pc x) (rewrite_fundef_body fd2 pc y)" - using assms by (auto simp: fundef.rel_sel) - -lemma rel_fundefs_rewrite_both: - assumes - rel_init: "rel_fundefs (Finca_get F1) (Fubx_get F2)" and - rel_new: "rel_fundef norm_eq fd1 fd2" - shows "rel_fundefs (Finca_get (Finca_add F1 f fd1)) (Fubx_get (Fubx_add F2 f fd2))" - unfolding rel_fundefs_def -proof - fix x - show "rel_option (rel_fundef norm_eq) (Finca_get (Finca_add F1 f fd1) x) (Fubx_get (Fubx_add F2 f fd2) x)" - proof (cases "x = f") +lemma is_valid_fun_call_map_entry_conv: + assumes "next_instr (Fubx_get F2) f l pc = Some instr" "\ is_fun_call instr" "\ is_fun_call instr'" + shows + "is_valid_fun_call (Fubx_get (Subx.Fenv.map_entry F2 f (\fd. rewrite_fundef_body fd l pc instr')))= + is_valid_fun_call (Fubx_get F2)" +proof (intro ext) + fix f' l' pc' \ g + show + "is_valid_fun_call (Fubx_get (Subx.Fenv.map_entry F2 f (\fd. rewrite_fundef_body fd l pc instr'))) f' l' pc' \ g = + is_valid_fun_call (Fubx_get F2) f' l' pc' \ g" + proof (cases "f = f'") case True then show ?thesis - using rel_new by simp + using assms + apply (cases "f = g") + by (auto simp: is_valid_fun_call_def next_instr_eq_Some_conv + instr_at_rewrite_fundef_body_conv if_split_eq1) next case False then show ?thesis - using rel_init - by (simp add: rel_fundefs_def) - qed -qed - -lemma rel_fundefs_generalize: - assumes - rel_F1_F2: "rel_fundefs (Finca_get F1) (Fubx_get F2)" and - F2_get_f: "Fubx_get F2 f = Some fd2" - shows "rel_fundefs (Finca_get F1) (Fubx_get (Fubx_add F2 f (Subx.generalize_fundef fd2)))" - unfolding rel_fundefs_def -proof - let ?F2' = "(Fubx_add F2 f (Subx.generalize_fundef fd2))" - fix x - show "rel_option (rel_fundef norm_eq) (Finca_get F1 x) (Fubx_get ?F2' x)" - proof (cases "x = f") - case True - then have F2'_get_x: "Fubx_get ?F2' x = Some (Subx.generalize_fundef fd2)" - using Subx.Fenv.get_add_eq by auto - obtain fd1 where "Finca_get F1 f = Some fd1" and "rel_fundef norm_eq fd1 fd2" - using rel_fundefs_Some2[OF rel_F1_F2 F2_get_f] by auto - thus ?thesis - unfolding F2'_get_x option_rel_Some2 - using True rel_fundef_generalize - by auto - next - case False - then have "Fubx_get ?F2' x = Fubx_get F2 x" - using Subx.Fenv.get_add_neq by auto - then show ?thesis - using rel_F1_F2 rel_fundefs_def - by fastforce + using assms + apply (cases "f = g") + by (auto simp: is_valid_fun_call_def next_instr_eq_Some_conv) qed qed -lemma rel_stacktraces_rewrite_fundef: - assumes - "rel_stacktraces (Fubx_get F2) xs ys opt" and - "Fubx_get F2 f = Some fd" and - "pc < length (body fd)" and - "\\. Subx.sp_instr (Fubx_get F2) (body fd ! pc) \ = Subx.sp_instr (Fubx_get F2) instr \" and - "\ is_fun_call (body fd ! pc)" - shows "rel_stacktraces - (Fubx_get (Fubx_add F2 f (rewrite_fundef_body fd pc instr))) xs ys opt" +lemma is_valid_fun_call_map_entry_neq_f_neq_l: + assumes "f \ g" "l \ l'" + shows + "is_valid_fun_call (Fubx_get (Subx.Fenv.map_entry F2 f (\fd. rewrite_fundef_body fd l pc instr'))) g l' = + is_valid_fun_call (Fubx_get F2) g l'" + apply (intro ext) + unfolding is_valid_fun_call_def + using assms + apply (simp add: next_instr_eq_Some_conv) + apply (rule iffI; simp) + unfolding Subx.Fenv.get_map_entry_conv + apply simp + apply (metis arity_rewrite_fundef_body) + apply safe + by simp + +inductive rel_stacktraces for F where + rel_stacktraces_Nil: + "rel_stacktraces F opt [] []" | + + rel_stacktraces_Cons: + "rel_stacktraces F (Some f) st1 st2 \ + \1 = map Subx.norm_unboxed \2 \ + R1 = map Subx.norm_unboxed R2 \ + list_all is_dyn_operand R2 \ + F f = Some fd2 \ map_of (body fd2) l = Some instrs \ + Subx.sp_instrs (map_option funtype \ F) (return fd2) (take pc instrs) [] (map typeof \2) \ + pred_option (is_valid_fun_call F f l pc \2) opt \ + rel_stacktraces F opt (Frame f l pc R1 \1 # st1) (Frame f l pc R2 \2 # st2)" + +lemma rel_stacktraces_map_entry_gneralize_fundefI[intro]: + assumes "rel_stacktraces (Fubx_get F2) opt st1 st2" + shows "rel_stacktraces (Fubx_get (Subx.Fenv.map_entry F2 f Subx.generalize_fundef)) + opt st1 (Subx.box_stack f st2)" using assms(1) -proof (induction xs ys opt rule: rel_stacktraces.induct) - case rel_stacktraces_Nil - then show ?case - by (auto intro: rel_stacktraces.intros) +proof (induction opt st1 st2 rule: rel_stacktraces.induct) + case (rel_stacktraces_Nil opt) + thus ?case + by (auto intro: rel_stacktraces.rel_stacktraces_Nil) next - case (rel_stacktraces_Cons st1 st2 g \1 \2 gd n opt) - have same_arities: "all_same_arities (Fubx_get F2) - (Fubx_get (Fubx_add F2 f (rewrite_fundef_body fd pc instr)))" - using all_same_arities_add[OF assms(2)] - by simp - + case (rel_stacktraces_Cons g st1 st2 \1 \2 R1 R2 gd2 l instrs pc opt) show ?case - proof (cases "g = f") + proof (cases "f = g") case True - then have "gd = fd" - using assms(2) rel_stacktraces_Cons.hyps(3) by auto - thus ?thesis - using True rel_stacktraces_Cons - apply (auto intro!: rel_stacktraces.intros - simp: sp_fundef_same_arities[OF same_arities, symmetric]) - apply (cases "n \ pc"; simp add: sp_fundef_def) - subgoal - using assms(3,4) - by (simp add: Subx.sp_rewrite_eq_Ok take_rewrite_swap) - subgoal - using rel_stacktraces_Cons.hyps(5) - proof (induction opt) - case (Some h) - hence "rewrite (body fd) pc instr ! n = Ubx.ICall h" - using \gd = fd\ - apply (simp add: is_valid_fun_call_def) - by (metis assms(5) instr.disc nth_rewrite_neq) - then show ?case - using Some \gd = fd\ - apply (simp add: is_valid_fun_call_def) - by (metis arity_rewrite_fundef_body assms(2) option.inject Subx.Fenv.get_add_eq Subx.Fenv.get_add_neq) - qed simp - done + then show ?thesis + using rel_stacktraces_Cons + apply auto + apply (rule rel_stacktraces.rel_stacktraces_Cons) + apply assumption + apply simp + apply (rule refl) + apply assumption + apply simp + by (auto simp add: take_map Subx.map_of_generalize_fundef_conv + intro!: Subx.sp_instrs_generalize0 + intro!: is_valid_funcall_map_entry_generalize_fundefI is_valid_fun_call_map_box_operandI + elim!: option.pred_mono_strong) next case False - thus ?thesis + then show ?thesis using rel_stacktraces_Cons - apply (auto intro!: rel_stacktraces.intros - simp: sp_fundef_same_arities[OF same_arities, symmetric]) - proof (induction opt) - case (Some h) - then show ?case - using Some assms(2) - by (cases "h = f"; simp add: is_valid_fun_call_def) - qed simp + by (auto intro: rel_stacktraces.intros is_valid_funcall_map_entry_generalize_fundefI + elim!: option.pred_mono_strong) + qed +qed + +lemma rel_stacktraces_map_entry_rewrite_fundef_body: + assumes + "rel_stacktraces (Fubx_get F2) opt st1 st2" and + "next_instr (Fubx_get F2) f l pc = Some instr" and + "\ret. Subx.sp_instr (map_option funtype \ Fubx_get F2) ret instr = + Subx.sp_instr (map_option funtype \ Fubx_get F2) ret instr'" and + "\ is_fun_call instr" "\ is_fun_call instr'" + shows "rel_stacktraces + (Fubx_get (Subx.Fenv.map_entry F2 f (\fd. rewrite_fundef_body fd l pc instr'))) opt st1 st2" + using assms(1) +proof (induction opt st1 st2 rule: rel_stacktraces.induct) + case (rel_stacktraces_Nil opt) + then show ?case + by (auto intro: rel_stacktraces.rel_stacktraces_Nil) +next + case (rel_stacktraces_Cons g st1 st2 \1 \2 R1 R2 gd2 l' instrs pc' opt) + show ?case (is "rel_stacktraces (Fubx_get ?F2') opt ?st1 ?st2") + proof (cases "f = g") + case True + show ?thesis + proof (cases "l' = l") + case True + show ?thesis + apply (rule rel_stacktraces.rel_stacktraces_Cons) + using rel_stacktraces_Cons.IH apply simp + using rel_stacktraces_Cons.hyps apply simp + using rel_stacktraces_Cons.hyps apply simp + using rel_stacktraces_Cons.hyps apply simp + using rel_stacktraces_Cons.hyps \f = g\ apply simp + using rel_stacktraces_Cons.hyps True apply simp + using rel_stacktraces_Cons.hyps apply simp + using rel_stacktraces_Cons.hyps assms \f = g\ True + apply (cases "pc' \ pc") [] + apply (auto simp add: take_update_swap intro!: Subx.sp_instrs_list_update + dest!: next_instrD instr_atD) [2] + using rel_stacktraces_Cons.hyps + unfolding is_valid_fun_call_map_entry_conv[OF assms(2,4,5)] + by simp + next + case False + show ?thesis + proof (rule rel_stacktraces.rel_stacktraces_Cons) + show "Fubx_get ?F2' g = Some (rewrite_fundef_body gd2 l pc instr')" + unfolding \f = g\ + using rel_stacktraces_Cons.hyps by simp + next + show "pred_option (is_valid_fun_call (Fubx_get ?F2') g l' pc' \2) opt" + unfolding is_valid_fun_call_map_entry_conv[OF assms(2,4,5)] + using rel_stacktraces_Cons.hyps by simp + qed (insert rel_stacktraces_Cons False, simp_all) + qed + next + case False + then show ?thesis + using rel_stacktraces_Cons + by (auto simp: is_valid_fun_call_map_entry_conv[OF assms(2,4,5)] + intro!: rel_stacktraces.rel_stacktraces_Cons) qed qed -section \Matching relation\ - -lemma sp_fundefs_get: - assumes "sp_fundefs F" and "F f = Some fd" - shows "sp_fundef F fd (body fd) = Ok [None]" - using assms by (simp add: sp_fundefs_def) - -lemma sp_fundefs_generalize: - assumes "sp_fundefs (Fubx_get F)" and "Fubx_get F f = Some fd" - shows "sp_fundefs (Fubx_get (Fubx_add F f (Subx.generalize_fundef fd)))" - unfolding sp_fundefs_def -proof (intro allI impI) - fix g gd - assume get_g: "Fubx_get (Fubx_add F f (Subx.generalize_fundef fd)) g = Some gd" - note sp_F_eq_sp_F' = all_same_arities_generalize_fundef[OF assms(2), THEN sp_same_arities, symmetric] - show "sp_fundef (Fubx_get (Fubx_add F f (Subx.generalize_fundef fd))) gd (body gd) = Ok [None]" - proof (cases "f = g") - case True - then have "gd = Subx.generalize_fundef fd" - using get_g by simp - then show ?thesis - using assms - by (simp add: sp_fundefs_def sp_fundef_def sp_F_eq_sp_F' Subx.sp_generalize2) - next - case False - then show ?thesis - using get_g assms - by (simp add: sp_fundefs_def sp_fundef_def sp_F_eq_sp_F') - qed -qed - -lemma sp_fundefs_add: - assumes - "sp_fundefs (Fubx_get F)" and - "sp_fundef (Fubx_get F) fd (body fd) = Ok [None]" and - "all_same_arities (Fubx_get F) (Fubx_get (Fubx_add F f fd))" - shows "sp_fundefs (Fubx_get (Fubx_add F f fd))" - unfolding sp_fundefs_def -proof (intro allI impI) - fix g gd - assume "Fubx_get (Fubx_add F f fd) g = Some gd" - show "sp_fundef (Fubx_get (Fubx_add F f fd)) gd (body gd) = Ok [None]" - proof (cases "f = g") - case True - then have "gd = fd" - using \Fubx_get (Fubx_add F f fd) g = Some gd\ by simp - then show ?thesis - unfolding sp_fundef_same_arities[OF assms(3), symmetric] - using assms(2) by simp - next - case False - then show ?thesis - unfolding sp_fundef_same_arities[OF assms(3), symmetric] - using \Fubx_get (Fubx_add F f fd) g = Some gd\ - using sp_fundefs_get[OF assms(1)] - by simp - qed -qed +section \Simulation relation\ inductive match (infix "\" 55) where - "rel_fundefs (Finca_get F1) (Fubx_get F2) \ - sp_fundefs (Fubx_get F2) \ - rel_stacktraces (Fubx_get F2) st1 st2 None \ - match (State F1 H st1) (State F2 H st2)" + matchI: "Subx.wf_state (State F2 H st2) \ + rel_fundefs (Finca_get F1) (Fubx_get F2) \ + rel_stacktraces (Fubx_get F2) None st1 st2 \ + match (State F1 H st1) (State F2 H st2)" + +lemmas matchI[consumes 0, case_names wf_state rel_fundefs rel_stacktraces] = match.intros(1) section \Backward simulation\ -lemma traverse_cast_Dyn_to_norm: "traverse cast_Dyn xs = Some ys \ norm_stack xs = ys" -proof (induction xs arbitrary: ys) - case Nil - then show ?case by simp -next - case (Cons x xs) - from Cons.prems show ?case - by (auto intro: Cons.IH elim: cast_Dyn.elims simp: Option.bind_eq_Some_conv) +lemma map_eq_append_map_drop: + "map f xs = ys @ map f (drop n xs) \ map f (take n xs) = ys" + by (metis append_same_eq append_take_drop_id map_append) + +lemma ap_map_list_cast_Dyn_to_map_norm: + assumes "ap_map_list cast_Dyn xs = Some ys" + shows "ys = map Subx.norm_unboxed xs" +proof - + from assms have "list_all2 (\x y. cast_Dyn x = Some y) xs ys" + by (simp add: ap_map_list_iff_list_all2) + thus ?thesis + by (induction xs ys rule: list.rel_induct) (auto dest: cast_inversions) qed -lemma traverse_cast_Dyn_to_all_Dyn: - "traverse cast_Dyn xs = Some ys \ list_all (\x. typeof x = None) xs" -proof (induction xs arbitrary: ys) - case Nil - then show ?case by simp -next - case (Cons x xs) - from Cons.prems show ?case - by (auto intro: Cons.IH elim: cast_Dyn.elims simp: Option.bind_eq_Some_conv) +lemma ap_map_list_cast_Dyn_to_all_Dyn: + assumes "ap_map_list cast_Dyn xs = Some ys" + shows "list_all (\x. typeof x = None) xs" +proof - + from assms have "list_all2 (\x y. cast_Dyn x = Some y) xs ys" + by (simp add: ap_map_list_iff_list_all2) + hence "list_all2 (\x y. typeof x = None) xs ys" + by (auto intro: list.rel_mono_strong cast_Dyn_eq_Some_imp_typeof) + thus ?thesis + by (induction xs ys rule: list.rel_induct) auto qed +lemma ap_map_list_cast_Dyn_map_typeof_replicate_conv: + assumes "ap_map_list cast_Dyn xs = Some ys" and "n = length xs" + shows "map typeof xs = replicate n None" + using assms + by (auto simp: list.pred_set intro!: replicate_eq_map[symmetric] + dest!: ap_map_list_cast_Dyn_to_all_Dyn) + +lemma cast_Dyn_eq_Some_conv_norm_unboxed[simp]: "cast_Dyn i = Some i' \ Subx.norm_unboxed i = i'" + by (cases i) simp_all + +lemma cast_Dyn_eq_Some_conv_typeof[simp]: "cast_Dyn i = Some i' \ typeof i = None" + by (cases i) simp_all + lemma backward_lockstep_simulation: - assumes "Subx.step s2 s2'" and "s1 \ s2" - shows "\s1'. Sinca.step s1 s1' \ s1' \ s2'" - using assms(2,1) + assumes "match s1 s2" and "Subx.step s2 s2'" + shows "\s1'. Sinca.step s1 s1' \ match s1' s2'" + using assms proof (induction s1 s2 rule: match.induct) - case (1 F1 F2 st1 st2 H) - have rel_F1_F2: "rel_fundefs (Finca_get F1) (Fubx_get F2)" using 1 by simp - have sp_F2: "sp_fundefs (Fubx_get F2)" using 1 by simp - - from "1"(3,4) show ?case - proof (induction st1 st2 "None :: 'fun option" rule: rel_stacktraces.induct) + case (matchI F2 H st2 F1 st1) + from matchI(3,1,2,4) show ?case + proof (induction "None :: 'fun option" st1 st2 rule: rel_stacktraces.induct) case rel_stacktraces_Nil hence False by (auto elim: Subx.step.cases) thus ?case by simp next - case (rel_stacktraces_Cons st1 st2 f \1 \2 fd2 pc) - have F2_f: "Fubx_get F2 f = Some fd2" - using rel_stacktraces_Cons by simp - have rel_st1_st2: "rel_stacktraces (Fubx_get F2) st1 st2 (Some f)" - using rel_stacktraces_Cons by simp - have sp_fundef_prefix: "sp_fundef (Fubx_get F2) fd2 (take pc (body fd2)) = Ok (map typeof \2)" - using rel_stacktraces_Cons by simp - have \1_def: "\1 = norm_stack \2" - using rel_stacktraces_Cons by simp + case (rel_stacktraces_Cons f st1 st2 \1 \2 R1 R2 fd2 l instrs pc) + note hyps = rel_stacktraces_Cons.hyps + note prems = rel_stacktraces_Cons.prems + have wf_state2: "Subx.wf_state (State F2 H (Frame f l pc R2 \2 # st2))" using prems by simp + have rel_F1_F2: "rel_fundefs (Finca_get F1) (Fubx_get F2)" using prems by simp + have rel_st1_st2: "rel_stacktraces (Fubx_get F2) (Some f) st1 st2" using hyps by simp + have \1_def: "\1 = map Subx.norm_unboxed \2" using hyps by simp + have R1_def: "R1 = map Subx.norm_unboxed R2" using hyps by simp + have all_dyn_R2: "list_all is_dyn_operand R2" using hyps by simp + have F2_f: "Fubx_get F2 f = Some fd2" using hyps by simp + have map_of_fd2_l: "map_of (body fd2) l = Some instrs" using hyps by simp + have sp_instrs_prefix: "Subx.sp_instrs (map_option funtype \ Fubx_get F2) (return fd2) + (take pc instrs) [] (map typeof \2)" + using hyps by simp - note sp_fundef_def[simp] - note sp_prefix = sp_fundef_prefix[unfolded sp_fundef_def] + note next_instr2 = rel_fundefs_next_instr2[OF rel_F1_F2] + note sp_instrs_prefix' = + Subx.sp_instrs_singletonI[THEN Subx.sp_instrs_appendI[OF sp_instrs_prefix]] - have sp_generalized: "Subx.sp (Fubx_get (Fubx_add F2 f (Subx.generalize_fundef fd2))) - (map Subx.generalize_instr (take pc (body fd2))) (replicate (arity fd2) None) = - Ok (map (\_. None) \2)" - using Subx.sp_generalize[OF sp_prefix, simplified] - using all_same_arities_generalize_fundef[OF F2_f] - by (simp add: sp_same_arities) + obtain fd1 where + F1_f: "Finca_get F1 f = Some fd1" and rel_fd1_fd2: "rel_fundef (=) norm_eq fd1 fd2" + using rel_fundefs_Some2[OF rel_F1_F2 F2_f] by auto + + have wf_F2: "Subx.wf_fundefs (Fubx_get F2)" + by (rule wf_state2[THEN Subx.wf_stateD, simplified]) + + have wf_fd2: "Subx.wf_fundef (map_option funtype \ Fubx_get F2) fd2" + using F2_f wf_F2[THEN Subx.wf_fundefsD, THEN spec, of f] + by simp + + have + instrs_neq_Nil: "instrs \ []" and + all_jumps_in_range: "list_all (Subx.jump_in_range (fst ` set (body fd2))) instrs" and + sp_instrs_instrs: "Subx.sp_instrs (map_option funtype \ Fubx_get F2) (return fd2) instrs [] []" + using list_all_map_of_SomeD[OF wf_fd2[THEN Subx.wf_fundef_all_wf_basic_blockD] map_of_fd2_l] + by (auto dest: Subx.wf_basic_blockD) + + have sp_instrs_instrs': "Subx.sp_instrs (map_option funtype \ Fubx_get F2) (return fd2) + (butlast instrs @ [instrs ! pc]) [] []" if pc_def: "pc = length instrs - 1" + unfolding pc_def last_conv_nth[OF instrs_neq_Nil, symmetric] + unfolding append_butlast_last_id[OF instrs_neq_Nil] + by (rule sp_instrs_instrs) + + have sp_instr_last: "Subx.sp_instr (map_option funtype \ Fubx_get F2) (return fd2) + (instrs ! pc) (map typeof \2) []" if pc_def: "pc = length instrs - 1" + using sp_instrs_instrs'[OF pc_def] + using sp_instrs_prefix[unfolded pc_def butlast_conv_take[symmetric]] + by (auto dest!: Subx.sp_instrs_appendD') + + from list_all_map_of_SomeD[OF wf_fd2[THEN Subx.wf_fundef_all_wf_basic_blockD] map_of_fd2_l] + have is_jump_nthD: "\n. is_jump (instrs ! n) \ n < length instrs \ n = length instrs - 1" + by (auto dest!: Subx.wf_basic_blockD + list_all_butlast_not_nthD[of "\i. \ is_jump i \ \ Ubx.instr.is_return i", simplified, OF _ disjI1]) + + note wf_s2' = Subx.wf_state_step_preservation[OF wf_state2 prems(3)] - from rel_stacktraces_Cons.prems(1) show ?case - proof (induction "State F2 H (Frame f pc \2 # st2)" s2' rule: Subx.step.induct) - case (step_push fd2' d) - then have [simp]: "fd2' = fd2" using F2_f by simp - obtain fd1 where fd1_thms: "Finca_get F1 f = Some fd1" "rel_fundef norm_eq fd1 fd2" - using rel_fundefs_Some2[OF rel_F1_F2 F2_f] by auto - show ?case (is "\x. ?STEP x \ ?MATCH x") - proof - - let ?s1' = "State F1 H (Frame f (Suc pc) (d # \1) # st1)" - have "?STEP ?s1'" - using step_push fd1_thms - by (auto intro: Sinca.step_push simp: rel_fundef_body_nth) - moreover have "?MATCH ?s1'" - using step_push rel_stacktraces_Cons rel_F1_F2 sp_F2 - by (auto intro!: match.intros intro: rel_stacktraces.intros simp: take_Suc_conv_app_nth) - ultimately show "?thesis" by blast - qed - next - case (step_push_ubx1 fd2' n) - then have [simp]: "fd2' = fd2" using F2_f by simp - obtain fd1 where fd1_thms: "Finca_get F1 f = Some fd1" "rel_fundef norm_eq fd1 fd2" - using rel_fundefs_Some2[OF rel_F1_F2 F2_f] by auto - show ?case (is "\x. ?STEP x \ ?MATCH x") - proof - - let ?s1' = "State F1 H (Frame f (Suc pc) (box_ubx1 n # \1) # st1)" - have "?STEP ?s1'" - using step_push_ubx1 fd1_thms - by (auto intro: Sinca.step_push simp: rel_fundef_body_nth) - moreover have "?MATCH ?s1'" - using step_push_ubx1 rel_stacktraces_Cons rel_F1_F2 sp_F2 - by (auto intro!: match.intros intro: rel_stacktraces.intros simp: take_Suc_conv_app_nth) - ultimately show "?thesis" by blast - qed - next - case (step_push_ubx2 fd2' b) - then have [simp]: "fd2' = fd2" using F2_f by simp - obtain fd1 where fd1_thms: "Finca_get F1 f = Some fd1" "rel_fundef norm_eq fd1 fd2" - using rel_fundefs_Some2[OF rel_F1_F2 F2_f] by auto - show ?case (is "\x. ?STEP x \ ?MATCH x") - proof - - let ?s1' = "State F1 H (Frame f (Suc pc) (box_ubx2 b # \1) # st1)" - have "?STEP ?s1'" - using step_push_ubx2 fd1_thms - by (auto intro: Sinca.step_push simp: rel_fundef_body_nth) - moreover have "?MATCH ?s1'" - using step_push_ubx2 rel_stacktraces_Cons rel_F1_F2 sp_F2 - by (auto intro!: match.intros intro: rel_stacktraces.intros simp: take_Suc_conv_app_nth) - ultimately show "?thesis" by blast - qed - next - case (step_pop fd2' x \2') - then have [simp]: "fd2' = fd2" using F2_f by simp - obtain fd1 where fd1_thms: "Finca_get F1 f = Some fd1" "rel_fundef norm_eq fd1 fd2" - using rel_fundefs_Some2[OF rel_F1_F2 F2_f] by auto - show ?case (is "\x. ?STEP x \ ?MATCH x") - proof - - let ?s1' = "State F1 H (Frame f (Suc pc) (norm_stack \2') # st1)" - have "?STEP ?s1'" - unfolding \1_def - using step_pop fd1_thms - by (auto intro!: Sinca.step_pop simp: rel_fundef_body_nth) - moreover have "?MATCH ?s1'" - using step_pop rel_stacktraces_Cons rel_F1_F2 sp_F2 - by (auto intro!: match.intros intro: rel_stacktraces.intros simp: take_Suc_conv_app_nth) - ultimately show "?thesis" by blast - qed - next - case (step_load fd2' x i i' d \2') - then have [simp]: "fd2' = fd2" using F2_f by simp - obtain fd1 where - F1_f: "Finca_get F1 f = Some fd1" and - rel_fd1_fd2: "rel_fundef norm_eq fd1 fd2" - using rel_fundefs_Some2[OF rel_F1_F2 F2_f] by auto - show ?case (is "\x. ?STEP x \ ?MATCH x") - proof - - let ?s1' = "State F1 H (Frame f (Suc pc) (d # norm_stack \2') # st1)" - have pc_in_range: "pc < length (body fd1)" - using rel_fd1_fd2 step_load(2) by simp - have "?STEP ?s1'" - using step_load rel_fundef_body_nth[OF rel_fd1_fd2 pc_in_range] - by (auto intro!: Sinca.step_load[OF F1_f pc_in_range, of x] - dest: cast_inversions(1) - simp: \1_def) - moreover have "?MATCH ?s1'" - using step_load refl rel_stacktraces_Cons rel_F1_F2 sp_F2 - by (auto intro!: match.intros intro: rel_stacktraces.intros(2)[OF rel_st1_st2] - dest!: cast_inversions(1) - simp: take_Suc_conv_app_nth[OF step_load(2), simplified]) - ultimately show "?thesis" by blast - qed - next - case (step_load_ubx_hit fd2' \ x i i' d blob \2') - then have [simp]: "fd2' = fd2" using F2_f by simp - obtain fd1 where - F1_f: "Finca_get F1 f = Some fd1" and - rel_fd1_fd2: "rel_fundef norm_eq fd1 fd2" - using rel_fundefs_Some2[OF rel_F1_F2 F2_f] by auto - show ?case (is "\x. ?STEP x \ ?MATCH x") - proof - - let ?s1' = "State F1 H (Frame f (Suc pc) (d # norm_stack \2') # st1)" - have pc_in_range: "pc < length (body fd1)" - using rel_fd1_fd2 step_load_ubx_hit(2) by simp - have "?STEP ?s1'" - using step_load_ubx_hit rel_fundef_body_nth[OF rel_fd1_fd2 pc_in_range] - by (auto intro!: Sinca.step_load[OF F1_f pc_in_range, of x] - dest: cast_inversions(1) - simp: rel_fundef_body_nth \1_def) - moreover have "?MATCH ?s1'" - using step_load_ubx_hit rel_stacktraces_Cons rel_F1_F2 sp_F2 - by (auto intro!: match.intros intro: rel_stacktraces.intros(2)[OF rel_st1_st2] - dest!: cast_inversions(1) - simp: take_Suc_conv_app_nth[OF step_load_ubx_hit(2), simplified]) - ultimately show "?thesis" by blast - qed - next - case (step_load_ubx_miss fd2' \ x i i' d F2' \2') - then have fd2'_def[simp]: "fd2' = fd2" using F2_f by simp - obtain fd1 where - F1_f: "Finca_get F1 f = Some fd1" and - rel_fd1_fd2: "rel_fundef norm_eq fd1 fd2" - using rel_fundefs_Some2[OF rel_F1_F2 F2_f] by auto - show ?case (is "\x. ?STEP x \ ?MATCH x") - proof - - let ?s1' = "State F1 H (Frame f (Suc pc) (d # norm_stack \2') # st1)" - have pc_in_range: "pc < length (body fd1)" - using rel_fd1_fd2 step_load_ubx_miss(2) by simp - have "?STEP ?s1'" - using step_load_ubx_miss F1_f rel_fd1_fd2 - by (auto intro!: Sinca.step_load[OF F1_f pc_in_range, of x] - dest!: cast_inversions(1) - simp: rel_fundef_body_nth \1_def) - moreover have "?MATCH ?s1'" - proof (rule match.intros) - show "rel_fundefs (Finca_get F1) (Fubx_get F2')" - unfolding step_load_ubx_miss.hyps(7)[symmetric] - using rel_fundefs_generalize[OF rel_F1_F2 F2_f] - by simp - next - show "sp_fundefs (Fubx_get F2')" - unfolding step_load_ubx_miss.hyps(7)[symmetric] - using sp_fundefs_generalize[OF sp_F2 F2_f] - by simp - next - show "rel_stacktraces (Fubx_get F2') - (Frame f (Suc pc) (d # norm_stack \2') # st1) - (Subx.box_stack f (Frame f (Suc pc) (OpDyn d # \2') # st2)) None" - using step_load_ubx_miss sp_fundef_prefix - by (auto intro!: rel_stacktraces_generalize[OF rel_st1_st2 F2_f] dest: cast_inversions) - qed - ultimately show "?thesis" by blast - qed - next - case (step_store fd2' x i i' y d H' \2') - then have [simp]: "fd2' = fd2" using F2_f by simp - obtain fd1 where fd1_thms: "Finca_get F1 f = Some fd1" "rel_fundef norm_eq fd1 fd2" - using rel_fundefs_Some2[OF rel_F1_F2 F2_f] by auto - show ?case (is "\x. ?STEP x \ ?MATCH x") - proof - - let ?s1' = "State F1 H' (Frame f (Suc pc) (norm_stack \2') # st1)" - have "?STEP ?s1'" - unfolding \1_def - using step_store fd1_thms - by (auto intro!: Sinca.step_store dest!: cast_inversions - simp: rel_fundef_body_nth) - moreover have "?MATCH ?s1'" - using step_store rel_stacktraces_Cons rel_F1_F2 sp_F2 - by (auto intro!: match.intros rel_stacktraces.intros dest!: cast_inversions - simp: take_Suc_conv_app_nth) - ultimately show "?thesis" by blast - qed - next - case (step_store_ubx fd2' \ x i i' blob d H' \2') - then have [simp]: "fd2' = fd2" using F2_f by simp - obtain fd1 where - F1_f: "Finca_get F1 f = Some fd1" and - rel_fd1_fd2: "rel_fundef norm_eq fd1 fd2" - using rel_fundefs_Some2[OF rel_F1_F2 F2_f] by auto - show ?case (is "\x. ?STEP x \ ?MATCH x") - proof - - let ?s1' = "State F1 H' (Frame f (Suc pc) (norm_stack \2') # st1)" - have pc_in_range: "pc < length (body fd1)" - using rel_fd1_fd2 step_store_ubx.hyps(2) by auto - have "?STEP ?s1'" - unfolding \1_def \i # blob # \2' = \2\[symmetric] - using step_store_ubx pc_in_range - by (auto intro!: Sinca.step_store[OF F1_f] - dest!: cast_inversions - simp: rel_fundef_body_nth[OF rel_fd1_fd2]) - moreover have "?MATCH ?s1'" - using step_store_ubx rel_stacktraces_Cons rel_F1_F2 sp_F2 - by (auto intro!: match.intros rel_stacktraces.intros - dest!: cast_inversions - simp: take_Suc_conv_app_nth) - ultimately show "?thesis" by blast + from prems(3) show ?case + using wf_s2' + proof (induction "State F2 H (Frame f l pc R2 \2 # st2)" s2' rule: Subx.step.induct) + case (step_push d) + let ?st1' = "Frame f l (Suc pc) R1 (d # \1) # st1" + let ?s1' = "State F1 H ?st1'" + show ?case (is "\x. ?STEP x \ ?MATCH x (State F2 H ?st2')") + proof (intro exI conjI) + show "?STEP ?s1'" + using step_push.hyps + by (auto intro!: Sinca.step_push dest: next_instr2) + next + have "rel_stacktraces (Fubx_get F2) None ?st1' ?st2'" + using step_push.hyps rel_stacktraces_Cons + using Subx.sp_instr.Push[THEN sp_instrs_prefix'] + by (auto simp: take_Suc_conv_app_nth + intro!: rel_stacktraces.intros dest!: next_instrD instr_atD) + thus "?MATCH ?s1' (State F2 H ?st2')" + using step_push.prems rel_F1_F2 + by (auto intro: match.intros) qed next - case (step_op fd2' op ar \2' x) - then have [simp]: "fd2' = fd2" using F2_f by simp - obtain fd1 where fd1_thms: "Finca_get F1 f = Some fd1" "rel_fundef norm_eq fd1 fd2" - using rel_fundefs_Some2[OF rel_F1_F2 F2_f] by auto - show ?case (is "\x. ?STEP x \ ?MATCH x") - proof - - let ?s1' = "State F1 H (Frame f (Suc pc) (x # drop ar (norm_stack \2)) # st1)" - have "?STEP ?s1'" - unfolding \1_def - using step_op fd1_thms take_norm_stack traverse_cast_Dyn_to_norm - by (auto intro!: Sinca.step_op simp: rel_fundef_body_nth) - moreover have "?MATCH ?s1'" - using step_op rel_stacktraces_Cons rel_F1_F2 sp_F2 - by (auto intro!: match.intros rel_stacktraces.intros - simp: take_Suc_conv_app_nth drop_norm_stack Let_def take_map drop_map - traverse_cast_Dyn_replicate) - ultimately show "?thesis" by blast - qed - next - case (step_op_inl fd2' op ar \2' opinl x F2') - then have [simp]: "fd2' = fd2" using F2_f by simp - obtain fd1 where - F1_f: "Finca_get F1 f = Some fd1" and - rel_fd1_fd2: "rel_fundef norm_eq fd1 fd2" - using rel_fundefs_Some2[OF rel_F1_F2 F2_f] by auto - show ?case (is "\x. ?STEP x \ ?MATCH x") - proof - - let ?F1' = "Finca_add F1 f (rewrite_fundef_body fd1 pc (Inca.IOpInl opinl))" - let ?s1' = "State ?F1' H (Frame f (Suc pc) (x # drop ar (norm_stack \2)) # st1)" - have "?STEP ?s1'" - unfolding \1_def - using step_op_inl F1_f rel_fd1_fd2 take_norm_stack traverse_cast_Dyn_to_norm - by (auto intro!: Sinca.step_op_inl simp: rel_fundef_body_nth) - moreover have "?MATCH ?s1'" - proof - show "rel_fundefs (Finca_get ?F1') (Fubx_get F2')" - using step_op_inl.hyps(9) - using rel_fundefs_rewrite_both[OF rel_F1_F2] - using rel_fundef_rewrite_both[OF rel_fd1_fd2] - by auto - next - show "sp_fundefs (Fubx_get F2')" - using step_op_inl.hyps all_same_arities_add sp_fundefs_get[OF sp_F2] - using Sinca.\\\_invertible - by (auto intro!: sp_fundefs_add[OF sp_F2] Subx.sp_rewrite_eq_Ok - simp: Subx.sp_instr_op Let_def) - next - have sp_F2_F2': "Subx.sp (Fubx_get F2) = Subx.sp (Fubx_get F2')" - using step_op_inl.hyps(1,9) - by (auto intro!: sp_same_arities all_same_arities_add) - - let ?fd2' = "rewrite_fundef_body fd2' pc (Ubx.IOpInl opinl)" - - show "rel_stacktraces (Fubx_get F2') - (Frame f (Suc pc) (x # drop ar (norm_stack \2)) # st1) - (Frame f (Suc pc) (OpDyn x # drop ar \2) # st2) None" - using step_op_inl - proof (intro rel_stacktraces.intros) - show "rel_stacktraces (Fubx_get F2') st1 st2 (Some f)" - using step_op_inl Sinca.\\\_invertible - by (auto intro!: rel_stacktraces_rewrite_fundef[OF rel_st1_st2] simp: Subx.sp_instr_op) - next - show "sp_fundef (Fubx_get F2') ?fd2' (take (Suc pc) (body ?fd2')) = - Ok (map typeof (OpDyn x # drop ar \2))" - using step_op_inl Sinca.\\\_invertible sp_prefix - using traverse_cast_Dyn_to_all_Dyn - by (auto simp: take_Suc_conv_app_nth[of pc] sp_F2_F2'[symmetric] Let_def - take_map drop_map traverse_cast_Dyn_replicate) - qed (auto simp add: drop_norm_stack) - qed - ultimately show "?thesis" by blast - qed - next - case (step_op_inl_hit fd2' opinl ar \2' x) - then have [simp]: "fd2' = fd2" using F2_f by simp - obtain fd1 where fd1_thms: "Finca_get F1 f = Some fd1" "rel_fundef norm_eq fd1 fd2" - using rel_fundefs_Some2[OF rel_F1_F2 F2_f] by auto - show ?case (is "\x. ?STEP x \ ?MATCH x") - proof - - let ?s1' = "State F1 H (Frame f (Suc pc) (x # drop ar (norm_stack \2)) # st1)" - have "?STEP ?s1'" - unfolding \1_def - using step_op_inl_hit fd1_thms take_norm_stack traverse_cast_Dyn_to_norm - by (auto intro!: Sinca.step_op_inl_hit simp: rel_fundef_body_nth) - moreover have "?MATCH ?s1'" - using step_op_inl_hit rel_stacktraces_Cons rel_F1_F2 sp_F2 - apply (auto intro!: match.intros rel_stacktraces.intros simp: take_Suc_conv_app_nth) - using drop_norm_stack apply blast - by (simp add: Let_def drop_map take_map traverse_cast_Dyn_replicate) - ultimately show "?thesis" by blast - qed - next - case (step_op_inl_miss fd2' opinl ar \2' x F2') - then have [simp]: "fd2' = fd2" using F2_f by simp - obtain fd1 where - F1_f: "Finca_get F1 f = Some fd1" and - rel_fd1_fd2: "rel_fundef norm_eq fd1 fd2" - using rel_fundefs_Some2[OF rel_F1_F2 F2_f] by auto - show ?case (is "\x. ?STEP x \ ?MATCH x") - proof - - let ?F1' = "Finca_add F1 f (rewrite_fundef_body fd1 pc (Inca.IOp (\
\\\\ opinl)))" - let ?s1' = "State ?F1' H (Frame f (Suc pc) (x # drop ar (norm_stack \2)) # st1)" - have "?STEP ?s1'" - unfolding \1_def - using step_op_inl_miss F1_f rel_fd1_fd2 take_norm_stack traverse_cast_Dyn_to_norm - by (auto intro!: Sinca.step_op_inl_miss simp: rel_fundef_body_nth) - moreover have "?MATCH ?s1'" - proof - show "rel_fundefs (Finca_get ?F1') (Fubx_get F2')" - using step_op_inl_miss - using rel_fundefs_rewrite_both[OF rel_F1_F2] - using rel_fundef_rewrite_both[OF rel_fd1_fd2] - by auto - next - show "sp_fundefs (Fubx_get F2')" - using step_op_inl_miss.hyps all_same_arities_add sp_fundefs_get[OF sp_F2] - by (auto intro!: sp_fundefs_add[OF sp_F2] Subx.sp_rewrite_eq_Ok - simp: Subx.sp_instr_op[symmetric]) - next - have sp_F2_F2': "Subx.sp (Fubx_get F2) = Subx.sp (Fubx_get F2')" - using step_op_inl_miss - by (auto intro!: sp_same_arities all_same_arities_add) - - let ?fd2' = "rewrite_fundef_body fd2' pc (Ubx.IOp (\
\\\\ opinl))" - - show "rel_stacktraces (Fubx_get F2') - (Frame f (Suc pc) (x # drop ar (norm_stack \2)) # st1) - (Frame f (Suc pc) (OpDyn x # drop ar \2) # st2) None" - proof - show "rel_stacktraces (Fubx_get F2') st1 st2 (Some f)" - using step_op_inl_miss.hyps - by (auto intro: rel_stacktraces_rewrite_fundef[OF rel_st1_st2] simp: Subx.sp_instr_op[symmetric]) - next - show "Fubx_get F2' f = Some ?fd2'" - using step_op_inl_miss Subx.Fenv.get_add_eq by blast - next - show "sp_fundef (Fubx_get F2') ?fd2' (take (Suc pc) (body ?fd2')) = - Ok (map typeof (OpDyn x # drop ar \2))" - using \pc < length (body fd2')\ - using sp_prefix step_op_inl_miss traverse_cast_Dyn_to_all_Dyn - by (auto simp add: take_Suc_conv_app_nth[of pc] sp_F2_F2'[symmetric] - traverse_cast_Dyn_replicate take_map drop_map) - qed (simp_all add: drop_norm_stack) - qed - ultimately show "?thesis" by blast + case (step_push_ubx1 n) + let ?st1' = "Frame f l (Suc pc) R1 (box_ubx1 n # \1) # st1" + let ?s1' = "State F1 H ?st1'" + show ?case (is "\x. ?STEP x \ ?MATCH x (State F2 H ?st2')") + proof (intro exI conjI) + show "?STEP ?s1'" + using step_push_ubx1.hyps + by (auto intro!: Sinca.step_push dest: next_instr2) + next + have "rel_stacktraces (Fubx_get F2) None ?st1' ?st2'" + using step_push_ubx1.hyps rel_stacktraces_Cons + using Subx.sp_instr.PushUbx1[THEN sp_instrs_prefix'] + by (auto simp: take_Suc_conv_app_nth + intro!: rel_stacktraces.intros dest!: next_instrD instr_atD) + thus "?MATCH ?s1' (State F2 H ?st2')" + using step_push_ubx1.prems rel_F1_F2 + by (auto intro!: match.intros) qed next - case (step_op_ubx fd2' opubx op ar x) - then have [simp]: "fd2' = fd2" using F2_f by simp - obtain fd1 where fd1_thms: "Finca_get F1 f = Some fd1" "rel_fundef norm_eq fd1 fd2" - using rel_fundefs_Some2[OF rel_F1_F2 F2_f] by auto - show ?case (is "\x. ?STEP x \ ?MATCH x") - proof - - let ?s1' = "State F1 H (Frame f (Suc pc) (Subx.norm_unboxed x # drop ar (norm_stack \2)) # st1)" - have "pc < length (body fd1)" - using fd1_thms(2) step_op_ubx.hyps(2) by auto - hence nth_fd1: "body fd1 ! pc = Inca.instr.IOpInl (\\\ opubx)" - using rel_fundef_body_nth[OF fd1_thms(2)] - using \body fd2' ! pc = IOpUbx opubx\ - by simp - have "\\\\\ (\\\ opubx) (take ar (map Subx.norm_unboxed \2))" - using step_op_ubx - by (auto intro: Sinca.\\\_\\\\\ Subx.\\\\\_to_\\\ simp: take_map) - hence "?STEP ?s1'" - unfolding norm_stack_def \1_def - using step_op_ubx fd1_thms(1) nth_fd1 \pc < length (body fd1)\ - using Subx.\\\\\_correct - by (auto intro!: Sinca.step_op_inl_hit simp: take_map) - moreover have "?MATCH ?s1'" - using step_op_ubx rel_stacktraces_Cons rel_F1_F2 sp_F2 Subx.\\\\\\\\_complete - by (auto intro!: match.intros rel_stacktraces.intros - simp: take_Suc_conv_app_nth drop_norm_stack Let_def drop_map take_map min.absorb2) - ultimately show "?thesis" by blast + case (step_push_ubx2 b) + let ?st1' = "Frame f l (Suc pc) R1 (box_ubx2 b # \1) # st1" + let ?s1' = "State F1 H ?st1'" + show ?case (is "\x. ?STEP x \ ?MATCH x (State F2 H ?st2')") + proof (intro exI conjI) + show "?STEP ?s1'" + using step_push_ubx2.hyps + by (auto intro!: Sinca.step_push dest: next_instr2) + next + have "rel_stacktraces (Fubx_get F2) None ?st1' ?st2'" + using step_push_ubx2.hyps rel_stacktraces_Cons + using Subx.sp_instr.PushUbx2[THEN sp_instrs_prefix'] + by (auto simp: take_Suc_conv_app_nth + intro!: rel_stacktraces.intros dest!: next_instrD instr_atD) + thus "?MATCH ?s1' (State F2 H ?st2')" + using step_push_ubx2.prems rel_F1_F2 + by (auto intro!: match.intros) qed next - case (step_cjump_true fd2' n x \2') - then have False - using F2_f sp_fundefs_get[OF sp_F2, unfolded sp_fundef_def, THEN Subx.sp_no_jump] nth_mem - by force - thus ?case by simp - next - case (step_cjump_false fd2' n x \2') - then have False - using F2_f sp_fundefs_get[OF sp_F2, unfolded sp_fundef_def, THEN Subx.sp_no_jump] nth_mem - by force - thus ?case by simp - next - case (step_fun_call f' fd2' pc' g gd2 ar \2' frame\<^sub>g) - then have [simp]: "f' = f" "pc' = pc" "fd2' = fd2" using F2_f by auto - obtain fd1 where "Finca_get F1 f = Some fd1" and rel_fd1_fd2: "rel_fundef norm_eq fd1 fd2" - using rel_fundefs_Some2[OF rel_F1_F2 F2_f] by auto - obtain gd1 where "Finca_get F1 g = Some gd1" and rel_gd1_gd2: "rel_fundef norm_eq gd1 gd2" - using step_fun_call rel_fundefs_Some2[OF rel_F1_F2] by blast - have pc_in_range: "pc < length (body fd1)" - using \pc' < length (body fd2')\ rel_fundef_body_length[OF rel_fd1_fd2] - by simp - show ?case (is "\x. ?STEP x \ ?MATCH x") - proof - - let ?s1' = "State F1 H (Frame g 0 (take (arity gd1) \1) # Frame f pc \1 # st1)" - have "?STEP ?s1'" - using step_fun_call rel_stacktraces_Cons.hyps(2) pc_in_range - using \Finca_get F1 f = Some fd1\ \Finca_get F1 g = Some gd1\ - using rel_fundef_body_nth[OF rel_fd1_fd2] - using rel_fundef_arities[OF rel_gd1_gd2, symmetric] - by (auto intro!: Sinca.step_fun_call) - moreover have "?MATCH ?s1'" - proof - show "rel_fundefs (Finca_get F1) (Fubx_get F2)" - using rel_F1_F2 . - next - show "sp_fundefs (Fubx_get F2)" - using sp_F2 . - next - show "rel_stacktraces (Fubx_get F2) - (Frame g 0 (take (arity gd1) \1) # Frame f pc \1 # st1) - (frame\<^sub>g # Frame f pc \2 # st2) None" - unfolding step_fun_call(9) - proof (rule rel_stacktraces.intros(2)) - show "rel_stacktraces (Fubx_get F2) (Frame f pc \1 # st1) (Frame f pc \2 # st2) (Some g)" - using rel_st1_st2 F2_f \1_def sp_prefix - using step_fun_call - by (auto intro!: rel_stacktraces.intros - elim!: list.pred_mono_strong simp: is_valid_fun_call_def) - next - show "take (arity gd1) \1 = norm_stack (take ar \2')" - using \1_def rel_fundef_arities[OF rel_gd1_gd2] - using step_fun_call by (simp add: take_norm_stack) - next - show "Fubx_get F2 g = Some gd2" - using step_fun_call by simp - next - show "sp_fundef (Fubx_get F2) gd2 (take 0 (body gd2)) = Ok (map typeof (take ar \2'))" - using step_fun_call - by (auto elim: replicate_eq_map) - qed simp_all - qed - ultimately show "?thesis" by blast + case (step_pop d \2') + let ?st1' = "Frame f l (Suc pc) R1 (map Subx.norm_unboxed \2') # st1" + let ?s1' = "State F1 H ?st1'" + show ?case (is "\x. ?STEP x \ ?MATCH x (State F2 H ?st2')") + proof (intro exI conjI) + show "?STEP ?s1'" + unfolding \1_def + using step_pop.hyps + by (auto intro!: Sinca.step_pop dest: next_instr2) + next + have "rel_stacktraces (Fubx_get F2) None ?st1' ?st2'" + using step_pop.hyps rel_stacktraces_Cons + by (auto simp: take_Suc_conv_app_nth + intro!: rel_stacktraces.intros sp_instrs_prefix' Subx.sp_instr.Pop + dest!: next_instrD instr_atD) + thus "?MATCH ?s1' (State F2 H ?st2')" + using step_pop.prems rel_F1_F2 + by (auto intro!: match.intros) qed next - case (step_fun_end f' fd2' \2\<^sub>g pc' \2' frame\<^sub>g g pc\<^sub>g frame\<^sub>g' st2') - then have [simp]: "f' = f" "fd2' = fd2" "pc' = pc" "\2' = \2" using F2_f by auto - obtain fd1 where "Finca_get F1 f = Some fd1" and rel_fd1_fd2: "rel_fundef norm_eq fd1 fd2" - using rel_fundefs_Some2[OF rel_F1_F2 F2_f] by auto - obtain gd2 \1\<^sub>g st1' where - st1_def: "st1 = Frame g pc\<^sub>g \1\<^sub>g # st1'" and - "\1\<^sub>g = norm_stack \2\<^sub>g" and - rel_st1'_st2': "rel_stacktraces (Fubx_get F2) st1' st2' (Some g)" and - "Fubx_get F2 g = Some gd2" and - sp_prefix_gd2: "Subx.sp (Fubx_get F2) (take pc\<^sub>g (body gd2)) - (replicate (arity gd2) None) = Ok (map typeof \2\<^sub>g)" and - pc\<^sub>g_in_range: "pc\<^sub>g < length (body gd2)" and - "body gd2 ! pc\<^sub>g = Ubx.ICall f" and - prefix_all_dyn_\2\<^sub>g: "list_all is_dyn_operand (take (arity fd2) \2\<^sub>g)" - using rel_st1_st2 step_fun_end - by (auto elim: rel_stacktraces.cases simp: is_valid_fun_call_def) - have pc_at_end: "pc = length (body fd1)" - using \pc' = length (body fd2')\ - using rel_fundef_body_length[OF rel_fd1_fd2] - by simp - show ?case (is "\x. ?STEP x \ ?MATCH x") - proof - - let ?s1' = "State F1 H (Frame g (Suc pc\<^sub>g) (\1 @ drop (arity fd1) \1\<^sub>g) # st1')" - have "?STEP ?s1'" - using step_fun_end.hyps(2) - using st1_def \\1\<^sub>g = norm_stack \2\<^sub>g\ - using \Finca_get F1 f = Some fd1\ pc_at_end - using rel_fundef_arities[OF rel_fd1_fd2, symmetric] - by (auto intro!: Sinca.step_fun_end) - moreover have "?MATCH ?s1'" - using rel_F1_F2 sp_F2 - proof - show "rel_stacktraces (Fubx_get F2) - (Frame g (Suc pc\<^sub>g) (\1 @ drop (arity fd1) \1\<^sub>g) # st1') - (frame\<^sub>g' # st2') None" - unfolding step_fun_end(6) - using rel_st1'_st2' - proof (rule rel_stacktraces.rel_stacktraces_Cons) - show "\1 @ drop (arity fd1) \1\<^sub>g = norm_stack (\2' @ drop (arity fd2') \2\<^sub>g)" - using \1_def \\1\<^sub>g = norm_stack \2\<^sub>g\ - using rel_fundef_arities[OF rel_fd1_fd2] - by (simp add: norm_stack_append drop_norm_stack) + case (step_get n d) + let ?st1' = "Frame f l (Suc pc) R1 (R1 ! n # map Subx.norm_unboxed \2) # st1" + let ?s1' = "State F1 H ?st1'" + show ?case (is "\x. ?STEP x \ ?MATCH x (State F2 H ?st2')") + proof (intro exI conjI) + show "?STEP ?s1'" + unfolding \1_def R1_def + using step_get.hyps + by (auto intro!: Sinca.step_get dest: next_instr2) + next + have "rel_stacktraces (Fubx_get F2) None ?st1' ?st2'" + using step_get.hyps rel_stacktraces_Cons + by (auto simp: take_Suc_conv_app_nth + intro!: rel_stacktraces.intros sp_instrs_prefix' Subx.sp_instr.Get + dest!: next_instrD instr_atD) + thus "?MATCH ?s1' (State F2 H ?st2')" + using step_get.prems rel_F1_F2 + by (auto intro!: match.intros) + qed + next + case (step_get_ubx_hit \ n d blob) + let ?st1' = "Frame f l (Suc pc) R1 (R1 ! n # map Subx.norm_unboxed \2) # st1" + let ?s1' = "State F1 H ?st1'" + show ?case (is "\x. ?STEP x \ ?MATCH x (State F2 H ?st2')") + proof (intro exI conjI) + show "?STEP ?s1'" + unfolding \1_def R1_def + using step_get_ubx_hit.hyps + by (auto intro!: Sinca.step_get dest: next_instr2) + next + have "rel_stacktraces (Fubx_get F2) None ?st1' ?st2'" + using step_get_ubx_hit.hyps rel_stacktraces_Cons + by (auto simp: take_Suc_conv_app_nth + intro!: rel_stacktraces.intros sp_instrs_prefix' Subx.sp_instr.GetUbx + dest!: next_instrD instr_atD) + thus "?MATCH ?s1' (State F2 H ?st2')" + using step_get_ubx_hit.prems rel_F1_F2 + by (auto intro!: match.intros) + qed + next + case (step_get_ubx_miss \ n d F2') + hence "R1 ! n = d" + by (simp add: R1_def) + let ?st1' = "Frame f l (Suc pc) R1 (R1 ! n # map Subx.norm_unboxed \2) # st1" + let ?s1' = "State F1 H ?st1'" + show ?case (is "\x. ?STEP x \ ?MATCH x (State F2' H ?st2')") + proof (intro exI conjI) + show "?STEP ?s1'" + unfolding \1_def R1_def + using step_get_ubx_miss.hyps + by (auto intro!: Sinca.step_get dest: next_instr2) + next + have "rel_stacktraces (Fubx_get F2') None ?st1' ?st2'" + apply simp + proof (rule rel_stacktraces.intros) + show "rel_stacktraces (Fubx_get F2') (Some f) st1 (Subx.box_stack f st2)" + unfolding step_get_ubx_miss.hyps + using rel_st1_st2 + by (rule rel_stacktraces_map_entry_gneralize_fundefI) + next + show "Fubx_get F2' f = Some (Subx.generalize_fundef fd2)" + unfolding step_get_ubx_miss.hyps + using F2_f + by simp + next + show "map_of (body (Subx.generalize_fundef fd2)) l = Some (map Subx.generalize_instr instrs)" + unfolding Subx.map_of_generalize_fundef_conv + unfolding map_of_fd2_l + by simp + next + show "Subx.sp_instrs (map_option funtype \ Fubx_get F2') + (return (Subx.generalize_fundef fd2)) + (take (Suc pc) (map Subx.generalize_instr instrs)) + [] (map typeof (OpDyn d # map Subx.box_operand \2))" + using step_get_ubx_miss.hyps F2_f map_of_fd2_l + by (auto simp: take_map take_Suc_conv_app_nth simp del: map_append + intro!: sp_instrs_prefix'[THEN Subx.sp_instrs_generalize0] Subx.sp_instr.GetUbx + dest!: next_instrD instr_atD) + qed (insert R1_def all_dyn_R2 \R1 ! n = d\, simp_all) + thus "?MATCH ?s1' (State F2' H ?st2')" + using step_get_ubx_miss.prems rel_F1_F2 + unfolding step_get_ubx_miss.hyps + by (auto intro!: match.intros rel_fundefs_generalizeI) + qed + next + case (step_set n blob d R2' \2') + let ?st1' = "Frame f l (Suc pc) (map Subx.norm_unboxed R2') (map Subx.norm_unboxed \2') # st1" + let ?s1' = "State F1 H ?st1'" + show ?case (is "\x. ?STEP x \ ?MATCH x (State F2 H ?st2')") + proof (intro exI conjI) + show "?STEP ?s1'" + unfolding \1_def R1_def + using step_set.hyps + by (auto simp: map_update intro!: Sinca.step_set dest!: next_instr2) + next + have "rel_stacktraces (Fubx_get F2) None ?st1' ?st2'" + using step_set.hyps rel_stacktraces_Cons + by (auto simp: take_Suc_conv_app_nth + intro!: rel_stacktraces.intros sp_instrs_prefix' Subx.sp_instr.Set + intro: list_all_list_updateI + dest!: next_instrD instr_atD) + thus "?MATCH ?s1' (State F2 H ?st2')" + using step_set.prems rel_F1_F2 + by (auto intro!: match.intros) + qed + next + case (step_set_ubx \ n blob d R2' \2') + let ?st1' = "Frame f l (Suc pc) (map Subx.norm_unboxed R2') (map Subx.norm_unboxed \2') # st1" + let ?s1' = "State F1 H ?st1'" + show ?case (is "\x. ?STEP x \ ?MATCH x (State F2 H ?st2')") + proof (intro exI conjI) + show "?STEP ?s1'" + unfolding \1_def R1_def + using step_set_ubx.hyps + by (auto simp: map_update intro!: Sinca.step_set dest!: next_instr2) + next + have "rel_stacktraces (Fubx_get F2) None ?st1' ?st2'" + using step_set_ubx.hyps rel_stacktraces_Cons + by (auto simp: take_Suc_conv_app_nth + intro!: rel_stacktraces.intros sp_instrs_prefix' Subx.sp_instr.SetUbx + intro: list_all_list_updateI + dest!: next_instrD instr_atD) + thus "?MATCH ?s1' (State F2 H ?st2')" + using step_set_ubx.prems rel_F1_F2 + by (auto intro!: match.intros) + qed + next + case (step_load x i i' d \2') + let ?st1' = "Frame f l (Suc pc) R1 (d # map Subx.norm_unboxed \2') # st1" + let ?s1' = "State F1 H ?st1'" + show ?case (is "\x. ?STEP x \ ?MATCH x (State F2 H ?st2')") + proof (intro exI conjI) + show "?STEP ?s1'" + unfolding \1_def + using step_load.hyps + by (auto intro!: Sinca.step_load dest!: next_instr2) + next + have "rel_stacktraces (Fubx_get F2) None ?st1' ?st2'" + using step_load.hyps rel_stacktraces_Cons + by (auto simp: take_Suc_conv_app_nth + intro!: rel_stacktraces.intros sp_instrs_prefix' Subx.sp_instr.Load + dest!: next_instrD instr_atD) + thus "?MATCH ?s1' (State F2 H ?st2')" + using step_load.prems rel_F1_F2 + by (auto intro!: match.intros) + qed + next + case (step_load_ubx_hit \ x i i' d blob \2') + let ?st1' = "Frame f l (Suc pc) R1 (d # map Subx.norm_unboxed \2') # st1" + let ?s1' = "State F1 H ?st1'" + show ?case (is "\x. ?STEP x \ ?MATCH x (State F2 H ?st2')") + proof (intro exI conjI) + show "?STEP ?s1'" + unfolding \1_def + using step_load_ubx_hit.hyps + by (auto intro!: Sinca.step_load dest!: next_instr2) + next + have "rel_stacktraces (Fubx_get F2) None ?st1' ?st2'" + using step_load_ubx_hit.hyps rel_stacktraces_Cons + by (auto simp: take_Suc_conv_app_nth + intro!: rel_stacktraces.intros sp_instrs_prefix' Subx.sp_instr.LoadUbx + dest!: next_instrD instr_atD) + thus "?MATCH ?s1' (State F2 H ?st2')" + using step_load_ubx_hit.prems rel_F1_F2 + by (auto intro!: match.intros) + qed + next + case (step_load_ubx_miss \ x i i' d F2' \2') + let ?st1' = "Frame f l (Suc pc) R1 (d # map Subx.norm_unboxed \2') # st1" + let ?s1' = "State F1 H ?st1'" + show ?case (is "\x. ?STEP x \ ?MATCH x (State F2' H ?st2')") + proof (intro exI conjI) + show "?STEP ?s1'" + unfolding \1_def + using step_load_ubx_miss.hyps + by (auto intro!: Sinca.step_load dest!: next_instr2) + next + have "rel_stacktraces (Fubx_get F2') None ?st1' ?st2'" + apply simp + proof (rule rel_stacktraces.intros) + show "rel_stacktraces (Fubx_get F2') (Some f) st1 (Subx.box_stack f st2)" + unfolding step_load_ubx_miss + using rel_st1_st2 + by (rule rel_stacktraces_map_entry_gneralize_fundefI) + next + show "Fubx_get F2' f = Some (Subx.generalize_fundef fd2)" + unfolding step_load_ubx_miss.hyps + using F2_f by (simp add: Subx.map_of_generalize_fundef_conv) + next + show "map_of (body (Subx.generalize_fundef fd2)) l = + Some (map Subx.generalize_instr instrs)" + unfolding Subx.map_of_generalize_fundef_conv + using step_load_ubx_miss.hyps F2_f map_of_fd2_l + by simp + next + show "Subx.sp_instrs (map_option funtype \ Fubx_get F2') (return (Subx.generalize_fundef fd2)) + (take (Suc pc) (map Subx.generalize_instr instrs)) + [] (map typeof (OpDyn d # map Subx.box_operand \2'))" + using step_load_ubx_miss.hyps F2_f map_of_fd2_l + by (auto simp: take_map take_Suc_conv_app_nth simp del: map_append + intro!: sp_instrs_prefix'[THEN Subx.sp_instrs_generalize0] Subx.sp_instr.LoadUbx + dest!: next_instrD instr_atD) + qed (insert all_dyn_R2, simp_all add: R1_def) + thus "?MATCH ?s1' (State F2' H ?st2')" + using step_load_ubx_miss.prems rel_F1_F2 + unfolding step_load_ubx_miss.hyps + by (auto intro: match.intros rel_fundefs_generalizeI) + qed + next + case (step_store x i i' y d H' \2') + let ?st1' = "Frame f l (Suc pc) R1 (map Subx.norm_unboxed \2') # st1" + let ?s1' = "State F1 H' ?st1'" + show ?case (is "\x. ?STEP x \ ?MATCH x (State F2 H' ?st2')") + proof (intro exI conjI) + show "?STEP ?s1'" + unfolding \1_def + using step_store.hyps + by (auto intro: Sinca.step_store dest!: next_instr2) + next + have "rel_stacktraces (Fubx_get F2) None ?st1' ?st2'" + using step_store.hyps rel_stacktraces_Cons + by (auto simp: take_Suc_conv_app_nth + intro!: rel_stacktraces.intros sp_instrs_prefix' Subx.sp_instr.Store + dest!: next_instrD instr_atD) + thus "?MATCH ?s1' (State F2 H' ?st2')" + using step_store.prems rel_F1_F2 + by (auto intro: match.intros) + qed + next + case (step_store_ubx \ x i i' blob d H' \2') + let ?st1' = "Frame f l (Suc pc) R1 (map Subx.norm_unboxed \2') # st1" + let ?s1' = "State F1 H' ?st1'" + show ?case (is "\x. ?STEP x \ ?MATCH x (State F2 H' ?st2')") + proof (intro exI conjI) + show "?STEP ?s1'" + unfolding \1_def + using step_store_ubx.hyps + by (auto intro: Sinca.step_store dest!: next_instr2) + next + have "rel_stacktraces (Fubx_get F2) None ?st1' ?st2'" + using step_store_ubx.hyps rel_stacktraces_Cons + by (auto simp: take_Suc_conv_app_nth + intro!: rel_stacktraces.intros sp_instrs_prefix' Subx.sp_instr.StoreUbx + dest!: next_instrD instr_atD) + thus "?MATCH ?s1' (State F2 H' ?st2')" + using step_store_ubx.prems rel_F1_F2 + by (auto intro: match.intros) + qed + next + case (step_op op ar \2' x) + let ?st1' = "Frame f l (Suc pc) R1 (x # drop ar (map Subx.norm_unboxed \2)) # st1" + let ?s1' = "State F1 H ?st1'" + show ?case (is "\x. ?STEP x \ ?MATCH x (State F2 H ?st2')") + proof (intro exI conjI) + show "?STEP ?s1'" + unfolding \1_def + using step_op.hyps + by (auto simp: take_map ap_map_list_cast_Dyn_to_map_norm[symmetric] + intro!: Sinca.step_op dest!: next_instr2) + next + have "rel_stacktraces (Fubx_get F2) None ?st1' ?st2'" + using step_op.hyps rel_stacktraces_Cons + by (auto simp: take_Suc_conv_app_nth drop_map map_eq_append_map_drop + simp: ap_map_list_cast_Dyn_map_typeof_replicate_conv + intro!: rel_stacktraces.intros sp_instrs_prefix' Subx.sp_instr.Op + dest!: next_instrD instr_atD) + thus "?MATCH ?s1' (State F2 H ?st2')" + using step_op.prems rel_F1_F2 + by (auto intro: match.intros) + qed + next + case (step_op_inl op ar \2' opinl x F2') + let ?F1' = "Sinca.Fenv.map_entry F1 f (\fd. rewrite_fundef_body fd l pc (Inca.IOpInl opinl))" + let ?st1' = "Frame f l (Suc pc) R1 (x # drop ar (map Subx.norm_unboxed \2)) # st1" + let ?s1' = "State ?F1' H ?st1'" + show ?case (is "\x. ?STEP x \ ?MATCH x (State F2' H ?st2')") + proof (intro exI conjI) + show "?STEP ?s1'" + unfolding \1_def + using step_op_inl.hyps + by (auto simp: take_map ap_map_list_cast_Dyn_to_map_norm[symmetric] + intro!: Sinca.step_op_inl dest!: next_instr2) + next + show "?MATCH ?s1' (State F2' H ?st2')" + using step_op_inl.prems + proof (rule match.intros) + show "rel_fundefs (Finca_get ?F1') (Fubx_get F2')" + unfolding step_op_inl.hyps + using rel_F1_F2 + by (auto intro: rel_fundefs_rewriteI) + next + let ?fd2' = "rewrite_fundef_body fd2 l pc (Ubx.instr.IOpInl opinl)" + let ?instrs' = "instrs[pc := Ubx.instr.IOpInl opinl]" + show "rel_stacktraces (Fubx_get F2') None ?st1' ?st2'" + proof (rule rel_stacktraces.intros) + show "rel_stacktraces (Fubx_get F2') (Some f) st1 st2" + using step_op_inl.hyps rel_st1_st2 Sinca.\\\_invertible + by (auto simp: Subx.sp_instr_Op_OpInl_conv + intro: rel_stacktraces_map_entry_rewrite_fundef_body) next - show "Fubx_get F2 g = Some gd2" - using \Fubx_get F2 g = Some gd2\ . + show "Fubx_get F2' f = Some ?fd2'" + unfolding step_op_inl.hyps + using F2_f by simp next - have "arity fd2 \ length \2\<^sub>g" - using step_fun_end.hyps(2) by simp - moreover have "list_all (\x. x = None) (take (arity fd2) (map typeof \2\<^sub>g))" - using prefix_all_dyn_\2\<^sub>g - by (auto elim!: list.pred_mono_strong simp: take_map list.pred_map is_dyn_operand_def) - moreover have "map typeof \2 = [None]" - using \pc' = length (body fd2')\ sp_prefix - using sp_fundefs_get[OF sp_F2 F2_f] - by simp - ultimately show "sp_fundef (Fubx_get F2) gd2 (take (Suc pc\<^sub>g) (body gd2)) = - Ok (map typeof (\2' @ drop (arity fd2') \2\<^sub>g))" - using sp_prefix_gd2 - using \body gd2 ! pc\<^sub>g = Ubx.ICall f\ F2_f - by (auto dest: list_all_eq_const_imp_replicate - simp: take_Suc_conv_app_nth[OF pc\<^sub>g_in_range] Let_def drop_map) - qed simp + show "map_of (body ?fd2') l = Some ?instrs'" + using map_of_fd2_l by simp + next + show "Subx.sp_instrs (map_option funtype \ Fubx_get F2') + (return (rewrite_fundef_body fd2 l pc (Ubx.instr.IOpInl opinl))) + (take (Suc pc) ?instrs') [] (map typeof (OpDyn x # drop ar \2))" + using rel_stacktraces_Cons step_op_inl.hyps + by (auto simp add: take_Suc_conv_app_nth Subx.Fenv.map_option_comp_map_entry + map_eq_append_map_drop ap_map_list_cast_Dyn_map_typeof_replicate_conv + Sinca.\\\_invertible + intro!: sp_instrs_prefix' Subx.sp_instr.OpInl + dest!: next_instrD instr_atD) + qed (insert all_dyn_R2 R1_def, simp_all add: drop_map) qed - ultimately show "?thesis" by blast + qed + next + case (step_op_inl_hit opinl ar \2' x) + let ?st1' = "Frame f l (Suc pc) R1 (x # drop ar (map Subx.norm_unboxed \2)) # st1" + let ?s1' = "State F1 H ?st1'" + show ?case (is "\x. ?STEP x \ ?MATCH x (State F2 H ?st2')") + proof (intro exI conjI) + show "?STEP ?s1'" + unfolding \1_def + using step_op_inl_hit.hyps + by (auto simp: take_map ap_map_list_cast_Dyn_to_map_norm[symmetric] + intro!: Sinca.step_op_inl_hit dest!: next_instr2) + next + show "?MATCH ?s1' (State F2 H ?st2')" + using step_op_inl_hit.prems rel_F1_F2 + proof (rule match.intros) + show "rel_stacktraces (Fubx_get F2) None ?st1' ?st2'" + using step_op_inl_hit.hyps rel_stacktraces_Cons + by (auto simp: take_Suc_conv_app_nth drop_map map_eq_append_map_drop + simp: ap_map_list_cast_Dyn_map_typeof_replicate_conv + intro!: rel_stacktraces.intros sp_instrs_prefix' Subx.sp_instr.OpInl + dest!: next_instrD instr_atD) + qed + qed + next + case (step_op_inl_miss opinl ar \2' x F2') + let ?F1' = "Sinca.Fenv.map_entry F1 f (\fd. rewrite_fundef_body fd l pc (Inca.IOp (\
\\\\ opinl)))" + let ?st1' = "Frame f l (Suc pc) R1 (x # drop ar (map Subx.norm_unboxed \2)) # st1" + let ?s1' = "State ?F1' H ?st1'" + show ?case (is "\x. ?STEP x \ ?MATCH x (State F2' H ?st2')") + proof (intro exI conjI) + show "?STEP ?s1'" + unfolding \1_def + using step_op_inl_miss.hyps + by (auto simp: take_map ap_map_list_cast_Dyn_to_map_norm[symmetric] + intro!: Sinca.step_op_inl_miss dest!: next_instr2) + next + show "?MATCH ?s1' (State F2' H ?st2')" + using step_op_inl_miss.prems + proof (rule match.intros) + show "rel_fundefs (Finca_get ?F1') (Fubx_get F2')" + unfolding step_op_inl_miss.hyps + using rel_F1_F2 + by (auto intro: rel_fundefs_rewriteI) + next + let ?fd2' = "rewrite_fundef_body fd2 l pc (Ubx.instr.IOp (\
\\\\ opinl))" + let ?instrs' = "instrs[pc := Ubx.instr.IOp (\
\\\\ opinl)]" + show "rel_stacktraces (Fubx_get F2') None ?st1' ?st2'" + proof (rule rel_stacktraces.intros) + show "rel_stacktraces (Fubx_get F2') (Some f) st1 st2" + using step_op_inl_miss.hyps rel_st1_st2 Sinca.\\\_invertible + by (auto intro: rel_stacktraces_map_entry_rewrite_fundef_body + Subx.sp_instr_Op_OpInl_conv[OF refl, symmetric]) + next + show "Fubx_get F2' f = Some ?fd2'" + unfolding step_op_inl_miss.hyps + using F2_f by simp + next + show "map_of (body ?fd2') l = Some ?instrs'" + using map_of_fd2_l by simp + next + show "Subx.sp_instrs (map_option funtype \ Fubx_get F2') + (return (rewrite_fundef_body fd2 l pc (Ubx.instr.IOp (\
\\\\ opinl)))) + (take (Suc pc) ?instrs') [] (map typeof (OpDyn x # drop ar \2))" + using rel_stacktraces_Cons step_op_inl_miss.hyps + by (auto simp add: take_Suc_conv_app_nth Subx.Fenv.map_option_comp_map_entry + map_eq_append_map_drop ap_map_list_cast_Dyn_map_typeof_replicate_conv + Sinca.\\\_invertible + intro!: sp_instrs_prefix' Subx.sp_instr.Op + dest!: next_instrD instr_atD) + qed (insert all_dyn_R2 R1_def, simp_all add: drop_map) + qed + qed + next + case (step_op_ubx opubx op ar x) + let ?st1' = "Frame f l (Suc pc) R1 (Subx.norm_unboxed x # drop ar (map Subx.norm_unboxed \2)) # st1" + let ?s1' = "State F1 H ?st1'" + show ?case (is "\x. ?STEP x \ ?MATCH x (State F2 H ?st2')") + proof (intro exI conjI) + show "?STEP ?s1'" + unfolding \1_def + using step_op_ubx.hyps + by (auto simp: take_map + intro!: Sinca.step_op_inl_hit + intro: Subx.\\\\\_to_\\\[THEN Sinca.\\\_\\\\\] Subx.\\\\\_correct + dest: next_instr2) + next + show "?MATCH ?s1' (State F2 H ?st2')" + using step_op_ubx.prems rel_F1_F2 + proof (rule match.intros) + show "rel_stacktraces (Fubx_get F2) None ?st1' ?st2'" + using step_op_ubx.hyps rel_stacktraces_Cons + by (auto simp: take_Suc_conv_app_nth drop_map map_eq_append_map_drop + intro!: rel_stacktraces.intros sp_instrs_prefix' Subx.sp_instr.OpUbx + dest!: next_instrD instr_atD + dest!: Subx.\\\\\\\\_complete) + qed + qed + next + case (step_cjump l\<^sub>t l\<^sub>f x d l' \2') + hence "instr_at fd2 l pc = Some (Ubx.instr.ICJump l\<^sub>t l\<^sub>f)" + using F2_f by (auto dest!: next_instrD) + hence pc_in_dom: "pc < length instrs" and nth_instrs_pc: "instrs ! pc = Ubx.instr.ICJump l\<^sub>t l\<^sub>f" + using map_of_fd2_l by (auto dest!: instr_atD) + hence "{l\<^sub>t, l\<^sub>f} \ fst ` set (body fd2)" + using all_jumps_in_range by (auto simp: list_all_length) + moreover have "l' \ {l\<^sub>t, l\<^sub>f}" + using step_cjump.hyps by auto + ultimately have "l' \ fst ` set (body fd2)" + by blast + then obtain instrs' where map_of_l': "map_of (body fd2) l' = Some instrs'" + by (auto dest: weak_map_of_SomeI) + + have pc_def: "pc = length instrs - 1" + using is_jump_nthD[OF _ pc_in_dom] nth_instrs_pc by simp + have \2'_eq_Nil: "\2' = []" + using sp_instr_last[OF pc_def] step_cjump.hyps + by (auto simp: nth_instrs_pc elim!: Subx.sp_instr.cases) + + let ?st1' = "Frame f l' 0 R1 (map Subx.norm_unboxed \2') # st1" + let ?s1' = "State F1 H ?st1'" + show ?case (is "\x. ?STEP x \ ?MATCH x (State F2 H ?st2')") + proof (intro exI conjI) + show "?STEP ?s1'" + unfolding \1_def + using step_cjump.hyps + by (auto intro!: Sinca.step_cjump dest: next_instr2) + next + show "?MATCH ?s1' (State F2 H ?st2')" + using step_cjump.prems rel_F1_F2 + proof (rule match.intros) + show "rel_stacktraces (Fubx_get F2) None ?st1' ?st2'" + using map_of_l' rel_stacktraces_Cons(1,3-5) + by (auto simp: \2'_eq_Nil intro!: rel_stacktraces.intros intro: Subx.sp_instrs.Nil) + qed + qed + next + case (step_call g gd2 frame\<^sub>g) + then obtain gd1 where + F1_g: "Finca_get F1 g = Some gd1" and rel_gd1_gd2: "rel_fundef (=) norm_eq gd1 gd2" + using rel_fundefs_Some2[OF rel_F1_F2] by auto + + have wf_gd2: "Subx.wf_fundef (map_option funtype \ Fubx_get F2) gd2" + using Subx.wf_fundefs_getD[OF wf_F2] step_call.hyps by simp + + obtain intrs\<^sub>g where gd2_fst_bblock: "map_of (body gd2) (fst (hd (body gd2))) = Some intrs\<^sub>g" + using Subx.wf_fundef_body_neq_NilD[OF wf_gd2] + by (metis hd_in_set map_of_eq_None_iff not_Some_eq prod.collapse prod_in_set_fst_image_conv) + + let ?frame\<^sub>g = "allocate_frame g gd1 (take (arity gd1) \1) uninitialized" + let ?st1' = "?frame\<^sub>g # Frame f l pc R1 \1 # st1" + let ?s1' = "State F1 H ?st1'" + show ?case (is "\x. ?STEP x \ ?MATCH x (State F2 H ?st2')") + proof (intro exI conjI) + show "?STEP ?s1'" + unfolding \1_def + using step_call.hyps F1_g rel_gd1_gd2 + by (auto simp: rel_fundef_arities intro!: Sinca.step_call dest: next_instr2) + next + show "?MATCH ?s1' (State F2 H ?st2')" + using step_call.prems rel_F1_F2 + proof (rule match.intros) + have FOO: "fst (hd (body gd1)) = fst (hd (body gd2))" + apply (rule rel_fundef_rel_fst_hd_bodies[OF rel_gd1_gd2]) + using Subx.wf_fundefs_getD[OF wf_F2 \Fubx_get F2 g = Some gd2\] + by (auto dest: Subx.wf_fundef_body_neq_NilD) + show "rel_stacktraces (Fubx_get F2) None ?st1' ?st2'" + unfolding step_call allocate_frame_def FOO + proof (rule rel_stacktraces.intros(2)) + show "rel_stacktraces (Fubx_get F2) (Some g) + (Frame f l pc R1 \1 # st1) (Frame f l pc R2 \2 # st2)" + using step_call rel_stacktraces_Cons + by (auto simp: is_valid_fun_call_def intro: rel_stacktraces.intros) + next + show "take (arity gd1) \1 @ replicate (fundef_locals gd1) uninitialized = + map Subx.norm_unboxed (take (arity gd2) \2 @ + replicate (fundef_locals gd2) (OpDyn uninitialized))" + using rel_gd1_gd2 + by (simp add: rel_fundef_arities rel_fundef_locals take_map \1_def) + next + show "list_all is_dyn_operand (take (arity gd2) \2 @ + replicate (fundef_locals gd2) (OpDyn uninitialized))" + using step_call.hyps by auto + qed (insert step_call gd2_fst_bblock, simp_all add: Subx.sp_instrs.Nil) + qed + qed + next + case (step_return fd2' \2\<^sub>g frame\<^sub>g' g l\<^sub>g pc\<^sub>g R2\<^sub>g st2') + hence fd2_fd2'[simp]: "fd2' = fd2" + using F2_f by simp + then obtain fd1 where + F1_f: "Finca_get F1 f = Some fd1" and rel_fd1_fd2: "rel_fundef (=) norm_eq fd1 fd2" + using F2_f rel_fundefs_Some2[OF rel_F1_F2] by auto + show ?case + using rel_st1_st2 unfolding \Frame g l\<^sub>g pc\<^sub>g R2\<^sub>g \2\<^sub>g # st2' = st2\[symmetric] + proof (cases rule: rel_stacktraces.cases) + case (rel_stacktraces_Cons st1' \1\<^sub>g R1\<^sub>g gd2 instrs) + hence is_valid_call_f: "is_valid_fun_call (Fubx_get F2) g l\<^sub>g pc\<^sub>g \2\<^sub>g f" + by simp + let ?s1' = "State F1 H (Frame g l\<^sub>g (Suc pc\<^sub>g) R1\<^sub>g (\1 @ drop (arity fd2) \1\<^sub>g) # st1')" + show ?thesis (is "\x. ?STEP x \ ?MATCH x (State F2 H ?st2')") + proof (intro exI conjI) + show "?STEP ?s1'" + unfolding rel_stacktraces_Cons + proof (rule Sinca.step.step_return) + show "next_instr (Finca_get F1) f l pc = Some Inca.instr.IReturn" + using \next_instr (Fubx_get F2) f l pc = Some Ubx.instr.IReturn\ + using rel_fundefs_next_instr2[OF rel_F1_F2] + by force + next + show "Finca_get F1 f = Some fd1" + by (rule F1_f) + qed (insert step_return.hyps rel_fd1_fd2, + simp_all add: \1_def rel_fundef_arities rel_fundef_return) + next + show "?MATCH ?s1' (State F2 H ?st2')" + unfolding step_return.hyps + proof (rule match.intros) + have "next_instr (Fubx_get F2) g l\<^sub>g pc\<^sub>g = Some (Ubx.instr.ICall f)" and + "arity fd2 \ length \2\<^sub>g" and "list_all is_dyn_operand (take (arity fd2) \2\<^sub>g)" + using is_valid_call_f[unfolded is_valid_fun_call_def] F2_f + by simp_all + hence + pc\<^sub>g_in_range: "pc\<^sub>g < length instrs" and + nth_instrs_pc\<^sub>g: "instrs ! pc\<^sub>g = Ubx.instr.ICall f" + using rel_stacktraces_Cons + by (auto dest!: next_instrD instr_atD) + have replicate_None: "replicate (arity fd2) None = map typeof (take (arity fd2) \2\<^sub>g)" + using \arity fd2 \ length \2\<^sub>g\ \list_all is_dyn_operand (take (arity fd2) \2\<^sub>g)\ + by (auto simp: is_dyn_operand_eq_typeof list_all_iff intro!: replicate_eq_map) + show "rel_stacktraces (Fubx_get F2) None + (Frame g l\<^sub>g (Suc pc\<^sub>g) R1\<^sub>g (\1 @ drop (arity fd2) \1\<^sub>g) # st1') + (Frame g l\<^sub>g (Suc pc\<^sub>g) R2\<^sub>g (\2 @ drop (arity fd2') \2\<^sub>g) # st2')" + using rel_stacktraces_Cons + apply (auto simp: \1_def drop_map take_Suc_conv_app_nth[OF pc\<^sub>g_in_range] nth_instrs_pc\<^sub>g + intro!: rel_stacktraces.intros elim!: Subx.sp_instrs_appendI) + apply (rule Subx.sp_instr.Call[of _ _ _ _ _ "map typeof (drop (arity fd2) \2\<^sub>g)"]) + apply (simp add: F2_f funtype_def) + apply (simp add: replicate_None map_append[symmetric]) + using \length \2 = return fd2'\ \list_all is_dyn_operand \2\ + by (auto simp: list.pred_set intro: replicate_eq_map[symmetric]) + qed (insert step_return rel_F1_F2, simp_all) + qed qed qed qed qed lemma match_final_backward: - "s1 \ s2 \ Subx.final s2 \ Sinca.final s1" -proof (induction s1 s2 rule: match.induct) - case (1 F1 F2 st1 st2 H) - obtain f fd2 pc \2 where - st2_def: "st2 = [Frame f pc \2]" and "Fubx_get F2 f = Some fd2" and "pc = length (body fd2)" - using \Subx.final (State F2 H st2)\ - by (auto elim!: Subx.final.cases) - obtain \1 where st1_def: "st1 = [Frame f pc \1]" - using \rel_stacktraces (Fubx_get F2) st1 st2 None\ - unfolding st2_def - by (auto elim!: rel_stacktraces.cases) - obtain fd1 where "Finca_get F1 f = Some fd1" and rel_fd1_fd2: "rel_fundef norm_eq fd1 fd2" - using \rel_fundefs (Finca_get F1) (Fubx_get F2)\ \Fubx_get F2 f = Some fd2\ - using rel_fundefs_Some2 - by fastforce - have "length (body fd1) = length (body fd2)" - using rel_fd1_fd2 by simp - thus ?case - unfolding st1_def - using \Finca_get F1 f = Some fd1\ \pc = length (body fd2)\ - by (auto intro: Sinca.final.intros) + assumes "match s1 s2" and final_s2: "final Fubx_get Ubx.IReturn s2" + shows "final Finca_get Inca.IReturn s1" + using \match s1 s2\ +proof (cases s1 s2 rule: match.cases) + case (matchI F2 H st2 F1 st1) + show ?thesis + using final_s2[unfolded matchI] + proof (cases _ _ "State F2 H st2" rule: final.cases) + case (finalI f l pc R \) + then show ?thesis + using matchI + by (auto intro!: final.intros elim!: rel_stacktraces.cases dest: rel_fundefs_next_instr2) + qed qed sublocale inca_to_ubx_simulation: - backward_simulation Sinca.step Subx.step Sinca.final Subx.final "\_ _. False" "\_. match" + backward_simulation Sinca.step Subx.step + "final Finca_get Inca.IReturn" + "final Fubx_get Ubx.IReturn" + "\_ _. False" "\_. match" using match_final_backward backward_lockstep_simulation lockstep_to_plus_backward_simulation[of match Subx.step Sinca.step] by unfold_locales auto section \Forward simulation\ -lemma traverse_cast_Dyn_eq_norm_stack: +lemma ap_map_list_cast_Dyn_eq_norm_stack: assumes "list_all (\x. x = None) (map typeof xs)" - shows "traverse cast_Dyn xs = Some (norm_stack xs)" + shows "ap_map_list cast_Dyn xs = Some (map Subx.norm_unboxed xs)" using assms proof (induction xs) case Nil - then show ?case by simp + thus ?case by simp next case (Cons x xs) from Cons.prems have typeof_x: "typeof x = None" and typeof_xs: "list_all (\x. x = None) (map typeof xs)" by simp_all obtain x' where "x = OpDyn x'" using typeof_unboxed_inversion(1)[OF typeof_x] by auto then show ?case using Cons.IH[OF typeof_xs] by simp qed lemma forward_lockstep_simulation: - assumes "Sinca.step s1 s1'" and "s1 \ s2" - shows "\s2'. Subx.step s2 s2' \ s1' \ s2'" - using assms(2,1) -proof (induction s1 s2 rule: match.induct) - case (1 F1 F2 st1 st2 H) - have rel_F1_F2: "rel_fundefs (Finca_get F1) (Fubx_get F2)" using 1 by simp - have sp_F2: "sp_fundefs (Fubx_get F2)" using 1 by simp - - note rel_fundefs_rewrite_both' = - rel_fundef_rewrite_both[THEN rel_fundefs_rewrite_both[OF rel_F1_F2]] - - from "1"(3,4) show ?case - proof (induction st1 st2 "None :: 'fun option" rule: rel_stacktraces.induct) - case rel_stacktraces_Nil - hence False by (auto elim: Sinca.step.cases) - thus ?case by simp - next - case (rel_stacktraces_Cons st1 st2 f \1 \2 fd2 pc) - have F2_f: "Fubx_get F2 f = Some fd2" - using rel_stacktraces_Cons by simp - have rel_st1_st2: "rel_stacktraces (Fubx_get F2) st1 st2 (Some f)" - using rel_stacktraces_Cons by simp - have sp_fundef_prefix: "sp_fundef (Fubx_get F2) fd2 (take pc (body fd2)) = Ok (map typeof \2)" - using rel_stacktraces_Cons by simp - have \1_def: "\1 = norm_stack \2" - using rel_stacktraces_Cons by simp - - note sp_fundef_def[simp] - note sp_prefix = sp_fundef_prefix[unfolded sp_fundef_def] + assumes "match s1 s2" and "Sinca.step s1 s1'" + shows "\s2'. Subx.step s2 s2' \ match s1' s2'" + using assms(1) +proof (cases s1 s2 rule: match.cases) + case (matchI F2 H st2 F1 st1) + have s2_def: "s2 = Global.state.State F2 H st2" using matchI by simp + have rel_F1_F2: "rel_fundefs (Finca_get F1) (Fubx_get F2)" using matchI by simp + have wf_s2: "Subx.wf_state s2" using matchI by simp + hence wf_F2: "Subx.wf_fundefs (Fubx_get F2)" by (auto simp: s2_def dest: Subx.wf_stateD) - note ex_F1_f = rel_fundefs_Some2[OF rel_F1_F2 F2_f] - note sp_fundef_full = sp_fundefs_get[OF sp_F2 \Fubx_get F2 f = Some fd2\] - note sp_full = sp_fundef_full[unfolded sp_fundef_def] - hence sp_sufix: "Subx.sp (Fubx_get F2) - (body fd2 ! pc # drop (Suc pc) (body fd2)) (map typeof \2) = Ok [None]" - if pc_in_range: "pc < length (body fd2)" - unfolding Cons_nth_drop_Suc[OF pc_in_range] - unfolding Subx.sp_eq_bind_take_drop[of _ "body fd2" _ pc] - unfolding sp_prefix - by simp + note wf_s2'I = Subx.wf_state_step_preservation[OF wf_s2] - from rel_stacktraces_Cons.prems(1) show ?case - proof (induction "State F1 H (Frame f pc \1 # st1)" s1' rule: Sinca.step.induct) - case (step_push fd1 d) - hence rel_fd1_fd2: "rel_fundef norm_eq fd1 fd2" - using ex_F1_f by simp - hence norm_instr_nth_body: "norm_instr (body fd2 ! pc) = body fd1 ! pc" - using step_push rel_fundef_body_nth by metis - have pc_in_range: "pc < length (body fd2)" - using step_push rel_fundef_body_length[OF rel_fd1_fd2] by simp - - show ?case (is "\x. ?STEP x \ ?MATCH x") - using step_push norm_instr_nth_body - proof (cases "body fd2 ! pc") + from \rel_stacktraces (Fubx_get F2) None st1 st2\ show ?thesis + proof (cases "Fubx_get F2" "None :: 'fun option" st1 st2 rule: rel_stacktraces.cases) + case rel_stacktraces_Nil + with matchI assms(2) show ?thesis by (auto elim: Sinca.step.cases) + next + case (rel_stacktraces_Cons f st1' st2' \1 \2 R1 R2 fd2 l instrs pc) + have rel_st1'_st2': "rel_stacktraces (Fubx_get F2) (Some f) st1' st2'" + using rel_stacktraces_Cons by simp + have st2_def: "st2 = Frame f l pc R2 \2 # st2'" using rel_stacktraces_Cons by simp + have \1_def: "\1 = map Subx.norm_unboxed \2" using rel_stacktraces_Cons by simp + have all_dyn_R2: "list_all is_dyn_operand R2" using rel_stacktraces_Cons by simp + have F2_f: "Fubx_get F2 f = Some fd2" using rel_stacktraces_Cons by simp + have map_of_fd2_l: "map_of (body fd2) l = Some instrs" using rel_stacktraces_Cons by simp + have sp_instrs_prefix: "Subx.sp_instrs (map_option funtype \ Fubx_get F2) (return fd2) + (take pc instrs) [] (map typeof \2)" + using rel_stacktraces_Cons by simp + note sp_instrs_prefix' = + Subx.sp_instrs_singletonI[THEN Subx.sp_instrs_appendI[OF sp_instrs_prefix]] + + have wf_fd2: "Subx.wf_fundef (map_option funtype \ Fubx_get F2) fd2" + using wf_F2 F2_f by (auto dest: Subx.wf_fundefs_getD) + hence sp_instrs_instrs: + "Subx.sp_instrs (map_option funtype \ Fubx_get F2) (return fd2) instrs [] []" + using wf_fd2[THEN Subx.wf_fundef_all_wf_basic_blockD] map_of_fd2_l + by (auto dest: list_all_map_of_SomeD[OF _ map_of_fd2_l] dest: Subx.wf_basic_blockD) + hence sp_instrs_sufix: "Subx.sp_instrs (map_option funtype \ Fubx_get F2) (return fd2) + (instrs ! pc # drop (Suc pc) instrs) (map typeof \2) []" if "pc < length instrs" + using that Subx.sp_instrs_appendD'[OF _ sp_instrs_prefix] + by (simp add: Cons_nth_drop_Suc) + + have + instrs_neq_Nil: "instrs \ []" and + all_jumps_in_range: "list_all (Subx.jump_in_range (fst ` set (body fd2))) instrs" and + sp_instrs_instrs: "Subx.sp_instrs (map_option funtype \ Fubx_get F2) (return fd2) instrs [] []" + using list_all_map_of_SomeD[OF wf_fd2[THEN Subx.wf_fundef_all_wf_basic_blockD] map_of_fd2_l] + by (auto dest: Subx.wf_basic_blockD) + + from assms(2)[unfolded matchI rel_stacktraces_Cons] show ?thesis + proof (induction "State F1 H (Frame f l pc (map Subx.norm_unboxed R2) (map Subx.norm_unboxed \2) # st1')" s1' rule: Sinca.step.induct) + case (step_push d) + then obtain instr2 where + next_instr2: "next_instr (Fubx_get F2) f l pc = Some instr2" and + norm_eq_instr1_instr2: "norm_eq (Inca.IPush d) instr2" + by (auto dest: rel_fundefs_next_instr1[OF rel_F1_F2]) + then show ?case (is "\x. ?STEP x \ ?MATCH (State F1 H ?st1') x") + proof (cases instr2) case (IPush d2) - then have d_def: "d = d2" - using norm_instr_nth_body step_push.hyps(3) by auto - let ?s2' = "State F2 H (Frame f (Suc pc) (OpDyn d2 # \2) # st2)" - have "?STEP ?s2'" - using IPush step_push d_def F2_f pc_in_range - by (auto intro: Subx.step_push) - moreover have "?MATCH ?s2'" - using rel_F1_F2 sp_F2 rel_st1_st2 F2_f \1_def - using IPush d_def sp_prefix take_Suc_conv_app_nth[OF pc_in_range] - by (auto intro!: match.intros rel_stacktraces.intros) - ultimately show ?thesis by auto - next - case (IPushUbx1 n) - then have d_def: "d = box_ubx1 n" - using norm_instr_nth_body step_push.hyps(3) by auto - let ?s2' = "State F2 H (Frame f (Suc pc) (OpUbx1 n # \2) # st2)" - have "?STEP ?s2'" - using IPushUbx1 step_push d_def F2_f pc_in_range - by (auto intro: Subx.step_push_ubx1) - moreover have "?MATCH ?s2'" - using rel_F1_F2 sp_F2 rel_st1_st2 F2_f \1_def - using IPushUbx1 d_def sp_prefix take_Suc_conv_app_nth[OF pc_in_range] - by (auto intro!: match.intros rel_stacktraces.intros) - ultimately show ?thesis by auto + let ?st2' = "Frame f l (Suc pc) R2 (OpDyn d # \2) # st2'" + let ?s2' = "State F2 H ?st2'" + show ?thesis + proof (intro exI conjI) + show "?STEP ?s2'" + using next_instr2 norm_eq_instr1_instr2 + unfolding IPush + by (auto simp: s2_def st2_def intro: Subx.step_push) + next + show "?MATCH (State F1 H ?st1') ?s2'" + proof (rule match.intros) + show "Subx.wf_state ?s2'" + using wf_F2 by (auto intro: Subx.wf_stateI) + next + show "rel_stacktraces (Fubx_get F2) None ?st1' ?st2'" + using next_instr2 rel_stacktraces_Cons + unfolding IPush + by (auto simp: next_instr_take_Suc_conv + intro!: rel_stacktraces.intros sp_instrs_prefix' intro: Subx.sp_instr.Push) + qed (simp_all add: rel_F1_F2) + qed next - case (IPushUbx2 b) - then have d_def: "d = box_ubx2 b" - using norm_instr_nth_body step_push.hyps(3) by auto - let ?s2' = "State F2 H (Frame f (Suc pc) (OpUbx2 b # \2) # st2)" - have "?STEP ?s2'" - using IPushUbx2 step_push d_def F2_f pc_in_range - by (auto intro: Subx.step_push_ubx2) - moreover have "?MATCH ?s2'" - using rel_F1_F2 sp_F2 rel_st1_st2 F2_f \1_def - using IPushUbx2 d_def sp_prefix take_Suc_conv_app_nth[OF pc_in_range] - by (auto intro!: match.intros rel_stacktraces.intros) - ultimately show ?thesis by auto - qed simp_all - next - case (step_pop fd1 x \1') - hence rel_fd1_fd2: "rel_fundef norm_eq fd1 fd2" - using ex_F1_f by simp - hence norm_instr_nth_body: "norm_instr (body fd2 ! pc) = body fd1 ! pc" - using step_pop rel_fundef_body_nth by metis - have pc_in_range: "pc < length (body fd2)" - using step_pop rel_fundef_body_length[OF rel_fd1_fd2] by simp - obtain x' \2' where \2_def: "\2 = x' # \2'" and "\1' = norm_stack \2'" - using step_pop \\1 = norm_stack \2\ norm_stack_def by auto - - show ?case (is "\x. ?STEP x \ ?MATCH x") - using step_pop norm_instr_nth_body - proof (cases "body fd2 ! pc") - case IPop - let ?s2' = "State F2 H (Frame f (Suc pc) \2' # st2)" - have "?STEP ?s2'" - using IPop step_pop F2_f pc_in_range \2_def - by (auto intro: Subx.step_pop) - moreover have "?MATCH ?s2'" - using rel_F1_F2 sp_F2 rel_st1_st2 F2_f \\1' = norm_stack \2'\ - using IPop sp_prefix take_Suc_conv_app_nth[OF pc_in_range] \2_def - by (auto intro!: match.intros rel_stacktraces.intros) - ultimately show ?thesis by auto - qed simp_all - next - case (step_load fd1 var i d \1') - hence rel_fd1_fd2: "rel_fundef norm_eq fd1 fd2" - using ex_F1_f by simp - hence norm_instr_nth_body: "norm_instr (body fd2 ! pc) = body fd1 ! pc" - using step_load rel_fundef_body_nth by metis - have pc_in_range: "pc < length (body fd2)" - using step_load rel_fundef_body_length[OF rel_fd1_fd2] by simp - obtain i' \2' where - \2_def: "\2 = i' # \2'" and norm_i': "Subx.norm_unboxed i' = i" and "norm_stack \2' = \1'" - using step_load \1_def norm_stack_def by auto - - show ?case (is "\x. ?STEP x \ ?MATCH x") - using step_load norm_instr_nth_body - proof (cases "body fd2 ! pc") - case (ILoad var') - then have [simp]: "var' = var" - using norm_instr_nth_body step_load.hyps(3) by auto - let ?s2' = "State F2 H (Frame f (Suc pc) (OpDyn d # \2') # st2)" - have "i' = OpDyn i" - using sp_sufix[OF pc_in_range, unfolded \2_def ILoad, simplified] - using norm_i' - by (cases i'; simp) - hence "?STEP ?s2'" - using ILoad step_load F2_f pc_in_range - by (auto intro!: Subx.step_load simp: \2_def) - moreover have "?MATCH ?s2'" - using rel_F1_F2 sp_F2 rel_st1_st2 F2_f \1_def - using ILoad sp_prefix take_Suc_conv_app_nth[OF pc_in_range] - by (auto intro!: match.intros rel_stacktraces.intros - simp: \norm_stack \2' = \1'\ \i' = OpDyn i\ \2_def) - ultimately show ?thesis by auto + case (IPushUbx1 u) + let ?st2' = "Frame f l (Suc pc) R2 (OpUbx1 u # \2) # st2'" + let ?s2' = "State F2 H ?st2'" + show ?thesis + proof (intro exI conjI) + show "?STEP ?s2'" + using next_instr2 norm_eq_instr1_instr2 + unfolding IPushUbx1 + by (auto simp: s2_def st2_def intro: Subx.step_push_ubx1) + next + show "?MATCH (State F1 H ?st1') ?s2'" + proof (rule match.intros) + show "Subx.wf_state ?s2'" + using wf_F2 by (auto intro: Subx.wf_stateI) + show "rel_stacktraces (Fubx_get F2) None ?st1' ?st2'" + using next_instr2 norm_eq_instr1_instr2 rel_stacktraces_Cons + unfolding IPushUbx1 + by (auto simp: next_instr_take_Suc_conv + intro!: rel_stacktraces.intros sp_instrs_prefix' intro: Subx.sp_instr.PushUbx1) + qed (simp_all add: rel_F1_F2) + qed next - case (ILoadUbx \ var') - then have [simp]: "var' = var" - using norm_instr_nth_body step_load.hyps(3) by auto - have [simp]: "i' = OpDyn i" - using sp_sufix[OF pc_in_range, unfolded \2_def ILoadUbx, simplified] - using norm_i' - by (cases i'; simp) + case (IPushUbx2 u) + let ?st2' = "Frame f l (Suc pc) R2 (OpUbx2 u # \2) # st2'" + let ?s2' = "State F2 H ?st2'" show ?thesis - proof (cases "Subx.unbox \ d") - case None - let ?F2' = "Fubx_add F2 f (Subx.generalize_fundef fd2)" - let ?frame = "Frame f (Suc pc) (OpDyn d # \2')" - let ?s2' = "State ?F2' H (Subx.box_stack f (?frame # st2))" - have "?STEP ?s2'" - using None ILoadUbx - using step_load pc_in_range F2_f - by (auto intro!: Subx.step_load_ubx_miss - simp add: \2_def - simp del: Subx.box_stack_Cons) - moreover have "?MATCH ?s2'" - using rel_fundefs_generalize[OF rel_F1_F2 F2_f] - using sp_fundefs_generalize[OF sp_F2 F2_f] - using sp_fundef_prefix ILoadUbx - using pc_in_range - by (auto intro!: match.intros rel_stacktraces_generalize[OF rel_st1_st2 F2_f] - simp: \2_def \norm_stack \2' = \1'\) - ultimately show ?thesis by auto + proof (intro exI conjI) + show "?STEP ?s2'" + using next_instr2 norm_eq_instr1_instr2 + unfolding IPushUbx2 + by (auto simp: s2_def st2_def intro: Subx.step_push_ubx2) next - case (Some blob) - let ?s2' = "State F2 H (Frame f (Suc pc) (blob # \2') # st2)" - have "?STEP ?s2'" - using Some ILoadUbx step_load F2_f pc_in_range - by (auto intro: Subx.step_load_ubx_hit simp: \2_def) - moreover have "?MATCH ?s2'" - using rel_F1_F2 sp_F2 rel_st1_st2 F2_f \1_def - using Some ILoadUbx sp_prefix take_Suc_conv_app_nth[OF pc_in_range] - by (auto intro!: match.intros rel_stacktraces.intros - simp: \norm_stack \2' = \1'\ \2_def) - ultimately show ?thesis by auto + show "?MATCH (State F1 H ?st1') ?s2'" + proof (rule match.intros) + show "Subx.wf_state ?s2'" + using wf_F2 by (auto intro: Subx.wf_stateI) + show "rel_stacktraces (Fubx_get F2) None ?st1' ?st2'" + using next_instr2 norm_eq_instr1_instr2 rel_stacktraces_Cons + unfolding IPushUbx2 + by (auto simp: next_instr_take_Suc_conv + intro!: rel_stacktraces.intros sp_instrs_prefix' intro: Subx.sp_instr.PushUbx2) + qed (simp_all add: rel_F1_F2) qed qed simp_all next - case (step_store fd1 var i x H' \1') - hence rel_fd1_fd2: "rel_fundef norm_eq fd1 fd2" - using ex_F1_f by simp - hence norm_instr_nth_body: "norm_instr (body fd2 ! pc) = body fd1 ! pc" - using step_store rel_fundef_body_nth by metis - have pc_in_range: "pc < length (body fd2)" - using step_store rel_fundef_body_length[OF rel_fd1_fd2] by simp - obtain i' x' \2' where - \2_def: "\2 = i' # x' # \2'" and - norm_i': "Subx.norm_unboxed i' = i" and - norm_x': "Subx.norm_unboxed x' = x" and - "norm_stack \2' = \1'" - using step_store \1_def norm_stack_def by auto - - show ?case (is "\x. ?STEP x \ ?MATCH x") - using step_store norm_instr_nth_body - proof (cases "body fd2 ! pc") - case (IStore var') - then have [simp]: "var' = var" - using norm_instr_nth_body step_store.hyps(3) by auto - note sp_sufix' = sp_sufix[OF pc_in_range, unfolded \2_def IStore, simplified] - have [simp]: "i' = OpDyn i" - using norm_i' sp_sufix' by (cases i'; simp) - have [simp]: "x' = OpDyn x" - using norm_x' sp_sufix' by (cases x'; simp) - let ?s2' = "State F2 (heap_add H (var, i) x) (Frame f (Suc pc) \2' # st2)" - have "?STEP ?s2'" - using IStore \2_def pc_in_range F2_f - by (auto intro: Subx.step_store) - moreover have "?MATCH ?s2'" - using rel_F1_F2 sp_F2 rel_st1_st2 F2_f \2_def - using \norm_stack \2' = \1'\ \heap_add H (var, i) x = H'\ - using IStore sp_prefix take_Suc_conv_app_nth[OF pc_in_range] - by (auto intro!: match.intros rel_stacktraces.intros) - ultimately show ?thesis by auto - next - case (IStoreUbx \ var') - then have [simp]: "var' = var" - using norm_instr_nth_body step_store.hyps(3) by auto - note sp_sufix' = sp_sufix[OF pc_in_range, unfolded \2_def IStoreUbx, simplified] - have [simp]: "i' = OpDyn i" - using norm_i' sp_sufix' by (cases i'; simp) - have typeof_x'[simp]: "typeof x' = Some \" - using sp_sufix' apply (auto simp add: Result.bind_eq_Ok_conv elim!: Subx.sp_instr.elims) - by (metis result.simps(4)) - let ?s2' = "State F2 (heap_add H (var, i) x) (Frame f (Suc pc) \2' # st2)" - have "?STEP ?s2'" - unfolding \2_def - using F2_f pc_in_range IStoreUbx - using Subx.typeof_and_norm_unboxed_imp_cast_and_box[OF typeof_x' norm_x'] - by (auto intro!: Subx.step_store_ubx) - moreover have "?MATCH ?s2'" - using rel_F1_F2 sp_F2 rel_st1_st2 F2_f \2_def - using \norm_stack \2' = \1'\ \heap_add H (var, i) x = H'\ - using IStoreUbx sp_prefix take_Suc_conv_app_nth[OF pc_in_range] - by (auto intro!: match.intros rel_stacktraces.intros) - ultimately show ?thesis by auto - qed simp_all - next - case (step_op fd1 op ar x) - hence rel_fd1_fd2: "rel_fundef norm_eq fd1 fd2" - using ex_F1_f by simp - hence norm_instr_nth_body: "norm_instr (body fd2 ! pc) = body fd1 ! pc" - using step_op rel_fundef_body_nth by metis - have pc_in_range: "pc < length (body fd2)" - using step_op rel_fundef_body_length[OF rel_fd1_fd2] by simp - - show ?case (is "\x. ?STEP x \ ?MATCH x") - using step_op norm_instr_nth_body - proof (cases "body fd2 ! pc") - case (IOp op') - then have "op' = op" - using norm_instr_nth_body step_op.hyps(3) by auto - hence "ar \ length \2" and - all_arg_Dyn: "list_all (\x. x = None) (map typeof (take ar \2))" - using IOp step_op sp_sufix[OF pc_in_range] - by (auto simp: take_map) - let ?s2' = "State F2 H (Frame f (Suc pc) (OpDyn x # drop ar \2) # st2)" - have "?STEP ?s2'" - using IOp step_op \op' = op\ \1_def pc_in_range F2_f - using traverse_cast_Dyn_eq_norm_stack[OF all_arg_Dyn] - by (auto intro!: Subx.step_op simp: take_norm_stack) - moreover have "?MATCH ?s2'" - using rel_F1_F2 sp_F2 rel_st1_st2 F2_f - using IOp sp_prefix take_Suc_conv_app_nth[OF pc_in_range] - using arg_cong[OF \1_def, of "drop ar"] \ar \ length \2\ - using \op' = op\ all_arg_Dyn \1_def \\\\\\ op = ar\ - by (auto intro!: match.intros rel_stacktraces.intros - dest: list_all_eq_const_imp_replicate - simp: drop_norm_stack[symmetric] Let_def take_map drop_map) - ultimately show ?thesis by auto + case (step_pop d \1') + then obtain u \2' where + \2_def: "\2 = u # \2'" and "d = Subx.norm_unboxed u" and + \1'_def: "\1' = map Subx.norm_unboxed \2'" + by auto + from step_pop obtain instr2 where + next_instr2: "next_instr (Fubx_get F2) f l pc = Some instr2" and + norm_eq_instr1_instr2: "norm_eq Inca.IPop instr2" + by (auto dest: rel_fundefs_next_instr1[OF rel_F1_F2]) + then show ?case (is "\x. ?STEP x \ ?MATCH (State F1 H ?st1') x") + proof (cases instr2) + case IPop + let ?st2' = "Frame f l (Suc pc) R2 \2' # st2'" + let ?s2' = "State F2 H ?st2'" + show ?thesis + proof (intro exI conjI) + show "?STEP ?s2'" + using next_instr2 norm_eq_instr1_instr2 + unfolding IPop + by (auto simp: s2_def st2_def \2_def intro: Subx.step_pop) + next + show "?MATCH (State F1 H ?st1') ?s2'" + proof (rule match.intros) + show "Subx.wf_state ?s2'" + using wf_F2 by (auto intro: Subx.wf_stateI) + next + show "rel_stacktraces (Fubx_get F2) None ?st1' ?st2'" + using next_instr2 rel_stacktraces_Cons + unfolding IPop + by (auto simp: \2_def \1'_def next_instr_take_Suc_conv + intro!: rel_stacktraces.intros sp_instrs_prefix' Subx.sp_instr.Pop) + qed (simp_all add: rel_F1_F2) + qed qed simp_all next - case (step_op_inl fd1 op ar opinl x F1') - hence rel_fd1_fd2: "rel_fundef norm_eq fd1 fd2" - using ex_F1_f by simp - hence norm_instr_nth_body: "norm_instr (body fd2 ! pc) = body fd1 ! pc" - using step_op_inl rel_fundef_body_nth by metis - have pc_in_range: "pc < length (body fd2)" - using step_op_inl rel_fundef_body_length[OF rel_fd1_fd2] by simp - - show ?case (is "\x. ?STEP x \ ?MATCH x") - using step_op_inl norm_instr_nth_body - proof (cases "body fd2 ! pc") - case (IOp op') - then have "op' = op" - using norm_instr_nth_body step_op_inl.hyps(3) by simp - hence "ar \ length \2" and - all_arg_Dyn: "list_all (\x. x = None) (map typeof (take ar \2))" - using IOp step_op_inl sp_sufix[OF pc_in_range] - by (auto simp: take_map) - let ?fd2' = "rewrite_fundef_body fd2 pc (Ubx.IOpInl opinl)" - let ?F2' = "Fubx_add F2 f ?fd2'" - let ?frame = "Frame f (Suc pc) (OpDyn x # drop ar \2)" - let ?s2' = "State ?F2' H (?frame # st2)" - have "?STEP ?s2'" - using IOp step_op_inl \op' = op\ - using \1_def pc_in_range \Fubx_get F2 f = Some fd2\ - using all_arg_Dyn take_norm_stack traverse_cast_Dyn_eq_norm_stack - by (auto intro!: Subx.step_op_inl) - moreover have "?MATCH ?s2'" - proof - show "rel_fundefs (Finca_get F1') (Fubx_get ?F2')" - using step_op_inl rel_fd1_fd2 - by (auto intro: rel_fundefs_rewrite_both') + case (step_get n d) + hence nth_R2_n: "R2 ! n = OpDyn d" + using all_dyn_R2 + by (metis Subx.norm_unboxed.simps(1) is_dyn_operand_def length_map list_all_length nth_map) + from step_get obtain instr2 where + next_instr2: "next_instr (Fubx_get F2) f l pc = Some instr2" and + norm_eq_instr1_instr2: "norm_eq (Inca.IGet n) instr2" + by (auto dest: rel_fundefs_next_instr1[OF rel_F1_F2]) + then show ?case (is "\x. ?STEP x \ ?MATCH (State F1 H ?st1') x") + proof (cases instr2) + case (IGet n') + hence "n' = n" using norm_eq_instr1_instr2 by simp + let ?st2' = "Frame f l (Suc pc) R2 (OpDyn d # \2) # st2'" + let ?s2' = "State F2 H ?st2'" + show ?thesis + proof (intro exI conjI) + show "?STEP ?s2'" + using step_get nth_R2_n + using next_instr2 norm_eq_instr1_instr2 + unfolding IGet + by (auto simp: s2_def st2_def intro: Subx.step_get) next - have "arity fd2 = arity ?fd2'" by simp - hence "all_same_arities (Fubx_get F2) (Fubx_get ?F2')" - using all_same_arities_add[OF F2_f] by simp - thus "sp_fundefs (Fubx_get ?F2')" - apply (auto intro!: sp_fundefs_add[OF sp_F2]) - apply (rule Subx.sp_rewrite_eq_Ok[OF pc_in_range _ sp_full]) - apply (rule allI) - unfolding IOp - apply (rule Subx.sp_instr_op) - using Sinca.\\\_invertible \op' = op\ step_op_inl.hyps(6) by blast + show "?MATCH (State F1 H ?st1') ?s2'" + proof (rule match.intros) + show "Subx.wf_state ?s2'" + using wf_F2 by (auto intro: Subx.wf_stateI) + next + show "rel_stacktraces (Fubx_get F2) None ?st1' ?st2'" + using next_instr2 rel_stacktraces_Cons + unfolding IGet + by (auto simp: next_instr_take_Suc_conv + intro!: rel_stacktraces.intros sp_instrs_prefix' intro: Subx.sp_instr.Get) + qed (simp_all add: rel_F1_F2) + qed + next + case (IGetUbx \ n') + hence "n' = n" using norm_eq_instr1_instr2 by simp + show ?thesis + proof (cases "Subx.unbox \ d") + case None + let ?F2' = "Subx.Fenv.map_entry F2 f Subx.generalize_fundef" + let ?st2' = "Subx.box_stack f (Frame f l (Suc pc) R2 (OpDyn d # \2) # st2')" + let ?s2' = "State ?F2' H ?st2'" + show ?thesis + proof (intro exI conjI) + show "?STEP ?s2'" + using step_get nth_R2_n None + using next_instr2 norm_eq_instr1_instr2 + unfolding IGetUbx + by (auto simp: s2_def st2_def intro: Subx.step_get_ubx_miss[simplified]) + next + show "?MATCH (State F1 H ?st1') ?s2'" + proof (rule match.intros) + show "Subx.wf_state ?s2'" + using wf_F2 + by (auto intro!: Subx.wf_stateI intro: Subx.wf_fundefs_generalize) + next + show "rel_fundefs (Finca_get F1) (Fubx_get ?F2')" + using rel_F1_F2 + by (auto intro: rel_fundefs_generalizeI) + next + have sp_instrs_gen: "Subx.sp_instrs (map_option funtype \ Fubx_get F2) (return fd2) + (take (Suc pc) (map Subx.generalize_instr instrs)) [] (map Map.empty (OpDyn d # \2))" + using rel_stacktraces_Cons step_get.hyps + using IGetUbx \n' = n\ + using next_instr_get_map_ofD[OF next_instr2 F2_f map_of_fd2_l] + by (auto simp: take_Suc_conv_app_nth take_map + intro!: Subx.sp_instrs_appendI + intro: Subx.sp_instrs_generalize0 Subx.sp_instr.Get) + then show "rel_stacktraces (Fubx_get ?F2') None ?st1' ?st2'" + apply simp + proof (rule rel_stacktraces.intros) + show "rel_stacktraces (Fubx_get ?F2') (Some f) st1' (Subx.box_stack f st2')" + using rel_st1'_st2' rel_stacktraces_map_entry_gneralize_fundefI by simp + next + show "Fubx_get ?F2' f = Some (Subx.generalize_fundef fd2)" + using F2_f by simp + next + show "map_of (body (Subx.generalize_fundef fd2)) l = + Some (map Subx.generalize_instr instrs)" + using map_of_fd2_l by (simp add: Subx.map_of_generalize_fundef_conv) + qed (insert all_dyn_R2 map_of_fd2_l, simp_all add: Subx.map_of_generalize_fundef_conv) + qed + qed next - have "rel_stacktraces (Fubx_get F2) - (Frame f (Suc pc) (x # drop ar \1) # st1) (?frame # st2) None" - using rel_st1_st2 sp_prefix F2_f - using IOp take_Suc_conv_app_nth[OF pc_in_range] - using arg_cong[OF \1_def, of "drop ar"] \ar \ length \2\ - using \op' = op\ step_op_inl.hyps(4) all_arg_Dyn - by (auto intro!: rel_stacktraces.intros - dest: list_all_eq_const_imp_replicate - simp: drop_norm_stack[symmetric] Let_def take_map drop_map) - thus "rel_stacktraces (Fubx_get ?F2') - (Frame f (Suc pc) (x # drop ar \1) # st1) (?frame # st2) None" - using F2_f pc_in_range IOp - using Sinca.\\\_invertible \op' = op\ step_op_inl.hyps(6) - by (auto intro!: rel_stacktraces_rewrite_fundef simp: Subx.sp_instr_op Let_def) + case (Some u) + let ?st2' = "Frame f l (Suc pc) R2 (u # \2) # st2'" + let ?s2' = "State F2 H ?st2'" + show ?thesis + proof (intro exI conjI) + show "?STEP ?s2'" + using step_get nth_R2_n Some + using next_instr2 norm_eq_instr1_instr2 + unfolding IGetUbx + by (auto simp: s2_def st2_def intro: Subx.step_get_ubx_hit) + next + show "?MATCH (State F1 H ?st1') ?s2'" + proof (rule match.intros) + show "Subx.wf_state ?s2'" + using wf_F2 by (auto intro: Subx.wf_stateI) + next + show "rel_stacktraces (Fubx_get F2) None ?st1' ?st2'" + using next_instr2 rel_stacktraces_Cons + using Some + unfolding IGetUbx + by (auto simp: next_instr_take_Suc_conv + intro!: rel_stacktraces.intros sp_instrs_prefix' intro: Subx.sp_instr.GetUbx) + qed (simp_all add: rel_F1_F2) + qed qed - ultimately show ?thesis by auto qed simp_all next - case (step_op_inl_hit fd1 opinl ar x) - hence rel_fd1_fd2: "rel_fundef norm_eq fd1 fd2" - using ex_F1_f by simp - hence norm_instr_nth_body: "norm_instr (body fd2 ! pc) = body fd1 ! pc" - using step_op_inl_hit rel_fundef_body_nth by metis - have pc_in_range: "pc < length (body fd2)" - using step_op_inl_hit rel_fundef_body_length[OF rel_fd1_fd2] by simp - - show ?case (is "\x. ?STEP x \ ?MATCH x") - using step_op_inl_hit norm_instr_nth_body - proof (cases "body fd2 ! pc") + case (step_set n R1' d \1') + then obtain u \2' where + \2_def: "\2 = u # \2'" and d_def: "d = Subx.norm_unboxed u" and + \1'_def: "\1' = map Subx.norm_unboxed \2'" + by auto + from step_set obtain instr2 where + next_instr2: "next_instr (Fubx_get F2) f l pc = Some instr2" and + norm_eq_instr1_instr2: "norm_eq (Inca.ISet n) instr2" + by (auto dest: rel_fundefs_next_instr1[OF rel_F1_F2]) + have pc_in_range: "pc < length instrs" and nth_instrs_pc: "instrs ! pc = instr2" + using next_instr_get_map_ofD[OF next_instr2 F2_f map_of_fd2_l] + by simp_all + from next_instr2 norm_eq_instr1_instr2 + show ?case (is "\x. ?STEP x \ ?MATCH (State F1 H ?st1') x") + proof (cases instr2) + case (ISet n') + hence "n' = n" using norm_eq_instr1_instr2 by simp + have typeof_u: "typeof u = None" + using sp_instrs_sufix[OF pc_in_range, unfolded nth_instrs_pc ISet, simplified] + by (auto simp: \2_def elim: Subx.sp_instrs.cases Subx.sp_instr.cases) + hence cast_Dyn_u: "cast_Dyn u = Some d" + by (auto simp add: d_def dest: Subx.typeof_and_norm_unboxed_imp_cast_Dyn) + let ?R2' = "R2[n := OpDyn d]" + let ?st2' = "Frame f l (Suc pc) ?R2' \2' # st2'" + let ?s2' = "State F2 H ?st2'" + show ?thesis + proof (intro exI conjI) + show "?STEP ?s2'" + using step_set.hyps cast_Dyn_u + using next_instr2 norm_eq_instr1_instr2 + unfolding ISet + by (auto simp: s2_def st2_def \2_def intro: Subx.step_set) + next + show "?MATCH (State F1 H ?st1') ?s2'" + proof (rule match.intros) + show "Subx.wf_state ?s2'" + using wf_F2 by (auto intro: Subx.wf_stateI) + next + show "rel_stacktraces (Fubx_get F2) None ?st1' ?st2'" + using next_instr2 rel_stacktraces_Cons + unfolding ISet + using step_set.hyps cast_Dyn_u + by (auto simp: \1'_def \2_def map_update next_instr_take_Suc_conv + intro!: rel_stacktraces.intros sp_instrs_prefix' Subx.sp_instr.Set + intro: list_all_list_updateI) + qed (simp_all add: rel_F1_F2) + qed + next + case (ISetUbx \ n') + hence "n' = n" using norm_eq_instr1_instr2 by simp + have typeof_u: "typeof u = Some \" + using sp_instrs_sufix[OF pc_in_range, unfolded nth_instrs_pc ISetUbx, simplified] + by (auto simp: \2_def elim: Subx.sp_instrs.cases Subx.sp_instr.cases) + hence cast_and_box_u: "Subx.cast_and_box \ u = Some d" + by (auto simp add: d_def dest: Subx.typeof_and_norm_unboxed_imp_cast_and_box) + let ?R2' = "R2[n := OpDyn d]" + let ?st2' = "Frame f l (Suc pc) ?R2' \2' # st2'" + let ?s2' = "State F2 H ?st2'" + show ?thesis + proof (intro exI conjI) + show "?STEP ?s2'" + using step_set cast_and_box_u + using next_instr2 norm_eq_instr1_instr2 + unfolding ISetUbx + by (auto simp: s2_def st2_def \2_def intro: Subx.step_set_ubx) + next + show "?MATCH (State F1 H ?st1') ?s2'" + proof (rule match.intros) + show "Subx.wf_state ?s2'" + using wf_F2 by (auto intro: Subx.wf_stateI) + next + show "rel_stacktraces (Fubx_get F2) None ?st1' ?st2'" + using next_instr2 rel_stacktraces_Cons + unfolding ISetUbx + using step_set.hyps cast_and_box_u + by (auto simp: \1'_def \2_def map_update next_instr_take_Suc_conv + intro!: rel_stacktraces.intros sp_instrs_prefix' Subx.sp_instr.SetUbx + intro: list_all_list_updateI) + qed (simp_all add: rel_F1_F2) + qed + qed simp_all + next + case (step_load x y d \1') + then obtain u \2' where + \2_def: "\2 = u # \2'" and d_def: "y = Subx.norm_unboxed u" and + \1'_def: "\1' = map Subx.norm_unboxed \2'" + by auto + from step_load obtain instr2 where + next_instr2: "next_instr (Fubx_get F2) f l pc = Some instr2" and + norm_eq_instr1_instr2: "norm_eq (Inca.ILoad x) instr2" + by (auto dest: rel_fundefs_next_instr1[OF rel_F1_F2]) + have pc_in_range: "pc < length instrs" and nth_instrs_pc: "instrs ! pc = instr2" + using next_instr_get_map_ofD[OF next_instr2 F2_f map_of_fd2_l] + by simp_all + from next_instr2 norm_eq_instr1_instr2 + show ?case (is "\x. ?STEP x \ ?MATCH (State F1 H ?st1') x") + proof (cases instr2) + case (ILoad x') + hence "x' = x" using norm_eq_instr1_instr2 by simp + have typeof_u: "typeof u = None" + using sp_instrs_sufix[OF pc_in_range, unfolded nth_instrs_pc ILoad, simplified] + by (auto simp: \2_def elim: Subx.sp_instrs.cases Subx.sp_instr.cases) + hence cast_Dyn_u: "cast_Dyn u = Some y" + by (auto simp add: d_def dest: Subx.typeof_and_norm_unboxed_imp_cast_Dyn) + let ?st2' = "Frame f l (Suc pc) R2 (OpDyn d # \2') # st2'" + let ?s2' = "State F2 H ?st2'" + show ?thesis + proof (intro exI conjI) + show "?STEP ?s2'" + using step_load.hyps cast_Dyn_u next_instr2 + unfolding ILoad \x' = x\ + by (auto simp: s2_def st2_def \2_def intro: Subx.step_load) + next + show "?MATCH (State F1 H ?st1') ?s2'" + proof (rule match.intros) + show "Subx.wf_state ?s2'" + using wf_F2 by (auto intro: Subx.wf_stateI) + next + show "rel_stacktraces (Fubx_get F2) None ?st1' ?st2'" + using next_instr2 rel_stacktraces_Cons + unfolding ILoad \x' = x\ + using step_load.hyps cast_Dyn_u + by (auto simp: \1'_def \2_def map_update next_instr_take_Suc_conv + intro!: rel_stacktraces.intros sp_instrs_prefix' Subx.sp_instr.Load + intro: list_all_list_updateI) + qed (simp_all add: rel_F1_F2) + qed + next + case (ILoadUbx \ x') + hence "x' = x" using norm_eq_instr1_instr2 by simp + have typeof_u: "typeof u = None" + using sp_instrs_sufix[OF pc_in_range, unfolded nth_instrs_pc ILoadUbx, simplified] + by (auto simp: \2_def elim: Subx.sp_instrs.cases Subx.sp_instr.cases) + hence cast_Dyn_u: "cast_Dyn u = Some y" + by (auto simp add: d_def dest: Subx.typeof_and_norm_unboxed_imp_cast_Dyn) + show ?thesis + proof (cases "Subx.unbox \ d") + case None + let ?F2' = "Subx.Fenv.map_entry F2 f Subx.generalize_fundef" + let ?st2' = "Subx.box_stack f (Frame f l (Suc pc) R2 (OpDyn d # \2') # st2')" + let ?s2' = "State ?F2' H ?st2'" + show ?thesis + proof (intro exI conjI) + show "?STEP ?s2'" + using step_load.hyps next_instr2 cast_Dyn_u None + unfolding ILoadUbx \x' = x\ + by (auto simp add: s2_def st2_def \2_def intro: Subx.step_load_ubx_miss[simplified]) + next + show "?MATCH (State F1 H ?st1') ?s2'" + proof (rule match.intros) + show "Subx.wf_state ?s2'" + using wf_F2 + by (auto intro!: Subx.wf_stateI intro: Subx.wf_fundefs_generalize) + next + show "rel_fundefs (Finca_get F1) (Fubx_get ?F2')" + using rel_F1_F2 + by (auto intro: rel_fundefs_generalizeI) + next + have "Subx.sp_instrs (map_option funtype \ Fubx_get F2) (return fd2) + (take (Suc pc) (map Subx.generalize_instr instrs)) [] (None # map Map.empty \2')" + using rel_stacktraces_Cons step_load.hyps + using pc_in_range nth_instrs_pc + using ILoadUbx \x' = x\ + by (auto simp: \2_def take_Suc_conv_app_nth take_map + intro!: Subx.sp_instrs_appendI + intro: Subx.sp_instrs_generalize0 Subx.sp_instr.Load) + thus "rel_stacktraces (Fubx_get ?F2') None ?st1' ?st2'" + apply simp + proof (rule rel_stacktraces.intros) + show "Fubx_get ?F2' f = Some (Subx.generalize_fundef fd2)" + using F2_f by simp + next + show "map_of (body (Subx.generalize_fundef fd2)) l = + Some (map Subx.generalize_instr instrs)" + using map_of_fd2_l + by (simp add: Subx.map_of_generalize_fundef_conv) + qed (insert rel_F1_F2 all_dyn_R2 F2_f map_of_fd2_l rel_st1'_st2', auto simp: \1'_def) + qed + qed + next + case (Some u2) + let ?st2' = "Frame f l (Suc pc) R2 (u2 # \2') # st2'" + let ?s2' = "State F2 H ?st2'" + show ?thesis + proof (intro exI conjI) + show "?STEP ?s2'" + using step_load.hyps next_instr2 cast_Dyn_u Some + unfolding ILoadUbx \x' = x\ + by (auto simp add: s2_def st2_def \2_def intro: Subx.step_load_ubx_hit[simplified]) + next + show "?MATCH (State F1 H ?st1') ?s2'" + proof (rule match.intros) + show "Subx.wf_state ?s2'" + using wf_F2 + by (auto intro!: Subx.wf_stateI intro: Subx.wf_fundefs_generalize) + next + show "rel_stacktraces (Fubx_get F2) None ?st1' ?st2'" + using next_instr2 rel_stacktraces_Cons + using Some typeof_u + unfolding ILoadUbx + by (auto simp: \2_def \1'_def next_instr_take_Suc_conv + intro!: rel_stacktraces.intros sp_instrs_prefix' Subx.sp_instr.LoadUbx) + qed (insert rel_F1_F2, simp_all) + qed + qed + qed simp_all + next + case (step_store x d1 d2 H' \1') + then obtain u1 u2 \2' where + \2_def: "\2 = u1 # u2 # \2'" and + d1_def: "d1 = Subx.norm_unboxed u1" and + d2_def: "d2 = Subx.norm_unboxed u2" and + \1'_def: "\1' = map Subx.norm_unboxed \2'" + by auto + from step_store obtain instr2 where + next_instr2: "next_instr (Fubx_get F2) f l pc = Some instr2" and + norm_eq_instr1_instr2: "norm_eq (Inca.instr.IStore x) instr2" + by (auto dest: rel_fundefs_next_instr1[OF rel_F1_F2]) + have pc_in_range: "pc < length instrs" and nth_instrs_pc: "instrs ! pc = instr2" + using next_instr_get_map_ofD[OF next_instr2 F2_f map_of_fd2_l] + by simp_all + from next_instr2 norm_eq_instr1_instr2 + show ?case (is "\x. ?STEP x \ ?MATCH (State F1 H' ?st1') x") + proof (cases instr2) + case (IStore x') + hence "x' = x" using norm_eq_instr1_instr2 by simp + have casts: "cast_Dyn u1 = Some d1" "cast_Dyn u2 = Some d2" + unfolding atomize_conj + using sp_instrs_sufix[OF pc_in_range, unfolded nth_instrs_pc IStore, simplified] + by (auto simp: d1_def d2_def \2_def elim: Subx.sp_instrs.cases Subx.sp_instr.cases + intro: Subx.typeof_and_norm_unboxed_imp_cast_Dyn) + let ?st2' = "Frame f l (Suc pc) R2 \2' # st2'" + let ?s2' = "State F2 H' ?st2'" + show ?thesis + proof (intro exI conjI) + show "?STEP ?s2'" + using step_store.hyps casts next_instr2 + unfolding IStore \x' = x\ + by (auto simp: s2_def st2_def \2_def intro: Subx.step_store) + next + show "?MATCH (State F1 H' ?st1') ?s2'" + proof (rule match.intros) + show "Subx.wf_state ?s2'" + using wf_F2 by (auto intro: Subx.wf_stateI) + next + show "rel_stacktraces (Fubx_get F2) None ?st1' ?st2'" + using next_instr2 rel_stacktraces_Cons + unfolding IStore \x' = x\ + using step_store.hyps casts + by (auto simp: \1'_def \2_def map_update next_instr_take_Suc_conv + intro!: rel_stacktraces.intros sp_instrs_prefix' Subx.sp_instr.Store + intro: list_all_list_updateI) + qed (insert rel_F1_F2, simp_all) + qed + next + case (IStoreUbx \ x') + hence "x' = x" using norm_eq_instr1_instr2 by simp + have casts: "cast_Dyn u1 = Some d1" "Subx.cast_and_box \ u2 = Some d2" + unfolding atomize_conj + using sp_instrs_sufix[OF pc_in_range, unfolded nth_instrs_pc IStoreUbx, simplified] + by (auto simp: d1_def d2_def \2_def elim: Subx.sp_instrs.cases Subx.sp_instr.cases + intro: Subx.typeof_and_norm_unboxed_imp_cast_Dyn + intro: Subx.typeof_and_norm_unboxed_imp_cast_and_box) + let ?st2' = "Frame f l (Suc pc) R2 \2' # st2'" + let ?s2' = "State F2 H' ?st2'" + show ?thesis + proof (intro exI conjI) + show "?STEP ?s2'" + using step_store.hyps casts next_instr2 + unfolding IStoreUbx \x' = x\ + by (auto simp: s2_def st2_def \2_def intro: Subx.step_store_ubx) + next + show "?MATCH (State F1 H' ?st1') ?s2'" + proof (rule match.intros) + show "Subx.wf_state ?s2'" + using wf_F2 by (auto intro: Subx.wf_stateI) + next + show "rel_stacktraces (Fubx_get F2) None ?st1' ?st2'" + using next_instr2 rel_stacktraces_Cons + unfolding IStoreUbx \x' = x\ + using step_store.hyps casts + by (auto simp: \1'_def \2_def map_update next_instr_take_Suc_conv + intro!: rel_stacktraces.intros sp_instrs_prefix' Subx.sp_instr.StoreUbx + intro: list_all_list_updateI) + qed (insert rel_F1_F2, simp_all) + qed + qed simp_all + next + case (step_op op ar x) + then obtain instr2 where + next_instr2: "next_instr (Fubx_get F2) f l pc = Some instr2" and + norm_eq_instr1_instr2: "norm_eq (Inca.IOp op) instr2" + by (auto dest: rel_fundefs_next_instr1[OF rel_F1_F2]) + have pc_in_range: "pc < length instrs" and nth_instrs_pc: "instrs ! pc = instr2" + using next_instr_get_map_ofD[OF next_instr2 F2_f map_of_fd2_l] + by simp_all + from next_instr2 norm_eq_instr1_instr2 + show ?case (is "\x. ?STEP x \ ?MATCH (State F1 H ?st1') x") + proof (cases instr2) + case (IOp op') + hence "op' = op" using norm_eq_instr1_instr2 by simp + have casts: + "ap_map_list cast_Dyn (take ar \2) = Some (take ar (map Subx.norm_unboxed \2))" + using sp_instrs_sufix[OF pc_in_range, unfolded nth_instrs_pc IOp, simplified] + using step_op.hyps + by (auto simp: \op' = op\ take_map + elim!: Subx.sp_instrs.cases[of _ _ "x # xs" for x xs, simplified] Subx.sp_instr.cases + intro!: ap_map_list_cast_Dyn_eq_norm_stack[of "take (\\\\\ op) \2"] + dest!: map_eq_append_replicate_conv) + let ?st2' = "Frame f l (Suc pc) R2 (OpDyn x # drop ar \2) # st2'" + let ?s2' = "State F2 H ?st2'" + show ?thesis + proof (intro exI conjI) + show "?STEP ?s2'" + using step_op.hyps casts next_instr2 + unfolding IOp \op' = op\ + by (auto simp: s2_def st2_def intro: Subx.step_op) + next + show "?MATCH (State F1 H ?st1') ?s2'" + proof (rule match.intros) + show "Subx.wf_state ?s2'" + using wf_F2 by (auto intro: Subx.wf_stateI) + next + show "rel_stacktraces (Fubx_get F2) None ?st1' ?st2'" + using step_op.hyps casts next_instr2 rel_stacktraces_Cons + unfolding IOp \op' = op\ + by (auto simp: min_absorb2 take_map[symmetric] drop_map[symmetric] + simp: next_instr_take_Suc_conv + intro!: rel_stacktraces.intros sp_instrs_prefix' Subx.sp_instr.Op + intro: list_all_list_updateI + dest!: ap_map_list_cast_Dyn_replicate[symmetric]) + qed (insert rel_F1_F2, simp_all) + qed + qed simp_all + next + case (step_op_inl op ar opinl x F1') + then obtain instr2 where + next_instr2: "next_instr (Fubx_get F2) f l pc = Some instr2" and + norm_eq_instr1_instr2: "norm_eq (Inca.IOp op) instr2" + by (auto dest: rel_fundefs_next_instr1[OF rel_F1_F2]) + have pc_in_range: "pc < length instrs" and nth_instrs_pc: "instrs ! pc = instr2" + using next_instr_get_map_ofD[OF next_instr2 F2_f map_of_fd2_l] + by simp_all + from next_instr2 norm_eq_instr1_instr2 + show ?case (is "\x. ?STEP x \ ?MATCH (State F1' H ?st1') x") + proof (cases instr2) + case (IOp op') + hence "op' = op" using norm_eq_instr1_instr2 by simp + have casts: + "ap_map_list cast_Dyn (take ar \2) = Some (take ar (map Subx.norm_unboxed \2))" + using sp_instrs_sufix[OF pc_in_range, unfolded nth_instrs_pc IOp, simplified] + using step_op_inl.hyps + by (auto simp: \op' = op\ take_map + elim!: Subx.sp_instrs.cases[of _ _ "x # xs" for x xs, simplified] Subx.sp_instr.cases + intro!: ap_map_list_cast_Dyn_eq_norm_stack[of "take (\\\\\ op) \2"] + dest!: map_eq_append_replicate_conv) + let ?st2' = "Frame f l (Suc pc) R2 (OpDyn x # drop ar \2) # st2'" + let ?F2' = "Subx.Fenv.map_entry F2 f (\fd. rewrite_fundef_body fd l pc (Ubx.IOpInl opinl))" + let ?s2' = "State ?F2' H ?st2'" + have step_s2_s2': "?STEP ?s2'" + using step_op_inl.hyps casts next_instr2 + unfolding IOp \op' = op\ + by (auto simp: s2_def st2_def intro: Subx.step_op_inl) + show ?thesis + proof (intro exI conjI) + show "?STEP ?s2'" by (rule step_s2_s2') + next + show "?MATCH (State F1' H ?st1') ?s2'" + proof (rule match.intros) + show "Subx.wf_state ?s2'" + by (rule Subx.wf_state_step_preservation[OF wf_s2 step_s2_s2']) + next + show "rel_fundefs (Finca_get F1') (Fubx_get ?F2')" + using rel_F1_F2 step_op_inl + by (auto intro: rel_fundefs_rewriteI) + next + have "rel_stacktraces (Fubx_get F2) None ?st1' ?st2'" + using rel_st1'_st2' + apply (rule rel_stacktraces.intros) + using step_op_inl.hyps casts next_instr2 rel_stacktraces_Cons + unfolding IOp \op' = op\ + by (auto simp: min_absorb2 take_map[symmetric] drop_map[symmetric] + simp: next_instr_take_Suc_conv + intro!: sp_instrs_prefix' Subx.sp_instr.Op + dest!: ap_map_list_cast_Dyn_replicate[symmetric]) + then show "rel_stacktraces (Fubx_get ?F2') None ?st1' ?st2'" + apply (rule rel_stacktraces_map_entry_rewrite_fundef_body) + apply (rule next_instr2) + unfolding IOp \op' = op\ + using Sinca.\\\_invertible Subx.sp_instr_Op_OpInl_conv step_op_inl.hyps(4) apply blast + apply simp + apply simp + done + qed + qed + qed simp_all + next + case (step_op_inl_hit opinl ar x) + then obtain instr2 where + next_instr2: "next_instr (Fubx_get F2) f l pc = Some instr2" and + norm_eq_instr1_instr2: "norm_eq (Inca.IOpInl opinl) instr2" + by (auto dest: rel_fundefs_next_instr1[OF rel_F1_F2]) + have pc_in_range: "pc < length instrs" and nth_instrs_pc: "instrs ! pc = instr2" + using next_instr_get_map_ofD[OF next_instr2 F2_f map_of_fd2_l] + by simp_all + from next_instr2 norm_eq_instr1_instr2 + show ?case (is "\x. ?STEP x \ ?MATCH (State F1 H ?st1') x") + proof (cases instr2) case (IOpInl opinl') - then have "opinl' = opinl" - using norm_instr_nth_body step_op_inl_hit.hyps(3) by simp - hence "ar \ length \2" and - all_arg_Dyn: "list_all (\x. x = None) (map typeof (take ar \2))" - using IOpInl step_op_inl_hit sp_sufix[OF pc_in_range] - by (auto simp: take_map) - - let ?s2' = "State F2 H (Frame f (Suc pc) (OpDyn x # drop ar \2) # st2)" - have "?STEP ?s2'" - using IOpInl step_op_inl_hit \opinl' = opinl\ - using \1_def pc_in_range F2_f all_arg_Dyn - by (auto intro: Subx.step_op_inl_hit traverse_cast_Dyn_eq_norm_stack - simp: take_norm_stack) - moreover have "?MATCH ?s2'" - using rel_F1_F2 sp_F2 rel_st1_st2 F2_f - using IOpInl sp_prefix take_Suc_conv_app_nth[OF pc_in_range] - using arg_cong[OF \1_def, of "drop ar"] - using \ar \ length \2\ \opinl' = opinl\ step_op_inl_hit.hyps(4,5) all_arg_Dyn - by (auto intro!: match.intros rel_stacktraces.intros - dest: list_all_eq_const_imp_replicate - simp: drop_norm_stack[symmetric] Let_def take_map drop_map) - ultimately show ?thesis by auto + hence "opinl' = opinl" using norm_eq_instr1_instr2 by simp + have casts: + "ap_map_list cast_Dyn (take ar \2) = Some (take ar (map Subx.norm_unboxed \2))" + using sp_instrs_sufix[OF pc_in_range, unfolded nth_instrs_pc IOpInl, simplified] + using step_op_inl_hit.hyps + by (auto simp: \opinl' = opinl\ take_map + elim!: Subx.sp_instrs.cases[of _ _ "x # xs" for x xs, simplified] Subx.sp_instr.cases + intro!: ap_map_list_cast_Dyn_eq_norm_stack[of "take (\\\\\ (\
\\\\ opinl)) \2"] + dest!: map_eq_append_replicate_conv) + let ?st2' = "Frame f l (Suc pc) R2 (OpDyn x # drop ar \2) # st2'" + let ?s2' = "State F2 H ?st2'" + have step_s2_s2': "?STEP ?s2'" + using step_op_inl_hit.hyps casts next_instr2 + unfolding IOpInl \opinl' = opinl\ + by (auto simp: s2_def st2_def intro: Subx.step_op_inl_hit) + show ?thesis + proof (intro exI conjI) + show "?STEP ?s2'" by (rule step_s2_s2') + next + show "?MATCH (State F1 H ?st1') ?s2'" + proof (rule match.intros) + show "Subx.wf_state ?s2'" + by (rule Subx.wf_state_step_preservation[OF wf_s2 step_s2_s2']) + next + show "rel_stacktraces (Fubx_get F2) None ?st1' ?st2'" + using step_op_inl_hit.hyps casts next_instr2 rel_stacktraces_Cons + unfolding IOpInl \opinl' = opinl\ + by (auto simp: min_absorb2 take_map[symmetric] drop_map[symmetric] + simp: next_instr_take_Suc_conv + intro!: rel_stacktraces.intros sp_instrs_prefix' Subx.sp_instr.OpInl + intro: list_all_list_updateI + dest!: ap_map_list_cast_Dyn_replicate[symmetric]) + qed (insert rel_F1_F2, simp_all) + qed next case (IOpUbx opubx) - then have "\\\ opubx = opinl" - using norm_instr_nth_body step_op_inl_hit.hyps(3) by simp - hence \\\\\_opubx[simp]: "\\\\\ (\
\\\\ (\\\ opubx)) = ar" - using step_op_inl_hit.hyps(4,5) by simp - note sp_sufix' = sp_sufix[OF pc_in_range, unfolded IOpUbx, simplified, unfolded Let_def, simplified] - obtain dom codom where typeof_opubx: "\\\\\\\\ opubx = (dom, codom)" - using prod.exhaust_sel by blast - hence dom_def: "dom = map typeof (take ar \2)" - unfolding \\\\\_opubx[unfolded Subx.\\\\\\\\_\\\\\, symmetric] - using sp_sufix' - by (auto simp: Let_def take_map) - - obtain x' where - eval_op: "\\\\\ opubx (take ar \2) = Some x'" and - typeof_x': "typeof x' = codom" - using typeof_opubx[unfolded dom_def, THEN Subx.\\\\\\\\_correct] - by auto + hence "opinl = \\\ opubx" using norm_eq_instr1_instr2 by simp + let ?ar = "\\\\\ (\
\\\\ (\\\ opubx))" - let ?s2' = "State F2 H (Frame f (Suc pc) (x' # drop ar \2) # st2)" - have "?STEP ?s2'" - using \1_def step_op_inl_hit eval_op \\\\\_opubx - using pc_in_range IOpUbx F2_f - by (auto intro!: Subx.step_op_ubx) - moreover have "?MATCH ?s2'" - proof - - have "x = Subx.norm_unboxed x'" - using Subx.\\\\\_correct[OF eval_op, unfolded \\\\ opubx = opinl\] - using step_op_inl_hit \1_def - by (simp add: take_map norm_stack_def) - thus ?thesis - using rel_F1_F2 sp_F2 rel_st1_st2 F2_f \1_def - using sp_prefix take_Suc_conv_app_nth[OF pc_in_range, unfolded IOpUbx] - using step_op_inl_hit - using typeof_opubx typeof_x' - by (auto intro!: match.intros rel_stacktraces.intros - simp: dom_def drop_norm_stack Let_def min.absorb2 take_map drop_map) + obtain codom where typeof_opubx: "\\\\\\\\ opubx = (map typeof (take ?ar \2), codom)" + using sp_instrs_sufix[OF pc_in_range, unfolded nth_instrs_pc IOpUbx, simplified] + by (cases "\\\\\\\\ opubx") + (auto simp: eq_append_conv_conj Subx.\\\\\\\\_\\\\\ take_map + dest!: Subx.sp_instrs_ConsD elim!: Subx.sp_instr.cases) + + obtain u where + eval_opubx: "\\\\\ opubx (take ?ar \2) = Some u" and typeof_u: "typeof u = codom" + using Subx.\\\\\\\\_correct[OF typeof_opubx] by auto + hence x_def: "x = Subx.norm_unboxed u" + using step_op_inl_hit.hyps + using Subx.\\\\\_correct[OF eval_opubx] + unfolding \opinl = \\\ opubx\ take_map + by simp + + let ?st2' = "Frame f l (Suc pc) R2 (u # drop ?ar \2) # st2'" + let ?s2' = "State F2 H ?st2'" + + have step_s2_s2': "?STEP ?s2'" + using step_op_inl_hit.hyps next_instr2 + unfolding IOpUbx \opinl = \\\ opubx\ + using eval_opubx + by (auto simp: s2_def st2_def intro!: Subx.step_op_ubx) + show ?thesis + proof (intro exI conjI) + show "?STEP ?s2'" by (rule step_s2_s2') + next + show "?MATCH (State F1 H ?st1') ?s2'" + proof (rule match.intros) + show "Subx.wf_state ?s2'" + by (rule Subx.wf_state_step_preservation[OF wf_s2 step_s2_s2']) + next + show "rel_stacktraces (Fubx_get F2) None ?st1' ?st2'" + using step_op_inl_hit.hyps next_instr2 rel_stacktraces_Cons + unfolding IOpUbx \opinl = \\\ opubx\ x_def + by (auto simp: typeof_opubx typeof_u + simp: min_absorb2 take_map[symmetric] drop_map[symmetric] + simp: next_instr_take_Suc_conv + intro!: rel_stacktraces.intros sp_instrs_prefix' Subx.sp_instr.OpUbx + dest!: ap_map_list_cast_Dyn_replicate[symmetric]) + qed (insert rel_F1_F2, simp_all) qed - ultimately show ?thesis by auto qed simp_all next - case (step_op_inl_miss fd1 opinl ar x F1') - hence rel_fd1_fd2: "rel_fundef norm_eq fd1 fd2" - using ex_F1_f by simp - hence norm_instr_nth_body: "norm_instr (body fd2 ! pc) = body fd1 ! pc" - using step_op_inl_miss rel_fundef_body_nth by metis - have pc_in_range: "pc < length (body fd2)" - using step_op_inl_miss rel_fundef_body_length[OF rel_fd1_fd2] by simp - - show ?case (is "\x. ?STEP x \ ?MATCH x") - using step_op_inl_miss norm_instr_nth_body - proof (cases "body fd2 ! pc") + case (step_op_inl_miss opinl ar x F1') + then obtain instr2 where + next_instr2: "next_instr (Fubx_get F2) f l pc = Some instr2" and + norm_eq_instr1_instr2: "norm_eq (Inca.IOpInl opinl) instr2" + by (auto dest: rel_fundefs_next_instr1[OF rel_F1_F2]) + have pc_in_range: "pc < length instrs" and nth_instrs_pc: "instrs ! pc = instr2" + using next_instr_get_map_ofD[OF next_instr2 F2_f map_of_fd2_l] + by simp_all + from next_instr2 norm_eq_instr1_instr2 + show ?case (is "\x. ?STEP x \ ?MATCH (State F1' H ?st1') x") + proof (cases instr2) case (IOpInl opinl') - then have "opinl' = opinl" - using norm_instr_nth_body step_op_inl_miss.hyps(3) by simp - hence "ar \ length \2" and - all_arg_Dyn: "list_all (\x. x = None) (map typeof (take ar \2))" - using IOpInl step_op_inl_miss sp_sufix[OF pc_in_range] - by (auto simp: take_map) - - let ?fd2' = "rewrite_fundef_body fd2 pc (Ubx.IOp (\
\\\\ opinl))" - let ?F2' = "Fubx_add F2 f ?fd2'" - let ?frame = "Frame f (Suc pc) (OpDyn x # drop ar \2)" - let ?s2' = "State ?F2' H (?frame # st2)" - have "?STEP ?s2'" - using IOpInl step_op_inl_miss \opinl' = opinl\ - using \1_def pc_in_range F2_f - using traverse_cast_Dyn_eq_norm_stack[OF all_arg_Dyn] - by (auto intro!: Subx.step_op_inl_miss simp: take_norm_stack) - moreover have "?MATCH ?s2'" - proof - show "rel_fundefs (Finca_get F1') (Fubx_get ?F2')" - using step_op_inl_miss rel_fd1_fd2 - by (auto intro: rel_fundefs_rewrite_both') + hence "opinl' = opinl" using norm_eq_instr1_instr2 by simp + have casts: + "ap_map_list cast_Dyn (take ar \2) = Some (take ar (map Subx.norm_unboxed \2))" + using sp_instrs_sufix[OF pc_in_range, unfolded nth_instrs_pc IOpInl, simplified] + using step_op_inl_miss.hyps + by (auto simp: \opinl' = opinl\ take_map + elim!: Subx.sp_instrs.cases[of _ _ "x # xs" for x xs, simplified] Subx.sp_instr.cases + intro!: ap_map_list_cast_Dyn_eq_norm_stack[of "take (\\\\\ (\
\\\\ opinl)) \2"] + dest!: map_eq_append_replicate_conv) + let ?st2' = "Frame f l (Suc pc) R2 (OpDyn x # drop ar \2) # st2'" + let ?F2' = "Subx.Fenv.map_entry F2 f (\fd. rewrite_fundef_body fd l pc (Ubx.IOp (\
\\\\ opinl)))" + let ?s2' = "State ?F2' H ?st2'" + have step_s2_s2': "?STEP ?s2'" + using step_op_inl_miss.hyps casts next_instr2 + unfolding IOpInl \opinl' = opinl\ + by (auto simp: s2_def st2_def intro: Subx.step_op_inl_miss) + show ?thesis + proof (intro exI conjI) + show "?STEP ?s2'" by (rule step_s2_s2') next - show "sp_fundefs (Fubx_get ?F2')" - using IOpInl \opinl' = opinl\ all_same_arities_add[OF F2_f] sp_full - by (auto intro: sp_fundefs_add[OF sp_F2] - simp: Subx.sp_instr_op[symmetric] Subx.sp_rewrite[OF pc_in_range]) - next - have "rel_stacktraces (Fubx_get F2) - (Frame f (Suc pc) (x # drop ar \1) # st1) (?frame # st2) None" - using rel_st1_st2 sp_prefix F2_f - using IOpInl \1_def take_Suc_conv_app_nth[OF pc_in_range] - using \ar \ length \2\ \opinl' = opinl\ step_op_inl_miss.hyps(4,5) - using all_arg_Dyn - by (auto intro!: rel_stacktraces.intros - dest: list_all_eq_const_imp_replicate - simp: drop_norm_stack Let_def take_map drop_map) - - thus "rel_stacktraces (Fubx_get ?F2') - (Frame f (Suc pc) (x # drop ar \1) # st1) (?frame # st2) None" - using F2_f pc_in_range IOpInl - by (auto intro: rel_stacktraces_rewrite_fundef - simp: \opinl' = opinl\ Let_def Subx.sp_instr_op[symmetric]) + show "?MATCH (State F1' H ?st1') ?s2'" + proof (rule match.intros) + show "Subx.wf_state ?s2'" + by (rule Subx.wf_state_step_preservation[OF wf_s2 step_s2_s2']) + next + show "rel_fundefs (Finca_get F1') (Fubx_get ?F2')" + using rel_F1_F2 step_op_inl_miss.hyps + by (auto intro: rel_fundefs_rewriteI) + next + have "rel_stacktraces (Fubx_get F2) None ?st1' ?st2'" + using rel_st1'_st2' + apply (rule rel_stacktraces.intros) + using step_op_inl_miss.hyps casts next_instr2 rel_stacktraces_Cons + unfolding IOpInl \opinl' = opinl\ + by (auto simp: min_absorb2 take_map[symmetric] drop_map[symmetric] + simp: next_instr_take_Suc_conv + intro!: sp_instrs_prefix' Subx.sp_instr.OpInl + dest!: ap_map_list_cast_Dyn_replicate[symmetric]) + then show "rel_stacktraces (Fubx_get ?F2') None ?st1' ?st2'" + apply (rule rel_stacktraces_map_entry_rewrite_fundef_body) + apply (rule next_instr2) + unfolding IOpInl \opinl' = opinl\ + using Sinca.\\\_invertible Subx.sp_instr_Op_OpInl_conv step_op_inl_miss.hyps apply metis + apply simp + apply simp + done + qed qed - ultimately show ?thesis by auto next case (IOpUbx opubx) - then have "\\\ opubx = opinl" - using norm_instr_nth_body step_op_inl_miss.hyps(3) by simp - hence \\\\\_opubx: "\\\\\ (\
\\\\ (\\\ opubx)) = ar" - using step_op_inl_miss.hyps(4,5) by simp - obtain dom codom where "\\\\\\\\ opubx = (dom, codom)" - using prod.exhaust_sel by blast - hence "\\\\\\\\ opubx = (map typeof (take ar \2), codom)" - unfolding \\\\\_opubx[unfolded Subx.\\\\\\\\_\\\\\, symmetric] - using sp_sufix[OF pc_in_range] - unfolding IOpUbx - by (auto simp: Let_def case_prod_beta take_map) - then obtain x where eval_op: "\\\\\ opubx (take ar \2) = Some x" - using Subx.\\\\\\\\_correct by blast - hence "\\\\\ opinl (take ar \1)" - unfolding \1_def norm_stack_def - using \\\\ opubx = opinl\ - by (auto intro: Sinca.\\\_\\\\\ Subx.\\\\\_to_\\\ simp: take_map) + hence "opinl = \\\ opubx" using norm_eq_instr1_instr2 by simp + let ?ar = "\\\\\ (\
\\\\ (\\\ opubx))" + + obtain codom where typeof_opubx: "\\\\\\\\ opubx = (map typeof (take ?ar \2), codom)" + using sp_instrs_sufix[OF pc_in_range, unfolded nth_instrs_pc IOpUbx, simplified] + by (cases "\\\\\\\\ opubx") + (auto simp: eq_append_conv_conj Subx.\\\\\\\\_\\\\\ take_map + dest!: Subx.sp_instrs_ConsD elim!: Subx.sp_instr.cases) + obtain u where "\\\\\ opubx (take ?ar \2) = Some u" + using Subx.\\\\\\\\_correct[OF typeof_opubx] by auto + hence "\\\\\ opinl (take ?ar \1)" + unfolding \1_def + by (auto simp: \opinl = \\\ opubx\ take_map dest: Subx.\\\\\_to_\\\[THEN Sinca.\\\_\\\\\]) hence False - using step_op_inl_miss by auto + using step_op_inl_miss.hyps + by (simp add: \1_def \opinl = \\\ opubx\) thus ?thesis by simp qed simp_all next - case (step_cjump_true fd1 n \1') - hence rel_fd1_fd2: "rel_fundef norm_eq fd1 fd2" - using ex_F1_f by simp - hence norm_instr_nth_body: "norm_instr (body fd2 ! pc) = body fd1 ! pc" - using step_cjump_true rel_fundef_body_nth by metis - have pc_in_range: "pc < length (body fd2)" - using step_cjump_true rel_fundef_body_length[OF rel_fd1_fd2] by simp + case (step_cjump l\<^sub>t l\<^sub>f d l' \1') + then obtain u \2' where + \2_def: "\2 = u # \2'" and + d_def: "d = Subx.norm_unboxed u" and + \1'_def: "\1' = map Subx.norm_unboxed \2'" + by auto + from step_cjump.hyps obtain instr2 where + next_instr2: "next_instr (Fubx_get F2) f l pc = Some instr2" and + norm_eq_instr1_instr2: "norm_eq (Inca.ICJump l\<^sub>t l\<^sub>f) instr2" + by (auto dest: rel_fundefs_next_instr1[OF rel_F1_F2]) + have pc_in_range: "pc < length instrs" and nth_instrs_pc: "instrs ! pc = instr2" + using next_instr_get_map_ofD[OF next_instr2 F2_f map_of_fd2_l] + by simp_all + + from next_instr2 norm_eq_instr1_instr2 + show ?case (is "\x. ?STEP x \ ?MATCH (State F1 H ?st1') x") + proof (cases instr2) + case (ICJump l\<^sub>t' l\<^sub>f') + hence "l\<^sub>t' = l\<^sub>t" and "l\<^sub>f' = l\<^sub>f" using norm_eq_instr1_instr2 by simp_all + hence "{l\<^sub>t, l\<^sub>f} \ fst ` set (body fd2)" + using all_jumps_in_range pc_in_range nth_instrs_pc ICJump by (auto simp: list_all_length) + moreover have "l' \ {l\<^sub>t, l\<^sub>f}" + using step_cjump.hyps by auto + ultimately have "l' \ fst ` set (body fd2)" + by blast + then obtain instrs' where map_of_l': "map_of (body fd2) l' = Some instrs'" + by (auto dest: weak_map_of_SomeI) + + have sp_instrs_instrs': "Subx.sp_instrs (map_option funtype \ Fubx_get F2) (return fd2) + (butlast instrs @ [instrs ! pc]) [] []" if pc_def: "pc = length instrs - 1" + unfolding pc_def last_conv_nth[OF instrs_neq_Nil, symmetric] + unfolding append_butlast_last_id[OF instrs_neq_Nil] + by (rule sp_instrs_instrs) - show ?case - using step_cjump_true norm_instr_nth_body - proof (cases "body fd2 ! pc") - case (ICJump n') - then have False - using F2_f sp_fundefs_get[OF sp_F2, unfolded sp_fundef_def, THEN Subx.sp_no_jump] - using nth_mem pc_in_range by fastforce - thus ?thesis by simp - qed simp_all - next - case (step_cjump_false fd1 n \1') - hence rel_fd1_fd2: "rel_fundef norm_eq fd1 fd2" - using ex_F1_f by simp - hence norm_instr_nth_body: "norm_instr (body fd2 ! pc) = body fd1 ! pc" - using step_cjump_false rel_fundef_body_nth by metis - have pc_in_range: "pc < length (body fd2)" - using step_cjump_false rel_fundef_body_length[OF rel_fd1_fd2] by simp - - show ?case - using step_cjump_false norm_instr_nth_body - proof (cases "body fd2 ! pc") - case (ICJump n') - then have False - using F2_f sp_fundefs_get[OF sp_F2, unfolded sp_fundef_def, THEN Subx.sp_no_jump] - using nth_mem pc_in_range by fastforce - thus ?thesis by simp + have sp_instr_last: "Subx.sp_instr (map_option funtype \ Fubx_get F2) (return fd2) + (instrs ! pc) (map typeof \2) []" if pc_def: "pc = length instrs - 1" + using sp_instrs_instrs'[OF pc_def] + using sp_instrs_prefix[unfolded pc_def butlast_conv_take[symmetric]] + by (auto dest!: Subx.sp_instrs_appendD') + + have is_jump_nthD: "\n. is_jump (instrs ! n) \ n < length instrs \ n = length instrs - 1" + using list_all_map_of_SomeD[OF wf_fd2[THEN Subx.wf_fundef_all_wf_basic_blockD] map_of_fd2_l] + by (auto dest!: Subx.wf_basic_blockD + list_all_butlast_not_nthD[of "\i. \ is_jump i \ \ Ubx.instr.is_return i", simplified, OF _ disjI1]) + + have pc_def: "pc = length instrs - 1" + using is_jump_nthD[OF _ pc_in_range] nth_instrs_pc ICJump by simp + have \2'_eq_Nil: "\2' = []" + using sp_instr_last[OF pc_def] step_cjump.hyps + by (auto simp: \2_def d_def nth_instrs_pc ICJump elim!: Subx.sp_instr.cases) + + have cast: "cast_Dyn u = Some d" + using sp_instrs_sufix[OF pc_in_range, unfolded nth_instrs_pc ICJump, simplified] + by (auto simp: \2_def d_def dest!: Subx.sp_instrs_ConsD elim!: Subx.sp_instr.cases + intro: Subx.typeof_and_norm_unboxed_imp_cast_Dyn) + + let ?st2' = "Frame f l' 0 R2 \2' # st2'" + let ?s2' = "State F2 H ?st2'" + + have step_s2_s2': "?STEP ?s2'" + using step_cjump.hyps cast next_instr2 + unfolding ICJump \l\<^sub>t' = l\<^sub>t\ \l\<^sub>f' = l\<^sub>f\ + by (auto simp: s2_def st2_def \2_def intro!: Subx.step_cjump) + show ?thesis + proof (intro exI conjI) + show "?STEP ?s2'" by (rule step_s2_s2') + next + show "?MATCH (State F1 H ?st1') ?s2'" + proof (rule match.intros) + show "Subx.wf_state ?s2'" + by (rule Subx.wf_state_step_preservation[OF wf_s2 step_s2_s2']) + next + show "rel_stacktraces (Fubx_get F2) None ?st1' ?st2'" + using step_cjump.hyps cast next_instr2 rel_stacktraces_Cons + using map_of_l' + unfolding ICJump \l\<^sub>t' = l\<^sub>t\ \l\<^sub>f' = l\<^sub>f\ \1'_def \2'_eq_Nil + by (auto simp: min_absorb2 take_map[symmetric] drop_map[symmetric] + simp: next_instr_take_Suc_conv + intro!: rel_stacktraces.intros sp_instrs_prefix' Subx.sp_instr.CJump + intro: Subx.sp_instrs.Nil + dest!: ap_map_list_cast_Dyn_replicate[symmetric]) + qed (insert rel_F1_F2, simp_all) + qed qed simp_all next - case (step_fun_call f' fd1 pc' g gd1 \1' frame\<^sub>g) - hence [simp]: "f' = f" "pc' = pc" "\1' = \1" by auto - hence rel_fd1_fd2: "rel_fundef norm_eq fd1 fd2" - using step_fun_call ex_F1_f by simp - obtain gd2 where "Fubx_get F2 g = Some gd2" and rel_gd1_gd2: "rel_fundef norm_eq gd1 gd2" - using step_fun_call rel_fundefs_Some1[OF rel_F1_F2] by blast - have pc_in_range: "pc < length (body fd2)" - using step_fun_call rel_fd1_fd2 by simp - - have nth_body2: "body fd2 ! pc = Ubx.ICall g" - using rel_fundef_body_nth[OF rel_fd1_fd2 \pc' < length (body fd1)\, symmetric] - unfolding \body fd1 ! pc' = Inca.ICall g\ - by (cases "body fd2 ! pc'"; simp) - - have prefix_\2_all_Dyn: "list_all (\x. x = None) (map typeof (take (arity gd2) \2))" - using \Fubx_get F2 g = Some gd2\ sp_sufix[OF pc_in_range] nth_body2 - by (auto simp: Let_def take_map) - hence all_dyn_args: "list_all is_dyn_operand (take (arity gd2) \2)" - by (auto elim: list.pred_mono_strong simp add: list.pred_map comp_def) + case (step_call g gd1 frame\<^sub>g) + then obtain instr2 where + next_instr2: "next_instr (Fubx_get F2) f l pc = Some instr2" and + norm_eq_instr1_instr2: "norm_eq (Inca.ICall g) instr2" + by (auto dest: rel_fundefs_next_instr1[OF rel_F1_F2]) + have pc_in_range: "pc < length instrs" and nth_instrs_pc: "instrs ! pc = instr2" + using next_instr_get_map_ofD[OF next_instr2 F2_f map_of_fd2_l] + by simp_all - show ?case (is "\x. ?STEP x \ ?MATCH x") - proof - - let ?args = "map OpDyn (norm_stack (take (arity gd2) \2))" - let ?frame = "Frame g 0 ?args" - let ?s1' = "State F2 H (?frame # Frame f pc \2 # st2)" - have "?STEP ?s1'" - proof - - have "arity gd2 \ length \2" - using \arity gd1 \ length \1'\ - using rel_fundef_arities[OF rel_gd1_gd2] \1_def - by simp - thus ?thesis - using pc_in_range nth_body2 - using \Fubx_get F2 g = Some gd2\ rel_stacktraces_Cons.hyps(3) - using all_dyn_args - by (auto intro!: Subx.step_fun_call[of _ _ fd2] nth_equalityI - simp: typeof_unboxed_eq_const list_all_length norm_stack_def) - qed - moreover have "?MATCH ?s1'" - using rel_F1_F2 sp_F2 - proof - show "rel_stacktraces (Fubx_get F2) - (frame\<^sub>g # Frame f pc \1 # st1) - (?frame # Frame f pc \2 # st2) None" - unfolding step_fun_call(7) - proof (rule rel_stacktraces.intros) - show "rel_stacktraces (Fubx_get F2) (Frame f pc \1 # st1) (Frame f pc \2 # st2) - (Some g)" - using pc_in_range nth_body2 rel_stacktraces_Cons - using \Fubx_get F2 g = Some gd2\ all_dyn_args - by (auto intro!: rel_stacktraces.intros simp: is_valid_fun_call_def) - next - show "take (arity gd1) \1' = norm_stack ?args" - using \1_def - using rel_fundef_arities[OF rel_gd1_gd2] - by (simp add: norm_stack_map take_norm_stack) + from step_call.hyps obtain gd2 where + F2_g: "Fubx_get F2 g = Some gd2" and rel_gd1_gd2: "rel_fundef (=) norm_eq gd1 gd2" + using rel_fundefs_Some1[OF rel_F1_F2] by auto + + have wf_gd2: "Subx.wf_fundef (map_option funtype \ Fubx_get F2) gd2" + by (rule Subx.wf_fundefs_getD[OF wf_F2 F2_g]) + + obtain intrs\<^sub>g where gd2_fst_bblock: "map_of (body gd2) (fst (hd (body gd2))) = Some intrs\<^sub>g" + using Subx.wf_fundef_body_neq_NilD[OF wf_gd2] + by (metis hd_in_set map_of_eq_None_iff not_Some_eq prod.collapse prod_in_set_fst_image_conv) + + from norm_eq_instr1_instr2 + show ?case (is "\x. ?STEP x \ ?MATCH (State F1 H ?st1') x") + proof (cases instr2) + case (ICall g') + hence "g' = g" using norm_eq_instr1_instr2 by simp + hence all_dyn_args: "list_all is_dyn_operand (take (arity gd2) \2)" + using sp_instrs_sufix[OF pc_in_range, unfolded nth_instrs_pc ICall, simplified] + using F2_g + by (auto simp: funtype_def eq_append_conv_conj take_map list.pred_set + dest!: Subx.sp_instrs_ConsD replicate_eq_impl_Ball_eq elim!: Subx.sp_instr.cases) + + let ?frame\<^sub>g = "allocate_frame g gd2 (take (arity gd2) \2) (OpDyn uninitialized)" + let ?st2' = "?frame\<^sub>g # Frame f l pc R2 \2 # st2'" + let ?s2' = "State F2 H ?st2'" + + have step_s2_s2': "?STEP ?s2'" + using step_call.hyps next_instr2 F2_g rel_gd1_gd2 all_dyn_args + unfolding ICall \g' = g\ + by (auto simp: s2_def st2_def rel_fundef_arities intro!: Subx.step_call) + show ?thesis + proof (intro exI conjI) + show "?STEP ?s2'" by (rule step_s2_s2') + next + show "?MATCH (State F1 H ?st1') ?s2'" + proof (rule match.intros) + show "Subx.wf_state ?s2'" + by (rule Subx.wf_state_step_preservation[OF wf_s2 step_s2_s2']) next - show "Fubx_get F2 g = Some gd2" - using \Fubx_get F2 g = Some gd2\ . - next - show "sp_fundef (Fubx_get F2) gd2 (take 0 (body gd2)) = Ok (map typeof ?args)" - using \1_def step_fun_call.hyps(5) - using rel_fundef_arities[OF rel_gd1_gd2] - by (simp add: map_replicate_const) - qed simp_all + have FOO: "fst (hd (body gd1)) = fst (hd (body gd2))" + apply (rule rel_fundef_rel_fst_hd_bodies[OF rel_gd1_gd2]) + using Subx.wf_fundefs_getD[OF wf_F2 \Fubx_get F2 g = Some gd2\] + by (auto dest: Subx.wf_fundef_body_neq_NilD) + show "rel_stacktraces (Fubx_get F2) None ?st1' ?st2'" + unfolding step_call.hyps allocate_frame_def FOO + proof (rule rel_stacktraces.intros) + show "Fubx_get F2 g = Some gd2" + by (rule F2_g) + next + show "rel_stacktraces (Fubx_get F2) (Some g) + (Frame f l pc (map Subx.norm_unboxed R2) (map Subx.norm_unboxed \2) # st1') + (Frame f l pc R2 \2 # st2')" + using step_call.hyps rel_stacktraces_Cons next_instr2 F2_g rel_gd1_gd2 all_dyn_args + unfolding ICall \g' = g\ + by (auto simp: is_valid_fun_call_def rel_fundef_arities intro!: rel_stacktraces.intros) + qed (insert rel_gd1_gd2 all_dyn_args gd2_fst_bblock, + simp_all add: take_map rel_fundef_arities rel_fundef_locals Subx.sp_instrs.Nil + list_all_replicateI) + qed (insert rel_F1_F2, simp_all) qed - ultimately show "?thesis" by blast - qed + qed simp_all next - case (step_fun_end f' fd1 \1\<^sub>g pc' \1' frame\<^sub>g g pc\<^sub>g frame\<^sub>g' st1') - hence [simp]: "f' = f" "pc' = pc" "\1' = \1" by auto - hence rel_fd1_fd2: "rel_fundef norm_eq fd1 fd2" - using step_fun_end ex_F1_f by simp - - note arities_fd1_fd2 = rel_fundef_arities[OF rel_fd1_fd2] - - obtain \2\<^sub>g st2' where - st2_def: "st2 = Frame g pc\<^sub>g \2\<^sub>g # st2'" - using step_fun_end rel_st1_st2 - by (auto elim: rel_stacktraces.cases) + case (step_return fd1 \1\<^sub>g frame\<^sub>g' g l\<^sub>g pc\<^sub>g R1\<^sub>g st1'') + then obtain instr2 where + next_instr2: "next_instr (Fubx_get F2) f l pc = Some instr2" and + norm_eq_instr1_instr2: "norm_eq Inca.IReturn instr2" + by (auto dest: rel_fundefs_next_instr1[OF rel_F1_F2]) + have pc_in_range: "pc < length instrs" and nth_instrs_pc: "instrs ! pc = instr2" + using next_instr_get_map_ofD[OF next_instr2 F2_f map_of_fd2_l] + by simp_all - hence all_dyn_prefix_\2\<^sub>g: "list_all (\x. x = None) (map typeof (take (arity fd2) \2\<^sub>g))" - using step_fun_end rel_st1_st2 F2_f - by (auto - elim!: rel_stacktraces.cases[of _ "_ # _"] list.pred_mono_strong - simp: list.pred_map is_valid_fun_call_def) + from step_return.hyps have rel_fd1_fd2: "rel_fundef (=) norm_eq fd1 fd2" + using rel_fundefsD[OF rel_F1_F2, of f] F2_f by simp - show ?case (is "\x. ?STEP x \ ?MATCH x") - proof - - let ?s1' = "State F2 H (Frame g (Suc pc\<^sub>g) (\2 @ drop (arity fd2) \2\<^sub>g) # st2')" - have "?STEP ?s1'" - unfolding st2_def - using step_fun_end st2_def rel_st1_st2 - using arities_fd1_fd2 - using rel_fundef_body_length[OF rel_fd1_fd2] - by (auto intro!: Subx.step_fun_end[OF \Fubx_get F2 f = Some fd2\] - elim!: rel_stacktraces.cases[of _ "_ # _"]) - moreover have "?MATCH ?s1'" - proof - - have "map typeof \2 = [None]" - using step_fun_end st2_def rel_st1_st2 - using rel_fd1_fd2 sp_full sp_prefix - by (auto elim!: rel_stacktraces.cases[of _ "_ # _"]) - - thus ?thesis - using step_fun_end st2_def rel_st1_st2 - using arities_fd1_fd2 \1_def F2_f all_dyn_prefix_\2\<^sub>g - by (auto intro!: match.intros rel_stacktraces.intros - intro: rel_F1_F2 sp_F2 - elim!: rel_stacktraces.cases[of _ "_ # _"] - dest: list_all_eq_const_imp_replicate - simp: take_Suc_conv_app_nth Let_def drop_norm_stack take_map drop_map - simp: is_valid_fun_call_def) + from norm_eq_instr1_instr2 + show ?case (is "\x. ?STEP x \ ?MATCH (State F1 H ?st1') x") + proof (cases instr2) + case IReturn + have map_typeof_\2: "map typeof \2 = replicate (return fd2) None" + using sp_instrs_sufix[OF pc_in_range, unfolded nth_instrs_pc IReturn, simplified] + by (auto elim: Subx.sp_instr.cases dest: Subx.sp_instrs_ConsD) + hence all_dyn_\2: "list_all is_dyn_operand \2" + by (auto simp: list.pred_set dest: replicate_eq_impl_Ball_eq[OF sym]) + show ?thesis + using rel_st1'_st2' unfolding \Frame g l\<^sub>g pc\<^sub>g R1\<^sub>g \1\<^sub>g # st1'' = st1'\[symmetric] + proof (cases rule: rel_stacktraces.cases) + case (rel_stacktraces_Cons st2'' \2\<^sub>g R2\<^sub>g gd2 instrs\<^sub>g) + let ?st2' = "Frame g l\<^sub>g (Suc pc\<^sub>g) R2\<^sub>g (\2 @ drop (arity fd2) \2\<^sub>g) # st2''" + let ?s2' = "State F2 H ?st2'" + have step_s2_s2': "?STEP ?s2'" + using step_return.hyps next_instr2 F2_f rel_fd1_fd2 rel_stacktraces_Cons all_dyn_\2 + unfolding IReturn + by (auto simp: s2_def st2_def rel_fundef_arities rel_fundef_return + intro: Subx.step_return) + show ?thesis + proof (intro exI conjI) + show "?STEP ?s2'" by (rule step_s2_s2') + next + show "?MATCH (State F1 H ?st1') ?s2'" + proof (rule match.intros) + show "Subx.wf_state ?s2'" + by (rule Subx.wf_state_step_preservation[OF wf_s2 step_s2_s2']) + next + have "Subx.sp_instr (map_option funtype \ Fubx_get F2) (return gd2) + (Ubx.instr.ICall f) (map typeof \2\<^sub>g) + (replicate (return fd2) None @ map typeof (drop (arity fd2) \2\<^sub>g))" + using rel_stacktraces_Cons F2_f + using replicate_eq_map[of "arity fd2" "take (arity fd2) \2\<^sub>g" typeof None] + by (auto simp: funtype_def is_valid_fun_call_def + simp: min_absorb2 list.pred_set take_map[symmetric] drop_map[symmetric] + intro!: Subx.sp_instr.Call[where \ = "map typeof (drop (arity fd2) \2\<^sub>g)"]) + then show "rel_stacktraces (Fubx_get F2) None ?st1' ?st2'" + unfolding step_return.hyps + using rel_stacktraces_Cons rel_fd1_fd2 F2_f + using map_typeof_\2 + by (auto simp: drop_map rel_fundef_arities is_valid_fun_call_def + simp: next_instr_take_Suc_conv funtype_def + intro!: rel_stacktraces.intros Subx.sp_instrs_appendI[where \ = "map typeof \2\<^sub>g"]) + qed (insert rel_F1_F2, simp_all) + qed qed - ultimately show ?thesis by auto - qed + qed simp_all qed qed qed lemma match_final_forward: - "s1 \ s2 \ Sinca.final s1 \ Subx.final s2" -proof (induction s1 s2 rule: match.induct) - case (1 F1 F2 st1 st2 H) - obtain f fd1 pc \1 where - st1_def: "st1 = [Frame f pc \1]" and - F1_f: "Finca_get F1 f = Some fd1" and - pc_def: "pc = length (body fd1)" - using \Sinca.final (State F1 H st1)\ - by (auto elim: Sinca.final.cases) - obtain \2 where st2_def: "st2 = [Frame f pc \2]" - using \rel_stacktraces (Fubx_get F2) st1 st2 None\ - unfolding st1_def - by (auto elim: rel_stacktraces.cases) - obtain fd2 where F2_f: "Fubx_get F2 f = Some fd2" and rel_fd1_fd2: "rel_fundef norm_eq fd1 fd2" - using rel_fundefs_Some1[OF \rel_fundefs (Finca_get F1) (Fubx_get F2)\ F1_f] - by auto - have "length (body fd1) = length (body fd2)" - using rel_fd1_fd2 by simp - thus ?case - unfolding st2_def pc_def - using F2_f by (auto intro: Subx.final.intros) + assumes "match s1 s2" and final_s1: "final Finca_get Inca.IReturn s1" + shows "final Fubx_get Ubx.IReturn s2" + using \match s1 s2\ +proof (cases s1 s2 rule: match.cases) + case (matchI F2 H st2 F1 st1) + show ?thesis + using final_s1[unfolded matchI] + proof (cases _ _ "State F1 H st1" rule: final.cases) + case (finalI f l pc R \) + then show ?thesis + using matchI + by (auto intro!: final.intros elim: rel_stacktraces.cases norm_instr.elims[OF sym] + dest: rel_fundefs_next_instr1) + qed qed sublocale inca_ubx_forward_simulation: - forward_simulation Sinca.step Subx.step Sinca.final Subx.final "\_ _. False" "\_. match" + forward_simulation Sinca.step Subx.step + "final Finca_get Inca.IReturn" + "final Fubx_get Ubx.IReturn" + "\_ _. False" "\_. match" using match_final_forward forward_lockstep_simulation using lockstep_to_plus_forward_simulation[of match Sinca.step _ Subx.step] by unfold_locales auto section \Bisimulation\ sublocale inca_ubx_bisimulation: - bisimulation Sinca.step Subx.step Sinca.final Subx.final "\_ _. False" "\_. match" + bisimulation Sinca.step Subx.step "final Finca_get Inca.IReturn" "final Fubx_get Ubx.IReturn" + "\_ _. False" "\_. match" by unfold_locales end end diff --git a/thys/Interpreter_Optimizations/List_util.thy b/thys/Interpreter_Optimizations/List_util.thy --- a/thys/Interpreter_Optimizations/List_util.thy +++ b/thys/Interpreter_Optimizations/List_util.thy @@ -1,107 +1,231 @@ theory List_util imports Main begin inductive same_length :: "'a list \ 'b list \ bool" where same_length_Nil: "same_length [] []" | same_length_Cons: "same_length xs ys \ same_length (x # xs) (y # ys)" code_pred same_length . lemma same_length_iff_eq_lengths: "same_length xs ys \ length xs = length ys" proof assume "same_length xs ys" then show "length xs = length ys" by (induction xs ys rule: same_length.induct) simp_all next assume "length xs = length ys" then show "same_length xs ys" proof (induction xs arbitrary: ys) case Nil then show ?case by (simp add: same_length_Nil) next case (Cons x xs) then show ?case by (metis length_Suc_conv same_length_Cons) qed qed lemma same_length_Cons: "same_length (x # xs) ys \ \y ys'. ys = y # ys'" "same_length xs (y # ys) \ \x xs'. xs = x # xs'" proof - assume "same_length (x # xs) ys" then show "\y ys'. ys = y # ys'" by (induction "x # xs" ys rule: same_length.induct) simp next assume "same_length xs (y # ys)" then show "\x xs'. xs = x # xs'" by (induction xs "y # ys" rule: same_length.induct) simp qed -inductive for_all2 for r where - for_all2_Nil: "for_all2 r [] []" | - for_all2_Cons: "r x y \ for_all2 r xs ys \ for_all2 r (x # xs) (y # ys)" - -code_pred for_all2 . - -declare for_all2_Nil[intro] -declare for_all2_Cons[intro] - -lemma for_all2_refl: "(\x. r x x) \ for_all2 r xs xs" - by (induction xs) auto - -lemma for_all2_same_length: "for_all2 r xs ys \ same_length xs ys" - by (induction rule: for_all2.induct) (auto intro: same_length.intros) - -lemma for_all2_ConsD: "for_all2 r (x # xs) (y # ys) \ r x y \ for_all2 r xs ys" - using for_all2.cases by blast - section \nth\_opt\ fun nth_opt where "nth_opt (x # _) 0 = Some x" | "nth_opt (_ # xs) (Suc n) = nth_opt xs n" | "nth_opt _ _ = None" lemma nth_opt_eq_Some_conv: "nth_opt xs n = Some x \ n < length xs \ xs ! n = x" by (induction xs n rule: nth_opt.induct; simp) lemmas nth_opt_eq_SomeD[dest] = nth_opt_eq_Some_conv[THEN iffD1] section \Generic lemmas\ +lemma list_rel_imp_pred1: + assumes + "list_all2 R xs ys" and + "\x y. (x, y) \ set (zip xs ys) \ R x y \ P x" + shows "list_all P xs" + using assms + by (induction xs ys rule: list.rel_induct) auto + +lemma list_rel_imp_pred2: + assumes + "list_all2 R xs ys" and + "\x y. (x, y) \ set (zip xs ys) \ R x y \ P y" + shows "list_all P ys" + using assms + by (induction xs ys rule: list.rel_induct) auto + +lemma eq_append_conv_conj: "(zs = xs @ ys) = (xs = take (length xs) zs \ ys = drop (length xs) zs)" + by (metis append_eq_conv_conj) + +lemma list_all_list_updateI: "list_all P xs \ P x \ list_all P (list_update xs n x)" + by (induction xs arbitrary: n) (auto simp add: nat.split_sels(2)) + +lemmas list_all2_update1_cong = list_all2_update_cong[of _ _ ys _ "ys ! i" i for ys i, simplified] +lemmas list_all2_update2_cong = list_all2_update_cong[of _ xs _ "xs ! i" _ i for xs i, simplified] + lemma map_list_update_id: "f (xs ! pc) = f instr \ map f (xs[pc := instr]) = map f xs" using list_update_id map_update by metis lemma list_all_eq_const_imp_replicate: assumes "list_all (\x. x = y) xs" shows "xs = replicate (length xs) y" using assms by (induction xs; simp) lemma list_all_eq_const_replicate_lhs[intro]: "list_all (\x. y = x) (replicate n y)" by (simp add: list_all_length) lemma list_all_eq_const_replicate_rhs[intro]: "list_all (\x. x = y) (replicate n y)" by (simp add: list_all_length) lemma replicate_eq_map: - assumes "list_all g (take n xs)" and "n \ length xs" and "\y. g y \ f y = x" - shows "replicate n x = map f (take n xs)" + assumes "n = length xs" and "\y. y \ set xs \ f y = x" + shows "replicate n x = map f xs" using assms proof (induction xs arbitrary: n) case Nil thus ?case by simp next case (Cons x xs) thus ?case by (cases n; auto) qed +lemma replicate_eq_impl_Ball_eq: + shows "replicate n c = xs \ (\x \ set xs. x = c)" + by (meson in_set_replicate) + +lemma rel_option_map_of: + assumes "list_all2 (rel_prod (=) R) xs ys" + shows "rel_option R (map_of xs l) (map_of ys l)" + using assms +proof (induction xs ys rule: list.rel_induct) + case Nil + thus ?case by simp +next + case (Cons x xs y ys) + from Cons.hyps have "fst x = fst y" and "R (snd x) (snd y)" + by (simp_all add: rel_prod_sel) + show ?case + proof (cases "l = fst y") + case True + then show ?thesis + by (simp add: \fst x = fst y\ \R (snd x) (snd y)\) + next + case False + then show ?thesis + using Cons.IH + by (simp add: \fst x = fst y\ ) + qed +qed + +lemma list_all2_rel_prod_nth: + assumes "list_all2 (rel_prod R1 R2) xs ys" and "n < length xs" + shows "R1 (fst (xs ! n)) (fst (ys ! n)) \ R2 (snd (xs ! n)) (snd (ys ! n))" + using assms +proof (induction n arbitrary: xs ys) + case 0 + then show ?case + using assms(1,2) + by (auto elim: list.rel_cases) +next + case (Suc n) + then obtain x xs' y ys' where xs_def[simp]: "xs = x # xs'" and ys_def[simp]: "ys = y # ys'" + by (auto elim: list.rel_cases) + show ?case + using Suc.prems Suc.IH[of xs' ys'] + by force +qed + +lemma list_all2_rel_prod_fst_hd: + assumes "list_all2 (rel_prod R1 R2) xs ys" and "xs \ [] \ ys \ []" + shows "R1 (fst (hd xs)) (fst (hd ys)) \ R2 (snd (hd xs)) (snd (hd ys))" + using assms + by (auto simp: rel_prod_sel elim: list.rel_cases) + +lemma list_all2_rel_prod_fst_last: + assumes "list_all2 (rel_prod R1 R2) xs ys" and "xs \ [] \ ys \ []" + shows "R1 (fst (last xs)) (fst (last ys)) \ R2 (snd (last xs)) (snd (last ys))" +proof - + have "xs \ []" and "ys \ []" + using assms by (auto elim: list.rel_cases) + moreover have "length xs = length ys" + by (rule assms(1)[THEN list_all2_lengthD]) + ultimately show ?thesis + using list_all2_rel_prod_nth[OF assms(1)] + by (simp add: last_conv_nth) +qed + +lemma list_all_nthD[intro]: "list_all P xs \ n < length xs \ P (xs ! n)" + by (simp add: list_all_length) + +lemma "list_all P xs \ \x\ set xs. P x" + using list_all_iff list.pred_set + by (simp add: list_all_iff) + + +lemma list_all_map_of_SomeD: + assumes "list_all P kvs" and "map_of kvs k = Some v" + shows "P (k, v)" + using assms + unfolding list.pred_set + by (auto dest: map_of_SomeD) + +lemma list_all_not_nthD:"list_all P xs \ \ P (xs ! n) \ length xs \ n" +proof (induction xs arbitrary: n) + case Nil + then show ?case + by simp +next + case (Cons x xs) + then show ?case + by (cases n) simp_all +qed + +lemma list_all_butlast_not_nthD: "list_all P (butlast xs) \ \ P (xs ! n) \ length xs \ Suc n" + using list_all_not_nthD[of _ "butlast xs" for xs, simplified] + by (smt (z3) One_nat_def Suc_pred le_Suc_eq length_butlast length_greater_0_conv less_Suc_eq list.pred_inject(1) list_all_not_nthD not_le nth_butlast) + +lemma list_all_replicateI[intro]: "P x \ list_all P (replicate n x)" + unfolding list.pred_set + by simp + +lemma map_eq_append_replicate_conv: + assumes "map f xs = replicate n x @ ys" + shows "map f (take n xs) = replicate n x" + using assms + by (metis append_eq_conv_conj length_replicate take_map) + +lemma map_eq_replicate_imp_list_all_const: + "map f xs = replicate n x \ n = length xs \ list_all (\y. f y = x) xs" + by (induction xs arbitrary: n) simp_all + +lemma map_eq_replicateI: "length xs = n \ (\x. x \ set xs \ f x = c) \ map f xs = replicate n c" + by (induction xs arbitrary: n) auto + + +section \Non-empty list\ + +type_synonym 'a nlist = "'a \ 'a list" + end diff --git a/thys/Interpreter_Optimizations/Map_Extra.thy b/thys/Interpreter_Optimizations/Map_Extra.thy new file mode 100644 --- /dev/null +++ b/thys/Interpreter_Optimizations/Map_Extra.thy @@ -0,0 +1,52 @@ +theory Map_Extra + imports Main "HOL-Library.Library" +begin + +lemmas map_of_eq_Some_imp_key_in_fst_dom[intro] = + domI[of "map_of xs" for xs, unfolded dom_map_of_conv_image_fst] + +lemma very_weak_map_of_SomeI: "k \ fst ` set kvs \ \v. map_of kvs k = Some v" + by (induction kvs) auto + +lemma map_of_fst_hd_neq_Nil[simp]: + assumes "xs \ []" + shows "map_of xs (fst (hd xs)) = Some (snd (hd xs))" + using assms + by (cases xs) simp_all + +definition map_merge where + "map_merge f m1 m2 x = + (case (m1 x, m2 x) of + (None, None) \ None + | (None, Some z) \ Some z + | (Some y, None) \ Some y + | (Some y, Some z) \ f y z)" + +lemma option_case_cancel[simp]: "(case opt of None \ x | Some _ \ x) = x" + using option.case_eq_if + by (simp add: option.case_eq_if) + +lemma map_le_map_merge_Some_const: + "f \\<^sub>m map_merge (\x y. Some x) f g" and + "g \\<^sub>m map_merge (\x y. Some y) f g" + unfolding map_le_def atomize_conj + find_theorems "_ &&& _" +proof (intro conjI ballI) + fix x + assume "x \ dom f" + then obtain y where "f x = Some y" + by blast + then show "f x = map_merge (\x y. Some x) f g x" + unfolding map_merge_def + by simp +next + fix x + assume "x \ dom g" + then obtain z where "g x = Some z" + by blast + then show "g x = map_merge (\x. Some) f g x" + unfolding map_merge_def + by (simp add: option.case_eq_if) +qed + +end \ No newline at end of file diff --git a/thys/Interpreter_Optimizations/Op.thy b/thys/Interpreter_Optimizations/Op.thy --- a/thys/Interpreter_Optimizations/Op.thy +++ b/thys/Interpreter_Optimizations/Op.thy @@ -1,14 +1,60 @@ theory Op imports Main begin section \n-ary operations\ locale nary_operations = fixes \\ :: "'op \ 'a list \ 'a" and \\\\\ :: "'op \ nat" assumes \\_\\\\\_domain: "length xs = \\\\\ op \ \y. \\ op xs = y" +(* locale nary_operations_inl = + nary_operations "\op args. fst (eval op args)" arity + for + eval :: "'op \ 'dyn list \ 'dyn \ 'op option" and arity +begin + +definition equiv_op where + "equiv_op x y \ fst \ eval x = fst \ eval y" + +lemma equivp_equiv_op[simp]: "equivp equiv_op" +proof (rule equivpI) + show "reflp equiv_op" + by (rule reflpI) (simp add: equiv_op_def) +next + show "symp equiv_op" + by (rule sympI) (simp add: equiv_op_def) +next + show "transp equiv_op" + by (rule transpI) (simp add: equiv_op_def) +qed + +end *) + +(* locale nary_operations0 = + fixes + eval :: "'op \ 'dyn list \ 'dyn \ 'op option" and + arity :: "'op \ nat" +begin + +definition equiv_op where + "equiv_op x y \ fst \ eval x = fst \ eval y" + +lemma equivp_equiv_op[simp]: "equivp equiv_op" +proof (rule equivpI) + show "reflp equiv_op" + by (rule reflpI) (simp add: equiv_op_def) +next + show "symp equiv_op" + by (rule sympI) (simp add: equiv_op_def) +next + show "transp equiv_op" + by (rule transpI) (simp add: equiv_op_def) +qed + +end *) + end \ No newline at end of file diff --git a/thys/Interpreter_Optimizations/OpUbx.thy b/thys/Interpreter_Optimizations/OpUbx.thy --- a/thys/Interpreter_Optimizations/OpUbx.thy +++ b/thys/Interpreter_Optimizations/OpUbx.thy @@ -1,45 +1,41 @@ theory OpUbx imports OpInl Unboxed begin section \n-ary operations\ locale nary_operations_ubx = nary_operations_inl \\ \\\\\ \\\\\ \\\ \\\\\ \
\\\\ + - unboxedval is_true is_false box_ubx1 unbox_ubx1 box_ubx2 unbox_ubx2 + unboxedval unitialized is_true is_false box_ubx1 unbox_ubx1 box_ubx2 unbox_ubx2 for \\ :: "'op \ 'dyn list \ 'dyn" and \\\\\ and \\\\\ and \\\ and \\\\\ and \
\\\\ :: "'opinl \ 'op" and - is_true :: "'dyn \ bool" and is_false and + unitialized :: 'dyn and is_true and is_false and box_ubx1 :: "'ubx1 \ 'dyn" and unbox_ubx1 and box_ubx2 :: "'ubx2 \ 'dyn" and unbox_ubx2 + fixes \\\\\ :: "'opubx \ ('dyn, 'ubx1, 'ubx2) unboxed list \ ('dyn, 'ubx1, 'ubx2) unboxed option" and \\\ :: "'opinl \ type option list \ 'opubx option" and \\\ :: "'opubx \ 'opinl" and \\\\\\\\ :: "'opubx \ type option list \ type option" assumes \\\_invertible: "\\\ opinl ts = Some opubx \ \\\ opubx = opinl" and \\\\\_correct: "\\\\\ opubx \ = Some z \ \\\\\ (\\\ opubx) (map norm_unboxed \) = norm_unboxed z" and \\\\\_to_\\\: "\\\\\ opubx \ = Some z \ \\\ (\
\\\\ (\\\ opubx)) (map norm_unboxed \) = Some (\\\ opubx)" and \\\\\\\\_\\\\\: "\\\\\ (\
\\\\ (\\\ opubx)) = length (fst (\\\\\\\\ opubx))" and \\\_opubx_type: "\\\ opinl ts = Some opubx \ fst (\\\\\\\\ opubx) = ts" and \\\\\\\\_correct: "\\\\\\\\ opubx = (map typeof xs, \) \ \y. \\\\\ opubx xs = Some y \ typeof y = \" and \\\\\\\\_complete: "\\\\\ opubx xs = Some y \ \\\\\\\\ opubx = (map typeof xs, typeof y)" -begin - -end - end \ No newline at end of file diff --git a/thys/Interpreter_Optimizations/Op_example.thy b/thys/Interpreter_Optimizations/Op_example.thy --- a/thys/Interpreter_Optimizations/Op_example.thy +++ b/thys/Interpreter_Optimizations/Op_example.thy @@ -1,241 +1,240 @@ theory Op_example - imports OpUbx Ubx_type_inference Global Unboxed_lemmas + imports OpUbx Global Unboxed_lemmas begin section \Dynamic values\ datatype dynamic = DNil | DBool bool | DNum nat definition is_true where - "is_true d = (d = DBool True)" + "is_true d \ (d = DBool True)" definition is_false where - "is_false d = (d = DBool False)" + "is_false d \ (d = DBool False)" -definition box_bool :: "bool \ dynamic" where - "box_bool = DBool" - -definition box_num :: "nat \ dynamic" where - "box_num = DNum" +interpretation dynval_dynamic: dynval DNil is_true is_false +proof unfold_locales + fix d + show "\ (is_true d \ is_false d)" + by (cases d) (simp_all add: is_true_def is_false_def) +qed fun unbox_num :: "dynamic \ nat option" where "unbox_num (DNum n) = Some n" | "unbox_num _ = None" fun unbox_bool :: "dynamic \ bool option" where "unbox_bool (DBool b) = Some b" | "unbox_bool _ = None" interpretation unboxed_dynamic: - unboxedval is_true is_false box_num unbox_num box_bool unbox_bool -proof (unfold_locales; (simp add: is_true_def is_false_def)?) + unboxedval DNil is_true is_false DNum unbox_num DBool unbox_bool +proof (unfold_locales) fix d n - show "unbox_num d = Some n \ box_num n = d" - by (cases d; simp add: box_num_def) + show "unbox_num d = Some n \ DNum n = d" + by (cases d; simp add: ) next fix d b - show "unbox_bool d = Some b \ box_bool b = d" - by (cases d; simp add: box_bool_def) + show "unbox_bool d = Some b \ DBool b = d" + by (cases d; simp add:) qed section \Normal operations\ datatype op = OpNeg | OpAdd | OpMul fun ar :: "op \ nat" where "ar OpNeg = 1" | "ar OpAdd = 2" | "ar OpMul = 2" fun eval_Neg :: "dynamic list \ dynamic" where "eval_Neg [DBool b] = DBool (\b)" | "eval_Neg [_] = DNil" fun eval_Add :: "dynamic list \ dynamic" where "eval_Add [DBool x, DBool y] = DBool (x \ y)" | "eval_Add [DNum x, DNum y] = DNum (x + y)" | "eval_Add [_, _] = DNil" fun eval_Mul :: "dynamic list \ dynamic" where "eval_Mul [DBool x, DBool y] = DBool (x \ y)" | "eval_Mul [DNum x, DNum y] = DNum (x * y)" | "eval_Mul [_, _] = DNil" fun eval :: "op \ dynamic list \ dynamic" where "eval OpNeg = eval_Neg" | "eval OpAdd = eval_Add" | "eval OpMul = eval_Mul" lemma eval_arith_domain: "length xs = ar op \ \y. eval op xs = y" by simp interpretation op_Op: nary_operations eval ar using eval_arith_domain by (unfold_locales; simp) section \Inlined operations\ datatype opinl = OpAddNum | OpMulNum fun inl :: "op \ dynamic list \ opinl option" where "inl OpAdd [DNum _, DNum _] = Some OpAddNum" | "inl OpMul [DNum _, DNum _] = Some OpMulNum" | "inl _ _ = None" inductive isinl :: "opinl \ dynamic list \ bool" where "isinl OpAddNum [DNum _, DNum _]" | "isinl OpMulNum [DNum _, DNum _]" fun deinl :: "opinl \ op" where "deinl OpAddNum = OpAdd" | "deinl OpMulNum = OpMul" lemma inl_inj: "inj inl" unfolding inj_def proof (intro allI impI) fix x y assume "inl x = inl y" thus "x = y" unfolding fun_eq_iff by (cases x; cases y; auto elim: inl.elims simp: HOL.eq_commute[of None]) qed lemma inl_invertible: "inl op xs = Some opinl \ deinl opinl = op" by (induction op xs rule: inl.induct; simp) fun eval_AddNum :: "dynamic list \ dynamic" where "eval_AddNum [DNum x, DNum y] = DNum (x + y)" | "eval_AddNum [DBool x, DBool y] = DBool (x \ y)" | "eval_AddNum [_, _] = DNil" fun eval_MulNum :: "dynamic list \ dynamic" where "eval_MulNum [DNum x, DNum y] = DNum (x * y)" | "eval_MulNum [DBool x, DBool y] = DBool (x \ y)" | "eval_MulNum [_, _] = DNil" fun eval_inl :: "opinl \ dynamic list \ dynamic" where "eval_inl OpAddNum = eval_AddNum" | "eval_inl OpMulNum = eval_MulNum" lemma eval_AddNum_correct: "length xs = 2 \ eval_AddNum xs = eval_Add xs" by (induction xs rule: eval_AddNum.induct; simp) lemma eval_MulNum_correct: "length xs = 2 \ eval_MulNum xs = eval_Mul xs" by (induction xs rule: eval_MulNum.induct; simp) lemma eval_inl_correct: "length xs = ar (deinl opinl) \ eval_inl opinl xs = eval (deinl opinl) xs" using eval_AddNum_correct eval_MulNum_correct by (cases opinl; simp) lemma inl_isinl: "inl op xs = Some opinl \ isinl opinl xs" by (induction op xs rule: inl.induct; auto intro: isinl.intros) interpretation op_OpInl: nary_operations_inl eval ar eval_inl inl isinl deinl using inl_invertible using eval_inl_correct using inl_isinl by (unfold_locales; simp) section \Unboxed operations\ datatype opubx = OpAddNumUbx fun ubx :: "opinl \ type option list \ opubx option" where "ubx OpAddNum [Some Ubx1, Some Ubx1] = Some OpAddNumUbx" | "ubx _ _ = None" fun deubx :: "opubx \ opinl" where "deubx OpAddNumUbx = OpAddNum" lemma ubx_invertible: "ubx opinl xs = Some opubx \ deubx opubx = opinl" by (induction opinl xs rule: ubx.induct; simp) fun eval_AddNumUbx :: "(dynamic, nat, bool) unboxed list \ (dynamic, nat, bool) unboxed option" where "eval_AddNumUbx [OpUbx1 x, OpUbx1 y] = Some (OpUbx1 (x + y))" | "eval_AddNumUbx _ = None" fun eval_ubx :: "opubx \ (dynamic, nat, bool) unboxed list \ (dynamic, nat, bool) unboxed option" where "eval_ubx OpAddNumUbx = eval_AddNumUbx" lemma eval_ubx_correct: "eval_ubx opubx xs = Some z \ eval_inl (deubx opubx) (map unboxed_dynamic.norm_unboxed xs) = unboxed_dynamic.norm_unboxed z" apply (cases opubx; simp) - apply (induction xs rule: eval_AddNumUbx.induct; auto) - by (simp add: box_num_def) + by (induction xs rule: eval_AddNumUbx.induct) auto lemma eval_ubx_to_inl: assumes "eval_ubx opubx \ = Some z" shows "inl (deinl (deubx opubx)) (map unboxed_dynamic.norm_unboxed \) = Some (deubx opubx)" proof (cases opubx) case OpAddNumUbx then have "eval_AddNumUbx \ = Some z" using assms by simp then show ?thesis using OpAddNumUbx - apply (induction \ rule: eval_AddNumUbx.induct; simp) - by (simp add: box_num_def) + by (induction \ rule: eval_AddNumUbx.induct) simp_all qed -subsection \Abstract interpretation\ +subsection \Typing\ fun typeof_opubx :: "opubx \ type option list \ type option" where "typeof_opubx OpAddNumUbx = ([Some Ubx1, Some Ubx1], Some Ubx1)" lemma ubx_imp_typeof_opubx: "ubx opinl ts = Some opubx \ fst (typeof_opubx opubx) = ts" by (induction opinl ts rule: ubx.induct; simp) lemma typeof_opubx_correct: "typeof_opubx opubx = (map typeof xs, codomain) \ \y. eval_ubx opubx xs = Some y \ typeof y = codomain" proof (induction opubx) case OpAddNumUbx obtain n1 n2 where xs_def: "xs = [OpUbx1 n1, OpUbx1 n2]" and "codomain = Some Ubx1" using OpAddNumUbx[symmetric] by (auto dest!: typeof_unboxed_inversion) then show ?case by simp qed lemma typeof_opubx_complete: "eval_ubx opubx xs = Some y \ typeof_opubx opubx = (map typeof xs, typeof y)" proof (induction opubx) case OpAddNumUbx then show ?case by (auto elim: eval_AddNumUbx.elims) qed lemma typeof_opubx_ar: "length (fst (typeof_opubx opubx)) = ar (deinl (deubx opubx))" by (induction opubx; simp) interpretation op_OpUbx: nary_operations_ubx eval ar eval_inl inl isinl deinl - is_true is_false box_num unbox_num box_bool unbox_bool + DNil is_true is_false DNum unbox_num DBool unbox_bool eval_ubx ubx deubx typeof_opubx using ubx_invertible using eval_ubx_correct using eval_ubx_to_inl using ubx_imp_typeof_opubx using typeof_opubx_correct using typeof_opubx_complete using typeof_opubx_ar by (unfold_locales; (simp; fail)?) end \ No newline at end of file diff --git a/thys/Interpreter_Optimizations/Option_Extra.thy b/thys/Interpreter_Optimizations/Option_Extra.thy new file mode 100644 --- /dev/null +++ b/thys/Interpreter_Optimizations/Option_Extra.thy @@ -0,0 +1,85 @@ +theory Option_Extra + imports Main +begin + +fun ap_option (infixl "\" 60) where + "(Some f) \ (Some x) = Some (f x)" | + "_ \ _ = None" + +lemma ap_option_eq_Some_conv: "f \ x = Some y \ (\f' x'. f = Some f' \ x = Some x' \ y = f' x')" + by (cases f; simp; cases x; auto) + +definition ap_map_prod where + "ap_map_prod f g p \ Some Pair \ f (fst p) \ g (snd p)" + +lemma ap_map_prod_eq_Some_conv: + "ap_map_prod f g p = Some p' \ (\x y. p = (x, y) \ (\x' y'. p' = (x', y') \ f x = Some x' \ g y = Some y'))" +proof (cases p) + case (Pair x y) + then show ?thesis + unfolding ap_map_prod_def + by (auto simp add: ap_map_prod_def ap_option_eq_Some_conv) +qed + +fun ap_map_list :: "('a \ 'b option) \ 'a list \ 'b list option" where + "ap_map_list _ [] = Some []" | + "ap_map_list f (x # xs) = Some (#) \ f x \ ap_map_list f xs" + +lemma length_ap_map_list: "ap_map_list f xs = Some ys \ length ys = length xs" + by (induction xs arbitrary: ys) (auto simp add: ap_option_eq_Some_conv) + +lemma ap_map_list_imp_rel_option_map_of: + assumes "ap_map_list f xs = Some ys" and + "\x y. (x, y) \ set (zip xs ys) \ f x = Some y \ fst x = fst y" + shows "rel_option (\x y. f (k, x) = Some (k, y)) (map_of xs k) (map_of ys k)" + using assms +proof (induction xs arbitrary: ys) + case Nil + thus ?case by simp +next + case (Cons x xs) + from Cons.prems show ?case + apply (auto simp add: ap_option_eq_Some_conv option_rel_Some1 elim!: Cons.IH) + apply (metis prod.collapse) + apply (metis prod.collapse) + by blast +qed + +lemma ap_map_list_ap_map_prod_imp_rel_option_map_of: + assumes "ap_map_list (ap_map_prod Some f) xs = Some ys" + shows "rel_option (\x y. f x = Some y) (map_of xs k) (map_of ys k)" + using assms +proof (induction xs arbitrary: ys) + case Nil + thus ?case by simp +next + case (Cons x xs) + from Cons.prems show ?case + by (auto simp: ap_option_eq_Some_conv ap_map_prod_eq_Some_conv elim: Cons.IH[simplified]) +qed + +lemma ex_ap_map_list_eq_SomeI: + assumes "list_all (\x. \y. f x = Some y) xs" + shows "\ys. ap_map_list f xs = Some ys" + using assms + by (induction xs) auto + +lemma ap_map_list_iff_list_all2: + "ap_map_list f xs = Some ys \ list_all2 (\x y. f x = Some y) xs ys" +proof (rule iffI) + show "ap_map_list f xs = Some ys \ list_all2 (\x y. f x = Some y) xs ys" + by (induction xs arbitrary: ys) (auto simp: ap_option_eq_Some_conv) +next + show "list_all2 (\x y. f x = Some y) xs ys \ ap_map_list f xs = Some ys" + by (induction xs ys rule: list.rel_induct) auto +qed + +lemma ap_map_list_map_conv: + assumes + "ap_map_list f xs = Some ys" and + "\x y. x \ set xs \ f x = Some y \ y = g x" + shows "ys = map g xs" + using assms + by (induction xs arbitrary: ys) (auto simp: ap_option_eq_Some_conv) + +end \ No newline at end of file diff --git a/thys/Interpreter_Optimizations/Option_applicative.thy b/thys/Interpreter_Optimizations/Option_applicative.thy deleted file mode 100644 --- a/thys/Interpreter_Optimizations/Option_applicative.thy +++ /dev/null @@ -1,35 +0,0 @@ -theory Option_applicative - imports Main -begin - -context begin - -local_setup \ - Local_Theory.map_background_naming (Name_Space.add_path "Option") -\ - -fun pure where - "pure x = Some x" - -fun ap :: "('a \ 'b) option \ 'a option \ 'b option" (infixl "<*>" 51) where - "Some f <*> Some x = Some (f x)" | - "_ <*> _ = None" - -end - -lemma identity: "pure id <*> x = x" - by (cases x) simp_all - -lemma homomorphism: "pure f <*> pure x = pure (f x)" - by simp - -lemma interchange: "f <*> pure x = pure (\g. g x) <*> f" - by (cases f) simp_all - -lemma composition: "pure (\) <*> u <*> v <*> w = u <*> (v <*> w)" - apply (cases u, simp) - apply (cases v, simp) - apply (cases w, simp) - by simp - -end \ No newline at end of file diff --git a/thys/Interpreter_Optimizations/ROOT b/thys/Interpreter_Optimizations/ROOT --- a/thys/Interpreter_Optimizations/ROOT +++ b/thys/Interpreter_Optimizations/ROOT @@ -1,14 +1,14 @@ chapter AFP -session Interpreter_Optimizations (AFP) = VeriComp + +session Interpreter_Optimizations (AFP) = "HOL-Library" + options [timeout = 600] sessions - "HOL-Library" + VeriComp theories Env_list Inca_to_Ubx_compiler Op_example Std_to_Inca_compiler document_files "root.bib" "root.tex" diff --git a/thys/Interpreter_Optimizations/Std.thy b/thys/Interpreter_Optimizations/Std.thy --- a/thys/Interpreter_Optimizations/Std.thy +++ b/thys/Interpreter_Optimizations/Std.thy @@ -1,108 +1,136 @@ theory Std imports List_util Global Op Env Env_list Dynamic "VeriComp.Language" begin -datatype ('dyn, 'var, 'fun, 'op) instr = + +datatype ('dyn, 'var, 'fun, 'label, 'op) instr = IPush 'dyn | IPop | + IGet nat | + ISet nat | ILoad 'var | IStore 'var | IOp 'op | - ICJump nat | - ICall 'fun + ICJump 'label 'label | + ICall 'fun | + is_return: IReturn locale std = Fenv: env F_empty F_get F_add F_to_list + Henv: env heap_empty heap_get heap_add heap_to_list + - dynval is_true is_false + + dynval uninitialized is_true is_false + nary_operations \\ \\\\\ for + \ \Functions environment\ F_empty and - F_get :: "'fenv \ 'fun \ ('dyn, 'var, 'fun, 'op) instr fundef option" and + F_get :: "'fenv \ 'fun \ ('label, ('dyn, 'var, 'fun, 'label, 'op) instr) fundef option" and F_add and F_to_list and + + \ \Memory heap\ heap_empty and heap_get :: "'henv \ 'var \ 'dyn \ 'dyn option" and heap_add and heap_to_list and - is_true :: "'dyn \ bool" and is_false and + + \ \Dynamic values\ + uninitialized :: 'dyn and is_true and is_false and + + \ \n-ary operations\ \\ :: "'op \ 'dyn list \ 'dyn" and \\\\\ begin -inductive final :: "('fenv, 'henv, ('fun, 'dyn) frame) state \ bool" where - "F_get F f = Some fd \ pc = length (body fd) \ final (State F H [Frame f pc \])" - inductive step (infix "\" 55) where step_push: - "F_get F f = Some fd \ pc < length (body fd) \ body fd ! pc = IPush d \ - State F H (Frame f pc \ # st) \ State F H (Frame f (Suc pc) (d # \) # st)" | + "next_instr (F_get F) f l pc = Some (IPush d) \ + State F H (Frame f l pc R \ # st) \ State F H (Frame f l (Suc pc) R (d # \) # st)" | step_pop: - "F_get F f = Some fd \ pc < length (body fd) \ body fd ! pc = IPop \ - State F H (Frame f pc (d # \) # st) \ State F H (Frame f (Suc pc) \ # st)" | + "next_instr (F_get F) f l pc = Some IPop \ + State F H (Frame f l pc R (d # \) # st) \ State F H (Frame f l (Suc pc) R \ # st)" | + + step_get: + "next_instr (F_get F) f l pc = Some (IGet n) \ + n < length R \ d = R ! n \ + State F H (Frame f l pc R \ # st) \ State F H (Frame f l (Suc pc) R (d # \) # st)" | + + step_set: + "next_instr (F_get F) f l pc = Some (ISet n) \ + n < length R \ R' = R[n := d] \ + State F H (Frame f l pc R (d # \) # st) \ State F H (Frame f l (Suc pc) R' \ # st)" | step_load: - "F_get F f = Some fd \ pc < length (body fd) \ body fd ! pc = ILoad x \ + "next_instr (F_get F) f l pc = Some (ILoad x) \ heap_get H (x, y) = Some d \ - State F H (Frame f pc (y # \) # st) \ State F H (Frame f (Suc pc) (d # \) # st)" | + State F H (Frame f l pc R (y # \) # st) \ State F H (Frame f l (Suc pc) R (d # \) # st)" | step_store: - "F_get F f = Some fd \ pc < length (body fd) \ body fd ! pc = IStore x \ + "next_instr (F_get F) f l pc = Some (IStore x) \ heap_add H (x, y) d = H' \ - State F H (Frame f pc (y # d # \) # st) \ State F H' (Frame f (Suc pc) \ # st)" | + State F H (Frame f l pc R (y # d # \) # st) \ State F H' (Frame f l (Suc pc) R \ # st)" | step_op: - "F_get F f = Some fd \ pc < length (body fd) \ body fd ! pc = IOp op \ + "next_instr (F_get F) f l pc = Some (IOp op) \ \\\\\ op = ar \ ar \ length \ \ \\ op (take ar \) = x \ - State F H (Frame f pc \ # st) \ State F H (Frame f (Suc pc) (x # drop ar \) # st)" | + State F H (Frame f l pc R \ # st) \ State F H (Frame f l (Suc pc) R (x # drop ar \) # st)" | - step_cjump_true: - "F_get F f = Some fd \ pc < length (body fd) \ body fd ! pc = ICJump n \ - is_true d \ - State F H (Frame f pc (d # \) # st) \ State F H (Frame f n \ # st)" | + step_cjump: + "next_instr (F_get F) f l pc = Some (ICJump l\<^sub>t l\<^sub>f) \ + is_true d \ l' = l\<^sub>t \ is_false d \ l' = l\<^sub>f \ + State F H (Frame f l pc R (d # \) # st) \ State F H (Frame f l' 0 R \ # st)" | - step_cjump_false: - "F_get F f = Some fd \ pc < length (body fd) \ body fd ! pc = ICJump n \ - is_false d \ - State F H (Frame f pc (d # \) # st) \ State F H (Frame f (Suc pc) \ # st)" | - - step_fun_call: - "F_get F f = Some fd \ pc < length (body fd) \ body fd ! pc = ICall g \ + step_call: + "next_instr (F_get F) f l pc = Some (ICall g) \ F_get F g = Some gd \ arity gd \ length \ \ - frame\<^sub>f = Frame f pc \ \ frame\<^sub>g = Frame g 0 (take (arity gd) \) \ - State F H (frame\<^sub>f # st) \ State F H (frame\<^sub>g # frame\<^sub>f # st)" | - - step_fun_end: - "F_get F g = Some gd \ arity gd \ length \\<^sub>f \ pc\<^sub>g = length (body gd) \ - frame\<^sub>g = Frame g pc\<^sub>g \\<^sub>g \ frame\<^sub>f = Frame f pc\<^sub>f \\<^sub>f \ - frame\<^sub>f' = Frame f (Suc pc\<^sub>f) (\\<^sub>g @ drop (arity gd) \\<^sub>f) \ - State F H (frame\<^sub>g # frame\<^sub>f # st) \ State F H (frame\<^sub>f' # st)" + frame\<^sub>g = allocate_frame g gd (take (arity gd) \) uninitialized \ + State F H (Frame f l pc R \ # st) \ State F H (frame\<^sub>g # Frame f l pc R \ # st)" | -lemma step_deterministic: "s1 \ s2 \ s1 \ s3 \ s2 = s3" - by (induction rule: step.cases; - auto elim!: step.cases dest: is_true_and_is_false_implies_False) + step_return: + "next_instr (F_get F) g l\<^sub>g pc\<^sub>g = Some IReturn \ + F_get F g = Some gd \ arity gd \ length \\<^sub>f \ + length \\<^sub>g = return gd \ + frame\<^sub>f' = Frame f l\<^sub>f (Suc pc\<^sub>f) R\<^sub>f (\\<^sub>g @ drop (arity gd) \\<^sub>f) \ + State F H (Frame g l\<^sub>g pc\<^sub>g R\<^sub>g \\<^sub>g # Frame f l\<^sub>f pc\<^sub>f R\<^sub>f \\<^sub>f # st) \ State F H (frame\<^sub>f' # st)" -lemma final_finished: "final s \ finished step s" +lemma step_deterministic: + assumes "s1 \ s2" and "s1 \ s3" + shows "s2 = s3" + using assms + apply (cases s1 s2 rule: step.cases) + apply (auto elim: step.cases dest: is_true_and_is_false_implies_False) [3] + apply (auto elim: step.cases dest: is_true_and_is_false_implies_False) [3] + apply (auto elim: step.cases dest: is_true_and_is_false_implies_False) [3] + apply (auto elim: step.cases dest: is_true_and_is_false_implies_False) [1] + done + +lemma step_right_unique: "right_unique step" + using step_deterministic + by (rule right_uniqueI) + +lemma final_finished: + assumes "final F_get IReturn s" + shows "finished step s" unfolding finished_def -proof - assume "final s" and "\x. step s x" - then obtain s' where "step s s'" - by auto +proof (rule notI) + assume "\x. step s x" + then obtain s' where "step s s'" by auto thus False - using \final s\ - by (auto elim!: step.cases final.cases) + using \final F_get IReturn s\ + by (auto + simp: state_ok_def + wf_fundefs_imp_wf_fundef[THEN wf_fundef_last_basic_block, THEN next_instr_length_instrs[rotated]] + elim!: step.cases final.cases) qed -sublocale semantics step final +sublocale semantics step "final F_get IReturn" using final_finished step_deterministic by unfold_locales -inductive load :: "('fenv, 'a, 'fun) prog \ ('fenv, 'a, ('fun, 'b) frame) state \ bool" where - "F_get F main = Some fd \ arity fd = 0 \ - load (Prog F H main) (State F H [Frame main 0 []])" +definition load where + "load \ Global.load F_get uninitialized" -sublocale language step final load +sublocale language step "final F_get IReturn" load by unfold_locales end end \ No newline at end of file diff --git a/thys/Interpreter_Optimizations/Std_to_Inca_compiler.thy b/thys/Interpreter_Optimizations/Std_to_Inca_compiler.thy --- a/thys/Interpreter_Optimizations/Std_to_Inca_compiler.thy +++ b/thys/Interpreter_Optimizations/Std_to_Inca_compiler.thy @@ -1,93 +1,277 @@ theory Std_to_Inca_compiler imports Std_to_Inca_simulation "VeriComp.Compiler" begin fun compile_instr where "compile_instr (Std.IPush d) = Inca.IPush d" | "compile_instr Std.IPop = Inca.IPop" | + "compile_instr (Std.IGet n) = Inca.IGet n" | + "compile_instr (Std.ISet n) = Inca.ISet n" | "compile_instr (Std.ILoad x) = Inca.ILoad x" | "compile_instr (Std.IStore x) = Inca.IStore x" | "compile_instr (Std.IOp op) = Inca.IOp op" | - "compile_instr (Std.ICJump n) = Inca.ICJump n" | - "compile_instr (Std.ICall f) = Inca.ICall f" + "compile_instr (Std.ICJump l\<^sub>t l\<^sub>f) = Inca.ICJump l\<^sub>t l\<^sub>f" | + "compile_instr (Std.ICall f) = Inca.ICall f" | + "compile_instr Std.IReturn = Inca.IReturn" fun compile_fundef where - "compile_fundef (Fundef xs ar) = Fundef (map compile_instr xs) ar" + "compile_fundef (Fundef [] _ _ _) = None" | + "compile_fundef (Fundef bblocks ar ret locals) = + (if distinct (map fst bblocks) then + Some (Fundef (map_ran (\_. map compile_instr) bblocks) ar ret locals) + else + None)" context std_inca_simulation begin +lemma lambda_eq_eq[simp]: "(\x y. y = x) = (=)" + apply (intro ext) + by blast + lemma norm_compile_instr: "norm_instr (compile_instr instr) = instr" by (cases instr) auto -lemma rel_compile_fundef: "rel_fundef norm_eq fd (compile_fundef fd)" -proof (cases fd) - case (Fundef xs ar) +lemma rel_compile_fundef: + assumes "compile_fundef fd1 = Some fd2" + shows "rel_fundef (=) norm_eq fd1 fd2" + using assms +proof (cases rule: compile_fundef.elims) + case (2 x xs ar) then show ?thesis - by (simp add: list.rel_map list.rel_eq norm_compile_instr) + by (auto intro: list.rel_refl + simp: const_eq_if_conv list.rel_map norm_compile_instr map_ran_def case_prod_beta + rel_prod_sel) +qed simp_all + +lemma rel_fundef_imp_fundef_ok_iff: + assumes "rel_fundef (=) norm_eq fd1 fd2" + shows "wf_fundef fd1 \ wf_fundef fd2" + unfolding wf_fundef_def + using assms + by (auto simp: list_all2_rel_prod_conv list.rel_eq elim: fundef.rel_cases) + +lemma rel_fundefs_imp_wf_fundefs_iff: + assumes rel_f_g: "rel_fundefs f g" + shows "wf_fundefs f \ wf_fundefs g" +proof (rule iffI) + assume "wf_fundefs f" show "wf_fundefs g" + proof (rule wf_fundefsI) + fix x fd + assume "g x = Some fd" thus "wf_fundef fd" + using \wf_fundefs f\ + by (meson rel_f_g rel_fundef_imp_fundef_ok_iff rel_fundefs_Some2 wf_fundefs_imp_wf_fundef) + qed +next + show "wf_fundefs g \ wf_fundefs f" + using rel_f_g + by (meson rel_fundef_imp_fundef_ok_iff rel_fundefs_Some1 wf_fundefs_imp_wf_fundef wf_fundefsI) qed +lemma compile_fundef_wf: + assumes "compile_fundef fd = Some fd'" + shows "wf_fundef fd'" + using assms +proof (cases rule: compile_fundef.elims) + case (2 x xs ar) + then show ?thesis + by (simp add: const_eq_if_conv wf_fundef_def map_ran_Cons_sel) +qed simp_all + definition compile_env where - "compile_env e = Sinca.Fenv.from_list (map (map_prod id compile_fundef) (Fstd_to_list e))" + "compile_env e \ + Some Sinca.Fenv.from_list \ ap_map_list (ap_map_prod Some compile_fundef) (Fstd_to_list e)" -lemma Finca_get_compile: "Finca_get (compile_env e) x = map_option compile_fundef (Fstd_get e x)" - using Sinca.Fenv.from_list_correct[symmetric] Sstd.Fenv.to_list_correct - by (simp add: compile_env_def map_prod_def map_of_map) +lemma compile_env_imp_rel_option: + assumes "compile_env F1 = Some F2" + shows "rel_option (\fd1 fd2. compile_fundef fd1 = Some fd2) (Fstd_get F1 f) (Finca_get F2 f)" +proof - + from assms(1) obtain xs where + ap_map_list_F1: "ap_map_list (ap_map_prod Some compile_fundef) (Fstd_to_list F1) = Some xs" and + F2_def: "F2 = Sinca.Fenv.from_list xs" + by (auto simp add: compile_env_def ap_option_eq_Some_conv) -lemma rel_fundefs_compile_env: "rel_fundefs (Fstd_get e) (Finca_get (compile_env e))" + show ?thesis + unfolding F2_def Sinca.Fenv.from_list_correct[symmetric] + unfolding Sinca.Fenv.from_list_correct + unfolding Sstd.Fenv.to_list_correct[symmetric] + using ap_map_list_F1 + by (simp add: ap_map_list_ap_map_prod_imp_rel_option_map_of) +qed + +lemma Finca_get_compile: + assumes compile_F1: "compile_env F1 = Some F2" + shows "Finca_get F2 f = Fstd_get F1 f \ compile_fundef" + using compile_env_imp_rel_option[OF assms(1), of f] + by (auto elim!: option.rel_cases) + +lemma compile_env_rel_fundefs: + assumes "compile_env F1 = Some F2" + shows "rel_fundefs (Fstd_get F1) (Finca_get F2)" unfolding rel_fundefs_def - unfolding Finca_get_compile - unfolding option.rel_map - by (auto intro: option.rel_refl rel_compile_fundef) +proof (rule allI) + fix f + show "rel_option (rel_fundef (=) norm_eq) (Fstd_get F1 f) (Finca_get F2 f)" + using compile_env_imp_rel_option[OF assms(1), of f] + by (auto elim!: option.rel_cases intro: rel_compile_fundef) +qed + +lemma compile_env_imp_wf_fundefs2: + assumes "compile_env F1 = Some F2" + shows "wf_fundefs (Finca_get F2)" + unfolding Finca_get_compile[OF assms(1)] + by (auto simp: bind_eq_Some_conv intro!: wf_fundefsI intro: compile_fundef_wf) fun compile where - "compile (Prog F H main) = Some (Prog (compile_env F) H main)" + "compile (Prog F1 H f) = Some Prog \ compile_env F1 \ Some H \ Some f" theorem compile_load: - assumes "compile p1 = Some p2" and "Sinca.load p2 s2" + assumes + compile_p1: "compile p1 = Some p2" and + load: "Sinca.load p2 s2" shows "\s1. Sstd.load p1 s1 \ match s1 s2" proof (cases p1) - case (Prog F H main) - then have p2_def: "p2 = Prog (compile_env F) H main" - using assms(1) by simp + case (Prog F1 H main) + with assms(1) obtain F2 where + compile_F1: "compile_env F1 = Some F2" and + p2_def: "p2 = Prog F2 H main" + by (auto simp: ap_option_eq_Some_conv) show ?thesis - using assms(2) unfolding p2_def - proof (cases _ s2 rule: Sinca.load.cases) - case (1 fd') - let ?s1 = "State F H [Frame main 0 []]" + using assms(2) unfolding p2_def Sinca.load_def + proof (cases _ _ _ s2 rule: load.cases) + case (1 fd2) + from 1 obtain fd1 where + F1_main: "Fstd_get F1 main = Some fd1" and compile_fd1: "compile_fundef fd1 = Some fd2" + using Finca_get_compile[OF compile_F1] by (auto simp: bind_eq_Some_conv) - from 1 obtain fd where - Fstd_main: "Fstd_get F main = Some fd" and "fd' = compile_fundef fd" - using Finca_get_compile by auto - with 1 have "arity fd = 0" - by (metis fundef.rel_sel rel_compile_fundef) - hence "Sstd.load p1 ?s1" - unfolding Prog - by (auto intro: Sstd.load.intros[OF Fstd_main]) - moreover have "?s1 \ s2" - unfolding 1 - using rel_fundefs_compile_env - by (auto intro!: match.intros) - ultimately show ?thesis by auto + note rel_fd1_fd2 = rel_compile_fundef[OF compile_fd1] + hence first_label_fd1_fd2: "fst (hd (body fd1)) = fst (hd (body fd2))" + using 1 + by (auto elim!: fundef.rel_cases dest: list_all2_rel_prod_fst_hd) + + have bodies_fd1_fd2: "body fd1 = [] \ body fd2 = []" + using rel_fundef_body_length[OF rel_fd1_fd2] + by auto + + let ?s1 = "State F1 H [allocate_frame main fd1 [] uninitialized]" + + show ?thesis + proof (intro exI conjI) + show "Sstd.load p1 ?s1" + unfolding Prog + using 1 F1_main rel_fd1_fd2 bodies_fd1_fd2 + by (auto simp: rel_fundef_arities Sstd.load_def intro!: load.intros) + next + note rel_F1_F2 = compile_env_rel_fundefs[OF compile_F1] + moreover have "wf_fundefs (Fstd_get F1)" + proof (rule wf_fundefsI) + fix f fd1 + assume "Fstd_get F1 f = Some fd1" + then show "wf_fundef fd1" + by (metis (mono_tags, lifting) Finca_get_compile bind.bind_lunit compile_F1 + rel_F1_F2 compile_fundef_wf rel_fundef_imp_fundef_ok_iff + rel_fundefs_Some1) + qed + ultimately show "match ?s1 s2" + unfolding 1 + using rel_fd1_fd2 + by (auto simp: allocate_frame_def first_label_fd1_fd2 rel_fundef_locals + intro: match.intros) + qed qed qed sublocale std_to_inca_compiler: - compiler Sstd.step Sinca.step Sstd.final Sinca.final Sstd.load Sinca.load - "\_ _. False" "\_. match" compile -using compile_load - by unfold_locales auto + compiler where + step1 = Sstd.step and final1 = "final Fstd_get Std.IReturn" and load1 = "Sstd.load" and + step2 = Sinca.step and final2 = "final Finca_get Inca.IReturn" and load2 = "Sinca.load" and + order = "\_ _. False" and match = "\_. match" and compile = compile + using compile_load + by unfold_locales auto + +section \Completeness of compilation\ + +lemma compile_fundef_complete: + assumes "wf_fundef fd1" + shows "\fd2. compile_fundef fd1 = Some fd2" +proof (cases fd1) + case (Fundef bbs ar) + then obtain bb bbs' where bbs_def: "bbs = bb # bbs'" + using assms(1) by (cases bbs; auto) + show ?thesis + using assms(1)[THEN wf_fundef_distinct_basic_blocksD] + unfolding Fundef bbs_def + by simp +qed + +lemma compile_env_complete: + assumes wf_F1: "wf_fundefs (Fstd_get F1)" + shows "\F2. compile_env F1 = Some F2" +proof - + show ?thesis + using wf_F1 + by (auto simp: compile_env_def ap_option_eq_Some_conv ap_map_prod_eq_Some_conv + intro!: ex_ap_map_list_eq_SomeI Sstd.Fenv.to_list_list_allI compile_fundef_complete + intro: wf_fundefs_imp_wf_fundef) +qed theorem compile_complete: - assumes "Sstd.load p\<^sub>1 s\<^sub>1" - shows "\p\<^sub>2 s\<^sub>2. compile p\<^sub>1 = Some p\<^sub>2 \ Sinca.load p\<^sub>2 s\<^sub>2 \ match s\<^sub>1 s\<^sub>2" - using assms - by (auto elim!: Sstd.load.cases - intro!: match.intros rel_fundefs_compile_env - simp add: Sinca.load.simps Finca_get_compile rel_fundef_arities[OF rel_compile_fundef]) + assumes + wf_p1: "wf_prog Fstd_get p1" and load_p1: "Sstd.load p1 s1" + shows "\p2 s2. compile p1 = Some p2 \ Sinca.load p2 s2 \ match s1 s2" +proof - + from load_p1 obtain F1 H main fd1 where + p1_def: "p1 = Prog F1 H main" and + s1_def: "s1 = Global.state.State F1 H [allocate_frame main fd1 [] uninitialized]" and + F1_main: "Fstd_get F1 main = Some fd1" and + arity_fd1: "arity fd1 = 0" + by (auto simp: Sstd.load_def elim: load.cases) + + obtain F2 where compile_F1: "compile_env F1 = Some F2" + using wf_p1[unfolded p1_def wf_prog_def, simplified] + using compile_env_complete[of F1] + by auto + + obtain fd2 where + F2_main: "Finca_get F2 main = Some fd2" and compile_fd1: "compile_fundef fd1 = Some fd2" + using compile_env_imp_rel_option[OF compile_F1, of main] + unfolding F1_main option_rel_Some1 + by auto + + hence rel_fd1_fd2: "rel_fundef (=) norm_eq fd1 fd2" + using rel_compile_fundef by auto + + note wf_fd2 = compile_fundef_wf[OF compile_fd1] + + let ?s2 = "State F2 H [allocate_frame main fd2 [] uninitialized]" + + show ?thesis + proof (intro exI conjI) + show "compile p1 = Some (Prog F2 H main)" + by (simp add: p1_def compile_F1) + next + show "Sinca.load (Prog F2 H main) ?s2" + using F2_main arity_fd1 + using rel_fundef_arities[OF rel_fd1_fd2] + using compile_fundef_wf[OF compile_fd1] + by (auto simp: wf_fundef_def Sinca.load_def intro: load.intros) + next + have "fst (hd (body fd1)) = fst (hd (body fd2))" + using rel_fd1_fd2 wf_fd2 + by (auto simp add: fundef.rel_sel wf_fundef_def list_all2_rel_prod_fst_hd) + moreover have "wf_fundefs (Fstd_get F1)" + using compile_env_imp_wf_fundefs2[OF compile_F1] + using compile_env_rel_fundefs[OF compile_F1, THEN rel_fundefs_imp_wf_fundefs_iff] + by simp + ultimately show "match s1 ?s2" + unfolding s1_def + using rel_fd1_fd2 compile_env_rel_fundefs[OF compile_F1] + by (auto simp: allocate_frame_def rel_fundef_locals intro!: match.intros) + qed +qed end end \ No newline at end of file diff --git a/thys/Interpreter_Optimizations/Std_to_Inca_simulation.thy b/thys/Interpreter_Optimizations/Std_to_Inca_simulation.thy --- a/thys/Interpreter_Optimizations/Std_to_Inca_simulation.thy +++ b/thys/Interpreter_Optimizations/Std_to_Inca_simulation.thy @@ -1,629 +1,681 @@ theory Std_to_Inca_simulation imports Global List_util Std Inca "VeriComp.Simulation" begin section \Generic definitions\ -print_locale std - locale std_inca_simulation = Sstd: std Fstd_empty Fstd_get Fstd_add Fstd_to_list heap_empty heap_get heap_add heap_to_list - is_true is_false + uninitialized is_true is_false \\ \\\\\ + Sinca: inca Finca_empty Finca_get Finca_add Finca_to_list heap_empty heap_get heap_add heap_to_list - is_true is_false + uninitialized is_true is_false \\ \\\\\ \\\\\ \\\ \\\\\ \
\\\\ for \ \Functions environments\ Fstd_empty and - Fstd_get :: "'fenv_std \ 'fun \ ('dyn, 'var, 'fun, 'op) Std.instr fundef option" and + Fstd_get :: "'fenv_std \ 'fun \ ('label, ('dyn, 'var, 'fun, 'label, 'op) Std.instr) fundef option" and Fstd_add and Fstd_to_list and Finca_empty and - Finca_get :: "'fenv_inca \ 'fun \ ('dyn, 'var, 'fun, 'op, 'opinl) Inca.instr fundef option" and + Finca_get :: "'fenv_inca \ 'fun \ ('label, ('dyn, 'var, 'fun, 'label, 'op, 'opinl) Inca.instr) fundef option" and Finca_add and Finca_to_list and \ \Memory heap\ - heap_empty and heap_get :: "'henv \ 'var \ 'dyn \ 'dyn option" and heap_add and heap_to_list and + heap_empty and + heap_get :: "'henv \ 'var \ 'dyn \ 'dyn option" and + heap_add and heap_to_list and \ \Dynamic values\ - is_true :: "'dyn \ bool" and is_false and + uninitialized :: 'dyn and is_true and is_false and \ \n-ary operations\ \\ :: "'op \ 'dyn list \ 'dyn" and \\\\\ and \\\\\ and \\\ and \\\\\ and \
\\\\ :: "'opinl \ 'op" begin fun norm_instr where "norm_instr (Inca.IPush d) = Std.IPush d" | "norm_instr Inca.IPop = Std.IPop" | + "norm_instr (Inca.IGet n) = Std.IGet n" | + "norm_instr (Inca.ISet n) = Std.ISet n" | "norm_instr (Inca.ILoad x) = Std.ILoad x" | "norm_instr (Inca.IStore x) = Std.IStore x" | "norm_instr (Inca.IOp op) = Std.IOp op" | "norm_instr (Inca.IOpInl opinl) = Std.IOp (\
\\\\ opinl)" | - "norm_instr (Inca.ICJump n) = Std.ICJump n" | - "norm_instr (Inca.ICall x) = Std.ICall x" + "norm_instr (Inca.ICJump l\<^sub>t l\<^sub>f) = Std.ICJump l\<^sub>t l\<^sub>f" | + "norm_instr (Inca.ICall x) = Std.ICall x" | + "norm_instr Inca.IReturn = Std.IReturn" -abbreviation (input) norm_eq where - "norm_eq x y \ x = norm_instr y" +abbreviation norm_eq where + "norm_eq x y \ norm_instr y = x" definition rel_fundefs where - "rel_fundefs f g = (\x. rel_option (rel_fundef (\x y. x = norm_instr y)) (f x) (g x))" + "rel_fundefs f g = (\x. rel_option (rel_fundef (=) norm_eq) (f x) (g x))" + +lemma rel_fundefsI: + assumes "\x. rel_option (rel_fundef (=) norm_eq) (F1 x) (F2 x)" + shows "rel_fundefs F1 F2" + using assms + by (simp add: rel_fundefs_def) + +lemma rel_fundefsD: + assumes "rel_fundefs F1 F2" + shows "rel_option (rel_fundef (=) norm_eq) (F1 x) (F2 x)" + using assms + by (simp add: rel_fundefs_def) + +lemma rel_fundefs_next_instr: + assumes rel_F1_F2: "rel_fundefs F1 F2" + shows "rel_option norm_eq (next_instr F1 f l pc) (next_instr F2 f l pc)" +proof - + have "rel_option (rel_fundef (=) norm_eq) (F1 f) (F2 f)" + using rel_F1_F2[unfolded rel_fundefs_def] by simp + thus ?thesis + proof (cases rule: option.rel_cases) + case None + then show ?thesis by (simp add: next_instr_def) + next + case (Some fd1 fd2) + hence "rel_option (list_all2 norm_eq) (map_of (body fd1) l) (map_of (body fd2) l)" + by (auto elim!: fundef.rel_cases intro: rel_option_map_of) + then show ?thesis + by (cases rule: option.rel_cases; + simp add: next_instr_def instr_at_def Some list_all2_conv_all_nth) + qed +qed + +lemma rel_fundefs_next_instr1: + assumes rel_F1_F2: "rel_fundefs F1 F2" and next_instr1: "next_instr F1 f l pc = Some instr1" + shows "\instr2. next_instr F2 f l pc = Some instr2 \ norm_eq instr1 instr2" + using rel_fundefs_next_instr[OF rel_F1_F2, of f l pc] + unfolding next_instr1 + unfolding option_rel_Some1 + by assumption + +lemma rel_fundefs_next_instr2: + assumes rel_F1_F2: "rel_fundefs F1 F2" and next_instr2: "next_instr F2 f l pc = Some instr2" + shows "\instr1. next_instr F1 f l pc = Some instr1 \ norm_eq instr1 instr2" + using rel_fundefs_next_instr[OF rel_F1_F2, of f l pc] + unfolding next_instr2 + unfolding option_rel_Some2 + by assumption + +lemma rel_fundefs_empty: "rel_fundefs (\_. None) (\_. None)" + by (simp add: rel_fundefs_def) + +lemma rel_fundefs_None1: + assumes "rel_fundefs f g" and "f x = None" + shows "g x = None" + by (metis assms rel_fundefs_def rel_option_None1) + +lemma rel_fundefs_None2: + assumes "rel_fundefs f g" and "g x = None" + shows "f x = None" + by (metis assms rel_fundefs_def rel_option_None2) lemma rel_fundefs_Some1: assumes "rel_fundefs f g" and "f x = Some y" - shows "\z. g x = Some z \ rel_fundef norm_eq y z" + shows "\z. g x = Some z \ rel_fundef (=) norm_eq y z" proof - - from assms(1) have "rel_option (rel_fundef norm_eq) (f x) (g x)" + from assms(1) have "rel_option (rel_fundef (=) norm_eq) (f x) (g x)" unfolding rel_fundefs_def by simp with assms(2) show ?thesis by (simp add: option_rel_Some1) qed lemma rel_fundefs_Some2: assumes "rel_fundefs f g" and "g x = Some y" - shows "\z. f x = Some z \ rel_fundef norm_eq z y" + shows "\z. f x = Some z \ rel_fundef (=) norm_eq z y" proof - - from assms(1) have "rel_option (rel_fundef norm_eq) (f x) (g x)" + from assms(1) have "rel_option (rel_fundef (=) norm_eq) (f x) (g x)" unfolding rel_fundefs_def by simp with assms(2) show ?thesis by (simp add: option_rel_Some2) qed -lemma rel_fundef_body_nth: - assumes "rel_fundef norm_eq fd1 fd2" and "pc < length (body fd1)" - shows "body fd1 ! pc = norm_instr (body fd2 ! pc)" - using assms - by (auto dest: list_all2_nthD simp: fundef.rel_sel) - -lemma rel_fundef_rewrite_body: - assumes - "rel_fundef norm_eq fd1 fd2" and - "norm_instr (body fd2 ! pc) = norm_instr instr" - shows "rel_fundef norm_eq fd1 (rewrite_fundef_body fd2 pc instr)" - using assms(1) -proof (cases rule: fundef.rel_cases) - case (Fundef xs ar' ys ar) - hence "length xs = length ys" - by (simp add: list_all2_conv_all_nth) - hence "length xs = length (rewrite ys pc instr)" - by simp - hence "list_all2 norm_eq xs (rewrite ys pc instr)" - proof (elim list_all2_all_nthI) - fix n - assume "n < length xs" - hence "n < length ys" - by (simp add: \length xs = length ys\) - thus "xs ! n = norm_instr (rewrite ys pc instr ! n)" - using list_all2_nthD[OF \list_all2 norm_eq xs ys\ \n < length xs\, symmetric] - using assms(2)[unfolded Fundef(2), simplified] - by (cases "pc = n"; simp) - qed - thus ?thesis - using Fundef by simp +lemma rel_fundefs_rel_option: + assumes "rel_fundefs f g" and "\x y. rel_fundef (=) norm_eq x y \ h x y" + shows "rel_option h (f z) (g z)" +proof - + have "rel_option (rel_fundef (=) norm_eq) (f z) (g z)" + using assms(1)[unfolded rel_fundefs_def] by simp + then show ?thesis + unfolding rel_option_unfold + by (auto simp add: assms(2)) qed -lemma rel_fundefs_rewrite: +lemma rel_fundefs_rewriteI2: assumes rel_F1_F2: "rel_fundefs (Fstd_get F1) (Finca_get F2)" and - F2_get_f: "Finca_get F2 f = Some fd2" and - F2_add_f: "Finca_add F2 f (rewrite_fundef_body fd2 pc instr) = F2'" and - norm_eq: "norm_instr (body fd2 ! pc) = norm_instr instr" - shows "rel_fundefs (Fstd_get F1) (Finca_get F2')" - unfolding rel_fundefs_def -proof + next_instr1: "next_instr (Fstd_get F1) f l pc = Some instr1" + "norm_eq instr1 instr2'" + shows "rel_fundefs (Fstd_get F1) + (Finca_get (Sinca.Fenv.map_entry F2 f (\fd. rewrite_fundef_body fd l pc instr2')))" + (is "rel_fundefs (Fstd_get ?F1') (Finca_get ?F2')") +proof (rule rel_fundefsI) fix x - show "rel_option (rel_fundef norm_eq) (Fstd_get F1 x) (Finca_get F2' x)" - proof (cases "x = f") + show "rel_option (rel_fundef (=) norm_eq) (Fstd_get ?F1' x) (Finca_get ?F2' x)" + proof (cases "f = x") case True - then have F2'_get_x: "Finca_get F2' x = Some (rewrite_fundef_body fd2 pc instr)" - using F2_add_f by auto - obtain fd1 where "Fstd_get F1 f = Some fd1" and "rel_fundef norm_eq fd1 fd2" - using rel_fundefs_Some2[OF rel_F1_F2 F2_get_f] by auto - thus ?thesis - unfolding F2'_get_x option_rel_Some2 - using True rel_fundef_rewrite_body[OF _ norm_eq] - by auto + show ?thesis + using rel_F1_F2[THEN rel_fundefsD, of x] True next_instr1 assms(3) + by (cases rule: option.rel_cases) + (auto intro!: rel_fundef_rewrite_body2' dest!: next_instrD) next case False - then have "Finca_get F2' x = Finca_get F2 x" - using F2_add_f by auto then show ?thesis - using rel_F1_F2 rel_fundefs_def - by fastforce + using rel_F1_F2[THEN rel_fundefsD, of x] by simp qed qed - section \Simulation relation\ inductive match (infix "\" 55) where - "rel_fundefs (Fstd_get F1) (Finca_get F2) \ (State F1 H st) \ (State F2 H st)" + "wf_fundefs (Fstd_get F1) \ + rel_fundefs (Fstd_get F1) (Finca_get F2) \ + (State F1 H st) \ (State F2 H st)" + section \Backward simulation\ lemma backward_lockstep_simulation: - assumes "Sinca.step s2 s2'" and "s1 \ s2" - shows "\s1'. Sstd.step s1 s1' \ s1' \ s2'" + assumes "Sinca.step s2 s2'" and "match s1 s2" + shows "\s1'. Sstd.step s1 s1' \ match s1' s2'" proof - from assms(2) obtain F1 F2 H st where s1_def: "s1 = State F1 H st" and s2_def: "s2 = State F2 H st" and + ok_F1: "wf_fundefs (Fstd_get F1)" and rel_F1_F2: "rel_fundefs (Fstd_get F1) (Finca_get F2)" by (auto elim: match.cases) + + note rel_F1_F2_next_instr = rel_fundefs_next_instr[OF rel_F1_F2] + from assms(1) show "?thesis" unfolding s1_def s2_def proof (induction "State F2 H st" s2' rule: Sinca.step.induct) - case (step_push f fd2 pc d \ st') - then obtain fd1 where "Fstd_get F1 f = Some fd1" and "rel_fundef norm_eq fd1 fd2" - using rel_fundefs_Some2[OF rel_F1_F2] by blast - show ?case (is "\x. ?STEP x \ ?MATCH x") - proof - - let ?s1' = "State F1 H (Frame f (Suc pc) (d # \) # st')" - have "?STEP ?s1'" - using step_push \Fstd_get F1 f = Some fd1\ \rel_fundef norm_eq fd1 fd2\ - by (auto intro!: Sstd.step_push simp: rel_fundef_body_nth) - moreover have "?MATCH ?s1'" - using rel_F1_F2 by (auto intro: match.intros) - ultimately show "?thesis" by blast - qed - next - case (step_pop f fd2 pc d \ st') - then obtain fd1 where "Fstd_get F1 f = Some fd1" and "rel_fundef norm_eq fd1 fd2" - using rel_fundefs_Some2[OF rel_F1_F2] by blast - show ?case (is "\x. ?STEP x \ ?MATCH x") - proof - - let ?s1' = "State F1 H (Frame f (Suc pc) \ # st')" - have "?STEP ?s1'" - using step_pop \Fstd_get F1 f = Some fd1\ \rel_fundef norm_eq fd1 fd2\ - by (auto intro!: Sstd.step_pop simp: rel_fundef_body_nth) - moreover have "?MATCH ?s1'" - using rel_F1_F2 by (auto intro: match.intros) - ultimately show "?thesis" by blast - qed - next - case (step_load f fd2 pc x y d \ st') - then obtain fd1 where "Fstd_get F1 f = Some fd1" and "rel_fundef norm_eq fd1 fd2" - using rel_fundefs_Some2[OF rel_F1_F2] by blast + case (step_push f l pc d R \ st') show ?case (is "\x. ?STEP x \ ?MATCH x") proof - - let ?s1' = "State F1 H (Frame f (Suc pc) (d # \) # st')" + let ?s1' = "State F1 H (Frame f l (Suc pc) R (d # \) # st')" have "?STEP ?s1'" - using step_load \Fstd_get F1 f = Some fd1\ \rel_fundef norm_eq fd1 fd2\ - by (auto intro!: Sstd.step_load simp: rel_fundef_body_nth) - moreover have "?MATCH ?s1'" - using rel_F1_F2 by (auto intro: match.intros) - ultimately show "?thesis" by blast - qed - next - case (step_store f fd2 pc x y d H' \ st') - then obtain fd1 where "Fstd_get F1 f = Some fd1" and "rel_fundef norm_eq fd1 fd2" - using rel_fundefs_Some2[OF rel_F1_F2] by blast - show ?case (is "\x. ?STEP x \ ?MATCH x") - proof - - let ?s1' = "State F1 H' (Frame f (Suc pc) \ # st')" - have "?STEP ?s1'" - using step_store \Fstd_get F1 f = Some fd1\ \rel_fundef norm_eq fd1 fd2\ - by (auto intro!: Sstd.step_store simp: rel_fundef_body_nth) + using step_push rel_F1_F2_next_instr[of f l pc] + by (auto intro!: Sstd.step_push simp: option_rel_Some2) moreover have "?MATCH ?s1'" - using rel_F1_F2 by (auto intro: match.intros) - ultimately show "?thesis" by blast - qed - next - case (step_op f fd2 pc op ar \ x st') - then obtain fd1 where "Fstd_get F1 f = Some fd1" and "rel_fundef norm_eq fd1 fd2" - using rel_fundefs_Some2[OF rel_F1_F2] by blast - show ?case (is "\x. ?STEP x \ ?MATCH x") - proof - - let ?s1' = "State F1 H (Frame f (Suc pc) (x # drop ar \) # st')" - have "?STEP ?s1'" - using step_op \Fstd_get F1 f = Some fd1\ \rel_fundef norm_eq fd1 fd2\ - by (auto intro!: Sstd.step_op simp: rel_fundef_body_nth) - moreover have "?MATCH ?s1'" - using rel_F1_F2 by (auto intro: match.intros) - ultimately show "?thesis" by blast - qed - next - case (step_op_inl f fd2 pc op ar \ opinl x F2' st') - then obtain fd1 where "Fstd_get F1 f = Some fd1" and "rel_fundef norm_eq fd1 fd2" - using rel_fundefs_Some2[OF rel_F1_F2] by blast - show ?case (is "\x. ?STEP x \ ?MATCH x") - proof - - let ?s1' = "State F1 H (Frame f (Suc pc) (x # drop ar \) # st')" - have "?STEP ?s1'" - using step_op_inl \Fstd_get F1 f = Some fd1\ \rel_fundef norm_eq fd1 fd2\ - using Sinca.\\\\\_correct Sinca.\\\_invertible - by (auto intro!: Sstd.step_op simp: rel_fundef_body_nth) - moreover have "?MATCH ?s1'" - using step_op_inl Sinca.\\\_invertible - by (auto intro!: match.intros rel_fundefs_rewrite[OF rel_F1_F2]) + using ok_F1 rel_F1_F2 by (auto intro: match.intros) ultimately show "?thesis" by blast qed next - case (step_op_inl_hit f fd2 pc opinl ar \ x st') - then obtain fd1 where "Fstd_get F1 f = Some fd1" and "rel_fundef norm_eq fd1 fd2" - using rel_fundefs_Some2[OF rel_F1_F2] by blast - show ?case (is "\x. ?STEP x \ ?MATCH x") - proof - - let ?s1' = "State F1 H (Frame f (Suc pc) (x # drop ar \) # st')" - have "?STEP ?s1'" - using step_op_inl_hit \Fstd_get F1 f = Some fd1\ \rel_fundef norm_eq fd1 fd2\ Sinca.\\\\\_correct - by (auto intro!: Sstd.step_op simp: rel_fundef_body_nth) - moreover have "?MATCH ?s1'" - using rel_F1_F2 by (auto intro: match.intros) - ultimately show "?thesis" by blast - qed - next - case (step_op_inl_miss f fd2 pc opinl ar \ x F' st') - then obtain fd1 where "Fstd_get F1 f = Some fd1" and "rel_fundef norm_eq fd1 fd2" - using rel_fundefs_Some2[OF rel_F1_F2] by blast + case (step_pop f l pc R d \ st') show ?case (is "\x. ?STEP x \ ?MATCH x") proof - - let ?s1' = "State F1 H (Frame f (Suc pc) (x # drop ar \) # st')" + let ?s1' = "State F1 H (Frame f l (Suc pc) R \ # st')" have "?STEP ?s1'" - using step_op_inl_miss \Fstd_get F1 f = Some fd1\ \rel_fundef norm_eq fd1 fd2\ Sinca.\\\\\_correct - by (auto intro!: Sstd.step_op simp: rel_fundef_body_nth) + using step_pop rel_F1_F2_next_instr[of f l pc] + by (auto intro!: Sstd.step_pop simp: option_rel_Some2) moreover have "?MATCH ?s1'" - using step_op_inl_miss rel_fundefs_rewrite[OF rel_F1_F2] - by (auto intro: match.intros) - ultimately show "?thesis" by blast - qed - next - case (step_cjump_true f fd2 pc n d \ st') - then obtain fd1 where "Fstd_get F1 f = Some fd1" and "rel_fundef norm_eq fd1 fd2" - using rel_fundefs_Some2[OF rel_F1_F2] by blast - show ?case (is "\x. ?STEP x \ ?MATCH x") - proof - - let ?s1' = "State F1 H (Frame f n \ # st')" - have "?STEP ?s1'" - using step_cjump_true \Fstd_get F1 f = Some fd1\ \rel_fundef norm_eq fd1 fd2\ - by (auto intro!: Sstd.step_cjump_true simp: rel_fundef_body_nth) - moreover have "?MATCH ?s1'" - using rel_F1_F2 by (auto intro: match.intros) + using ok_F1 rel_F1_F2 by (auto intro: match.intros) ultimately show "?thesis" by blast qed next - case (step_cjump_false f fd2 pc n d \ st') - then obtain fd1 where "Fstd_get F1 f = Some fd1" and "rel_fundef norm_eq fd1 fd2" - using rel_fundefs_Some2[OF rel_F1_F2] by blast + case (step_get f l pc n R d \ st') show ?case (is "\x. ?STEP x \ ?MATCH x") proof - - let ?s1' = "State F1 H (Frame f (Suc pc) \ # st')" + let ?s1' = "State F1 H (Frame f l (Suc pc) R (d # \) # st')" have "?STEP ?s1'" - using step_cjump_false \Fstd_get F1 f = Some fd1\ \rel_fundef norm_eq fd1 fd2\ - by (auto intro!: Sstd.step_cjump_false simp: rel_fundef_body_nth) + using step_get rel_F1_F2_next_instr[of f l pc] + by (auto intro: Sstd.step_get simp: option_rel_Some2) moreover have "?MATCH ?s1'" - using rel_F1_F2 by (auto intro: match.intros) + using ok_F1 rel_F1_F2 by (auto intro: match.intros) ultimately show "?thesis" by blast qed next - case (step_fun_call f fd2 pc g gd2 \ frame\<^sub>f frame\<^sub>g st') - obtain fd1 where "Fstd_get F1 f = Some fd1" and rel_fd1_fd2: "rel_fundef norm_eq fd1 fd2" - using step_fun_call rel_fundefs_Some2[OF rel_F1_F2] by blast - obtain gd1 where "Fstd_get F1 g = Some gd1" and rel_gd1_gd2: "rel_fundef norm_eq gd1 gd2" - using step_fun_call rel_fundefs_Some2[OF rel_F1_F2] by blast - have pc_in_range: "pc < length (body fd1)" - using \pc < length (body fd2)\ rel_fundef_body_length[OF rel_fd1_fd2] - by simp - have arity_gd1_gd2: "arity gd2 = arity gd1" - using rel_fundef_arities[OF rel_gd1_gd2, symmetric] . + case (step_set f l pc n R R' d \ st') show ?case (is "\x. ?STEP x \ ?MATCH x") proof - - let ?\g = "take (arity gd1) \" - let ?s1' = "State F1 H (Frame g 0 ?\g # Frame f pc \ # st')" + let ?s1' = "State F1 H (Frame f l (Suc pc) R' \ # st')" have "?STEP ?s1'" - using step_fun_call pc_in_range - using \Fstd_get F1 f = Some fd1\ \Fstd_get F1 g = Some gd1\ - using rel_fundef_body_nth[OF rel_fd1_fd2 pc_in_range] arity_gd1_gd2 - by (auto intro: Sstd.step_fun_call) + using step_set rel_F1_F2_next_instr[of f l pc] + by (auto intro: Sstd.step_set simp: option_rel_Some2) moreover have "?MATCH ?s1'" - using step_fun_call rel_F1_F2 arity_gd1_gd2 by (auto intro: match.intros) + using ok_F1 rel_F1_F2 by (auto intro: match.intros) + ultimately show "?thesis" by blast + qed + next + case (step_load f l pc x y d R \ st') + show ?case (is "\x. ?STEP x \ ?MATCH x") + proof - + let ?s1' = "State F1 H (Frame f l (Suc pc) R (d # \) # st')" + have "?STEP ?s1'" + using step_load rel_F1_F2_next_instr[of f l pc] + by (auto intro!: Sstd.step_load simp: option_rel_Some2) + moreover have "?MATCH ?s1'" + using ok_F1 rel_F1_F2 by (auto intro: match.intros) + ultimately show "?thesis" by blast + qed + next + case (step_store f l pc x y d H' R \ st') + show ?case (is "\x. ?STEP x \ ?MATCH x") + proof - + let ?s1' = "State F1 H' (Frame f l (Suc pc) R \ # st')" + have "?STEP ?s1'" + using step_store rel_F1_F2_next_instr[of f l pc] + by (auto intro!: Sstd.step_store simp: option_rel_Some2) + moreover have "?MATCH ?s1'" + using ok_F1 rel_F1_F2 by (auto intro: match.intros) + ultimately show "?thesis" by blast + qed + next + case (step_op f l pc op ar \ x R st') + show ?case (is "\x. ?STEP x \ ?MATCH x") + proof - + let ?s1' = "State F1 H (Frame f l (Suc pc) R (x # drop ar \) # st')" + have "?STEP ?s1'" + using step_op rel_F1_F2_next_instr[of f l pc] + by (auto intro!: Sstd.step_op simp: option_rel_Some2) + moreover have "?MATCH ?s1'" + using ok_F1 rel_F1_F2 by (auto intro: match.intros) + ultimately show "?thesis" by blast + qed + next + case (step_op_inl f l pc op ar \ opinl x F2' R st') + then obtain instr1 where + next_instr1: "next_instr (Fstd_get F1) f l pc = Some instr1" and + norm_eq_instr1_instr2: "norm_eq instr1 (Inca.IOp op)" + using rel_F1_F2_next_instr[of f l pc] by (simp add: option_rel_Some2) + then show ?case (is "\x. ?STEP x \ ?MATCH x") + proof (cases instr1) + case (IOp op') + hence "op' = op" using norm_eq_instr1_instr2 by simp + let ?s1' = "State F1 H (Frame f l (Suc pc) R (x # drop ar \) # st')" + have "?STEP ?s1'" + using step_op_inl next_instr1 IOp \op' = op\ + using Sinca.\\\\\_correct Sinca.\\\_invertible + by (auto intro: Sstd.step_op) + moreover have "?MATCH ?s1'" + using ok_F1 step_op_inl next_instr1 IOp \op' = op\ + using Sinca.\\\_invertible + by (auto intro!: match.intros intro: rel_fundefs_rewriteI2[OF rel_F1_F2]) + ultimately show "?thesis" by blast + qed simp_all + next + case (step_op_inl_hit f l pc opinl ar \ x R st') + show ?case (is "\x. ?STEP x \ ?MATCH x") + proof - + let ?s1' = "State F1 H (Frame f l (Suc pc) R (x # drop ar \) # st')" + have "?STEP ?s1'" + using step_op_inl_hit rel_F1_F2_next_instr[of f l pc] + using Sinca.\\\\\_correct + by (auto intro: Sstd.step_op simp: option_rel_Some2) + moreover have "?MATCH ?s1'" + using ok_F1 rel_F1_F2 by (auto intro: match.intros) + ultimately show "?thesis" by blast + qed + next + case (step_op_inl_miss f l pc opinl ar \ x F' R st') + then obtain instr1 where + next_instr1: "next_instr (Fstd_get F1) f l pc = Some instr1" and + norm_eq_instr1_instr2: "norm_eq instr1 (Inca.IOpInl opinl)" + using rel_F1_F2_next_instr[of f l pc] by (simp add: option_rel_Some2) + then show ?case (is "\x. ?STEP x \ ?MATCH x") + proof (cases instr1) + case (IOp op) + hence "op = \
\\\\ opinl" using norm_eq_instr1_instr2 by simp + let ?s1' = "State F1 H (Frame f l (Suc pc) R (x # drop ar \) # st')" + have "?STEP ?s1'" + using step_op_inl_miss next_instr1 IOp \op = \
\\\\ opinl\ + using Sinca.\\\\\_correct + by (auto intro: Sstd.step_op) + moreover have "?MATCH ?s1'" + using ok_F1 step_op_inl_miss next_instr1 IOp \op = \
\\\\ opinl\ + using Sinca.\\\_invertible + by (auto intro!: match.intros intro: rel_fundefs_rewriteI2[OF rel_F1_F2]) + ultimately show "?thesis" by blast + qed simp_all + next + case (step_cjump f l pc l\<^sub>t l\<^sub>f d l' R \ st') + show ?case (is "\x. ?STEP x \ ?MATCH x") + proof - + let ?s1' = "State F1 H (Frame f l' 0 R \ # st')" + have "?STEP ?s1'" + using step_cjump rel_F1_F2_next_instr[of f l pc] + by (auto intro: Sstd.step_cjump simp: option_rel_Some2) + moreover have "?MATCH ?s1'" + using ok_F1 rel_F1_F2 by (auto intro: match.intros) + ultimately show "?thesis" by blast + qed + next + case (step_call f l\<^sub>f pc\<^sub>f g gd2 \\<^sub>f frame\<^sub>g R\<^sub>f st') + then obtain gd1 where + F1_g: "Fstd_get F1 g = Some gd1" and + rel_gd1_gd2: "rel_fundef (=) norm_eq gd1 gd2" + using rel_fundefs_Some2[OF rel_F1_F2] by blast + + have same_fst_label_gd1_gd2: "fst (hd (body gd1)) = fst (hd (body gd2))" + using wf_fundefs_imp_wf_fundef[OF ok_F1 F1_g, unfolded wf_fundef_def] rel_gd1_gd2 + by (auto elim!: fundef.rel_cases dest: list_all2_rel_prod_fst_hd) + + show ?case (is "\x. ?STEP x \ ?MATCH x") + proof - + let ?\g = "take (arity gd1) \\<^sub>f" + let ?lg = "(fst (hd (body gd1)))" + let ?s1' = "State F1 H (frame\<^sub>g # Frame f l\<^sub>f pc\<^sub>f R\<^sub>f \\<^sub>f # st')" + have "?STEP ?s1'" + using step_call rel_F1_F2_next_instr[of f l\<^sub>f pc\<^sub>f] + using F1_g same_fst_label_gd1_gd2 rel_gd1_gd2 + by (auto intro!: Sstd.step_call + simp: option_rel_Some2 allocate_frame_def rel_fundef_arities rel_fundef_locals) + moreover have "?MATCH ?s1'" + using ok_F1 rel_F1_F2 + by (auto intro: match.intros) ultimately show ?thesis by auto qed next - case (step_fun_end g gd2 \\<^sub>f pc\<^sub>g frame\<^sub>g \\<^sub>g frame\<^sub>f f pc\<^sub>f frame\<^sub>f' st') - then obtain gd1 where "Fstd_get F1 g = Some gd1" and rel_gd1_gd2: "rel_fundef norm_eq gd1 gd2" + case (step_return g l\<^sub>g pc\<^sub>g gd2 \\<^sub>f \\<^sub>g frame\<^sub>f' f l\<^sub>f pc\<^sub>f R\<^sub>f R\<^sub>g st') + then obtain gd1 where + F1_g: "Fstd_get F1 g = Some gd1" and + rel_gd1_gd2: "rel_fundef (=) norm_eq gd1 gd2" using rel_fundefs_Some2[OF rel_F1_F2] by blast - have arity_gd1_gd2: "arity gd2 = arity gd1" - using rel_fundef_arities[OF rel_gd1_gd2, symmetric] . - show ?case (is "\x. ?STEP x \ ?MATCH x") - proof - - let ?s1' = "State F1 H (Frame f (Suc pc\<^sub>f) (\\<^sub>g @ drop (arity gd1) \\<^sub>f) # st')" + obtain instr1 where + next_instr1: "next_instr (Fstd_get F1) g l\<^sub>g pc\<^sub>g = Some instr1" and + norm_eq_instr1_instr2: "norm_eq instr1 Inca.IReturn" + using step_return rel_F1_F2_next_instr[of g l\<^sub>g pc\<^sub>g] by (simp add: option_rel_Some2) + then show ?case (is "\x. ?STEP x \ ?MATCH x") + proof (cases instr1) + case IReturn + let ?s1' = "State F1 H (frame\<^sub>f' # st')" have "?STEP ?s1'" - using step_fun_end \Fstd_get F1 g = Some gd1\ - using rel_fundef_body_length[OF rel_gd1_gd2] arity_gd1_gd2 - by (auto intro: Sstd.step_fun_end) + using step_return next_instr1 IReturn + using F1_g rel_gd1_gd2 + by (auto intro!: Sstd.step_return simp: rel_fundef_arities rel_fundef_return) moreover have "?MATCH ?s1'" - using step_fun_end rel_F1_F2 arity_gd1_gd2 by (auto intro: match.intros) + using ok_F1 rel_F1_F2 by (auto intro: match.intros) ultimately show ?thesis by auto - qed + qed simp_all qed qed lemma match_final_backward: - "s1 \ s2 \ Sinca.final s2 \ Sstd.final s1" -proof (induction s1 s2 rule: match.induct) + assumes "match s1 s2" and final_s2: "final Finca_get Inca.IReturn s2" + shows "final Fstd_get Std.IReturn s1" + using \match s1 s2\ +proof (cases s1 s2 rule: match.cases) case (1 F1 F2 H st) - then obtain f fd2 pc \ where - st_def: "st = [Frame f pc \]" and - "Finca_get F2 f = Some fd2" and - pc_def: "pc = length (body fd2)" - by (auto elim: Sinca.final.cases) - then obtain fd1 where "Fstd_get F1 f = Some fd1" and "rel_fundef norm_eq fd1 fd2" - using 1 rel_fundefs_Some2 by fast - then show ?case - unfolding st_def - using pc_def rel_fundef_body_length[OF \rel_fundef norm_eq fd1 fd2\] - by (auto intro: Sstd.final.intros) + show ?thesis + using final_s2[unfolded 1] + proof (cases _ _ "State F2 H st" rule: final.cases) + case (finalI f l pc R \) + then show ?thesis + using 1 + by (auto intro!: final.intros dest: rel_fundefs_next_instr2) + qed qed sublocale std_inca_simulation: - backward_simulation Sstd.step Sinca.step Sstd.final Sinca.final "\_ _. False" "\_. match" + backward_simulation where + step1 = Sstd.step and final1 = "final Fstd_get Std.IReturn" and + step2 = Sinca.step and final2 = "final Finca_get Inca.IReturn" and + order = "\_ _. False" and match = "\_. match" using match_final_backward backward_lockstep_simulation lockstep_to_plus_backward_simulation[of match Sinca.step Sstd.step] by unfold_locales auto + section \Forward simulation\ lemma forward_lockstep_simulation: - assumes "Sstd.step s1 s1'" and "s1 \ s2" + assumes "Sstd.step s1 s1'" and "match s1 s2" shows "\s2'. Sinca.step s2 s2' \ s1' \ s2'" proof - from assms(2) obtain F1 F2 H st where s1_def: "s1 = State F1 H st" and s2_def: "s2 = State F2 H st" and + ok_F1: "wf_fundefs (Fstd_get F1)" and rel_F1_F2: "rel_fundefs (Fstd_get F1) (Finca_get F2)" by (auto elim: match.cases) + + note rel_F1_F2_next_instr = rel_fundefs_next_instr[OF rel_F1_F2] + from assms(1) show "?thesis" unfolding s1_def s2_def proof(induction "State F1 H st" s1' rule: Sstd.step.induct) - case (step_push f fd1 pc d \ st') - then obtain fd2 where "Finca_get F2 f = Some fd2" and "rel_fundef norm_eq fd1 fd2" - using rel_fundefs_Some1[OF rel_F1_F2] by blast - show ?case (is "\x. ?STEP x \ ?MATCH x") - proof - - let ?s1' = "State F2 H (Frame f (Suc pc) (d # \) # st')" - have "?STEP ?s1'" - using step_push \Finca_get F2 f = Some fd2\ \rel_fundef norm_eq fd1 fd2\ - by (auto intro!: Sinca.step_push elim!: norm_instr.elims simp: rel_fundef_body_nth) - moreover have "?MATCH ?s1'" - using rel_F1_F2 by (auto intro: match.intros) - ultimately show "?thesis" by blast - qed - next - case (step_pop f fd1 pc d \ st') - then obtain fd2 where "Finca_get F2 f = Some fd2" and "rel_fundef norm_eq fd1 fd2" - using rel_fundefs_Some1[OF rel_F1_F2] by blast + case (step_push f l pc d R \ st') show ?case (is "\x. ?STEP x \ ?MATCH x") proof - - let ?s1' = "State F2 H (Frame f (Suc pc) \ # st')" + let ?s1' = "State F2 H (Frame f l (Suc pc) R (d # \) # st')" have "?STEP ?s1'" - using step_pop \Finca_get F2 f = Some fd2\ \rel_fundef norm_eq fd1 fd2\ - by (auto intro!: Sinca.step_pop elim!: norm_instr.elims simp: rel_fundef_body_nth) + using step_push rel_F1_F2_next_instr[of f l pc] + by (auto intro!: Sinca.step_push elim: norm_instr.elims simp: option_rel_Some1) moreover have "?MATCH ?s1'" - using rel_F1_F2 by (auto intro: match.intros) - ultimately show "?thesis" by blast - qed - next - case (step_load f fd1 pc x y d \ st') - then obtain fd2 where "Finca_get F2 f = Some fd2" and "rel_fundef norm_eq fd1 fd2" - using rel_fundefs_Some1[OF rel_F1_F2] by blast - show ?case (is "\x. ?STEP x \ ?MATCH x") - proof - - let ?s1' = "State F2 H (Frame f (Suc pc) (d # \) # st')" - have "?STEP ?s1'" - using step_load \Finca_get F2 f = Some fd2\ \rel_fundef norm_eq fd1 fd2\ - by (auto intro!: Sinca.step_load elim!: norm_instr.elims simp: rel_fundef_body_nth) - moreover have "?MATCH ?s1'" - using rel_F1_F2 by (auto intro: match.intros) + using ok_F1 rel_F1_F2 by (auto intro: match.intros) ultimately show "?thesis" by blast qed next - case (step_store f fd1 pc x y d H' \ st') - then obtain fd2 where "Finca_get F2 f = Some fd2" and "rel_fundef norm_eq fd1 fd2" - using rel_fundefs_Some1[OF rel_F1_F2] by blast + case (step_pop f l pc R d \ st') show ?case (is "\x. ?STEP x \ ?MATCH x") proof - - let ?s1' = "State F2 H' (Frame f (Suc pc) \ # st')" + let ?s1' = "State F2 H (Frame f l (Suc pc) R \ # st')" have "?STEP ?s1'" - using step_store \Finca_get F2 f = Some fd2\ \rel_fundef norm_eq fd1 fd2\ - by (auto intro!: Sinca.step_store elim!: norm_instr.elims simp: rel_fundef_body_nth) + using step_pop rel_F1_F2_next_instr[of f l pc] + by (auto intro: Sinca.step_pop elim: norm_instr.elims simp: option_rel_Some1) moreover have "?MATCH ?s1'" - using rel_F1_F2 by (auto intro: match.intros) + using ok_F1 rel_F1_F2 by (auto intro: match.intros) ultimately show "?thesis" by blast qed next - case (step_op f fd1 pc op ar \ x st') - then obtain fd2 where F2_get_f[intro]: "Finca_get F2 f = Some fd2" and "rel_fundef norm_eq fd1 fd2" - using rel_fundefs_Some1[OF rel_F1_F2] by blast - have pc_in_range[intro]: "pc < length (body fd2)" - using step_op \rel_fundef norm_eq fd1 fd2\ rel_fundef_body_length by fastforce - have norm_body2: "norm_instr (body fd2 ! pc) = Std.IOp op" - using step_op rel_fundef_body_nth[OF \rel_fundef norm_eq fd1 fd2\, of pc] - by simp - then show ?case (is "\x. ?STEP x \ ?MATCH x") - proof (cases "body fd2 ! pc") + case (step_get f l pc n R d \ st') + show ?case (is "\x. ?STEP x \ ?MATCH x") + proof - + let ?s1' = "State F2 H (Frame f l (Suc pc) R (d # \) # st')" + have "?STEP ?s1'" + using step_get rel_F1_F2_next_instr[of f l pc] + by (auto intro: Sinca.step_get elim: norm_instr.elims simp: option_rel_Some1) + moreover have "?MATCH ?s1'" + using ok_F1 rel_F1_F2 by (auto intro: match.intros) + ultimately show "?thesis" by blast + qed + next + case (step_set f l pc n R R' d \ st') + show ?case (is "\x. ?STEP x \ ?MATCH x") + proof - + let ?s1' = "State F2 H (Frame f l (Suc pc) R' \ # st')" + have "?STEP ?s1'" + using step_set rel_F1_F2_next_instr[of f l pc] + by (auto intro: Sinca.step_set elim: norm_instr.elims simp: option_rel_Some1) + moreover have "?MATCH ?s1'" + using ok_F1 rel_F1_F2 by (auto intro: match.intros) + ultimately show "?thesis" by blast + qed + next + case (step_load f l pc x y d R \ st') + show ?case (is "\x. ?STEP x \ ?MATCH x") + proof - + let ?s1' = "State F2 H (Frame f l (Suc pc) R (d # \) # st')" + have "?STEP ?s1'" + using step_load rel_F1_F2_next_instr[of f l pc] + by (auto intro: Sinca.step_load elim: norm_instr.elims simp: option_rel_Some1) + moreover have "?MATCH ?s1'" + using ok_F1 rel_F1_F2 by (auto intro: match.intros) + ultimately show "?thesis" by blast + qed + next + case (step_store f l pc x y d H' R \ st') + show ?case (is "\x. ?STEP x \ ?MATCH x") + proof - + let ?s1' = "State F2 H' (Frame f l (Suc pc) R \ # st')" + have "?STEP ?s1'" + using step_store rel_F1_F2_next_instr[of f l pc] + by (auto intro: Sinca.step_store elim: norm_instr.elims simp: option_rel_Some1) + moreover have "?MATCH ?s1'" + using ok_F1 rel_F1_F2 by (auto intro: match.intros) + ultimately show "?thesis" by blast + qed + next + case (step_op f l pc op ar \ x R st') + then obtain instr2 where + next_instr2: "next_instr (Finca_get F2) f l pc = Some instr2" and + norm_eq_instr2: "norm_eq (Std.IOp op) instr2" + using rel_F1_F2_next_instr[of f l pc] by (auto simp: option_rel_Some1) + thus ?case (is "\x. ?STEP x \ ?MATCH x") + proof (cases instr2) case (IOp op') - then have "op' = op" using norm_body2 by simp + then have "op' = op" using norm_eq_instr2 by simp show ?thesis proof (cases "\\\ op' (take ar \)") case None - let ?s2' = "State F2 H (Frame f (Suc pc) (x # drop ar \) # st')" + let ?s2' = "State F2 H (Frame f l (Suc pc) R (x # drop ar \) # st')" have "?STEP ?s2'" - using None IOp step_op \op' = op\ + using None IOp step_op \op' = op\ next_instr2 by (auto intro: Sinca.step_op) moreover have "?MATCH ?s2'" - using rel_F1_F2 by (auto intro: match.intros) + using ok_F1 rel_F1_F2 by (auto intro: match.intros) ultimately show ?thesis by auto next case (Some opinl) - let ?fd2' = "rewrite_fundef_body fd2 pc (IOpInl opinl)" - let ?s2' = "State (Finca_add F2 f ?fd2') H (Frame f (Suc pc) (x # drop ar \) # st')" + let ?F2' = "Sinca.Fenv.map_entry F2 f (\fd. rewrite_fundef_body fd l pc (IOpInl opinl))" + let ?s2' = "State ?F2' H (Frame f l (Suc pc) R (x # drop ar \) # st')" have "?STEP ?s2'" using Some IOp step_op \op' = op\ - using Sinca.\\\\\_correct Sinca.\\\_invertible - by (auto intro: Sinca.step_op_inl) + using Sinca.\\\\\_correct Sinca.\\\_invertible next_instr2 + by (auto intro: Sinca.step_op_inl) moreover have "?MATCH ?s2'" - using Some IOp Sinca.\\\_invertible - by (auto intro!: match.intros rel_fundefs_rewrite[OF rel_F1_F2]) + using ok_F1 step_op IOp \op' = op\ Some + by (auto intro!: match.intros intro: rel_fundefs_rewriteI2[OF rel_F1_F2] + simp: Sinca.\\\_invertible) ultimately show ?thesis by auto qed next - case (IOpInl op') - then have "\
\\\\ op' = op" using norm_body2 by simp + case (IOpInl opinl) + hence "op = \
\\\\ opinl" using norm_eq_instr2 by simp show ?thesis - proof (cases "\\\\\ op' (take ar \)") + proof (cases "\\\\\ opinl (take ar \)") case True - let ?s2' = "State F2 H (Frame f (Suc pc) (x # drop ar \) # st')" + let ?s2' = "State F2 H (Frame f l (Suc pc) R (x # drop ar \) # st')" have "?STEP ?s2'" - using True IOpInl step_op \\
\\\\ op' = op\ Sinca.\\\\\_correct - by (auto intro!: Sinca.step_op_inl_hit) + using step_op IOpInl \op = \
\\\\ opinl\ True + using next_instr2 Sinca.\\\\\_correct + by (auto intro: Sinca.step_op_inl_hit) moreover have "?MATCH ?s2'" - using rel_F1_F2 by (auto intro: match.intros) + using ok_F1 rel_F1_F2 by (auto intro: match.intros) ultimately show ?thesis by auto next case False - let ?fd2' = "rewrite_fundef_body fd2 pc (IOp (\
\\\\ op'))" - let ?s2' = "State (Finca_add F2 f ?fd2') H (Frame f (Suc pc) (x # drop ar \) # st')" + let ?F2' = "Sinca.Fenv.map_entry F2 f (\fd. rewrite_fundef_body fd l pc (IOp (\
\\\\ opinl)))" + let ?s2' = "State ?F2' H (Frame f l (Suc pc) R (x # drop ar \) # st')" have "?STEP ?s2'" - using False IOpInl step_op \\
\\\\ op' = op\ Sinca.\\\\\_correct - by (auto intro!: Sinca.step_op_inl_miss) + using step_op IOpInl \op = \
\\\\ opinl\ False + using next_instr2 Sinca.\\\\\_correct + by (auto intro: Sinca.step_op_inl_miss) moreover have "?MATCH ?s2'" - using IOpInl - by (auto intro!: match.intros intro: rel_fundefs_rewrite[OF rel_F1_F2]) + using ok_F1 step_op \op = \
\\\\ opinl\ + by (auto intro!: match.intros intro: rel_fundefs_rewriteI2[OF rel_F1_F2]) ultimately show ?thesis by auto qed qed simp_all next - case (step_cjump_true f fd1 pc n d \ st') - then obtain fd2 where F2_get_f[intro]: "Finca_get F2 f = Some fd2" and "rel_fundef norm_eq fd1 fd2" - using rel_fundefs_Some1[OF rel_F1_F2] by blast - have pc_in_range[intro]: "pc < length (body fd2)" - using \rel_fundef norm_eq fd1 fd2\ rel_fundef_body_length step_cjump_true.hyps(2) by fastforce - have norm_body2: "norm_instr (body fd2 ! pc) = Std.ICJump n" - using step_cjump_true rel_fundef_body_nth[OF \rel_fundef norm_eq fd1 fd2\, of pc] - by simp + case (step_cjump f l pc l\<^sub>t l\<^sub>f d l' R \ st') show ?case (is "\x. ?STEP x \ ?MATCH x") proof - - let ?s1' = "State F2 H (Frame f n \ # st')" + let ?s1' = "State F2 H (Frame f l' 0 R \ # st')" have "?STEP ?s1'" - using step_cjump_true norm_body2 - by (auto intro: Sinca.step_cjump_true elim: norm_instr.elims) + using step_cjump rel_F1_F2_next_instr[of f l pc] + by (auto intro: Sinca.step_cjump elim: norm_instr.elims simp: option_rel_Some1) moreover have "?MATCH ?s1'" - using step_cjump_true rel_F1_F2 by (auto intro: match.intros) - ultimately show "?thesis" by blast - qed - next - case (step_cjump_false f fd1 pc n d \ st') - then obtain fd2 where F2_get_f[intro]: "Finca_get F2 f = Some fd2" and "rel_fundef norm_eq fd1 fd2" - using rel_fundefs_Some1[OF rel_F1_F2] by blast - have pc_in_range[intro]: "pc < length (body fd2)" - using \rel_fundef norm_eq fd1 fd2\ rel_fundef_body_length step_cjump_false by fastforce - have norm_body2: "norm_instr (body fd2 ! pc) = Std.ICJump n" - using step_cjump_false rel_fundef_body_nth[OF \rel_fundef norm_eq fd1 fd2\, of pc] - by simp - show ?case (is "\x. ?STEP x \ ?MATCH x") - proof - - let ?s1' = "State F2 H (Frame f (Suc pc) \ # st')" - have "?STEP ?s1'" - using step_cjump_false norm_body2 - by (auto intro: Sinca.step_cjump_false elim: norm_instr.elims) - moreover have "?MATCH ?s1'" - using step_cjump_false rel_F1_F2 by (auto intro: match.intros) + using ok_F1 rel_F1_F2 by (auto intro: match.intros) ultimately show "?thesis" by blast qed next - case (step_fun_call f fd1 pc g gd1 \ frame\<^sub>f frame\<^sub>g st') - then obtain fd2 where "Finca_get F2 f = Some fd2" and "rel_fundef norm_eq fd1 fd2" + case (step_call f l\<^sub>f pc\<^sub>f g gd1 \\<^sub>f frame\<^sub>g R\<^sub>f st') + then obtain gd2 where + F2_g: "Finca_get F2 g = Some gd2" and + rel_gd1_gd2: "rel_fundef (=) norm_eq gd1 gd2" using rel_fundefs_Some1[OF rel_F1_F2] by blast - obtain gd2 where "Finca_get F2 g = Some gd2" and rel_gd1_gd2: "rel_fundef norm_eq gd1 gd2" - using step_fun_call rel_fundefs_Some1[OF rel_F1_F2] by blast - have pc_in_range: "pc < length (body fd2)" - using step_fun_call \rel_fundef norm_eq fd1 fd2\ rel_fundef_body_length by fastforce - have arity_gd1_gd2: "arity gd1 = arity gd2" - using rel_fundef_arities[OF rel_gd1_gd2] . + + have same_fst_label_gd1_gd2: "fst (hd (body gd1)) = fst (hd (body gd2))" + using rel_gd1_gd2 + using wf_fundefs_imp_wf_fundef[OF ok_F1 \Fstd_get F1 g = Some gd1\, unfolded wf_fundef_def] + by (auto elim: fundef.rel_cases dest!: list_all2_rel_prod_fst_hd) + show ?case (is "\x. ?STEP x \ ?MATCH x") proof - - let ?\g = "take (arity gd2) \" - let ?s2' = "State F2 H (Frame g 0 ?\g # Frame f pc \ # st')" + let ?\g = "take (arity gd2) \\<^sub>f" + let ?s2' = "State F2 H (frame\<^sub>g # Frame f l\<^sub>f pc\<^sub>f R\<^sub>f \\<^sub>f # st')" have "?STEP ?s2'" - proof - - have "arity gd2 \ length \" - using \arity gd1 \ length \\ arity_gd1_gd2 by simp - moreover have "body fd2 ! pc = Inca.ICall g" - using step_fun_call rel_fundef_body_nth[OF \rel_fundef norm_eq fd1 fd2\, of pc] - by (auto elim: norm_instr.elims) - ultimately show ?thesis - using step_fun_call - using \Finca_get F2 g = Some gd2\ \Finca_get F2 f = Some fd2\ - using pc_in_range \body fd2 ! pc = Inca.ICall g\ - by (auto intro: Sinca.step_fun_call) - qed + using step_call F2_g rel_gd1_gd2 rel_F1_F2_next_instr[of f l\<^sub>f pc\<^sub>f] + using same_fst_label_gd1_gd2 + by (auto intro!: Sinca.step_call elim: norm_instr.elims + simp: option_rel_Some1 rel_fundef_arities rel_fundef_locals allocate_frame_def) moreover have "?MATCH ?s2'" - using step_fun_call rel_F1_F2 arity_gd1_gd2 by (auto intro: match.intros) + using ok_F1 rel_F1_F2 + by (auto intro: match.intros) ultimately show ?thesis by auto qed next - case (step_fun_end g gd1 \\<^sub>f pc\<^sub>g frame\<^sub>g \\<^sub>g frame\<^sub>f f pc\<^sub>f frame\<^sub>f' st') - then obtain gd2 where "Finca_get F2 g = Some gd2" and rel_gd1_gd2: "rel_fundef norm_eq gd1 gd2" + case (step_return g l\<^sub>g pc\<^sub>g gd1 \\<^sub>f \\<^sub>g frame\<^sub>f' f l\<^sub>f pc\<^sub>f R\<^sub>f R\<^sub>g st') + then obtain gd2 where + F2_g: "Finca_get F2 g = Some gd2" and + rel_gd1_gd2: "rel_fundef (=) norm_eq gd1 gd2" using rel_fundefs_Some1[OF rel_F1_F2] by blast - have pc_at_end: "pc\<^sub>g = length (body gd2)" - using rel_fundef_body_length[OF rel_gd1_gd2] step_fun_end by fastforce - have arity_gd1_gd2: "arity gd1 = arity gd2" - using rel_fundef_arities[OF rel_gd1_gd2] . - show ?case (is "\x. ?STEP x \ ?MATCH x") - proof - - let ?s1' = "State F2 H (Frame f (Suc pc\<^sub>f) (\\<^sub>g @ drop (arity gd2) \\<^sub>f) # st')" + obtain instr2 where + next_instr2: "next_instr (Finca_get F2) g l\<^sub>g pc\<^sub>g = Some instr2" and + norm_eq_instr1_instr2: "norm_eq Std.IReturn instr2" + using step_return rel_F1_F2_next_instr[of g l\<^sub>g pc\<^sub>g] by (auto simp: option_rel_Some1) + thus ?case (is "\x. ?STEP x \ ?MATCH x") + proof (cases instr2) + case IReturn + let ?s1' = "State F2 H (frame\<^sub>f' # st')" have "?STEP ?s1'" - proof - - have "arity gd2 \ length \\<^sub>f" - using \arity gd1 \ length \\<^sub>f\ arity_gd1_gd2 by simp - then show ?thesis - using step_fun_end pc_at_end \Finca_get F2 g = Some gd2\ - by (auto intro: Sinca.step_fun_end) - qed + using step_return next_instr2 IReturn + using F2_g rel_gd1_gd2 + by (auto intro!: Sinca.step_return simp: rel_fundef_arities rel_fundef_return) moreover have "?MATCH ?s1'" - using step_fun_end rel_F1_F2 arity_gd1_gd2 by (auto intro: match.intros) - ultimately show "?thesis" by auto - qed + using ok_F1 rel_F1_F2 by (auto intro: match.intros) + ultimately show ?thesis by auto + qed simp_all qed qed -lemma forward_match_final: - "s1 \ s2 \ Sstd.final s1 \ Sinca.final s2" -proof (induction s1 s2 rule: match.induct) +lemma match_final_forward: + assumes "match s1 s2" and final_s1: "final Fstd_get Std.IReturn s1" + shows "final Finca_get Inca.IReturn s2" + using \match s1 s2\ +proof (cases s1 s2 rule: match.cases) case (1 F1 F2 H st) - then obtain f fd1 pc \ where - st_def: "st = [Frame f pc \]" and - "Fstd_get F1 f = Some fd1" and - pc_def: "pc = length (body fd1)" - by (auto elim: Sstd.final.cases) - then obtain fd2 where "Finca_get F2 f = Some fd2" and "rel_fundef norm_eq fd1 fd2" - using 1 rel_fundefs_Some1 by fast - then show ?case - unfolding st_def - using pc_def rel_fundef_body_length[OF \rel_fundef norm_eq fd1 fd2\] - by (auto intro: Sinca.final.intros) + show ?thesis + using final_s1[unfolded 1] + proof (cases _ _ "State F1 H st" rule: final.cases) + case (finalI f l pc R \) + then show ?thesis + using 1 + by (auto intro: final.intros elim: norm_instr.elims dest: rel_fundefs_next_instr1) + qed qed sublocale std_inca_forward_simulation: - forward_simulation Sstd.step Sinca.step Sstd.final Sinca.final "\_ _. False" "\_. match" - using forward_match_final forward_lockstep_simulation + forward_simulation where + step1 = Sstd.step and final1 = "final Fstd_get Std.IReturn" and + step2 = Sinca.step and final2 = "final Finca_get Inca.IReturn" and + order = "\_ _. False" and match = "\_. match" + using match_final_forward forward_lockstep_simulation lockstep_to_plus_forward_simulation[of match Sstd.step _ Sinca.step] by unfold_locales auto section \Bisimulation\ sublocale std_inca_bisimulation: - bisimulation Sstd.step Sinca.step Sstd.final Sinca.final "\_ _. False" "\_. match" + bisimulation where + step1 = Sstd.step and final1 = "final Fstd_get Std.IReturn" and + step2 = Sinca.step and final2 = "final Finca_get Inca.IReturn" and + order = "\_ _. False" and match = "\_. match" by unfold_locales end end \ No newline at end of file diff --git a/thys/Interpreter_Optimizations/Ubx.thy b/thys/Interpreter_Optimizations/Ubx.thy --- a/thys/Interpreter_Optimizations/Ubx.thy +++ b/thys/Interpreter_Optimizations/Ubx.thy @@ -1,218 +1,317 @@ theory Ubx imports Global OpUbx Env "VeriComp.Language" begin section \Unboxed caching\ -datatype ('dyn, 'var, 'fun, 'op, 'opinl, 'opubx, 'num, 'bool) instr = +subsection \Static representation\ + +datatype ('dyn, 'var, 'fun, 'label, 'op, 'opinl, 'opubx, 'num, 'bool) instr = IPush 'dyn | IPushUbx1 'num | IPushUbx2 'bool | IPop | + IGet nat | IGetUbx type nat | + ISet nat | ISetUbx type nat | ILoad 'var | ILoadUbx type 'var | IStore 'var | IStoreUbx type 'var | - IOp 'op | - IOpInl 'opinl | - IOpUbx 'opubx | - is_jump: ICJump nat | - is_fun_call: ICall 'fun + IOp 'op | IOpInl 'opinl | IOpUbx 'opubx | + is_jump: ICJump 'label 'label | + is_fun_call: ICall 'fun | + is_return: IReturn locale ubx = Fenv: env F_empty F_get F_add F_to_list + Henv: env heap_empty heap_get heap_add heap_to_list + nary_operations_ubx \\ \\\\\ \\\\\ \\\ \\\\\ \
\\\\ - is_true is_false box_ubx1 unbox_ubx1 box_ubx2 unbox_ubx2 + uninitialized is_true is_false box_ubx1 unbox_ubx1 box_ubx2 unbox_ubx2 \\\\\ \\\ \\\ \\\\\\\\ for \ \Functions environment\ F_empty and - F_get :: "'fenv \ 'fun \ ('dyn, 'var, 'fun, 'op, 'opinl, 'opubx, 'num, 'bool) instr fundef option" and + F_get :: "'fenv \ 'fun \ ('label, ('dyn, 'var, 'fun, 'label, 'op, 'opinl, 'opubx, 'num, 'bool) instr) fundef option" and F_add and F_to_list and \ \Memory heap\ heap_empty and heap_get :: "'henv \ 'var \ 'dyn \ 'dyn option" and heap_add and heap_to_list and + \ \Dynamic values\ + uninitialized :: 'dyn and is_true and is_false and + \ \Unboxed values\ - is_true :: "'dyn \ bool" and is_false and box_ubx1 and unbox_ubx1 and box_ubx2 and unbox_ubx2 and \ \n-ary operations\ \\ :: "'op \ 'dyn list \ 'dyn" and \\\\\ and \\\\\ and \\\ and \\\\\ and \
\\\\ :: "'opinl \ 'op" and \\\\\ :: "'opubx \ ('dyn, 'num, 'bool) unboxed list \ ('dyn, 'num, 'bool) unboxed option" and \\\ :: "'opinl \ type option list \ 'opubx option" and \\\ :: "'opubx \ 'opinl" and \\\\\\\\ begin + +lemmas map_option_funtype_comp_map_entry_rewrite_fundef_body[simp] = + Fenv.map_option_comp_map_entry[of _ funtype "\fd. rewrite_fundef_body fd l pc instr" for l pc instr, simplified] + fun generalize_instr where "generalize_instr (IPushUbx1 n) = IPush (box_ubx1 n)" | "generalize_instr (IPushUbx2 b) = IPush (box_ubx2 b)" | + "generalize_instr (IGetUbx _ n) = IGet n" | + "generalize_instr (ISetUbx _ n) = ISet n" | "generalize_instr (ILoadUbx _ x) = ILoad x" | "generalize_instr (IStoreUbx _ x) = IStore x" | "generalize_instr (IOpUbx opubx) = IOpInl (\\\ opubx)" | "generalize_instr instr = instr" +lemma instr_sel_generalize_instr[simp]: + "is_jump (generalize_instr instr) \ is_jump instr" + "is_fun_call (generalize_instr instr) \ is_fun_call instr" + "is_return (generalize_instr instr) \ is_return instr" + unfolding atomize_conj + by (cases instr; simp) + +lemma instr_sel_generalize_instr_comp[simp]: + "is_jump \ generalize_instr = is_jump" and + "is_fun_call \ generalize_instr = is_fun_call" and + "is_return \ generalize_instr = is_return" + unfolding atomize_conj + using instr_sel_generalize_instr by auto + fun generalize_fundef where - "generalize_fundef (Fundef xs ar) = Fundef (map generalize_instr xs) ar" + "generalize_fundef (Fundef xs ar ret locals) = + Fundef (map_ran (\_. map generalize_instr) xs) ar ret locals" lemma generalize_instr_idempotent[simp]: "generalize_instr (generalize_instr x) = generalize_instr x" by (cases x) simp_all lemma generalize_instr_idempotent_comp[simp]: "generalize_instr \ generalize_instr = generalize_instr" by fastforce -lemma generalize_fundef_length[simp]: "length (body (generalize_fundef fd)) = length (body fd)" +lemma length_body_generalize_fundef[simp]: "length (body (generalize_fundef fd)) = length (body fd)" + by (cases fd) (simp add: map_ran_def) + +(* lemma body_generalize_fundef[simp]: "body (generalize_fundef fd) = map generalize_instr (body fd)" + by (cases fd) simp *) + +lemma arity_generalize_fundef[simp]: "arity (generalize_fundef fd) = arity fd" by (cases fd) simp -lemma body_generalize_fundef[simp]: "body (generalize_fundef fd) = map generalize_instr (body fd)" +lemma return_generalize_fundef[simp]: "return (generalize_fundef fd) = return fd" by (cases fd) simp -lemma arity_generalize_fundef[simp]: "arity (generalize_fundef fd2) = arity fd2" - by (cases fd2) simp +lemma fundef_locals_generalize[simp]: "fundef_locals (generalize_fundef fd) = fundef_locals fd" + by (cases fd; simp) -inductive final where - "F_get F f = Some fd \ pc = length (body fd) \ final (State F H [Frame f pc \])" +lemma funtype_generalize_fundef[simp]: "funtype (generalize_fundef fd) = funtype fd" + by (cases fd; simp add: funtype_def) + +lemmas map_option_comp_map_entry_generalize_fundef[simp] = + Fenv.map_option_comp_map_entry[of _ funtype generalize_fundef, simplified] + +lemma image_fst_set_body_generalize_fundef[simp]: + "fst ` set (body (generalize_fundef fd)) = fst ` set (body fd)" + by (cases fd) (simp add: dom_map_ran) + +lemma map_of_generalize_fundef_conv: + "map_of (body (generalize_fundef fd)) l = map_option (map generalize_instr) (map_of (body fd) l)" + by (cases fd) (simp add: map_ran_conv) + +lemma map_option_arity_get_map_entry_generalize_fundef[simp]: + "map_option arity \ F_get (Fenv.map_entry F2 f generalize_fundef) = + map_option arity \ F_get F2" + by (auto intro: Fenv.map_option_comp_map_entry) + +lemma instr_at_generalize_fundef_conv: + "instr_at (generalize_fundef fd) l = map_option generalize_instr o instr_at fd l" +proof (intro ext) + fix n + show "instr_at (generalize_fundef fd) l n = (map_option generalize_instr \ instr_at fd l) n" + proof (cases fd) + case (Fundef bblocks ar locals) + then show ?thesis + by (cases "map_of bblocks l") (simp_all add: instr_at_def map_ran_conv) + qed +qed subsection \Semantics\ inductive step (infix "\" 55) where step_push: - "F_get F f = Some fd \ pc < length (body fd) \ body fd ! pc = IPush d \ - State F H (Frame f pc \ # st) \ State F H (Frame f (Suc pc) (OpDyn d # \) # st)" | + "next_instr (F_get F) f l pc = Some (IPush d) \ + State F H (Frame f l pc R \ # st) \ State F H (Frame f l (Suc pc) R (OpDyn d # \) # st)" | step_push_ubx1: - "F_get F f = Some fd \ pc < length (body fd) \ body fd ! pc = IPushUbx1 n \ - State F H (Frame f pc \ # st) \ State F H (Frame f (Suc pc) (OpUbx1 n # \) # st)" | + "next_instr (F_get F) f l pc = Some (IPushUbx1 n) \ + State F H (Frame f l pc R \ # st) \ State F H (Frame f l (Suc pc) R (OpUbx1 n # \) # st)" | step_push_ubx2: - "F_get F f = Some fd \ pc < length (body fd) \ body fd ! pc = IPushUbx2 b \ - State F H (Frame f pc \ # st) \ State F H (Frame f (Suc pc) (OpUbx2 b # \) # st)" | + "next_instr (F_get F) f l pc = Some (IPushUbx2 b) \ + State F H (Frame f l pc R \ # st) \ State F H (Frame f l (Suc pc) R (OpUbx2 b # \) # st)" | step_pop: - "F_get F f = Some fd \ pc < length (body fd) \ body fd ! pc = IPop \ - State F H (Frame f pc (x # \) # st) \ State F H (Frame f (Suc pc) \ # st)" | + "next_instr (F_get F) f l pc = Some IPop \ + State F H (Frame f l pc R (x # \) # st) \ State F H (Frame f l (Suc pc) R \ # st)" | + + step_get: + "next_instr (F_get F) f l pc = Some (IGet n) \ + n < length R \ cast_Dyn (R ! n) = Some d \ + State F H (Frame f l pc R \ # st) \ State F H (Frame f l (Suc pc) R (OpDyn d # \) # st)" | + + step_get_ubx_hit: + "next_instr (F_get F) f l pc = Some (IGetUbx \ n) \ + n < length R \ cast_Dyn (R ! n) = Some d \ unbox \ d = Some blob \ + State F H (Frame f l pc R \ # st) \ State F H (Frame f l (Suc pc) R (blob # \) # st)" | + + step_get_ubx_miss: + "next_instr (F_get F) f l pc = Some (IGetUbx \ n) \ + n < length R \ cast_Dyn (R ! n) = Some d \ unbox \ d = None \ + F' = Fenv.map_entry F f generalize_fundef \ + State F H (Frame f l pc R \ # st) \ State F' H (box_stack f (Frame f l (Suc pc) R (OpDyn d # \) # st))" | + + step_set: + "next_instr (F_get F) f l pc = Some (ISet n) \ + n < length R \ cast_Dyn blob = Some d \ R' = R[n := OpDyn d] \ + State F H (Frame f l pc R (blob # \) # st) \ State F H (Frame f l (Suc pc) R' \ # st)" | + + step_set_ubx: + "next_instr (F_get F) f l pc = Some (ISetUbx \ n) \ + n < length R \ cast_and_box \ blob = Some d \ R' = R[n := OpDyn d] \ + State F H (Frame f l pc R (blob # \) # st) \ State F H (Frame f l (Suc pc) R' \ # st)" | step_load: - "F_get F f = Some fd \ pc < length (body fd) \ body fd ! pc = ILoad x \ + "next_instr (F_get F) f l pc = Some (ILoad x) \ cast_Dyn i = Some i' \ heap_get H (x, i') = Some d \ - State F H (Frame f pc (i # \) # st) \ State F H (Frame f (Suc pc) (OpDyn d # \) # st)" | + State F H (Frame f l pc R (i # \) # st) \ State F H (Frame f l (Suc pc) R (OpDyn d # \) # st)" | step_load_ubx_hit: - "F_get F f = Some fd \ pc < length (body fd) \ body fd ! pc = ILoadUbx \ x \ + "next_instr (F_get F) f l pc = Some (ILoadUbx \ x) \ cast_Dyn i = Some i' \ heap_get H (x, i') = Some d \ unbox \ d = Some blob \ - State F H (Frame f pc (i # \) # st) \ State F H (Frame f (Suc pc) (blob # \) # st)" | + State F H (Frame f l pc R (i # \) # st) \ State F H (Frame f l (Suc pc) R (blob # \) # st)" | step_load_ubx_miss: - "F_get F f = Some fd \ pc < length (body fd) \ body fd ! pc = ILoadUbx \ x \ + "next_instr (F_get F) f l pc = Some (ILoadUbx \ x) \ cast_Dyn i = Some i' \ heap_get H (x, i') = Some d \ unbox \ d = None \ - F_add F f (generalize_fundef fd) = F' \ - State F H (Frame f pc (i # \) # st) \ State F' H (box_stack f (Frame f (Suc pc) (OpDyn d # \) # st))" | + F' = Fenv.map_entry F f generalize_fundef \ + State F H (Frame f l pc R (i # \) # st) \ State F' H (box_stack f (Frame f l (Suc pc) R (OpDyn d # \) # st))" | step_store: - "F_get F f = Some fd \ pc < length (body fd) \ body fd ! pc = IStore x \ + "next_instr (F_get F) f l pc = Some (IStore x) \ cast_Dyn i = Some i' \ cast_Dyn y = Some d \ heap_add H (x, i') d = H' \ - State F H (Frame f pc (i # y # \) # st) \ State F H' (Frame f (Suc pc) \ # st)" | + State F H (Frame f l pc R (i # y # \) # st) \ State F H' (Frame f l (Suc pc) R \ # st)" | step_store_ubx: - "F_get F f = Some fd \ pc < length (body fd) \ body fd ! pc = IStoreUbx \ x \ + "next_instr (F_get F) f l pc = Some (IStoreUbx \ x) \ cast_Dyn i = Some i' \ cast_and_box \ blob = Some d \ heap_add H (x, i') d = H' \ - State F H (Frame f pc (i # blob # \) # st) \ State F H' (Frame f (Suc pc) \ # st)" | + State F H (Frame f l pc R (i # blob # \) # st) \ State F H' (Frame f l (Suc pc) R \ # st)" | step_op: - "F_get F f = Some fd \ pc < length (body fd) \ body fd ! pc = IOp op \ + "next_instr (F_get F) f l pc = Some (IOp op) \ \\\\\ op = ar \ ar \ length \ \ - traverse cast_Dyn (take ar \) = Some \' \ + ap_map_list cast_Dyn (take ar \) = Some \' \ \\\ op \' = None \ \\ op \' = x \ - State F H (Frame f pc \ # st) \ State F H (Frame f (Suc pc) (OpDyn x # drop ar \) # st)" | + State F H (Frame f l pc R \ # st) \ State F H (Frame f l (Suc pc) R (OpDyn x # drop ar \) # st)" | step_op_inl: - "F_get F f = Some fd \ pc < length (body fd) \ body fd ! pc = IOp op \ + "next_instr (F_get F) f l pc = Some (IOp op) \ \\\\\ op = ar \ ar \ length \ \ - traverse cast_Dyn (take ar \) = Some \' \ + ap_map_list cast_Dyn (take ar \) = Some \' \ \\\ op \' = Some opinl \ \\\\\ opinl \' = x \ - F_add F f (rewrite_fundef_body fd pc (IOpInl opinl)) = F' \ - State F H (Frame f pc \ # st) \ State F' H (Frame f (Suc pc) (OpDyn x # drop ar \) # st)" | + F' = Fenv.map_entry F f (\fd. rewrite_fundef_body fd l pc (IOpInl opinl)) \ + State F H (Frame f l pc R \ # st) \ State F' H (Frame f l (Suc pc) R (OpDyn x # drop ar \) # st)" | step_op_inl_hit: - "F_get F f = Some fd \ pc < length (body fd) \ body fd ! pc = IOpInl opinl \ + "next_instr (F_get F) f l pc = Some (IOpInl opinl) \ \\\\\ (\
\\\\ opinl) = ar \ ar \ length \ \ - traverse cast_Dyn (take ar \) = Some \' \ + ap_map_list cast_Dyn (take ar \) = Some \' \ \\\\\ opinl \' \ \\\\\ opinl \' = x \ - State F H (Frame f pc \ # st) \ State F H (Frame f (Suc pc) (OpDyn x # drop ar \) # st)" | + State F H (Frame f l pc R \ # st) \ State F H (Frame f l (Suc pc) R (OpDyn x # drop ar \) # st)" | step_op_inl_miss: - "F_get F f = Some fd \ pc < length (body fd) \ body fd ! pc = IOpInl opinl \ + "next_instr (F_get F) f l pc = Some (IOpInl opinl) \ \\\\\ (\
\\\\ opinl) = ar \ ar \ length \ \ - traverse cast_Dyn (take ar \) = Some \' \ + ap_map_list cast_Dyn (take ar \) = Some \' \ \ \\\\\ opinl \' \ \\\\\ opinl \' = x \ - F_add F f (rewrite_fundef_body fd pc (IOp (\
\\\\ opinl))) = F' \ - State F H (Frame f pc \ # st) \ State F' H (Frame f (Suc pc) (OpDyn x # drop ar \) # st)" | + F' = Fenv.map_entry F f (\fd. rewrite_fundef_body fd l pc (IOp (\
\\\\ opinl))) \ + State F H (Frame f l pc R \ # st) \ State F' H (Frame f l (Suc pc) R (OpDyn x # drop ar \) # st)" | step_op_ubx: - "F_get F f = Some fd \ pc < length (body fd) \ body fd ! pc = IOpUbx opubx \ + "next_instr (F_get F) f l pc = Some (IOpUbx opubx) \ \
\\\\ (\\\ opubx) = op \ \\\\\ op = ar \ ar \ length \ \ \\\\\ opubx (take ar \) = Some x \ - State F H (Frame f pc \ # st) \ State F H (Frame f (Suc pc) (x # drop ar \) # st)" | - - step_cjump_true: - "F_get F f = Some fd \ pc < length (body fd) \ body fd ! pc = ICJump n \ - cast_Dyn y = Some d \ is_true d \ - State F H (Frame f pc (y # \) # st) \ State F H (Frame f n \ # st)" | - - step_cjump_false: - "F_get F f = Some fd \ pc < length (body fd) \ body fd ! pc = ICJump n \ - cast_Dyn y = Some d \ is_false d \ - State F H (Frame f pc (y # \) # st) \ State F H (Frame f (Suc pc) \ # st)" | - - step_fun_call: - "F_get F f = Some fd \ pc < length (body fd) \ body fd ! pc = ICall g \ - F_get F g = Some gd \ arity gd = ar \ ar \ length \ \ - frame\<^sub>f = Frame f pc \ \ list_all is_dyn_operand (take ar \) \ - frame\<^sub>g = Frame g 0 (take ar \) \ - State F H (frame\<^sub>f # st) \ State F H (frame\<^sub>g # frame\<^sub>f # st)" | + State F H (Frame f l pc R \ # st) \ State F H (Frame f l (Suc pc) R (x # drop ar \) # st)" | - step_fun_end: - "F_get F g = Some gd \ arity gd \ length \\<^sub>f \ pc\<^sub>g = length (body gd) \ - frame\<^sub>g = Frame g pc\<^sub>g \\<^sub>g \ frame\<^sub>f = Frame f pc\<^sub>f \\<^sub>f \ - frame\<^sub>f' = Frame f (Suc pc\<^sub>f) (\\<^sub>g @ drop (arity gd) \\<^sub>f) \ - State F H (frame\<^sub>g # frame\<^sub>f # st) \ State F H (frame\<^sub>f' # st)" + step_cjump: + "next_instr (F_get F) f l pc = Some (ICJump l\<^sub>t l\<^sub>f) \ + cast_Dyn y = Some d \ is_true d \ l' = l\<^sub>t \ is_false d \ l' = l\<^sub>f \ + State F H (Frame f l pc R (y # \) # st) \ State F H (Frame f l' 0 R \ # st)" | -theorem step_deterministic: "s1 \ s2 \ s1 \ s3 \ s2 = s3" - by (induction rule: step.induct; - erule step.cases; safe; - auto dest: is_true_and_is_false_implies_False) + step_call: + "next_instr (F_get F) f l pc = Some (ICall g) \ + F_get F g = Some gd \ arity gd \ length \ \ + list_all is_dyn_operand (take (arity gd) \) \ + frame\<^sub>g = allocate_frame g gd (take (arity gd) \) (OpDyn uninitialized) \ + State F H (Frame f l pc R\<^sub>f \ # st) \ State F H (frame\<^sub>g # Frame f l pc R\<^sub>f \ # st)" | -lemma final_finished: "final s \ finished step s" + step_return: + "next_instr (F_get F) g l\<^sub>g pc\<^sub>g = Some IReturn \ + F_get F g = Some gd \ arity gd \ length \\<^sub>f \ + length \\<^sub>g = return gd \ list_all is_dyn_operand \\<^sub>g \ + frame\<^sub>f' = Frame f l\<^sub>f (Suc pc\<^sub>f) R\<^sub>f (\\<^sub>g @ drop (arity gd) \\<^sub>f) \ + State F H (Frame g l\<^sub>g pc\<^sub>g R\<^sub>g \\<^sub>g # Frame f l\<^sub>f pc\<^sub>f R\<^sub>f \\<^sub>f # st) \ State F H (frame\<^sub>f' # st)" + +lemma step_deterministic: + assumes "s1 \ s2" and "s1 \ s3" + shows "s2 = s3" + using assms + apply (cases s1 s2 rule: step.cases) + apply (auto simp: next_instr_length_instrs elim: step.cases dest: is_true_and_is_false_implies_False) [3] + apply (auto simp: next_instr_length_instrs elim: step.cases dest: is_true_and_is_false_implies_False) [3] + apply (auto simp: next_instr_length_instrs elim: step.cases dest: is_true_and_is_false_implies_False) [3] + apply (auto simp: next_instr_length_instrs elim: step.cases dest: is_true_and_is_false_implies_False) [3] + apply (auto simp: next_instr_length_instrs elim: step.cases dest: is_true_and_is_false_implies_False) [3] + apply (auto simp: next_instr_length_instrs elim: step.cases dest: is_true_and_is_false_implies_False) [3] + apply (auto simp: next_instr_length_instrs elim!: step.cases dest: is_true_and_is_false_implies_False) [3] + apply (auto simp: next_instr_length_instrs elim!: step.cases dest: is_true_and_is_false_implies_False) [1] + done + +lemma step_right_unique: "right_unique step" + using step_deterministic + by (rule right_uniqueI) + +lemma final_finished: + assumes "final F_get IReturn s" + shows "finished step s" unfolding finished_def -proof - assume "final s" and "\x. step s x" - then obtain s' where "step s s'" - by auto +proof (rule notI) + assume "\x. step s x" + then obtain s' where "step s s'" by auto thus False - using \final s\ - by (auto elim!: step.cases final.cases) + using \final F_get IReturn s\ + by (auto + simp: state_ok_def + wf_fundefs_imp_wf_fundef[THEN wf_fundef_last_basic_block, THEN next_instr_length_instrs[rotated]] + elim!: step.cases final.cases) qed -sublocale ubx_sem: semantics step final +sublocale ubx_sem: semantics step "final F_get IReturn" using final_finished step_deterministic by unfold_locales -inductive load :: "('fenv, 'a, 'fun) prog \ ('fenv, 'a, ('fun, 'b) frame) state \ bool" where - "F_get F main = Some fd \ arity fd = 0 \ - load (Prog F H main) (State F H [Frame main 0 []])" +definition load where + "load \ Global.load F_get (OpDyn uninitialized)" -sublocale inca_lang: language step final load +sublocale inca_lang: language step "final F_get IReturn" load by unfold_locales end end \ No newline at end of file diff --git a/thys/Interpreter_Optimizations/Ubx_Verification.thy b/thys/Interpreter_Optimizations/Ubx_Verification.thy new file mode 100644 --- /dev/null +++ b/thys/Interpreter_Optimizations/Ubx_Verification.thy @@ -0,0 +1,540 @@ +theory Ubx_Verification + imports "HOL-Library.Sublist" Ubx Map_Extra +begin + +lemma f_g_eq_f_imp_f_comp_g_eq_f[simp]: "(\x. f (g x) = f x) \ (f \ g) = f" + by (simp add: comp_def) + +context ubx begin + +inductive sp_instr for F ret where + Push: + "sp_instr F ret (IPush d) \ (None # \)" | + PushUbx1: + "sp_instr F ret (IPushUbx1 u) \ (Some Ubx1 # \)" | + PushUbx2: + "sp_instr F ret (IPushUbx2 u) \ (Some Ubx2 # \)" | + Pop: + "sp_instr F ret IPop (\ # \) \" | + Get: + "sp_instr F ret (IGet n) \ (None # \)" | + GetUbx: + "sp_instr F ret (IGetUbx \ n) \ (Some \ # \)" | + Set: + "sp_instr F ret (ISet n) (None # \) \" | + SetUbx: + "sp_instr F ret (ISetUbx \ n) (Some \ # \) \" | + Load: + "sp_instr F ret (ILoad x) (None # \) (None # \)" | + LoadUbx: + "sp_instr F ret (ILoadUbx \ x) (None # \) (Some \ # \)" | + Store: + "sp_instr F ret (IStore x) (None # None # \) \" | + StoreUbx: + "sp_instr F ret (IStoreUbx \ x) (None # Some \ # \) \" | + Op: + "\i = (replicate (\\\\\ op) None @ \) \ + sp_instr F ret (IOp op) \i (None # \)" | + OpInl: + "\i = (replicate (\\\\\ (\
\\\\ opinl)) None @ \) \ + sp_instr F ret (IOpInl opinl) \i (None # \)" | + OpUbx: + "\i = fst (\\\\\\\\ opubx) @ \ \ result = snd (\\\\\\\\ opubx) \ + sp_instr F ret (IOpUbx opubx) \i (result # \)" | + CJump: + "sp_instr F ret (ICJump l\<^sub>t l\<^sub>f) [None] []" | + Call: + "F f = Some (ar, r) \ \i = replicate ar None @ \ \ \o = replicate r None @ \ \ + sp_instr F ret (ICall f) \i \o" | + Return: "\i = replicate ret None \ + sp_instr F ret IReturn \i []" + +inductive sp_instrs for F ret where + Nil: + "sp_instrs F ret [] \ \" | + Cons: + "sp_instr F ret instr \i \ \ sp_instrs F ret instrs \ \o \ + sp_instrs F ret (instr # instrs) \i \o" + +lemmas sp_instrs_ConsE = sp_instrs.cases[of _ _ "x # xs" for x xs, simplified] + +lemma sp_instrs_ConsD: + assumes "sp_instrs F ret (instr # instrs) \i \o" + shows "\\. sp_instr F ret instr \i \ \ sp_instrs F ret instrs \ \o" + using assms + by (auto elim: sp_instrs_ConsE) + +lemma sp_instr_deterministic: + assumes + "sp_instr F ret instr \i \o" and + "sp_instr F ret instr \i \o'" + shows "\o = \o'" + using assms by (auto elim!: sp_instr.cases) + +lemma sp_instr_right_unique: "right_unique (\(instr, \i) \o. sp_instr F ret instr \i \o)" + by (auto intro!: right_uniqueI intro: sp_instr_deterministic) + +lemma sp_instrs_deterministic: + assumes + "sp_instrs F ret instr \i \o" and + "sp_instrs F ret instr \i \o'" + shows "\o = \o'" + using assms +proof (induction instr \i \o arbitrary: \o' rule: sp_instrs.induct) + case (Nil \) + then show ?case + by (auto elim: sp_instrs.cases) +next + case (Cons instr \i \ instrs \o) + from Cons.prems obtain \' where "sp_instr F ret instr \i \'" and "sp_instrs F ret instrs \' \o'" + by (auto elim: sp_instrs.cases) + hence "\' = \" and "sp_instrs F ret instrs \' \o'" + using Cons.hyps + by (auto intro: sp_instr_deterministic) + then show ?case + by (auto intro: Cons.IH) +qed + +fun fun_call_in_range where + "fun_call_in_range F (ICall f) \ f \ dom F" | + "fun_call_in_range F instr \ True" + +lemma fun_call_in_range_generalize_instr[simp]: + "fun_call_in_range F (generalize_instr instr) \ fun_call_in_range F instr" + by (induction instr; simp) + +lemma sp_instr_complete: + assumes "fun_call_in_range F instr" + shows "\\i \o. sp_instr F ret instr \i \o" + using assms + by (cases instr; auto intro: sp_instr.intros) + +lemma sp_instr_Op2I: + assumes "\\\\\ op = 2" + shows "sp_instr F ret (IOp op) (None # None # \) (None # \)" + using sp_instr.Op[of _ op] + unfolding assms numeral_2_eq_2 + by auto + +lemma + assumes + F_def: "F = Map.empty" and + arity_op: "\\\\\ op = 2" + shows "sp_instrs F 1 [IPush x, IPush y, IOp op, IReturn] [] []" + by (auto simp: arity_op numeral_2_eq_2 + intro!: sp_instrs.intros + intro: sp_instr.intros sp_instr_Op2I) + +lemma sp_intrs_NilD[dest]: "sp_instrs F ret [] \i \o \ \i = \o" + by (auto elim: sp_instrs.cases) + +lemma sp_instrs_list_update: + assumes + "sp_instrs F ret instrs \i \o" and + "sp_instr F ret (instrs!n) = sp_instr F ret instr" + shows "sp_instrs F ret (instrs[n := instr]) \i \o" + using assms +proof (induction instrs \i \o arbitrary: n rule: sp_instrs.induct) + case (Nil \) + thus ?case by (auto intro: sp_instrs.Nil) +next + case (Cons instr \i \ instrs \o) + show ?case + proof (cases n) + case 0 + thus ?thesis + using Cons.hyps Cons.prems[unfolded 0, simplified] + by (auto intro: sp_instrs.Cons) + next + case (Suc n') + then show ?thesis + using Cons.hyps Cons.prems[unfolded Suc, simplified, THEN Cons.IH] + by (auto intro: sp_instrs.Cons) + qed +qed + +lemma sp_instrs_appendD: + assumes "sp_instrs F ret (instrs1 @ instrs2) \i \o" + shows "\\. sp_instrs F ret instrs1 \i \ \ sp_instrs F ret instrs2 \ \o" + using assms +proof (induction instrs1 arbitrary: \i) + case Nil + thus ?case by (auto intro: sp_instrs.Nil) +next + case (Cons instr instrs1) + then obtain \ where "sp_instr F ret instr \i \" and "sp_instrs F ret (instrs1 @ instrs2) \ \o" + by (auto elim: sp_instrs.cases) + thus ?case + by (auto intro: sp_instrs.Cons dest: Cons.IH) +qed + +lemma sp_instrs_appendD': + assumes "sp_instrs F ret (instrs1 @ instrs2) \i \o" and "sp_instrs F ret instrs1 \i \" + shows "sp_instrs F ret instrs2 \ \o" +proof - + obtain \' where + sp_instrs1: "sp_instrs F ret instrs1 \i \'" and + sp_instrs2: "sp_instrs F ret instrs2 \' \o" + using sp_instrs_appendD[OF assms(1)] + by auto + have "\' = \" + using sp_instrs_deterministic[OF assms(2) sp_instrs1] by simp + thus ?thesis + using sp_instrs2 by simp +qed + +lemma sp_instrs_appendI[intro]: + assumes "sp_instrs F ret instrs1 \i \" and "sp_instrs F ret instrs2 \ \o" + shows "sp_instrs F ret (instrs1 @ instrs2) \i \o" + using assms + by (induction instrs1 \i \ rule: sp_instrs.induct) (auto intro: sp_instrs.Cons) + +lemma sp_instrs_singleton_conv[simp]: + "sp_instrs F ret [instr] \i \o \ sp_instr F ret instr \i \o" + by (auto intro: sp_instrs.intros elim: sp_instrs.cases) + +lemma sp_instrs_singletonI: + assumes "sp_instr F ret instr \i \o" + shows "sp_instrs F ret [instr] \i \o" + using assms by (auto intro: sp_instrs.intros) + +fun local_var_in_range where + "local_var_in_range n (IGet k) \ k < n" | + "local_var_in_range n (IGetUbx \ k) \ k < n" | + "local_var_in_range n (ISet k) \ k < n" | + "local_var_in_range n (ISetUbx \ k) \ k < n" | + "local_var_in_range _ _ \ True" + +lemma local_var_in_range_generalize_instr[simp]: + "local_var_in_range n (generalize_instr instr) \ local_var_in_range n instr" + by (cases instr; simp) + +lemma local_var_in_range_comp_generalize_instr[simp]: + "local_var_in_range n \ generalize_instr = local_var_in_range n" + using local_var_in_range_generalize_instr + by auto + +fun jump_in_range where + "jump_in_range L (ICJump l\<^sub>t l\<^sub>f) \ {l\<^sub>t, l\<^sub>f} \ L" | + "jump_in_range L _ \ True" + +inductive wf_basic_block for F L ret n where + "instrs \ [] \ + list_all (local_var_in_range n) instrs \ + list_all (fun_call_in_range F) instrs \ + list_all (jump_in_range L) instrs \ + list_all (\i. \ is_jump i \ \ is_return i) (butlast instrs) \ + sp_instrs F ret instrs [] [] \ + wf_basic_block F L ret n (label, instrs)" + +lemmas wf_basic_blockI = wf_basic_block.simps[THEN iffD2] +lemmas wf_basic_blockD = wf_basic_block.simps[THEN iffD1] + +definition wf_fundef where + "wf_fundef F fd \ + body fd \ [] \ + list_all + (wf_basic_block F (fst ` set (body fd)) (return fd) (arity fd + fundef_locals fd)) + (body fd)" + +lemmas wf_fundefI = wf_fundef_def[THEN iffD2, OF conjI] +lemmas wf_fundefD = wf_fundef_def[THEN iffD1] +lemmas wf_fundef_body_neq_NilD = wf_fundefD[THEN conjunct1] +lemmas wf_fundef_all_wf_basic_blockD = wf_fundefD[THEN conjunct2] + +definition wf_fundefs where + "wf_fundefs F \ (\f. pred_option (wf_fundef (map_option funtype \ F)) (F f))" + +lemmas wf_fundefsI = wf_fundefs_def[THEN iffD2] +lemmas wf_fundefsD = wf_fundefs_def[THEN iffD1] + +lemma wf_fundefs_getD: + shows "wf_fundefs F \ F f = Some fd \ wf_fundef (map_option funtype \ F) fd" + by (auto dest!: wf_fundefsD[THEN spec, of _ f]) + +definition wf_prog where + "wf_prog p \ wf_fundefs (F_get (prog_fundefs p))" + +definition wf_state where + "wf_state s \ wf_fundefs (F_get (state_fundefs s))" + +lemmas wf_stateI = wf_state_def[THEN iffD2] +lemmas wf_stateD = wf_state_def[THEN iffD1] + +lemma sp_instr_generalize0: + assumes "sp_instr F ret instr \i \o" and + "\i' = map (\_. None) \i" and "\o' = map (\_. None) \o" + shows "sp_instr F ret (generalize_instr instr) \i' \o'" + using assms +proof (induction instr \i \o rule: sp_instr.induct) + case (OpUbx \i opubx \ result) + then show ?case + apply simp + unfolding map_replicate_const + unfolding \\\\\\\\_\\\\\[symmetric] + by (auto intro: sp_instr.OpInl) +qed (auto simp: intro: sp_instr.intros) + +lemma sp_instrs_generalize0: + assumes "sp_instrs F ret instrs \i \o" and + "\i' = map (\_. None) \i" and "\o' = map (\_. None) \o" + shows "sp_instrs F ret (map generalize_instr instrs) \i' \o'" + using assms +proof (induction instrs \i \o arbitrary: \i' \o' rule: sp_instrs.induct) + case (Nil \) + then show ?case + by (auto intro: sp_instrs.Nil) +next + case (Cons instr \i \ instrs \o) + then show ?case + by (auto intro: sp_instrs.Cons sp_instr_generalize0) +qed + +lemmas sp_instr_generalize = sp_instr_generalize0[OF _ refl refl] +lemmas sp_instr_generalize_Nil_Nil = sp_instr_generalize[of _ _ _ "[]" "[]", simplified] +lemmas sp_instrs_generalize = sp_instrs_generalize0[OF _ refl refl] +lemmas sp_instrs_generalize_Nil_Nil = sp_instrs_generalize[of _ _ _ "[]" "[]", simplified] + +lemma jump_in_range_generalize_instr[simp]: + "jump_in_range L (generalize_instr instr) \ jump_in_range L instr" + by (cases instr) simp_all + +lemma wf_basic_block_map_generalize_instr: + assumes "wf_basic_block F L ret n (label, instrs)" + shows "wf_basic_block F L ret n (label, map generalize_instr instrs)" + using assms + by (auto simp add: wf_basic_block.simps list.pred_map map_butlast[symmetric] + intro: sp_instrs_generalize_Nil_Nil) + +lemma list_all_wf_basic_block_generalize_fundef: + assumes "list_all (wf_basic_block F L ret n) (body fd)" + shows "list_all (wf_basic_block F L ret n) (body (generalize_fundef fd))" +proof (cases fd) + case (Fundef bblocks ar ret locals) + then show ?thesis + using assms + by (auto simp: map_ran_def list.pred_map comp_def case_prod_beta + intro: wf_basic_block_map_generalize_instr + elim!: list.pred_mono_strong) +qed + +lemma wf_fundefs_map_entry: + assumes wf_F: "wf_fundefs (F_get F)" and + same_funtype: "\fd. fd \ ran (F_get F) \ funtype (f fd) = funtype fd" and + same_arity: "\fd. F_get F x = Some fd \ arity (f fd) = arity fd" and + same_return: "\fd. F_get F x = Some fd \ return (f fd) = return fd" and + same_body_length: "\fd. F_get F x = Some fd \ length (body (f fd)) = length (body fd)" and + same_locals: "\fd. F_get F x = Some fd \ fundef_locals (f fd) = fundef_locals fd" and + same_labels: "\fd. F_get F x = Some fd \ fst ` set (body (f fd)) = fst ` set (body fd)" and + list_all_wf_basic_block_f: "\fd. + F_get F x = Some fd \ + list_all (wf_basic_block (map_option funtype \ F_get F) (fst ` set (body fd)) (return fd) + (arity fd + fundef_locals fd)) (body fd) \ + list_all (wf_basic_block (map_option funtype \ F_get F) (fst ` set (body fd)) (return fd) + (arity fd + fundef_locals fd)) (body (f fd))" + shows "wf_fundefs (F_get (Fenv.map_entry F x f))" +proof (intro wf_fundefsI allI) + fix y + let ?F' = "F_get (Fenv.map_entry F x f)" + show "pred_option (wf_fundef (map_option funtype \ ?F')) (?F' y)" + proof (cases "F_get F y") + case None + then show ?thesis + by (simp add: Fenv.get_map_entry_conv) + next + case (Some gd) + hence wf_gd: "wf_fundef (map_option funtype \ F_get F) gd" + using wf_F[THEN wf_fundefsD, THEN spec, of y] by simp + then show ?thesis + proof (cases "x = y") + case True + moreover have "wf_fundef (map_option funtype \ F_get (Fenv.map_entry F y f)) (f gd)" + proof (rule wf_fundefI) + show "body (f gd) \ []" + using same_body_length[unfolded True, OF Some] + using wf_fundef_body_neq_NilD[OF wf_gd] + by auto + next + show "list_all (wf_basic_block (map_option funtype \ F_get (Fenv.map_entry F y f)) + (fst ` set (body (f gd))) (return (f gd)) (arity (f gd) + fundef_locals (f gd))) (body (f gd))" + using Some True + using wf_gd[THEN wf_fundefD] Some[unfolded True] + by (auto simp: Fenv.map_option_comp_map_entry ranI assms(2-8) intro!: wf_fundefI) + qed + ultimately show ?thesis + by (simp add: Some) + next + case False + then show ?thesis + unfolding Fenv.get_map_entry_neq[OF False] + unfolding Some option.pred_inject + using wf_gd[THEN wf_fundefD] + using same_funtype + by (auto simp: Fenv.map_option_comp_map_entry intro!: wf_fundefI) + qed + qed +qed + +lemma wf_fundefs_generalize: + assumes wf_F: "wf_fundefs (F_get F)" + shows "wf_fundefs (F_get (Fenv.map_entry F f generalize_fundef))" + using wf_F + apply (elim wf_fundefs_map_entry) + by (auto elim: list_all_wf_basic_block_generalize_fundef) + +lemma list_all_wf_basic_block_rewrite_fundef_body: + assumes + "list_all (wf_basic_block F L ret n) (body fd)" and + "instr_at fd l pc = Some instr" and + sp_instr_eq: "sp_instr F ret instr = sp_instr F ret instr'" and + local_var_in_range_iff: "local_var_in_range n instr' \ local_var_in_range n instr" and + fun_call_in_range_iff: "fun_call_in_range F instr' \ fun_call_in_range F instr" and + jump_in_range_iff: "jump_in_range L instr' \ jump_in_range L instr" and + is_jump_iff: "is_jump instr' \ is_jump instr" and + is_return_iff: "is_return instr' \ is_return instr" + shows "list_all (wf_basic_block F L ret n) (body (rewrite_fundef_body fd l pc instr'))" +proof (cases fd) + case (Fundef bblocks ar ret' locals) + have wf_bblocks: "list_all (wf_basic_block F L ret n) bblocks" + using assms(1) unfolding Fundef by simp + + moreover have "wf_basic_block F L ret n (l, instrs[pc := instr'])" + if "map_of bblocks l = Some instrs" for instrs + proof - + have wf_basic_block_l_instrs: "wf_basic_block F L ret n (l, instrs)" + by (rule list_all_map_of_SomeD[OF wf_bblocks that]) + + have nth_instrs_pc: "pc < length instrs" "instrs ! pc = instr" + using assms(2)[unfolded instr_at_def Fundef, simplified, unfolded that, simplified] + by simp_all + + show ?thesis + proof (rule wf_basic_blockI, simp, intro conjI) + show "instrs \ []" + using wf_basic_block_l_instrs + by (auto elim: wf_basic_block.cases) + next + show "list_all (local_var_in_range n) (instrs[pc := instr'])" + using wf_basic_block_l_instrs nth_instrs_pc + by (auto simp: local_var_in_range_iff + elim!: wf_basic_block.cases intro!: list_all_list_updateI) + next + show "list_all (\i. \ is_jump i \ \ is_return i) (butlast (instrs[pc := instr']))" + using wf_basic_block_l_instrs nth_instrs_pc + apply (auto simp: butlast_list_update elim!: wf_basic_block.cases) + apply (rule list_all_list_updateI) + apply assumption + by (simp add: is_jump_iff is_return_iff list_all_length nth_butlast) + next + show "sp_instrs F ret (instrs[pc := instr']) [] []" + using wf_basic_block_l_instrs nth_instrs_pc sp_instr_eq + by (auto elim!: wf_basic_block.cases elim!: sp_instrs_list_update) + next + show "list_all (fun_call_in_range F) (instrs[pc := instr'])" + using wf_basic_block_l_instrs nth_instrs_pc + by (auto simp: fun_call_in_range_iff + elim!: wf_basic_block.cases intro!: list_all_list_updateI) + next + show "list_all (jump_in_range L) (instrs[pc := instr'])" + using wf_basic_block_l_instrs nth_instrs_pc + by (auto simp: jump_in_range_iff elim!: wf_basic_block.cases intro!: list_all_list_updateI) + qed + qed + + with Fundef show ?thesis + using assms(1,2) + by (auto simp add: rewrite_fundef_body_def map_entry_map_of_Some_conv + intro: list_all_updateI dest!: instr_atD) +qed + +lemma wf_fundefs_rewrite_body: + assumes "wf_fundefs (F_get F)" and + "next_instr (F_get F) f l pc = Some instr" and + sp_instr_eq: "\ret. + sp_instr (map_option funtype \ F_get F) ret instr' = + sp_instr (map_option funtype \ F_get F) ret instr" and + local_var_in_range_iff: "\n. local_var_in_range n instr' \ local_var_in_range n instr" and + fun_call_in_range_iff: + "fun_call_in_range (map_option funtype \ F_get F) instr' \ + fun_call_in_range (map_option funtype \ F_get F) instr" and + jump_in_range_iff: "\L. jump_in_range L instr' \ jump_in_range L instr" and + is_jump_iff: "is_jump instr' \ is_jump instr" and + is_return_iff: "is_return instr' \ is_return instr" + shows "wf_fundefs (F_get (Fenv.map_entry F f (\fd. rewrite_fundef_body fd l pc instr')))" +proof - + obtain fd where F_f: "F_get F f = Some fd" and instr_at_pc: "instr_at fd l pc = Some instr" + using assms(2)[THEN next_instrD] + by auto + + show ?thesis + using assms(1) + proof (rule wf_fundefs_map_entry) + fix fd + show "fst ` set (body (rewrite_fundef_body fd l pc instr')) = fst ` set (body fd)" + by (cases fd) (simp add: rewrite_fundef_body_def dom_map_entry) + next + fix gd + assume "F_get F f = Some gd" + hence "gd = fd" + using F_f by simp + then show + "list_all (wf_basic_block (map_option funtype \ F_get F) (fst ` set (body gd)) (return gd) + (arity gd + fundef_locals gd)) (body gd) \ + list_all (wf_basic_block (map_option funtype \ F_get F) (fst ` set (body gd)) (return gd) + (arity gd + fundef_locals gd)) (body (rewrite_fundef_body gd l pc instr'))" + using assms(1,3-) + using F_f instr_at_pc + by (elim list_all_wf_basic_block_rewrite_fundef_body) simp_all + qed simp_all +qed + +lemma sp_instr_Op_OpInl_conv: + assumes "op = \
\\\\ opinl" + shows "sp_instr F ret (IOp op) = sp_instr F ret (IOpInl opinl)" + by (auto simp: assms elim: sp_instr.cases intro: sp_instr.OpInl sp_instr.Op) + +lemma wf_state_step_preservation: + assumes "wf_state s" and "step s s'" + shows "wf_state s'" + using assms(2,1) +proof (cases s s' rule: step.cases) + case (step_op_inl F f l pc op ar \ \' opinl x F' H R st) + then show ?thesis + using assms(1) + by (auto intro!: wf_stateI intro: sp_instr_Op_OpInl_conv[symmetric] + elim!: wf_fundefs_rewrite_body dest!: wf_stateD \\\_invertible) +next + case (step_op_inl_miss F f l pc opinl ar \ \' x F' H R st) + then show ?thesis + using assms(1) + by (auto intro!: wf_stateI intro: sp_instr_Op_OpInl_conv + elim!: wf_fundefs_rewrite_body dest!: wf_stateD) +qed (auto simp: box_stack_def + intro!: wf_stateI wf_fundefs_generalize + intro: sp_instr.intros + dest!: wf_stateD) + +(* lemma + assumes "wf_state s" + shows "\s'. step s s'" +proof (cases s) + case (State F H frames) + hence "frames \ []" and wf_fundefs_bblocks: "wf_fundefs (F_get F)" + using assms(1)[THEN wf_stateD] by simp_all + then obtain f l pc locals \ frames' where "frames = Frame f l pc locals \ # frames'" + unfolding neq_Nil_conv by (metis frame.exhaust) + + show ?thesis + proof (cases "next_instr (F_get F) f l pc") + case None + then show ?thesis sorry + next + case (Some a) + then show ?thesis sorry + qed +qed *) + + +end + +end \ No newline at end of file diff --git a/thys/Interpreter_Optimizations/Ubx_type_inference.thy b/thys/Interpreter_Optimizations/Ubx_type_inference.thy deleted file mode 100644 --- a/thys/Interpreter_Optimizations/Ubx_type_inference.thy +++ /dev/null @@ -1,210 +0,0 @@ -theory Ubx_type_inference - imports Result Ubx - "HOL-Library.Monad_Syntax" -begin - -section \Type of operations\ - -locale ubx_sp = - ubx - F_empty F_get F_add F_to_list - heap_empty heap_get heap_add heap_to_list - is_true is_false box_ubx1 unbox_ubx1 box_ubx2 unbox_ubx2 - \\ \\\\\ \\\\\ \\\ \\\\\ \
\\\\ \\\\\ \\\ \\\ \\\\\\\\ - for - \ \Functions environment\ - F_empty and F_get and F_add and F_to_list and - - \ \Memory heap\ - heap_empty and heap_get and heap_add and heap_to_list and - - \ \Unboxed values\ - is_true and is_false and - box_ubx1 and unbox_ubx1 and - box_ubx2 and unbox_ubx2 and - - \ \n-ary operations\ - \\ and \\\\\ and \\\\\ and \\\ and \\\\\ and \
\\\\ and \\\\\ and \\\ and \\\ and \\\\\\\\ -begin - - -section \Strongest postcondition of instructions\ - -fun sp_gen_pop_push where - "sp_gen_pop_push (domain, codom) \ = ( - let ar = length domain in - if ar \ length \ \ take ar \ = domain then - Ok (codom # drop ar \) - else - Error () - )" - -fun sp_instr :: "('fun \ _ fundef option) \ _ \ type option list \ (unit, type option list) result" where - "sp_instr _ (IPush _) \ = Ok (None # \)" | - "sp_instr _ (IPushUbx1 _) \ = Ok (Some Ubx1 # \)" | - "sp_instr _ (IPushUbx2 _) \ = Ok (Some Ubx2 # \)" | - "sp_instr _ IPop (_ # \) = Ok \" | - "sp_instr _ (ILoad _) (None # \) = Ok (None # \)" | - "sp_instr _ (ILoadUbx \ _) (None # \) = Ok (Some \ # \)" | - "sp_instr _ (IStore _) (None # None # \) = Ok \" | - "sp_instr _ (IStoreUbx \\<^sub>1 _) (None # Some \\<^sub>2 # \) = (if \\<^sub>1 = \\<^sub>2 then Ok \ else Error ())" | - "sp_instr _ (IOp op) \ = sp_gen_pop_push (replicate (\\\\\ op) None, None) \" | - "sp_instr _ (IOpInl opinl) \ = sp_gen_pop_push (replicate (\\\\\ (\
\\\\ opinl)) None, None) \" | - "sp_instr _ (IOpUbx opubx) \ = sp_gen_pop_push (\\\\\\\\ opubx) \" | - "sp_instr F (ICall f) \ = do { - fd \ Result.of_option () (F f); - sp_gen_pop_push (replicate (arity fd) None, None) \ - }" | - "sp_instr _ _ _ = Error ()" - -lemma sp_instr_no_jump: "sp_instr F instr \ = Ok type \ \ is_jump instr" - by (induction instr \ rule: sp_instr.induct) simp_all - -lemma map_constant[simp]: "\x \ set xs. x = y \ map (\_. y) xs = xs" - by (simp add: map_idI) - -lemma sp_generalize_instr: - assumes "sp_instr F x \ = Ok \'" - shows "sp_instr F (generalize_instr x) (map (\_. None) \) = Ok (map (\_. None) \')" - using assms - apply (induction F x \ rule: sp_instr.induct; - (auto simp: Let_def take_map drop_map; fail)?) - subgoal for _ opubx - apply (cases "\\\\\\\\ opubx") - by (auto simp add: Let_def take_map drop_map map_replicate_const \\\\\\\\_\\\\\) - done - -lemma map_typeof_box: "map typeof (map box_operand \) = replicate (length \) None" -proof (induction \) - case Nil - then show ?case by simp -next - case (Cons x xs) - then show ?case - by (cases x) simp_all -qed - -section \Strongest postcondition of function definitions\ - -fun sp :: "('fun \ _ fundef option) \ _ list \ type option list \ (unit, type option list) result" where - "sp _ [] \ = Ok \" | - "sp F (instr # p) \ = do { - \' \ sp_instr F instr \; - sp F p \' - }" - -lemma sp_no_jump: "sp F xs \ = Ok type \ \x \ set xs. \ is_jump x" - by (induction xs arbitrary: \; auto dest: sp_instr_no_jump) - -lemma sp_append: "sp F (xs @ ys) \ = sp F xs \ \ sp F ys" -proof (induction xs arbitrary: \) - case Nil - then show ?case by simp -next - case (Cons x xs) - then show ?case - by (cases "sp_instr F x \"; simp) -qed - -lemmas sp_eq_bind_take_drop = - sp_append[of _ "take n xs" "drop n xs" for n xs, unfolded append_take_drop_id] - -lemma sp_generalize: - assumes "sp F xs \ = Ok \'" - shows "sp F (map generalize_instr xs) (map (\_. None) \) = Ok (map (\_. None) \')" -using assms -proof (induction xs arbitrary: \ \') - case Nil - then show ?case using assms(1) by simp -next - case (Cons x xs) - show ?case - proof (cases "sp_instr F x \") - case (Error x1) - then show ?thesis - using Cons.prems by auto - next - case (Ok \'') - show ?thesis - using Cons sp_generalize_instr[OF Ok] - by (simp add: Ok) - qed -qed - -lemma sp_generalize2: - assumes "sp F xs (replicate n None) = Ok \'" - shows "sp F (map generalize_instr xs) (replicate n None) = Ok (map (\_. None) \')" - using assms - using sp_generalize by fastforce - -lemma comp_K[simp]: "(\_. x) \ f = (\_. x)" - by auto - -lemma is_ok_sp_generalize: - assumes "is_ok (sp F xs (map (\_. None) \))" - shows "is_ok (sp F (map generalize_instr xs) (map (\_. None) \))" -proof - - from assms obtain res where 0: "sp F xs (map (\_. None) \) = Ok res" - by (auto simp add: is_ok_def) - show ?thesis - using sp_generalize[OF 0] by simp -qed - -lemma is_ok_sp_generalize2: - assumes "is_ok (sp F xs (replicate n None))" - shows "is_ok (sp F (map generalize_instr xs) (replicate n None))" - using assms is_ok_sp_generalize - by (metis Ex_list_of_length map_replicate_const) - -lemma sp_instr_op: - assumes "\
\\\\ opinl = op" - shows "sp_instr F (IOp op) \ = sp_instr F (IOpInl opinl) \" - by (rule sp_instr.cases[of "(F, (IOp op), \)"]; simp add: Let_def assms) - -lemma sp_list_update: - assumes - "n < length xs" and - "\\. sp_instr F (xs ! n) \ = sp_instr F x \" - shows "sp F (xs[n := x]) = sp F xs" -proof (intro ext) - fix ys - have "sp F (take n xs @ x # drop (Suc n) xs) ys = sp F (take n xs @ xs ! n # drop (Suc n) xs) ys" - unfolding sp_append - using assms(2) - by (cases "sp F (take n xs) ys"; simp) - thus "sp F (xs[n := x]) ys = sp F xs ys" - unfolding id_take_nth_drop[OF assms(1), symmetric] - unfolding upd_conv_take_nth_drop[OF assms(1), symmetric] - by assumption -qed - -lemma sp_list_update_eq_Ok: - assumes - "n < length xs" and - "\\. sp_instr F (xs ! n) \ = sp_instr F x \" and - "sp F xs ys = Ok zs" - shows "sp F (xs[n := x]) ys = Ok zs" - unfolding sp_list_update[OF assms(1,2)] - using assms(3) - by assumption - -lemma is_ok_sp_list_update: - assumes - "is_ok (sp F xs types)" and - "pc < length xs" and - "\\. sp_instr F (xs ! pc) \ = sp_instr F instr \" - shows "is_ok (sp F (xs[pc := instr]) types)" -proof - - from assms(1) obtain types' where "sp F xs types = Ok types'" - unfolding is_ok_def by auto - thus ?thesis - using sp_list_update[OF assms(2,3)] by simp -qed - -lemmas sp_rewrite = sp_list_update[folded rewrite_def] -lemmas sp_rewrite_eq_Ok = sp_list_update_eq_Ok[folded rewrite_def] -lemmas is_ok_sp_rewrite = is_ok_sp_list_update[folded rewrite_def] - -end - -end \ No newline at end of file diff --git a/thys/Interpreter_Optimizations/Unboxed.thy b/thys/Interpreter_Optimizations/Unboxed.thy --- a/thys/Interpreter_Optimizations/Unboxed.thy +++ b/thys/Interpreter_Optimizations/Unboxed.thy @@ -1,66 +1,66 @@ theory Unboxed imports Global Dynamic begin datatype type = Ubx1 | Ubx2 datatype ('dyn, 'ubx1, 'ubx2) unboxed = is_dyn_operand: OpDyn 'dyn | OpUbx1 'ubx1 | OpUbx2 'ubx2 fun typeof where "typeof (OpDyn _) = None" | "typeof (OpUbx1 _) = Some Ubx1" | "typeof (OpUbx2 _) = Some Ubx2" fun cast_Dyn where "cast_Dyn (OpDyn d) = Some d" | "cast_Dyn _ = None" fun cast_Ubx1 where "cast_Ubx1 (OpUbx1 x) = Some x" | "cast_Ubx1 _ = None" fun cast_Ubx2 where "cast_Ubx2 (OpUbx2 x) = Some x" | "cast_Ubx2 _ = None" -locale unboxedval = dynval is_true is_false - for is_true :: "'dyn \ bool" and is_false + +locale unboxedval = dynval uninitialized is_true is_false + for uninitialized :: 'dyn and is_true and is_false + fixes box_ubx1 :: "'ubx1 \ 'dyn" and unbox_ubx1 :: "'dyn \ 'ubx1 option" and box_ubx2 :: "'ubx2 \ 'dyn" and unbox_ubx2 :: "'dyn \ 'ubx2 option" assumes box_unbox_inverse: "unbox_ubx1 d = Some u1 \ box_ubx1 u1 = d" "unbox_ubx2 d = Some u2 \ box_ubx2 u2 = d" begin fun unbox :: "type \ 'dyn \ ('dyn, 'ubx1, 'ubx2) unboxed option" where "unbox Ubx1 = map_option OpUbx1 \ unbox_ubx1" | "unbox Ubx2 = map_option OpUbx2 \ unbox_ubx2" fun cast_and_box where "cast_and_box Ubx1 = map_option box_ubx1 \ cast_Ubx1" | "cast_and_box Ubx2 = map_option box_ubx2 \ cast_Ubx2" fun norm_unboxed where "norm_unboxed (OpDyn d) = d" | "norm_unboxed (OpUbx1 x) = box_ubx1 x" | "norm_unboxed (OpUbx2 x) = box_ubx2 x" fun box_operand where "box_operand (OpDyn d) = OpDyn d" | "box_operand (OpUbx1 x) = OpDyn (box_ubx1 x)" | "box_operand (OpUbx2 x) = OpDyn (box_ubx2 x)" fun box_frame where - "box_frame f (Frame g pc \) = Frame g pc (if f = g then map box_operand \ else \)" + "box_frame f (Frame g l pc R \) = Frame g l pc R (if f = g then map box_operand \ else \)" definition box_stack where "box_stack f \ map (box_frame f)" end end \ No newline at end of file diff --git a/thys/Interpreter_Optimizations/Unboxed_lemmas.thy b/thys/Interpreter_Optimizations/Unboxed_lemmas.thy --- a/thys/Interpreter_Optimizations/Unboxed_lemmas.thy +++ b/thys/Interpreter_Optimizations/Unboxed_lemmas.thy @@ -1,90 +1,103 @@ theory Unboxed_lemmas imports Unboxed begin +lemma cast_Dyn_eq_Some_imp_typeof: "cast_Dyn u = Some d \ typeof u = None" + by (auto elim: cast_Dyn.elims) + lemma typeof_bind_OpDyn[simp]: "typeof \ OpDyn = (\_. None)" by auto lemma is_dyn_operand_eq_typeof: "is_dyn_operand = (\x. typeof x = None)" proof (intro ext) fix x show "is_dyn_operand x = (typeof x = None)" by (cases x; simp) qed lemma is_dyn_operand_eq_typeof_Dyn[simp]: "is_dyn_operand x \ typeof x = None" by (cases x; simp) lemma typeof_unboxed_eq_const: fixes x shows "typeof x = None \ (\d. x = OpDyn d)" "typeof x = Some Ubx1 \ (\n. x = OpUbx1 n)" "typeof x = Some Ubx2 \ (\b. x = OpUbx2 b)" by (cases x; simp)+ lemmas typeof_unboxed_inversion = typeof_unboxed_eq_const[THEN iffD1] lemma cast_inversions: "cast_Dyn x = Some d \ x = OpDyn d" "cast_Ubx1 x = Some n \ x = OpUbx1 n" "cast_Ubx2 x = Some b \ x = OpUbx2 b" by (cases x; simp)+ -lemma traverse_cast_Dyn_replicate: - assumes "traverse cast_Dyn xs = Some ys" +lemma ap_map_list_cast_Dyn_replicate: + assumes "ap_map_list cast_Dyn xs = Some ys" shows "map typeof xs = replicate (length xs) None" using assms proof (induction xs arbitrary: ys) case Nil then show ?case by simp next case (Cons x xs) from Cons.prems show ?case - by (auto intro: Cons.IH dest: cast_inversions(1) simp add: bind_eq_Some_conv) + by (auto intro: Cons.IH dest: cast_inversions(1) simp: ap_option_eq_Some_conv) qed context unboxedval begin lemma unbox_typeof[simp]: "unbox \ d = Some blob \ typeof blob = Some \" by (cases \; auto) lemma cast_and_box_imp_typeof[simp]: "cast_and_box \ blob = Some d \ typeof blob = Some \" using cast_inversions[of blob] by (induction \; auto dest: cast_inversions[of blob]) lemma norm_unbox_inverse[simp]: "unbox \ d = Some blob \ norm_unboxed blob = d" using box_unbox_inverse by (cases \; auto) lemma norm_cast_and_box_inverse[simp]: "cast_and_box \ blob = Some d \ norm_unboxed blob = d" by (induction \; auto elim: cast_Dyn.elims cast_Ubx1.elims cast_Ubx2.elims) +lemma typeof_and_norm_unboxed_imp_cast_Dyn: + assumes "typeof x' = None" "norm_unboxed x' = x" + shows "cast_Dyn x' = Some x" + using assms + unfolding typeof_unboxed_eq_const + by auto + lemma typeof_and_norm_unboxed_imp_cast_and_box: assumes "typeof x' = Some \" "norm_unboxed x' = x" shows "cast_and_box \ x' = Some x" using assms by (induction \; induction x'; simp) lemma norm_unboxed_bind_OpDyn[simp]: "norm_unboxed \ OpDyn = id" by auto lemmas box_stack_Nil[simp] = list.map(1)[of "box_frame f" for f, folded box_stack_def] lemmas box_stack_Cons[simp] = list.map(2)[of "box_frame f" for f, folded box_stack_def] -lemma typeof_box_operand[simp]: "typeof (box_operand x) = None" - by (cases x; simp) +lemma typeof_box_operand[simp]: "typeof (box_operand u) = None" + by (cases u) simp_all + +lemma typeof_box_operand_comp[simp]: "typeof \ box_operand = (\_. None)" + by auto lemma is_dyn_box_operand: "is_dyn_operand (box_operand x)" by (cases x) simp_all lemma is_dyn_operand_comp_box_operand[simp]: "is_dyn_operand \ box_operand = (\_. True)" using is_dyn_box_operand by auto lemma norm_box_operand[simp]: "norm_unboxed (box_operand x) = norm_unboxed x" by (cases x) simp_all end end \ No newline at end of file