(* * Copyright 2014, NICTA * * This software may be distributed and modified according to the terms of * the BSD 2-Clause license. Note that NO WARRANTY is provided. * See "LICENSE_BSD2.txt" for details. * * @TAG(NICTA_BSD) *) theory GraphProof imports TailrecPre GraphLangLemmas "Lib.SplitRule" begin declare sep_false_def[symmetric, simp del] lemma exec_graph_step_Gamma_deterministic: assumes stacks: "(stack, stack') \ exec_graph_step Gamma" "(stack, stack'') \ exec_graph_step (adds ++ Gamma)" and adds: "\gf \ ran Gamma. \nn fname inps outps. (Call nn fname inps outps) \ ran (function_graph gf) \ fname \ dom adds" shows "stack'' = stack'" proof - have adds': "\fname fn fname' fn'. adds fname = Some fn \ Gamma fname' = Some fn' \ \x nn inps outps. function_graph fn' x \ Some (Call nn fname inps outps)" using adds by (metis ranI domI) show ?thesis using stacks by (auto simp: all_exec_graph_step_cases split: graph_function.split_asm dest: adds') qed lemmas exec_graph_step_deterministic = exec_graph_step_Gamma_deterministic[where adds=Map.empty, simplified] lemma exec_graph_n_Gamma_deterministic: "(stack, stack') \ exec_graph_n Gamma n \ (stack, stack'') \ exec_graph_n (adds ++ Gamma) n \ \gf \ ran Gamma. \nn fname inps outps. (Call nn fname inps outps) \ ran (function_graph gf) \ fname \ dom adds \ stack'' = stack'" using exec_graph_step_Gamma_deterministic apply (induct n arbitrary: stack' stack'', simp_all add: exec_graph_n_def) apply blast done lemmas exec_graph_n_deterministic = exec_graph_n_Gamma_deterministic[where adds=Map.empty, simplified] lemma step_implies_continuing: "(stack, stack') \ exec_graph_step Gamma \ continuing stack" by (simp add: exec_graph_step_def continuing_def split: list.split_asm prod.split_asm next_node.split_asm) lemma exec_trace_Gamma_deterministic: "trace \ exec_trace Gamma f \ trace' \ exec_trace (adds ++ Gamma) f \ \gf \ ran Gamma. \nn fname inps outps. (Call nn fname inps outps) \ ran (function_graph gf) \ fname \ dom adds \ trace 0 = trace' 0 \ trace n = trace' n" apply (induct n) apply simp apply (clarsimp simp: exec_trace_def nat_trace_rel_def) apply (drule_tac x=n in spec)+ apply (case_tac "trace' n", clarsimp+) apply (case_tac "trace (Suc n)", simp_all) apply (auto dest: step_implies_continuing)[1] apply (clarsimp simp: step_implies_continuing) apply (metis exec_graph_step_Gamma_deterministic) done lemmas exec_trace_deterministic = exec_trace_Gamma_deterministic[where adds=Map.empty, simplified] lemma exec_trace_nat_trace: "tr \ exec_trace Gamma f \ tr \ nat_trace_rel continuing (exec_graph_step Gamma)" by (clarsimp simp add: exec_trace_def) abbreviation(input) "exec_double_trace Gamma1 f1 Gamma2 f2 trace1 trace2 \ (trace1 \ exec_trace Gamma1 f1 \ trace2 \ exec_trace Gamma2 f2)" definition trace_refine :: "bool \ (state \ state \ bool) \ trace \ trace \ bool" where "trace_refine precise rel tr tr' = (case (trace_end tr, trace_end tr') of (None, None) \ True | (_, None) \ \ precise | (_, Some [(Err, _, _)]) \ True | (Some [(Ret, st, _)], Some [(Ret, st', _)]) \ rel st st' | _ \ False)" definition "switch_val f x y = (\z. if f z then y z else x z)" definition "tuple_switch vs = switch_val id (\_. fst vs) (\_. snd vs)" lemma tuple_switch_simps[simp]: "tuple_switch vs True = snd vs" "tuple_switch vs False = fst vs" by (auto simp: tuple_switch_def switch_val_def) definition "function_outputs_st f st = acc_vars (function_outputs f) st" definition trace_refine2 :: "bool \ graph_function \ graph_function \ ((bool \ bool \ variable list) \ bool) \ trace \ trace \ bool" where "trace_refine2 precise gfs orel tr tr' = trace_refine precise (\st st'. orel (switch_val fst (tuple_switch (acc_vars (function_inputs (fst gfs)) (fst (snd (hd (the (tr 0))))), function_outputs_st (fst gfs) st) o snd) (tuple_switch (acc_vars (function_inputs (snd gfs)) (fst (snd (hd (the (tr' 0))))), function_outputs_st (snd gfs) st') o snd) )) tr tr'" definition graph_function_refine :: "bool \ (string \ graph_function option) \ string \ (string \ graph_function option) \ string \ ((bool \ variable list) \ bool) \ ((bool \ bool \ variable list) \ bool) \ bool" where "graph_function_refine precise Gamma1 f1 Gamma2 f2 irel orel = (\xs ys tr tr' gf gf'. tr \ exec_trace Gamma1 f1 \ tr' \ exec_trace Gamma2 f2 \ Gamma1 f1 = Some gf \ tr 0 = Some [init_state f1 gf xs] \ Gamma2 f2 = Some gf' \ tr' 0 = Some [init_state f2 gf' ys] \ length xs = length (function_inputs gf) \ length ys = length (function_inputs gf') \ irel (tuple_switch (xs, ys)) \ trace_refine precise (\st st'. orel (switch_val fst (tuple_switch (xs, function_outputs_st gf st) o snd) (tuple_switch (ys, function_outputs_st gf' st') o snd) )) tr tr')" lemma graph_function_refine_trace: "graph_function_refine prec Gamma1 f1 Gamma2 f2 irel orel = (\tr tr' xs ys gf gf'. exec_double_trace Gamma1 f1 Gamma2 f2 tr tr' \ tr 0 \ None \ tr' 0 \ None \ Gamma1 f1 = Some gf \ Gamma2 f2 = Some gf' \ the (tr 0) = [init_state f1 gf xs] \ the (tr' 0) = [init_state f2 gf' ys] \ length xs = length (function_inputs gf) \ length ys = length (function_inputs gf') \ irel (tuple_switch (xs, ys)) \ trace_refine prec (\st st'. orel (switch_val fst (tuple_switch (xs, function_outputs_st gf st) o snd) (tuple_switch (ys, function_outputs_st gf' st') o snd) )) tr tr')" apply (clarsimp simp: graph_function_refine_def) apply (simp only: ex_simps[symmetric] all_simps[symmetric] cong: conj_cong) apply auto done definition trace_addr :: "trace \ nat \ next_node option" where "trace_addr tr n = (case tr n of Some [(nn, _, _)] \ Some nn | _ \ None)" lemma trace_addr_SomeD: "trace_addr tr n = Some nn \ \st g. tr n = Some [(nn, st, g)]" by (simp add: trace_addr_def split: option.split_asm list.split_asm prod.split_asm) lemma trace_addr_SomeI: "\x. tr n = Some [(nn, x)] \ trace_addr tr n = Some nn" by (clarsimp simp add: trace_addr_def) type_synonym restrs = "nat \ nat set" definition restrs_condition :: "trace \ restrs \ nat \ bool" where "restrs_condition tr restrs n = (\m. card {i. i < n \ trace_addr tr i = Some (NextNode m)} \ restrs m)" definition succ_restrs :: "next_node \ restrs \ restrs" where "succ_restrs nn rs = (case nn of NextNode n \ rs (n := {x. x - 1 \ rs n}) | _ \ rs)" abbreviation "succ_restrs' n \ succ_restrs (NextNode n)" definition pred_restrs :: "next_node \ restrs \ restrs" where "pred_restrs nn rs = (case nn of NextNode n \ rs (n := {x. Suc x \ rs n}) | _ \ rs)" abbreviation "pred_restrs' n \ pred_restrs (NextNode n)" definition trace_bottom_addr :: "trace \ nat \ next_node option" where "trace_bottom_addr tr n = (case tr n of Some [] \ None | Some xs \ Some (fst (List.last xs)) | None \ None)" definition double_trace_imp :: "(trace \ trace \ bool) \ (trace \ trace \ bool) \ (trace \ trace \ bool)" where "double_trace_imp P Q = (\(tr, tr'). P (tr, tr') \ Q (tr, tr'))" type_synonym visit_addr = "bool \ next_node \ restrs" definition restrs_eventually_condition :: "trace \ restrs \ bool" where "restrs_eventually_condition tr restrs = (\n. tr n \ None \ (\m \ n. restrs_condition tr restrs m))" definition visit :: "trace \ next_node \ restrs \ state option" where "visit tr n restrs = (if \i. restrs_condition tr restrs i \ trace_addr tr i = Some n then Some (fst (snd (hd (the (tr (LEAST i. restrs_condition tr restrs i \ trace_addr tr i = Some n)))))) else None)" definition pc :: "next_node \ restrs \ trace \ bool" where "pc n restrs tr = (visit tr n restrs \ None)" abbreviation "pc' n \ pc (NextNode n)" definition double_visit :: "visit_addr \ trace \ trace \ state option" where "double_visit = (\(side, nn, restrs) (tr, tr'). visit (if side then tr' else tr) nn restrs)" definition double_pc :: "visit_addr \ trace \ trace \ bool" where "double_pc = (\(side, nn, restrs) (tr, tr'). (visit (if side then tr' else tr) nn restrs) \ None)" definition related_pair :: "visit_addr \ (state \ 'a) \ visit_addr \ (state \ 'b) \ ('a \ 'b \ bool) \ (trace \ trace) \ bool" where "related_pair v1 expr1 v2 expr2 rel trs = rel (expr1 (the (double_visit v1 trs)), expr2 (the (double_visit v2 trs)))" abbreviation "equals v1 expr1 v2 expr2 \ related_pair v1 expr1 v2 expr2 (split (=))" abbreviation "equals' n1 restrs1 expr1 n2 restrs2 expr2 \ equals (False, NextNode n1, restrs1) expr1 (True, NextNode n2, restrs2) expr2" definition restr_trace_refine :: "bool \ (string \ graph_function option) \ string \ (string \ graph_function option) \ string \ restrs \ restrs \ ((bool \ bool \ variable list) \ bool) \ trace \ trace \ bool" where "restr_trace_refine precise Gamma1 f1 Gamma2 f2 restrs1 restrs2 orel tr tr' = (\gf gf'. exec_double_trace Gamma1 f1 Gamma2 f2 tr tr' \ Gamma1 f1 = Some gf \ Gamma2 f2 = Some gf' \ (\xs. the (tr 0) = [init_state f1 gf xs] \ length xs = length (function_inputs gf)) \ (\ys. the (tr' 0) = [init_state f2 gf' ys] \ length ys = length (function_inputs gf')) \ restrs_eventually_condition tr restrs1 \ restrs_eventually_condition tr' restrs2 \ trace_refine2 precise (gf, gf') orel tr tr')" definition restrs_list :: "(nat \ nat list) list \ restrs" where "restrs_list rs = (\i. fold (\(n, xs) S. if i = n then set xs \ S else S) rs UNIV)" definition fill_in_below :: "nat list \ nat list" where "fill_in_below xs = [0 ..< fold max (map Suc xs) 0]" definition restrs_visit :: "(nat \ nat list) list \ next_node \ graph_function \ (nat \ nat list) list" where "restrs_visit xs nn gf = map (\(m, xsf). if (nn, NextNode m) \ rtrancl (reachable_step' gf) then (m, (fill_in_below xsf)) else (m, xsf)) xs" definition eqs :: "(('a \ (variable list \ variable)) \ ('a \ (variable list \ variable))) list \ ('a \ nat option) \ ('a \ variable list) \ bool" where "eqs xs lens vs = ((\((laddr, lval), (raddr, rval)) \ set xs. lval (vs laddr) = rval (vs raddr)) \ (\addr. lens addr \ None \ length (vs addr) = the (lens addr)))" definition visit_restrs_preds_raw :: "next_node \ restrs \ (next_node \ restrs) list \ bool" where "visit_restrs_preds_raw nn restrs preds = (\tr i. trace_addr tr i = Some nn \ restrs_condition tr restrs i \ (\j. trace_addr tr j = Some nn \ restrs_condition tr restrs j \ j = i) \ (\j nn' restrs'. j < i \ (nn', restrs') \ set preds \ trace_addr tr j = Some nn' \ restrs_condition tr restrs' j))" definition visit_restrs_preds :: "visit_addr \ visit_addr list \ bool" where "visit_restrs_preds vis preds = (case vis of (side, nn, restrs) \ (fst ` set preds \ {side}) \ visit_restrs_preds_raw nn restrs (map snd preds))" lemma visit_known: "tr i = Some [(nn, st, g)] \ restrs_condition tr restrs i \ (\j < i. trace_addr tr j = Some nn \ \ restrs_condition tr restrs j) \ visit tr nn restrs = Some st" apply (clarsimp simp: visit_def) apply (subst Least_equality, blast intro: trace_addr_SomeI) apply (blast intro: linorder_not_le[THEN iffD1]) apply (auto simp: trace_addr_SomeI) done lemma fold_invariant: "(\x \ set xs. \ s. g (f x s) = g s) \ g s = v \ g (fold f xs s) = v" by (induct xs arbitrary: s, simp_all) lemma var_acc_var_upd: "var_acc s (var_upd s' v st) = (if s = s' then v else var_acc s st)" by (cases st, simp add: var_acc_def var_upd_def) lemma var_acc_save_vals_distinct: "distinct xs \ length xs = length vs \ map (\i. var_acc i (save_vals xs vs st)) xs = vs" apply (induct xs arbitrary: vs st) apply simp apply (case_tac vs, simp_all add: save_vals_def) apply (rule fold_invariant) apply (clarsimp simp: var_acc_var_upd elim!: in_set_zipE) apply (simp add: var_acc_var_upd) done definition wf_graph_function :: "graph_function \ nat \ nat \ bool" where "wf_graph_function f ilen olen = (case f of GraphFunction inputs outputs graph ep \ distinct inputs \ distinct outputs \ length inputs = ilen \ length outputs = olen \ (\nn i. (nn, NextNode i) \ reachable_step' f \ i \ dom graph) \ ep \ dom graph)" lemma wf_graph_function_acc_vars_save_vals: "\ wf_graph_function f ilen olen; xs = function_inputs f; length vs = ilen \ \ acc_vars xs (save_vals xs vs st) = vs" by (cases f, simp add: acc_vars_def wf_graph_function_def var_acc_save_vals_distinct) lemma wf_graph_function_length_function_inputs: "wf_graph_function f ilen olen \ length (function_inputs f) = ilen" by (cases f, simp_all add: wf_graph_function_def) lemma graph_function_refine_trace2: "Gamma1 f1 = Some gf \ wf_graph_function gf ilen1 olen1 \ Gamma2 f2 = Some gf' \ wf_graph_function gf' ilen2 olen2 \ graph_function_refine prec Gamma1 f1 Gamma2 f2 irel orel = (\tr tr'. exec_double_trace Gamma1 f1 Gamma2 f2 tr tr' \ Gamma1 f1 = Some gf \ Gamma2 f2 = Some gf' \ (\xs ys. the (tr 0) = [init_state f1 gf xs] \ length xs = length (function_inputs gf) \ the (tr' 0) = [init_state f2 gf' ys] \ length ys = length (function_inputs gf') \ irel (tuple_switch (xs, ys))) \ trace_refine2 prec (gf, gf') orel tr tr')" apply (clarsimp simp: graph_function_refine_trace trace_refine2_def) apply (frule wf_graph_function_length_function_inputs[where f=gf]) apply (frule wf_graph_function_length_function_inputs[where f=gf']) apply (simp add: imp_conjL) apply (rule arg_cong[where f=All] ext imp_cong[OF refl])+ apply (clarsimp dest!: exec_trace_init) apply (safe, simp_all add: init_state_def wf_graph_function_acc_vars_save_vals) done lemma exec_trace_step_reachable_induct: "tr \ exec_trace Gamma fn \ Gamma fn = Some gf \ trace_bottom_addr tr (Suc i) = Some addr \ (\n. trace_bottom_addr tr i = Some (NextNode n) \ n \ dom (function_graph gf)) \ trace_bottom_addr tr i \ None \ (trace_bottom_addr tr i = Some addr \ trace_addr tr (Suc i) = None \ trace_addr tr (Suc i) = Some addr \ (the (trace_bottom_addr tr i), addr) \ reachable_step' gf)" apply (frule_tac i=i in exec_trace_invariant') apply (frule_tac i=i in exec_trace_step_cases) apply (simp add: all_exec_graph_step_cases, safe dest!: trace_addr_SomeD, simp_all add: trace_bottom_addr_def del: last.simps) apply (auto simp: exec_graph_invariant_Cons reachable_step_def get_state_function_call_def split: graph_function.split_asm, auto split: next_node.splits option.split node.splits if_split_asm simp: trace_addr_def neq_Nil_conv) done lemma exec_trace_step_reachable_induct2: "tr \ exec_trace Gamma fn \ Gamma fn = Some gf \ wf_graph_function gf ilen olen \ trace_bottom_addr tr (Suc i) = Some addr \ (\n. trace_bottom_addr tr i = Some (NextNode n) \ n \ dom (function_graph gf)) \ trace_bottom_addr tr i \ None \ (trace_bottom_addr tr i = Some addr \ trace_addr tr (Suc i) = None \ trace_addr tr (Suc i) = Some addr \ (the (trace_bottom_addr tr i), addr) \ reachable_step' gf)" apply (induct i arbitrary: addr) apply (frule(2) exec_trace_step_reachable_induct) apply (clarsimp simp: exec_trace_def trace_bottom_addr_def wf_graph_function_def split: graph_function.split_asm) apply (frule(2) exec_trace_step_reachable_induct) apply atomize apply (case_tac "trace_bottom_addr tr (Suc i)", simp_all) apply clarsimp apply (frule_tac i=i in exec_trace_step_reachable_induct, simp) apply clarsimp apply (erule disjE, simp_all) apply (simp add: wf_graph_function_def split: graph_function.split_asm) apply auto done lemma Collect_less_Suc: "{i. i < Suc n \ P i} = {i. i < n \ P i} \ (if P n then {n} else {})" by (auto simp: less_Suc_eq) lemma trace_addr_trace_bottom_addr_eq: "trace_addr tr i = Some addr \ trace_bottom_addr tr i = Some addr" by (clarsimp simp: trace_bottom_addr_def dest!: trace_addr_SomeD) lemma reachable_trace_induct: "tr \ exec_trace Gamma fname \ Gamma fname = Some f \ wf_graph_function f ilen olen \ trace_addr tr i = Some nn \ i + k \ j \ (\nn'. trace_bottom_addr tr (i + k) = Some nn' \ (Some nn' \ trace_addr tr ` {i .. i + k}) \ (nn, nn') \ rtrancl (reachable_step' f \ {(x, y). Some x \ trace_addr tr ` {i ..< j}}) \ (k \ 0 \ trace_addr tr (i + k) \ None \ (nn, nn') \ trancl (reachable_step' f \ {(x, y). Some x \ trace_addr tr ` {i ..< j}})))" apply (induct k) apply (clarsimp simp: trace_bottom_addr_def dest!: trace_addr_SomeD) apply clarsimp apply (frule(3) exec_trace_step_reachable_induct2) apply clarsimp apply (subgoal_tac "i + k \ {i ..< j} \ Suc (i + k) \ {i .. Suc (i + k)}") apply (rule conjI) apply (elim disjE conjE) apply simp apply (blast intro: sym) apply (rule conjI) apply (elim disjE conjE) apply simp apply (erule rtrancl_into_rtrancl) apply simp apply clarify apply (elim disjE conjE) apply simp apply (erule rtrancl_into_trancl1) apply simp apply simp done lemma reachable_trace: "tr \ exec_trace Gamma fname \ Gamma fname = Some f \ wf_graph_function f ilen olen \ trace_addr tr i = Some nn \ trace_addr tr j = Some nn' \ i < j \ (nn, nn') \ trancl (reachable_step' f \ {(x, y). Some x \ trace_addr tr ` {i ..< j}})" apply (frule(3) reachable_trace_induct[where k="j - i" and j=j]) apply simp apply (clarsimp simp: trace_bottom_addr_def dest!: trace_addr_SomeD) done lemma reachable_trace_eq: "tr \ exec_trace Gamma fname \ Gamma fname = Some f \ wf_graph_function f ilen olen \ trace_addr tr i = Some nn \ trace_addr tr j = Some nn' \ i \ j \ (nn, nn') \ rtrancl (reachable_step' f \ {(x, y). Some x \ trace_addr tr ` {i ..< j}})" apply (cases "i = j", simp_all) apply (rule trancl_into_rtrancl) apply (fastforce elim: reachable_trace) done lemma restrs_single_visit_neq: "tr \ exec_trace Gamma fname \ Gamma fname = Some f \ wf_graph_function f ilen olen \ trace_addr tr i = Some nn \ restrs_condition tr restrs i \ trace_addr tr j = Some nn \ restrs_condition tr restrs j \ \x. NextNode x \ set cuts \ (\y. restrs x \ {y}) \ (nn, nn) \ trancl (reachable_step' f \ {(x, y). x \ set cuts}) \ i < j \ False" apply (frule(5) reachable_trace) apply (erule notE, erule trancl_mono) apply clarsimp apply (case_tac a, simp_all) apply (drule spec, drule(1) mp, clarsimp) apply (clarsimp simp: restrs_condition_def) apply (drule spec, drule(1) subsetD[rotated])+ apply clarsimp apply (frule card_seteq[rotated -1, OF eq_imp_le], simp, clarsimp)[1] apply (blast dest: linorder_not_less[THEN iffD2]) apply (simp add: reachable_step_def)+ done lemma restrs_single_visit: "tr \ exec_trace Gamma fname \ Gamma fname = Some f \ wf_graph_function f ilen olen \ trace_addr tr i = Some nn \ restrs_condition tr restrs i \ trace_addr tr j = Some nn \ restrs_condition tr restrs j \ \x. NextNode x \ set cuts \ (\y. restrs x \ {y}) \ (nn, nn) \ trancl (reachable_step' f \ {(x, y). x \ set cuts}) \ i = j" by (metis restrs_single_visit_neq linorder_neq_iff) lemma restrs_single_visit2: "trace_addr tr i = Some nn \ trace_addr tr j = Some nn \ tr \ exec_trace Gamma fname \ Gamma fname = Some f \ wf_graph_function f ilen olen \ restrs_condition tr restrs i \ restrs_condition tr restrs j \ \x. NextNode x \ set cuts \ (\y. restrs x \ {y}) \ (nn, nn) \ trancl (reachable_step' f \ {(x, y). x \ set cuts}) \ i = j" by (metis restrs_single_visit) lemma not_trancl_converse_step: "converse stp `` {y} \ S \ \z \ S. (x, z) \ rtrancl (stp \ constraint) \ (x, y) \ trancl (stp \ constraint)" using tranclD2 by fastforce lemma reachable_order_mono: "(nn, nn') \ rtrancl R \ (\(nn, nn') \ R. order nn \ order nn') \ order nn \ (order nn' :: 'a :: linorder)" apply (induct rule: rtrancl.induct) apply simp apply (blast intro: order_trans) done lemma not_reachable_by_order: "(\(nn, nn') \ R. order nn \ order nn') \ order nn > (order nn' :: 'a :: linorder) \ (nn, nn') \ rtrancl R" by (metis reachable_order_mono linorder_not_less) lemma Collect_less_nat: "(n :: nat) - 1 = m \ {i. i < n \ P i} = (if n = 0 then {} else {i. i < m \ P i} \ (if P m then {m} else {}))" by (cases n, simp_all add: Collect_less_Suc) lemma test_restrs_condition: "\ graph = [ 1 \ Call (NextNode 2) ''foo'' [] [], 2 \ Basic (NextNode 3) [], 3 \ Basic Ret [] ]; Gamma ''bar'' = Some (GraphFunction [] [] graph 1); tr \ exec_trace Gamma ''bar''; Gamma ''foo'' = Some (GraphFunction [] [] [ 1 \ Basic Ret []] 1) \ \ restrs_condition tr (restrs_list [(1, [1])]) 5" apply (frule exec_trace_init, elim exE conjE) apply (frule_tac i=0 in exec_trace_step_cases, clarsimp) apply (clarsimp simp: all_exec_graph_step_cases) apply (frule_tac i=1 in exec_trace_step_cases, clarsimp) apply (clarsimp simp: all_exec_graph_step_cases upd_vars_def save_vals_def init_vars_def) apply (frule_tac i=2 in exec_trace_step_cases, clarsimp) apply (clarsimp simp: all_exec_graph_step_cases upd_vars_def save_vals_def init_vars_def TWO return_vars_def) apply (frule_tac i=3 in exec_trace_step_cases, clarsimp) apply (clarsimp simp: all_exec_graph_step_cases upd_vars_def save_vals_def init_vars_def) apply (frule_tac i=4 in exec_trace_step_cases, clarsimp) apply (clarsimp simp: all_exec_graph_step_cases upd_vars_def save_vals_def init_vars_def) apply (frule_tac i=5 in exec_trace_step_cases, clarsimp) apply (clarsimp simp: all_exec_graph_step_cases upd_vars_def save_vals_def init_vars_def) apply (simp add: restrs_condition_def Collect_less_nat Collect_conv_if trace_addr_def restrs_list_def) done primrec first_reached_prop :: "visit_addr list \ (visit_addr \ (trace \ trace) \ bool) \ (trace \ trace) \ bool" where "first_reached_prop [] p trs = False" | "first_reached_prop (v # vs) p trs = (if double_pc v trs then p v trs else first_reached_prop vs p trs)" definition get_function_call :: "(graph_function \ graph_function) \ visit_addr \ (next_node \ string \ (state \ variable) list \ string list) option" where "get_function_call gfs x = (case x of (side, NextNode n, restrs) \ (case function_graph (tuple_switch gfs side) n of Some (Call nn fname inputs outputs) \ Some (nn, fname, inputs, outputs) | _ \ None) | _ \ None)" definition func_inputs_match :: "(graph_function \ graph_function) \ ((bool \ variable list) \ bool) \ visit_addr \ visit_addr \ (trace \ trace) \ bool" where "func_inputs_match gfs rel vis1 vis2 trs = (case (get_function_call gfs vis1, get_function_call gfs vis2) of (Some (_, _, inputs1, _), Some (_, _, inputs2, _)) \ rel (\x. map (\f. f (the (double_visit (tuple_switch (vis1, vis2) x) trs))) (tuple_switch (inputs1, inputs2) x)) | _ \ False)" definition succ_visit :: "next_node \ visit_addr \ visit_addr" where "succ_visit nn vis = (case vis of (side, nn', restrs) \ (side, nn', succ_restrs nn restrs))" lemma fst_succ_visit[simp]: "fst (succ_visit nn v) = fst v" by (simp add: succ_visit_def split_def split: next_node.split) lemma wf_graph_function_init_extract: "\ wf_graph_function f ilen olen; tr \ exec_trace Gamma fn; Gamma fn = Some f; the (tr 0) = [init_state fn f xs]; \n. 0 \ restrs n; length xs = ilen \ \ acc_vars (function_inputs f) (the (visit tr (NextNode (entry_point f)) restrs)) = xs" apply (cases f) apply (frule wf_graph_function_length_function_inputs) apply (frule exec_trace_init) apply (clarsimp simp: visit_known restrs_condition_def init_state_def wf_graph_function_acc_vars_save_vals) done lemma exec_trace_reachable_step: "tr \ exec_trace Gamma fn \ Gamma fn = Some f \ trace_addr tr (Suc i) = Some nn' \ nn' \ Err \ \nn. (nn, nn') \ reachable_step' f" apply (clarsimp dest!: trace_addr_SomeD) apply (frule(1) exec_trace_invariant) apply (subgoal_tac "\stack. tr i = Some stack \ exec_graph_invariant Gamma fn stack") prefer 2 apply (metis exec_trace_invariant) apply (frule_tac i=i in exec_trace_step_cases) apply (rule_tac x="the (trace_bottom_addr tr i)" in exI) apply (clarsimp simp: all_exec_graph_step_cases) apply (safe, simp_all add: reachable_step_def)[1] apply (auto simp: exec_graph_invariant_def trace_addr_def get_state_function_call_def trace_bottom_addr_def split: graph_function.split_asm next_node.split option.splits node.splits next_node.split) done lemma init_pc: "tr \ exec_trace Gamma fn \ Gamma fn = Some f \ tr 0 = Some [init_state fn f xs] \ \n. 0 \ restrs n \ wf_graph_function f ilen olen \ pc' (entry_point f) restrs tr" apply (clarsimp simp: pc_def visit_def) apply (rule_tac x=0 in exI) apply (simp add: trace_addr_SomeI init_state_def restrs_condition_def split: graph_function.split) done lemma eqs_eq_conj_len_assert: "(eqs xs lens vs) = ((eqs xs lens $ vs) \ (\addr. lens addr \ None \ length (vs addr) = the (lens addr)))" by (auto simp: eqs_def) lemma length_acc_vars[simp]: "length (acc_vars vs st) = length vs" by (simp add: acc_vars_def) lemma restrs_eventually_init: "tr \ exec_trace Gamma fn \ Gamma fn = Some f \ converse (reachable_step' f) `` {NextNode (entry_point f)} \ set [] \ rs = (restrs_list [(entry_point f, [1])]) \ restrs_eventually_condition tr rs" apply (clarsimp simp: restrs_eventually_condition_def) apply (rule_tac x=1 in exI) apply (rule conjI) apply (clarsimp simp: exec_trace_def nat_trace_rel_def) apply (clarsimp simp: restrs_condition_def restrs_list_def) apply (subst arg_cong[where f=card and y="{0}"], simp_all) apply (safe, simp_all) apply (case_tac x, simp_all) apply (drule(2) exec_trace_reachable_step, simp) apply auto[1] apply (clarsimp simp: exec_trace_def trace_addr_def) done lemma graph_function_restr_trace_refine: assumes wfs: "wf_graph_function f1 il1 ol1" "wf_graph_function f2 il2 ol2" and inps: "converse (reachable_step' f1) `` {NextNode (entry_point f1)} \ set []" "converse (reachable_step' f2) `` {NextNode (entry_point f2)} \ set []" and gfs: "Gamma1 fn1 = Some f1" "Gamma2 fn2 = Some f2" notes wf_facts = wf_graph_function_init_extract wf_graph_function_init_extract wf_graph_function_length_function_inputs shows "graph_function_refine prec Gamma1 fn1 Gamma2 fn2 (eqs ieqs (Some o ils)) orel = (\tr tr'. related_pair (False, NextNode (entry_point f1), restrs_list []) (acc_vars (function_inputs f1)) (True, NextNode (entry_point f2), restrs_list []) (acc_vars (function_inputs f2)) (eqs ieqs (Some o ils) o tuple_switch) (tr, tr') \ exec_double_trace Gamma1 fn1 Gamma2 fn2 tr tr' \ restr_trace_refine prec Gamma1 fn1 Gamma2 fn2 (restrs_list [(entry_point f1, [1])]) (restrs_list [(entry_point f2, [1])]) orel tr tr')" using wfs apply (simp add: graph_function_refine_trace2 gfs) apply (simp add: restr_trace_refine_def gfs restrs_eventually_init restrs_list_def split_def related_pair_def double_visit_def wf_graph_function_init_extract) apply (rule arg_cong[where f=All] ext)+ apply (auto simp: wf_graph_function_init_extract wf_graph_function_length_function_inputs gfs restrs_list_def cong: if_cong elim!: restrs_eventually_init[where Gamma=Gamma1, OF _ _ inps(1)] restrs_eventually_init[where Gamma=Gamma2, OF _ _ inps(2)]; metis) done lemma restrs_list_Int: "restrs_list rs = (\j. \(n, xsf) \ set rs. (if j = n then set xsf else UNIV))" apply (induct rs rule: rev_induct) apply (simp add: restrs_list_def) apply (clarsimp simp add: restrs_list_def fun_eq_iff) done lemma restrs_list_Int2: "restrs_list rs = (\j. \(_, xsf) \ {x \ set rs. fst x = j}. set xsf)" by (auto simp add: restrs_list_Int fun_eq_iff set_eq_iff) lemma restrs_list_Cons: "restrs_list ((n, xs) # rs) = (\j. if j = n then set xs \ restrs_list rs j else restrs_list rs j)" by (simp add: restrs_list_Int fun_eq_iff) lemma restrs_eventually_Cons: "restrs_eventually_condition tr ((K UNIV) (n := set xs)) \ restrs_eventually_condition tr (restrs_list rs) \ restrs_eventually_condition tr (restrs_list ((n, xs) # rs))" apply (clarsimp simp: restrs_eventually_condition_def) apply (rule_tac x="max na naa" in exI) apply (simp add: restrs_condition_def restrs_list_Cons split: if_split_asm) apply (simp add: max_def) done lemma ex_greatest_nat_le: "P j \ j \ n \ \k \ n. (\i \ {Suc k .. n}. \ P i) \ P k" apply (cases "P n") apply (rule_tac x=n in exI, simp) apply (cut_tac P="\j. P (n - j)" and n = "n - j" in ex_least_nat_le, simp+) apply (clarify, rule_tac x="n - k" in exI) apply clarsimp apply (drule_tac x="n - i" in spec, simp) done lemma neq_counting_le: assumes neq: "\i. f i \ bound" assumes mono: "\i. f (Suc i) \ Suc (f i)" shows "f 0 < bound \ f i < bound" proof (induct i) case 0 show ?case using 0 by simp next case (Suc n) show ?case using neq[rule_format, of n] neq[rule_format, of "Suc n"] mono[rule_format, of n] Suc by simp qed lemma restrs_eventually_upt: "\m. card {i. i < m \ trace_addr tr i = Some (NextNode n)} = i \ tr m \ None \ \m. card {i. i < m \ trace_addr tr i = Some (NextNode n)} \ j \ restrs_eventually_condition tr ((\_. UNIV) (n := {i ..< j}))" apply (clarsimp simp: restrs_eventually_condition_def restrs_condition_def) apply (rule exI, rule conjI, erule exI) apply (clarsimp simp: Suc_le_eq) apply (intro conjI impI allI) apply (rule card_mono) apply simp apply clarsimp apply (frule neq_counting_le, simp_all) apply (clarsimp simp: Collect_less_Suc) apply (drule_tac x=0 in spec, simp) done lemma set_fill_in_below: "set (fill_in_below xs) = {k. \l. l : set xs \ k \ l}" apply (induct xs rule: rev_induct, simp_all add: fill_in_below_def) apply (simp add: set_eq_iff less_max_iff_disj less_Suc_eq_le) apply auto done lemma restrs_visit_abstract: assumes dist: "distinct (map fst rs)" shows "restrs_list (restrs_visit rs nn gf) = (\j. if (nn, NextNode j) \ rtrancl (reachable_step' gf) then {k. \l. l \ restrs_list rs j \ k \ l} else restrs_list rs j)" proof - have P: "\j. {x \ set (restrs_visit rs nn gf). fst x = j} = (if (nn, NextNode j) \ rtrancl (reachable_step' gf) then (\(k, xsf). (k, fill_in_below xsf)) ` {x \ set rs. fst x = j} else {x \ set rs. fst x = j})" by (force simp: restrs_visit_def elim: image_eqI[rotated]) show ?thesis apply (rule ext, clarsimp simp: restrs_list_Int2 split_def P) apply (cut_tac dist[THEN set_map_of_compr]) apply (case_tac "\b. (j, b) \ set rs") apply (auto simp: set_fill_in_below) done qed lemma last_upd_stack: "List.last (upd_stack nn stf xs) = (if length xs = 1 then (nn, stf (fst (snd (hd xs))), snd (snd (hd xs))) else List.last xs)" by (cases xs, simp_all) lemma not_reachable_visits_same: "tr \ exec_trace Gamma fn \ Gamma fn = Some gf \ trace_addr tr i = Some n \ (n, m) \ rtrancl (reachable_step' gf) \ wf_graph_function gf ilen olen \ j > i \ {k. k < j \ trace_addr tr k = Some m} = {k. k < i \ trace_addr tr k = Some m}" using reachable_trace_eq[THEN subsetD[OF rtrancl_mono, OF Int_lower1]] apply (safe, simp_all) apply (rule ccontr, clarsimp simp: linorder_not_less) done lemma not_reachable_visits_same_symm: "tr \ exec_trace Gamma fn \ Gamma fn = Some gf \ trace_addr tr i = Some n \ trace_addr tr j = Some n \ (n, m) \ rtrancl (reachable_step' gf) \ wf_graph_function gf ilen olen \ {k. k < j \ trace_addr tr k = Some m} = {k. k < i \ trace_addr tr k = Some m}" using linorder_neq_iff[where x=i and y=j] not_reachable_visits_same by blast lemma restrs_eventually_at_visit: "restrs_eventually_condition tr (restrs_list rs) \ trace_addr tr i = Some nn \ tr \ exec_trace Gamma fn \ Gamma fn = Some gf \ distinct (map fst rs) \ wf_graph_function gf ilen olen \ restrs_condition tr (restrs_list (restrs_visit rs nn gf)) i" apply (clarsimp simp: restrs_eventually_condition_def restrs_condition_def restrs_visit_abstract) apply (case_tac "i \ n") apply (drule spec, drule(1) mp) apply auto[1] apply (simp add: linorder_not_le) apply (drule spec, drule mp, rule order_refl) apply safe apply (fastforce intro: card_mono) apply (simp add: not_reachable_visits_same[symmetric]) done lemma fold_double_trace_imp: "fold double_trace_imp hyps hyp trs = ((\h \ set hyps. h trs) \ hyp trs)" apply (induct hyps arbitrary: hyp, simp_all) apply (auto simp add: double_trace_imp_def) done lemma exec_trace_addr_Suc: "tr \ exec_trace Gamma f \ trace_addr tr n = Some (NextNode m) \ tr (Suc n) \ None" apply (drule_tac i=n in exec_trace_step_cases) apply (auto dest!: trace_addr_SomeD) done lemma num_visits_equals_j_first: "card {i. i < m \ trace_addr tr i = Some n} = j \ j \ 0 \ \m'. trace_addr tr m' = Some n \ card {i. i < m' \ trace_addr tr i = Some n} = j - 1" apply (frule_tac P="\m. card {i. i < m \ trace_addr tr i = Some n} = j" in ex_least_nat_le[rotated]) apply simp apply clarify apply (rule_tac x="k - 1" in exI) apply (case_tac k) apply (simp del: Collect_empty_eq) apply (simp add: Collect_less_Suc split: if_split_asm) done lemma ex_least_nat: "\n. P n \ \n :: nat. P n \ (\i < n. \ P i)" apply clarsimp apply (case_tac "n = 0") apply fastforce apply (cut_tac P="\i. i \ 0 \ P i" in ex_least_nat_le, auto) done lemma visit_eqs: "(visit trace nn restrs = None) = (\i. trace_addr trace i = Some nn \ \ restrs_condition trace restrs i)" "(visit trace nn restrs = Some st) = (\i. restrs_condition trace restrs i \ trace_addr trace i = Some nn \ st = fst (snd (hd (the (trace i)))) \ (\j < i. restrs_condition trace restrs j \ trace_addr trace j \ Some nn))" apply (simp_all add: visit_def) apply auto[1] apply (safe elim!: exE[OF ex_least_nat] del: exE, simp_all) apply (rule exI, rule conjI, assumption, simp) apply (rule arg_cong[OF Least_equality], simp) apply (blast intro: linorder_not_le[THEN iffD1]) apply (rule arg_cong[OF Least_equality], simp) apply (blast intro: linorder_not_le[THEN iffD1]) done lemmas visit_eqs' = visit_eqs(1) arg_cong[where f=Not, OF visit_eqs(1), simplified] theorem restr_trace_refine_Restr1: "j \ 0 \ distinct (map fst rs1) \ wf_graph_function f1 ilen olen \ Gamma1 fn1 = Some f1 \ i \ 0 \ pc' n (restrs_list ((n, [i - 1]) # (restrs_visit rs1 (NextNode n) f1))) tr \ \ pc' n (restrs_list ((n, [j - 1]) # (restrs_visit rs1 (NextNode n) f1))) tr \ restr_trace_refine prec Gamma1 fn1 Gamma2 fn2 (restrs_list ((n, [i ..< j]) # rs1)) rs2 orel tr tr' \ restr_trace_refine prec Gamma1 fn1 Gamma2 fn2 (restrs_list rs1) rs2 orel tr tr'" apply (clarsimp simp: restr_trace_refine_def) apply (subst(asm) restrs_eventually_Cons, simp_all add: K_def) prefer 2 apply auto[1] apply (rule restrs_eventually_upt) apply (clarsimp simp: fold_double_trace_imp double_trace_imp_def pc_def) apply (case_tac "i = 0") apply (rule_tac x=0 in exI) apply (clarsimp simp: dest!: exec_trace_init) apply simp apply (clarsimp simp: visit_eqs) apply (simp(no_asm_use) add: restrs_list_Cons restrs_condition_def) apply (drule_tac x=n in spec)+ apply simp apply (rule_tac x="Suc ia" in exI, simp add: Collect_less_Suc) apply (rule exec_trace_addr_Suc[simplified], assumption) apply simp apply (clarsimp simp: fold_double_trace_imp double_trace_imp_def pc_def) apply (thin_tac "0 < i \ q" for q) apply (clarsimp simp: visit_eqs) apply (frule num_visits_equals_j_first[OF refl, simplified neq0_conv]) apply clarsimp apply (drule(5) restrs_eventually_at_visit) apply (drule_tac x=m' in spec, clarsimp) apply (clarsimp simp: restrs_list_Cons restrs_condition_def split: if_split_asm) done theorem restr_trace_refine_Restr2: "j \ 0 \ distinct (map fst rs2) \ wf_graph_function f2 ilen olen \ Gamma2 fn2 = Some f2 \ i \ 0 \ pc' n (restrs_list ((n, [i - 1]) # (restrs_visit rs2 (NextNode n) f2))) tr' \ \ pc' n (restrs_list ((n, [j - 1]) # (restrs_visit rs2 (NextNode n) f2))) tr' \ restr_trace_refine prec Gamma1 fn1 Gamma2 fn2 rs1 (restrs_list ((n, [i ..< j]) # rs2)) orel tr tr' \ restr_trace_refine prec Gamma1 fn1 Gamma2 fn2 rs1 (restrs_list rs2) orel tr tr'" apply (clarsimp simp: restr_trace_refine_def) apply (subst(asm) restrs_eventually_Cons, simp_all add: K_def) prefer 2 apply auto[1] apply (rule restrs_eventually_upt) apply (clarsimp simp: fold_double_trace_imp double_trace_imp_def pc_def) apply (case_tac "i = 0") apply (rule_tac x=0 in exI) apply (clarsimp dest!: exec_trace_init) apply simp apply (clarsimp simp: visit_eqs restrs_list_Cons restrs_condition_def) apply (drule_tac x=n in spec)+ apply simp apply (rule_tac x="Suc ia" in exI, simp add: Collect_less_Suc) apply (rule exec_trace_addr_Suc[simplified], assumption) apply simp apply (clarsimp simp: fold_double_trace_imp double_trace_imp_def pc_def) apply (thin_tac "0 < i \ q" for q) apply (clarsimp simp: visit_eqs) apply (frule num_visits_equals_j_first[OF refl, simplified neq0_conv]) apply clarsimp apply (drule(5) restrs_eventually_at_visit) apply (drule_tac x=m' in spec, clarsimp) apply (clarsimp simp: restrs_list_Cons restrs_condition_def split: if_split_asm) done lemma pc_Ret_Err_trace_end: "er \ {Ret, Err} \ pc er restrs tr \ tr \ exec_trace Gamma f \ \st g. trace_end tr = Some [(er, st, g)]" apply (clarsimp simp: pc_def visit_eqs trace_end_def dest!: trace_addr_SomeD) apply (frule_tac i=i in exec_trace_step_cases, simp add: all_exec_graph_step_cases) apply (cut_tac tr=tr and n="Suc i" in trace_None_dom_subset) apply (auto simp add: exec_trace_def)[2] apply (subst Max_eqI[rotated 2], erule domI) apply (simp add: finite_subset) apply (simp add: subset_iff less_Suc_eq_le) apply auto[1] done lemmas pc_Ret_trace_end = pc_Ret_Err_trace_end[where er=Ret, simplified] lemma exec_trace_end_SomeD: "trace_end tr = Some v \ tr \ exec_trace Gamma f \ \n. tr n = Some v \ tr (Suc n) = None \ (\nn st g. v = [(nn, st, g)] \ nn \ {Ret, Err})" apply (frule exec_trace_nat_trace) apply (drule(1) trace_end_SomeD) apply clarsimp apply (frule_tac i=n in exec_trace_Nil) apply (case_tac "fst (hd v)", auto simp add: continuing_def split: list.split_asm) done lemma reachable_from_Ret: "((Ret, nn) \ reachable_step' gf)" by (simp add: reachable_step_def) lemma trace_end_visit_Ret: "tr n = Some [(Ret, st, g)] \ tr (Suc n) = None \ tr \ exec_trace Gamma gf \ restrs_eventually_condition tr rs \ visit tr Ret rs = Some st" apply (rule visit_known, assumption) apply (clarsimp simp: restrs_eventually_condition_def) apply (drule spec, erule mp) apply (clarsimp simp: exec_trace_def) apply (drule(1) trace_None_dom_subset) apply auto[1] apply (clarsimp dest!: trace_addr_SomeD) apply (frule_tac i=j in exec_trace_step_cases, simp add: all_exec_graph_step_cases) apply (clarsimp simp: exec_trace_def) apply (auto dest!: trace_None_dom_subset) done definition "output_rel orel gfs restrs trs = orel (\(x, y). acc_vars (tuple_switch (function_inputs, function_outputs) y (tuple_switch gfs x)) (the (double_visit (x, tuple_switch (NextNode (entry_point (tuple_switch gfs x)), Ret) y, tuple_switch (restrs_list [], tuple_switch restrs x) y) trs)))" theorem restr_trace_refine_Leaf: "wf_graph_function f1 ilen1 olen1 \ Gamma1 fn1 = Some f1 \ wf_graph_function f2 ilen2 olen2 \ Gamma2 fn2 = Some f2 \ pc Ret rs1 tr \ prec \ pc Ret rs2 tr' \ output_rel orel (f1, f2) (rs1, rs2) (tr, tr') \ restr_trace_refine prec Gamma1 fn1 Gamma2 fn2 rs1 rs2 orel tr tr'" apply (clarsimp simp only: restr_trace_refine_def) apply (clarsimp simp: trace_refine2_def trace_refine_def) apply (frule(1) pc_Ret_trace_end) apply (case_tac "trace_end tr'") apply (clarsimp split: option.split) apply (cut_tac tr=tr' in pc_Ret_trace_end, blast, assumption+) apply clarsimp apply (clarsimp simp: output_rel_def) apply (drule(1) exec_trace_end_SomeD)+ apply (safe del: impCE, simp_all) apply (erule rsubst[where P=orel], rule ext) apply (clarsimp simp: tuple_switch_def switch_val_def double_visit_def wf_graph_function_init_extract restrs_list_def wf_graph_function_length_function_inputs function_outputs_st_def trace_end_visit_Ret init_state_def wf_graph_function_acc_vars_save_vals) done lemma first_reached_propD: "first_reached_prop addrs propn trs \ \addr \ set addrs. propn addr trs \ (\propn. first_reached_prop addrs propn trs = propn addr trs)" by (induct addrs, simp_all split: if_split_asm) lemma double_pc_reds: "double_pc (False, nn, restrs) trs = pc nn restrs (fst trs)" "double_pc (True, nn, restrs) trs = pc nn restrs (snd trs)" by (simp_all add: double_pc_def pc_def split_def) definition merge_opt :: "'a option \ 'a option \ 'a option" where "merge_opt opt opt' = case_option opt' Some opt" lemma merge_opt_simps[simp]: "merge_opt (Some x) v = Some x" "merge_opt None v = v" by (simp_all add: merge_opt_def) lemma fold_merge_opt_Nones_eq: "(\v \ set xs. v = None) \ fold merge_opt xs v = v" by (induct xs, simp_all) lemma set_zip_rev: "length xs = length ys \ set (zip xs ys) = set (zip (rev xs) (rev ys))" by (simp add: zip_rev) lemma exec_trace_non_Call: "\ tr \ exec_trace Gamma fn; Gamma fn = Some f; trace_bottom_addr tr i = Some (NextNode n); function_graph f n = Some node; case node of Call _ _ _ _ \ False | _ \ True \ \ trace_addr tr i = Some (NextNode n)" apply (clarsimp simp: trace_bottom_addr_def split: option.split_asm) apply (frule(1) exec_trace_invariant) apply (clarsimp simp: exec_graph_invariant_def neq_Nil_conv) apply (case_tac "tl (the (tr i))" rule: rev_cases) apply (clarsimp simp: trace_addr_SomeI) apply (clarsimp simp: get_state_function_call_def split: node.split_asm) done lemma visit_immediate_pred: "\ tr \ exec_trace Gamma fn; Gamma fn = Some f; wf_graph_function f ilen olen; trace_addr tr i = Some nn; nn \ NextNode (entry_point f); converse (reachable_step' f) `` {nn} \ S \ \ \i' nn'. i = Suc i' \ nn' \ S \ trace_bottom_addr tr i' = Some nn'" apply (case_tac i) apply (clarsimp dest!: exec_trace_init trace_addr_SomeD) apply (rename_tac i') apply (frule trace_addr_SomeD, clarsimp) apply (frule(1) exec_trace_invariant) apply (frule_tac i=i' in exec_trace_step_reachable_induct2, assumption+) apply (simp add: trace_bottom_addr_def) apply auto done lemma pred_restrs: "\ tr \ exec_trace Gamma f; trace_bottom_addr tr i = Some nn \ \ restrs_condition tr restrs (Suc i) = restrs_condition tr (if trace_addr tr i = None then restrs else pred_restrs nn restrs) i" apply (clarsimp simp: restrs_condition_def Collect_less_Suc all_conj_distrib pred_restrs_def split: if_split_asm next_node.split) apply (auto simp: trace_addr_trace_bottom_addr_eq) done lemma visit_merge: assumes tr: "tr \ exec_trace Gamma fn" "Gamma fn = Some f" and wf: "wf_graph_function f ilen olen" and ns: "nn \ NextNode (entry_point f)" "\n \ set ns. graph n = Some (Basic nn [])" "converse (reachable_step graph) `` {nn} \ NextNode ` set ns" and geq: "function_graph f = graph" and cut: "\x. NextNode x \ set cuts \ (\y. restrs x \ {y})" "\n \ set ns. (nn, NextNode n) \ rtrancl (reachable_step graph \ {(x, y). x \ set cuts})" shows "visit tr nn restrs = fold merge_opt (map (\n. visit tr (NextNode n) (pred_restrs' n restrs)) ns) None" proof - note ns = ns[folded geq] note cut = cut[folded geq] have step_after: "\n i. n \ set ns \ trace_bottom_addr tr i = Some (NextNode n) \ \st. tr i = Some [(NextNode n, st, fn)] \ tr (Suc i) = Some [(nn, st, fn)] \ trace_addr tr (Suc i) = Some nn \ restrs_condition tr restrs (Suc i) = restrs_condition tr (pred_restrs' n restrs) i" apply (drule exec_trace_non_Call[OF tr], (simp add: ns)+) apply (frule ns[rule_format], cut_tac tr(2)) apply (frule trace_addr_SomeD, clarsimp) apply (frule exec_trace_invariant[OF tr(1)]) apply (cut_tac i=i in exec_trace_step_cases[OF tr(1)]) apply (clarsimp simp: all_exec_graph_step_cases exec_graph_invariant_Cons upd_vars_def save_vals_def) apply (simp add: pred_restrs[OF tr(1)] trace_addr_SomeI trace_bottom_addr_def K_def) done have step_after_single: "\n i. n \ set ns \ trace_bottom_addr tr i = Some (NextNode n) \ restrs_condition tr restrs (Suc i) \ (\n' j. n' \ set ns \ trace_addr tr j = Some (NextNode n') \ restrs_condition tr (pred_restrs' n' restrs) j \ j = i)" apply clarsimp apply (frule step_after, erule trace_addr_trace_bottom_addr_eq) apply (frule(1) step_after) apply clarsimp apply (drule(2) restrs_single_visit[OF tr wf _ _ _ _ cut(1)], simp_all) apply (rule not_trancl_converse_step, rule ns) apply (simp add: cut) done have visit_after: "\n v. n \ set ns \ visit tr (NextNode n) (pred_restrs' n restrs) = Some v \ visit tr nn restrs \ None" apply (clarsimp simp: visit_eqs) apply (drule_tac i=i in step_after, simp add: trace_addr_trace_bottom_addr_eq) apply (rule_tac x="Suc i" in exI) apply clarsimp done show ?thesis apply (rule sym, cases "visit tr nn restrs", simp_all) apply (rule fold_merge_opt_Nones_eq) apply (rule ccontr, clarsimp simp: visit_after) apply (clarsimp simp: visit_eqs) apply (frule visit_immediate_pred[OF tr wf _ ns(1, 3)]) apply clarsimp apply (frule(1) step_after, clarsimp) apply (frule(2) step_after_single) apply (drule in_set_conv_decomp_last[THEN iffD1]) apply clarsimp apply (rule trans, rule fold_merge_opt_Nones_eq) apply (rule ccontr, clarsimp simp: visit_eqs pc_def ball_Un) apply (simp add: trace_addr_SomeI) apply (subst visit_known, assumption, simp_all) apply clarsimp done qed lemma visit_merge_restrs: assumes tr: "tr \ exec_trace Gamma fn" "Gamma fn = Some gf" and geq: "function_graph gf = graph" assumes indep: "opt \ set (opt2 # opts)" assumes reach: "(nn, NextNode addr) \ rtrancl (reachable_step graph)" and wf: "wf_graph_function gf ilen olen" fixes rs defines "rs1 \ restrs_list ((addr, [opt]) # rs)" and "rs2 \ restrs_list ((addr, opt2 # opts) # rs)" and "rs3 \ restrs_list ((addr, opt # opt2 # opts) # rs)" shows "visit tr nn rs3 = merge_opt (visit tr nn rs1) (visit tr nn rs2)" proof - let ?card = "\i. card {j. j < i \ trace_addr tr j = Some (NextNode addr)}" obtain n where vis: "\i. trace_addr tr i = Some nn \ ?card i = n" apply (case_tac "\i. trace_addr tr i = Some nn") apply clarsimp apply (erule_tac x="?card i" in meta_allE) apply (erule meta_mp) apply clarsimp apply (auto dest: not_reachable_visits_same_symm[OF tr _ _ reach[folded geq] wf]) done have indep2: "\i. restrs_condition tr rs1 i \ \ restrs_condition tr rs2 i" using indep apply (clarsimp, clarsimp simp: restrs_condition_def rs1_def rs2_def restrs_list_Cons) apply (drule_tac x=addr in spec)+ apply (clarsimp simp: restrs_list_def) done have disj: "\i. restrs_condition tr rs3 i = (restrs_condition tr rs1 i \ restrs_condition tr rs2 i)" apply (clarsimp simp: restrs_condition_def rs1_def rs2_def rs3_def restrs_list_Cons) apply blast done have indep3: "\j i. trace_addr tr i = Some nn \ restrs_condition tr rs1 i \ trace_addr tr j = Some nn \ \ restrs_condition tr rs2 j" using indep apply (clarsimp, clarsimp simp: restrs_condition_def rs1_def rs2_def restrs_list_Cons) apply (drule_tac x=addr in spec)+ apply (clarsimp simp: restrs_list_def vis) done show ?thesis unfolding visit_def by (auto simp: disj dest: indep3[rule_format]; metis) qed theorem visit_explode_restr: assumes gf1: "tr \ exec_trace Gamma fn" "Gamma fn = Some gf" "function_graph gf = graph" and gf2: "(nn, NextNode addr) \ (reachable_step graph)\<^sup>*" "wf_graph_function gf ilen olen" and rs: "restrs_list rs addr = set xs" "filter (\(addr', xs). addr' \ addr) rs = rs'" and xs: "distinct xs" shows "visit tr nn (restrs_list rs) = fold (\x. merge_opt (visit tr nn (restrs_list ((addr, [x]) # rs')))) xs None" proof - from rs have Q1: "restrs_list rs addr = restrs_list ((addr, rev xs) # rs') addr" apply (clarsimp simp: restrs_list_def fun_eq_iff) apply (rule fold_invariant[where g="\x. x", symmetric], auto) done { fix addr' from rs(2) have "addr \ addr' \ restrs_list rs addr' = restrs_list ((addr, rev xs) # rs') addr'" apply clarsimp apply hypsubst_thin apply (induct rs, simp_all add: restrs_list_Cons) apply (clarsimp simp: restrs_list_Cons) done } hence Q: "restrs_list rs = restrs_list ((addr, rev xs) # rs')" using Q1 apply (simp add: fun_eq_iff) apply metis done show "?thesis" using xs unfolding Q apply (induct xs rule: rev_induct) apply (simp add: restrs_list_Cons visit_def restrs_condition_def) apply metis apply (case_tac xs rule: rev_cases) apply (simp add: merge_opt_def) apply (clarsimp simp: visit_merge_restrs[OF gf1 _ gf2]) done qed lemma visit_impossible: "tr \ exec_trace Gamma fn \ Gamma fn = Some gf \ function_graph gf = graph \ wf_graph_function gf ilen olen \ 0 \ restrs n \ (NextNode n, nn) \ rtrancl (reachable_step graph) \ visit tr nn restrs = None" apply (clarsimp simp: visit_eqs restrs_condition_def) apply (rule_tac x=n in exI) apply (subst arg_cong[where f=card and y="{}"], simp_all) apply clarsimp apply (drule(5) reachable_trace) apply (drule trancl_mono[OF _ Int_lower1]) apply (simp add: trancl_into_rtrancl) done lemma visit_inconsistent: "restrs i = {} \ visit tr nn restrs = None" by (auto simp add: visit_def restrs_condition_def) lemma visit_immediate_pred_step: "tr i = Some [(nn, st, fn')] \ tr \ exec_trace Gamma fn \ Gamma fn = Some gf \ wf_graph_function gf ilen olen \ converse (reachable_step (function_graph gf)) `` {nn} \ {NextNode n} \ nn \ NextNode (entry_point gf) \ (case (function_graph gf n) of None \ False | Some (Call _ _ _ _) \ False | _ \ True) \ \i' st'. i = Suc i' \ fn' = fn \ tr i' = Some [(NextNode n, st', fn)] \ (([(NextNode n, st', fn)], [(nn, st, fn)]) \ exec_graph_step Gamma)" apply (frule(2) visit_immediate_pred[where i=i], simp_all) apply (simp add: trace_addr_def) apply (clarsimp split: option.split_asm) apply (frule(4) exec_trace_non_Call) apply (clarsimp dest!: trace_addr_SomeD) apply (frule_tac i=i' in exec_trace_step_cases, clarsimp) apply (frule_tac i=i' in exec_trace_invariant') apply (frule_tac i=i in exec_trace_invariant') apply (clarsimp simp: exec_graph_invariant_def) done lemma visit_Basic: "tr \ exec_trace Gamma fn \ Gamma fn = Some gf \ function_graph gf = graph \ wf_graph_function gf ilen olen \ graph n = Some (Basic nn upds) \ converse (reachable_step graph) `` {nn} \ {NextNode n} \ nn \ NextNode (entry_point gf) \ visit tr nn restrs = option_map (upd_vars upds) (visit tr (NextNode n) (pred_restrs' n restrs))" apply (cases "visit tr (NextNode n) (pred_restrs' n restrs)", simp_all) apply (clarsimp simp: visit_eqs) apply (frule(5) visit_immediate_pred) apply (clarsimp simp: pred_restrs) apply (frule(3) exec_trace_non_Call, simp) apply simp apply (clarsimp simp: visit_eqs) apply (frule trace_addr_SomeD, clarsimp) apply (frule_tac i=i in exec_trace_step_cases) apply (frule(1) exec_trace_invariant) apply (rule_tac x="Suc i" in exI) apply (clarsimp simp: all_exec_graph_step_cases exec_graph_invariant_Cons K_def) apply (simp add: pred_restrs[unfolded trace_bottom_addr_def] trace_addr_SomeI) apply clarsimp apply (frule(2) visit_immediate_pred, simp+) apply clarsimp apply (frule(2) exec_trace_non_Call, simp+) apply (clarsimp simp: pred_restrs split: if_split_asm) done definition "option_guard guard opt = (case opt of None \ None | Some v \ if guard v then Some v else None)" lemmas option_guard_simps[simp] = option_guard_def[split_simps option.split] lemma visit_Cond: "tr \ exec_trace Gamma fn \ Gamma fn = Some gf \ function_graph gf = graph \ wf_graph_function gf ilen olen \ graph n = Some (Cond l r cond) \ converse (reachable_step graph) `` {nn} \ {NextNode n} \ nn \ NextNode (entry_point gf) \ \x. NextNode x \ set cuts \ (\y. pred_restrs' n restrs x \ {y}) \ (l, NextNode n) \ rtrancl (reachable_step graph \ {(x, y). x \ set cuts}) \ (r, NextNode n) \ rtrancl (reachable_step graph \ {(x, y). x \ set cuts}) \ visit tr nn restrs = option_guard (\st. (nn = l \ cond st) \ (nn = r \ \ cond st)) (visit tr (NextNode n) (pred_restrs' n restrs))" apply (cases "visit tr (NextNode n) (pred_restrs' n restrs)"; simp) apply (clarsimp simp: visit_eqs) apply (frule(5) visit_immediate_pred) apply (clarsimp simp: pred_restrs) apply (frule(3) exec_trace_non_Call, simp) apply simp apply (clarsimp simp: visit_eqs simp del: imp_disjL) apply (frule trace_addr_SomeD, clarsimp simp del: imp_disjL) apply (frule_tac i=i in exec_trace_step_cases) apply (frule(1) exec_trace_invariant) apply (clarsimp simp: all_exec_graph_step_cases exec_graph_invariant_Cons simp del: imp_disjL) apply (intro conjI impI) apply (rule_tac x="Suc i" in exI) apply (clarsimp simp: pred_restrs trace_bottom_addr_def) apply (rule conjI, simp add: trace_addr_def) apply clarsimp apply (frule(3) visit_immediate_pred, simp, simp) apply (clarsimp simp: pred_restrs) apply (frule(2) exec_trace_non_Call, simp+)[1] apply (rule_tac x="Suc i" in exI) apply (clarsimp simp: pred_restrs trace_bottom_addr_def) apply (rule conjI, simp add: trace_addr_def) apply clarsimp apply (frule(3) visit_immediate_pred, simp, simp) apply (clarsimp simp: pred_restrs) apply (frule(2) exec_trace_non_Call, simp+)[1] apply clarsimp apply (frule(3) visit_immediate_pred, simp, simp) apply (clarsimp simp: pred_restrs) apply (frule(2) exec_trace_non_Call, simp+) apply (drule(3) restrs_single_visit2, simp+) apply (clarsimp simp: reachable_step_def[THEN eqset_imp_iff] dest!: tranclD) apply auto[1] apply (clarsimp dest!: trace_addr_SomeD split: if_split_asm) done lemma exec_trace_pc_Call: "pc' n restrs tr \ tr \ exec_trace Gamma fn \ Gamma fn = Some f \ function_graph f n = Some (node.Call nn fname inps outps) \ Gamma fname = Some g \ (\x. restrs_condition tr restrs x \ trace_addr tr x = Some (NextNode n) \ (trace_drop_n (Suc x) 1 tr \ exec_trace Gamma fname) \ visit tr (NextNode n) restrs = Some (fst (snd (hd (the (tr x))))))" apply (clarsimp simp: pc_def visit_eqs elim!: exE[OF ex_least_nat] del: exE) apply (rule exI, rule conjI, assumption, simp) apply (rule conjI[rotated], blast) apply (clarsimp dest!: trace_addr_SomeD) apply (frule(1) exec_trace_invariant) apply (simp add: exec_graph_invariant_def) apply (erule(4) exec_trace_drop_n[simplified]) done lemma exec_trace_step: "tr \ exec_trace Gamma f \ tr i = Some stack \ continuing stack \ \stack'. tr (Suc i) = Some stack' \ (stack, stack') \ exec_graph_step Gamma" apply (frule_tac i=i in exec_trace_step_cases) apply auto done lemma trace_bottom_addr_trace_addr_prev: "\ trace_bottom_addr tr i = Some addr; restrs_condition tr restrs i; tr \ exec_trace Gamma fn; Gamma fn = Some f; wf_graph_function f ilen olen \ \ \j. trace_addr tr j = Some addr \ (i = j \ trace_addr tr i = None) \ (trace_addr tr ` {Suc j .. i} \ {None}) \ (trace_bottom_addr tr ` {Suc j .. i} \ {Some addr}) \ j \ i \ restrs_condition tr (if i = j then restrs else pred_restrs addr restrs) j" apply (induct i) apply (clarsimp dest!: exec_trace_init) apply (simp add: trace_addr_def trace_bottom_addr_def) apply clarsimp apply (case_tac "trace_addr tr (Suc i)") defer apply (auto simp: trace_addr_trace_bottom_addr_eq)[1] apply (subgoal_tac "trace_bottom_addr tr i = Some addr") apply (case_tac "trace_addr tr i") apply simp apply (drule meta_mp) apply (simp add: restrs_condition_def Collect_less_Suc) apply (clarsimp, clarsimp split: if_split_asm) apply (fastforce simp add: atLeastAtMostSuc_conv) apply (simp add: trace_addr_trace_bottom_addr_eq pred_restrs) apply (rule_tac x=i in exI, simp) apply (frule_tac i=i in exec_trace_step_cases) apply (elim disjE conjE) apply (simp add: trace_bottom_addr_def) apply (clarsimp simp: trace_bottom_addr_def) apply (clarsimp simp: all_exec_graph_step_cases) apply (elim disjE conjE exE, simp_all add: trace_addr_def trace_bottom_addr_def split: list.split_asm graph_function.split_asm if_split_asm) done lemma visit_extended_pred: "\ trace_addr tr i = Some addr; restrs_condition tr restrs i; tr \ exec_trace Gamma fn; Gamma fn = Some f; wf_graph_function f ilen olen; converse (reachable_step' f) `` {addr} \ S; addr \ NextNode (entry_point f) \ \ \j nn'. trace_addr tr j = Some nn' \ nn' \ S \ j < i \ (trace_addr tr ` {Suc j ..< i} \ {None}) \ (trace_bottom_addr tr ` {Suc j ..< i} \ {Some nn'}) \ restrs_condition tr (pred_restrs nn' restrs) j" apply (frule(5) visit_immediate_pred) apply (clarsimp simp: pred_restrs) apply (frule(4) trace_bottom_addr_trace_addr_prev) apply clarsimp apply (rule_tac x=j in exI, simp) apply (clarsimp split: if_split_asm) apply (simp add: atLeastLessThanSuc_atLeastAtMost) done lemma if_x_None_eq_Some: "((if P then x else None) = Some y) = (P \ x = Some y)" by simp lemma subtract_le_nat: "((a :: nat) \ a - b) = (a = 0 \ b = 0)" by arith lemma bottom_addr_only: "trace_addr tr i = None \ trace_bottom_addr tr i = Some nn \ \x x' xs. tr i = Some (x # x' # xs) \ nn = fst (last (x' # xs))" apply (clarsimp simp: trace_addr_def trace_bottom_addr_def split: option.split_asm list.split_asm) apply blast done lemma extended_pred_trace_drop_n: "trace_addr tr i = Some (NextNode n) \ trace_addr tr j = Some nn \ tr \ exec_trace Gamma fn \ Gamma fn = Some f \ wf_graph_function f ilen olen \ function_graph f n = Some (Call nn fname inputs outputs) \ Gamma fname = Some f' \ i < j \ nn \ Err \ trace_addr tr ` {Suc i ..< j} \ {None} \ trace_bottom_addr tr ` {Suc i ..< j} \ {Some (NextNode n)} \ trace_drop_n (Suc i) 1 tr \ exec_trace Gamma fname \ Suc i < j \ (\st g. trace_drop_n (Suc i) 1 tr (j - Suc (Suc i)) = Some [(Ret, st, g)]) \ trace_drop_n (Suc i) 1 tr (j - Suc i) = None" apply (clarsimp dest!: trace_addr_SomeD) apply (frule_tac i=i in exec_trace_invariant, simp) apply (simp add: exec_graph_invariant_Cons) apply (frule(3) exec_trace_drop_n, simp+) apply (rule context_conjI) apply (rule ccontr, simp add: linorder_not_less le_Suc_eq) apply (frule_tac i=i in exec_trace_step_cases, simp) apply (clarsimp simp: all_exec_graph_step_cases) apply (subgoal_tac "Suc (j - 2) = j - 1", simp_all) apply (rule context_conjI) prefer 2 apply (clarsimp simp: trace_drop_n_def if_x_None_eq_Some) apply (rule_tac x="j - Suc (Suc i)" in exI, simp) apply (frule subsetD, rule_tac x="j - 1" and f="trace_addr tr" in imageI, simp) apply (frule subsetD, rule_tac x="j - 1" and f="trace_bottom_addr tr" in imageI, simp) apply (frule_tac i="j - 1" in exec_trace_step_cases) apply clarsimp apply (simp add: all_exec_graph_step_cases, safe; (solves \simp add: trace_addr_def\)?) apply (simp(no_asm) add: trace_drop_n_def if_x_None_eq_Some) apply clarsimp apply (cut_tac tr=tr and i="Suc (i + ja)" in bottom_addr_only) apply (simp add: image_subset_iff) apply (simp add: image_subset_iff) apply (cut_tac tr=tr and i="Suc (Suc (i + ja))" in bottom_addr_only) apply (simp add: image_subset_iff) apply (simp add: image_subset_iff) apply (clarsimp simp: continuing_def split: list.split next_node.split) apply (case_tac "drop 2 (the (tr (Suc (i + ja))))", simp_all) apply (frule_tac i="Suc (i + ja)" in exec_trace_step_cases) apply (frule_tac i="Suc (i + ja)" in exec_trace_invariant, simp) apply (auto simp: all_exec_graph_step_cases exec_graph_invariant_Cons split: graph_function.split_asm) done lemma restrs_condition_no_change: "restrs_condition tr restrs i \ j \ i \ (\k \ {i ..< j}. trace_addr tr k = None) \ restrs_condition tr restrs j" apply (clarsimp simp: restrs_condition_def) apply (rule_tac P="\S. card S \ SS" for SS in subst[rotated]) apply (erule spec) apply (auto intro: linorder_not_le[THEN iffD1, OF notI]) done lemma trace_end_exec_SomeI: "tr \ exec_trace Gamma fn \ tr i = Some stk \ tr (Suc i) = None \ trace_end tr = Some stk" apply (clarsimp simp: trace_end_def exI[where x="Suc i"]) apply (drule(1) exec_trace_None_dom_subset) apply (subst Max_eqI[where x=i], auto elim: finite_subset) done definition function_call_trace :: "nat \ restrs \ trace \ trace option" where "function_call_trace n restrs tr = (if pc' n restrs tr then Some (trace_drop_n (Suc (Least (\i. trace_addr tr i = Some (NextNode n) \ restrs_condition tr restrs i))) 1 tr) else None)" lemma function_call_trace_eq: assumes tr: "tr \ exec_trace Gamma fname" "Gamma fname = Some f" "wf_graph_function f ilen olen" and i: "trace_addr tr i = Some (NextNode n)" "restrs_condition tr restrs i" and cut: "\x. NextNode x \ set cuts \ (\y. restrs x \ {y})" "(NextNode n, NextNode n) \ trancl (reachable_step' f \ {(x, y). x \ set cuts})" shows "function_call_trace n restrs tr = Some (trace_drop_n (Suc i) 1 tr)" proof - have P: "\j < i. trace_addr tr j = Some (NextNode n) \ \ restrs_condition tr restrs j" using restrs_single_visit[OF tr _ _ i cut] by clarsimp note eq = Least_equality[where x=i, OF _ ccontr, unfolded linorder_not_le] show ?thesis using i apply (clarsimp simp: function_call_trace_def pc_def visit_known P dest!: trace_addr_SomeD) apply (subst eq, auto simp: i P) done qed definition function_call_trace_embed :: "state option \ trace \ trace option \ (string \ graph_function option) \ string \ (state \ variable) list \ bool" where "function_call_trace_embed vis tr tr' Gamma f inps = (((vis = None) = (tr' = None)) \ (\tr''. tr' = Some tr'' \ tr'' \ exec_trace Gamma f \ (trace_end tr'' = None \ trace_end tr = None) \ (option_map (fst o hd) (trace_end tr'') = Some Err \ option_map (fst o hd) (trace_end tr) = Some Err) \ length inps = length (function_inputs (the (Gamma f))) \ tr'' 0 = Some [init_state f (the (Gamma f)) (map (\i. i (the vis)) inps)]))" lemma exec_trace_Err_propagate: "tr \ exec_trace Gamma f \ tr i = Some ((Err, st, fname) # xs) \ j \ length xs \ tr (i + j) = Some (upd_stack Err id (drop j ((Err, st, fname) # xs)))" apply (induct j arbitrary: xs) apply simp apply atomize apply clarsimp apply (cut_tac xs="drop j ((Err, st, fname) # xs)" in neq_Nil_conv) apply clarsimp apply (frule_tac i="i + j" in exec_trace_step_cases) apply (auto simp: exec_graph_step_def drop_eq_Cons Cons_eq_append_conv split: list.split_asm) done lemma trace_end_trace_drop_n_Err: "option_map (fst o hd) (trace_end (trace_drop_n i j tr)) = Some Err \ tr \ exec_trace Gamma f \ trace_drop_n i j tr \ exec_trace Gamma f' \ option_map (fst o hd) (trace_end tr) = Some Err" apply clarsimp apply (drule(1) exec_trace_end_SomeD) apply (clarsimp simp: trace_drop_n_def[THEN fun_cong] if_x_None_eq_Some drop_eq_Cons rev_swap) apply (frule(1) exec_trace_Err_propagate[OF _ _ order_refl]) apply (case_tac "tl (the (tr (n + i)))" rule: rev_cases) apply (frule_tac i="n + i" in exec_trace_step_cases) apply (clarsimp simp: exec_graph_step_def trace_end_exec_SomeI) apply (simp add: rev_swap) apply (frule_tac i="n + i + length (the (tr (n + i))) - 2" in exec_trace_step_cases) apply (frule_tac i="n + i + length (the (tr (n + i))) - 1" in exec_trace_step_cases) apply (clarsimp simp: exec_graph_step_def trace_end_exec_SomeI) done lemma trace_end_Nil: "tr \ exec_trace Gamma f \ trace_end tr \ Some []" by (auto dest: exec_trace_end_SomeD simp: exec_trace_Nil) lemma visit_Call_loop_lemma: "(nn, NextNode n) \ rtrancl (reachable_step' f \ S) \ function_graph f n = Some (Call nn fname inputs outputs) \ converse (reachable_step' f) `` {nn} \ {NextNode n} \ (nn, nn) \ trancl (reachable_step' f \ S) \ (NextNode n, NextNode n) \ trancl (reachable_step' f \ S)" apply (simp add: not_trancl_converse_step) apply (clarsimp simp: reachable_step_def[THEN eqset_imp_iff] dest!: tranclD) apply (erule disjE, simp_all) apply (erule converse_rtranclE, simp_all add: reachable_step_def) done lemma pred_restrs_list: "pred_restrs nn (restrs_list xs) = restrs_list (map (\(i, ns). (i, if nn = NextNode i then map (\x. x - 1) (filter ((\) 0) ns) else ns)) xs)" apply (clarsimp simp: pred_restrs_def split: next_node.split) apply (rule sym) apply (induct xs; simp, rule ext) apply (simp add: restrs_list_def) apply (clarsimp simp: restrs_list_Cons split del: if_split cong: if_cong) apply (auto intro: image_eqI[rotated]) done lemma pred_restrs_cut: "(\y. restrs x \ {y}) \ (\y. pred_restrs nn restrs x \ {y})" apply (clarsimp simp: pred_restrs_def split: next_node.split) apply blast done lemma pred_restrs_cut2: "\x. NextNode x \ set cuts \ (\y. restrs x \ {y}) \ \x. NextNode x \ set cuts \ (\y. pred_restrs nn restrs x \ {y})" by (metis pred_restrs_cut) lemma visit_Call: "tr \ exec_trace Gamma fn \ Gamma fn = Some f \ wf_graph_function f ilen olen \ function_graph f n = Some (Call nn fname inputs outputs) \ Gamma fname = Some f' \ length inputs = length (function_inputs f') \ converse (reachable_step' f) `` {nn} \ {NextNode n} \ nn \ NextNode (entry_point f) \ nn \ Err \ (nn, NextNode n) \ rtrancl (reachable_step' f \ {(x, y). x \ set cuts}) \ \x. NextNode x \ set cuts \ (\y. restrs x \ {y}) \ let v = visit tr nn restrs; v' = visit tr (NextNode n) (pred_restrs' n restrs); tr' = function_call_trace n (pred_restrs' n restrs) tr in function_call_trace_embed v' tr tr' Gamma fname inputs \ v = option_map (\st. return_vars (function_outputs f') outputs (fst (snd (hd (the (trace_end (the tr')))))) st) (option_guard (\_. tr' \ None \ trace_end (the tr') \ None \ fst (hd (the (trace_end (the tr')))) = Ret) v')" apply (cases "visit tr (NextNode n) (pred_restrs' n restrs)", simp_all split del: if_split) apply (clarsimp simp: visit_eqs function_call_trace_embed_def) apply (rule context_conjI) apply (clarsimp simp: function_call_trace_def pc_def visit_eqs) apply clarsimp apply (frule(6) visit_extended_pred) apply clarsimp apply (drule visit_eqs[THEN iffD1]) apply (drule(2) visit_Call_loop_lemma) apply (frule_tac nn="NextNode n" in pred_restrs_cut2) apply (clarsimp simp: Let_def split del: if_split) apply (frule trace_addr_SomeD, clarsimp split del: if_split) apply (frule(1) exec_trace_invariant) apply (clarsimp simp: exec_graph_invariant_Cons split del: if_split) apply (frule(4) exec_trace_drop_n) apply (clarsimp simp: function_call_trace_eq visit_eqs Let_def simp del: imp_disjL split del: if_split) apply (rule conjI) apply (clarsimp simp: function_call_trace_embed_def simp del: map_option_eq_Some) apply (rule conjI, blast intro: trace_end_trace_drop_n_None) apply (rule conjI, metis trace_end_trace_drop_n_Err) apply (clarsimp simp: trace_drop_n_def) apply (frule_tac i=i in exec_trace_step_cases) apply (clarsimp simp: exec_graph_step_def init_state_def init_vars_def split: graph_function.split_asm dest!: trace_addr_SomeD) apply (simp del: imp_disjL) apply (rule conjI) apply (clarsimp simp: visit_eqs) apply (drule(1) exec_trace_end_SomeD, clarsimp) apply (frule(4) exec_trace_drop_n_rest[rule_format], simp) apply (frule_tac i="Suc i + na" in exec_trace_step_cases) apply (clarsimp simp: exec_graph_step_def split: graph_function.split_asm) apply (subgoal_tac "restrs_condition tr restrs (Suc (Suc i + na))") apply (rule exI, rule conjI, assumption) apply (clarsimp simp: trace_addr_SomeI) apply (drule(2) restrs_single_visit2[OF trace_addr_SomeI, OF exI], simp+)[1] apply (rule_tac i="Suc i" in restrs_condition_no_change) apply (simp add: pred_restrs[unfolded trace_bottom_addr_def]) apply simp apply clarsimp apply (rule ccontr, clarsimp dest!: trace_addr_SomeD) apply (case_tac "trace_drop_n (Suc i) 1 tr (k - Suc i)") apply simp apply (drule(1) exec_trace_None_dom_subset)+ apply auto[1] apply (frule(2) exec_trace_drop_n_rest[rule_format], fastforce, assumption+) apply (clarsimp simp: exec_trace_Nil) apply (intro impI allI notI conjI) apply (clarsimp simp: visit_eqs) apply (rename_tac k) apply (frule(6) visit_extended_pred) apply clarsimp apply (drule(8) restrs_single_visit[rotated 3]) apply clarsimp apply (frule(10) extended_pred_trace_drop_n) apply (clarsimp simp: trace_end_exec_SomeI Suc_diff_Suc) done definition "double_function_call_trace vis trs = (case vis of (side, NextNode n, restrs) \ function_call_trace n restrs (tuple_switch trs side) | _ \ None)" lemma double_visit_None: "(double_visit vis trs = None) = (\ double_pc vis trs)" by (simp add: double_pc_def double_visit_def split: prod.split) lemma restr_trace_refine_Call_single: "\ fst ccall \ (\nn outps. get_function_call gfs ccall = Some (nn, cfname, cinps, outps)) \ Gamma1 cfname = Some cf \ wf_graph_function cf cilen colen \ fst acall \ (\nn outps. get_function_call gfs acall = Some (nn, afname, ainps, outps)) \ Gamma2 afname = Some af \ wf_graph_function af ailen aolen \ graph_function_refine prec Gamma1 cfname Gamma2 afname irel orel' \ double_pc ccall (tr, tr') \ func_inputs_match gfs irel ccall acall (tr, tr') \ double_pc ccall (tr, tr') \ double_pc acall (tr, tr') \ function_call_trace_embed (double_visit ccall (tr, tr')) tr (ctr (tr, tr')) Gamma1 cfname cinps \ function_call_trace_embed (double_visit acall (tr, tr')) tr' (ctr' (tr, tr')) Gamma2 afname ainps \ (double_pc ccall (tr, tr') \ ctr (tr, tr') \ None \ ctr' (tr, tr') \ None \ trace_end (the (ctr (tr, tr'))) \ None \ trace_refine2 prec (cf, af) orel' (the (ctr (tr, tr'))) (the (ctr' (tr, tr')))) \ restr_trace_refine prec Gamma1 f1 Gamma2 f2 rs1 rs2 orel tr tr' \ restr_trace_refine prec Gamma1 f1 Gamma2 f2 rs1 rs2 orel tr tr'" apply (clarsimp simp: restr_trace_refine_def) apply (case_tac "\ double_pc ccall (tr, tr')") apply (auto simp: double_trace_imp_def)[1] apply (clarsimp simp: function_call_trace_embed_def double_trace_imp_def double_visit_None) apply (drule(4) graph_function_refine_trace2[THEN iffD1, rotated -1, rule_format]) apply simp apply (rule conjI, assumption)+ apply simp apply (rule exI, rule conjI, rule refl, simp)+ apply (clarsimp simp: func_inputs_match_def) apply (erule rsubst[where P=irel]) apply (rule ext, simp add: tuple_switch_def switch_val_def) apply (case_tac "trace_end (the (ctr (tr, tr')))") apply (case_tac "trace_end (the (ctr' (tr, tr')))") apply (clarsimp simp: trace_refine2_def trace_refine_def) apply (clarsimp simp: trace_refine2_def trace_refine_def) apply (clarsimp simp: list.split[where P="Not", THEN consume_Not_eq] next_node.split[where P="Not", THEN consume_Not_eq] trace_end_Nil) apply (drule(1) exec_trace_end_SomeD)+ apply clarsimp apply clarsimp apply fastforce done definition split_visit_rel where "split_visit_rel rel ccall acall j trs = (double_pc (ccall j) trs \ double_pc (acall j) trs \ rel j (\(x, y). the (double_visit ((tuple_switch (ccall, acall) x) (tuple_switch (0, j) y)) trs)))" lemma not_finite_two: "\ finite S \ \x y. x \ S \ y \ S \ x \ y" apply (case_tac "\x. x \ S") apply (rule ccontr, clarsimp) apply (erule notE, rule_tac B="{x}" in finite_subset) apply auto done definition induct_var :: "next_node \ nat \ bool" where "induct_var nn n = True" lemma infinite_subset: "\ finite S \ S \ T \ \ finite T" by (metis finite_subset) lemma restr_trace_refine_Split_orig: fixes rs1 rs2 hyps Gamma1 f1 Gamma2 f2 defines "H trs \ exec_double_trace Gamma1 f1 Gamma2 f2 (fst trs) (snd trs) \ restrs_eventually_condition (fst trs) (restrs_list rs1) \ restrs_eventually_condition (snd trs) (restrs_list rs2)" assumes Suc: "\i. double_pc (ccall (Suc i)) (tr, tr') \ double_pc (ccall i) (tr, tr')" and init: "\i. i < k \ H (tr, tr') \ double_pc (ccall i) (tr, tr') \ split_visit_rel rel ccall acall i (tr, tr')" and induct: "\i. H (tr, tr') \ double_pc (ccall (i + k)) (tr, tr') \ (\j < k. split_visit_rel rel ccall acall (i + j) (tr, tr')) \ split_visit_rel rel ccall acall (i + k) (tr, tr')" and distinct: "\tr i j k. restrs_condition tr (snd (snd (ccall i))) k \ restrs_condition tr (snd (snd (ccall j))) k \ i = j" "\tr i j k. restrs_condition tr (snd (snd (acall i))) k \ restrs_condition tr (snd (snd (acall j))) k \ i = j" and sides: "\i. \ fst (ccall i) \ fst (acall i)" and k: "k \ 0" and init_case: "\ double_pc (ccall k) (tr, tr') \ restr_trace_refine prec Gamma1 f1 Gamma2 f2 (restrs_list rs1) (restrs_list rs2) orel tr tr'" and induct_case: "\i. induct_var (fst (snd (ccall 0))) i \ induct_var (fst (snd (acall 0))) i \ \ double_pc (ccall (i + k)) (tr, tr') \ \j < k. split_visit_rel rel ccall acall (i + j) (tr, tr') \ restr_trace_refine prec Gamma1 f1 Gamma2 f2 (restrs_list rs1) (restrs_list rs2) orel tr tr'" shows "restr_trace_refine prec Gamma1 f1 Gamma2 f2 (restrs_list rs1) (restrs_list rs2) orel tr tr'" proof - have pc_ccall_less: "double_pc (ccall n) (tr, tr') \ m \ n \ double_pc (ccall m) (tr, tr')" for n m apply (induct_tac n, simp_all) apply clarsimp apply (frule Suc[rule_format]) apply (clarsimp simp: linorder_not_le Suc_le_eq[symmetric]) done {fix i assume le[simp]: "i \ k" note induct[rule_format, where i="i - k", simplified] } note induct_minus = this {fix j have "double_pc (ccall j) (tr, tr') \ H (tr, tr') \ split_visit_rel rel ccall acall j (tr, tr')" apply (induct j rule: nat_less_induct) apply (case_tac "n < k") apply (simp add: init) apply clarsimp apply (rule induct_minus, simp+) apply (simp add: field_simps pc_ccall_less) done } note rel_induct = this note H_intro = H_def[THEN meta_eq_to_obj_eq, THEN iffD2, OF conjI] have Suc_Max: "finite S \ Suc (Max S) \ S" for S by (auto dest: Max_ge) have unreached: "finite {n. \j x \ {n. \j \ double_pc (ccall (Max {n. \j exec_trace G f \ \i j k. restrs_condition tr (snd (vis i)) k \ restrs_condition tr (snd (vis j)) k \ i = j \ \ finite {n. pc (fst (vis n)) (snd (vis n)) tr} \ trace_end tr = None" for tr G f and vis :: "'b \ next_node \ (nat \ nat set)" apply (clarsimp simp: trace_end_def) apply (drule(1) exec_trace_None_dom_subset) apply (drule_tac R="\n i. restrs_condition tr (snd (vis n)) i" in pigeonhole_infinite_rel) apply (erule finite_subset, simp) apply (auto simp: pc_def visit_eqs dest!: trace_addr_SomeD)[1] apply (auto dest!: not_finite_two) done have infinite: "exec_double_trace Gamma1 f1 Gamma2 f2 tr tr' \ H (tr, tr') \ \ finite {n. \j trace_end tr = None \ trace_end tr' = None" apply (case_tac "\n. \ double_pc (ccall n) (tr, tr')") apply clarsimp apply (erule notE, rule_tac k1=n in finite_subset[OF _ finite_lessThan]) apply clarsimp apply (drule_tac x=0 in spec, simp add: k[simplified]) apply (rule ccontr, simp add: pc_ccall_less) apply clarsimp apply (rule conjI) apply (erule_tac vis="\i. snd (ccall i)" in infinite_helper) apply (simp add: distinct) apply (simp add: double_pc_def[folded pc_def] split_def sides) apply (erule_tac vis="\i. snd (acall i)" in infinite_helper) apply (simp add: distinct) apply clarsimp apply (drule_tac A="UNIV" in finite_subset[rotated]) apply clarsimp apply (rename_tac n, drule_tac x="n" in spec) apply (drule(1) rel_induct[rule_format, OF conjI]) apply (clarsimp simp: split_visit_rel_def) apply (simp add: double_pc_def[folded pc_def] sides split_def) apply simp done show ?thesis apply (clarsimp simp: restr_trace_refine_def) apply (case_tac "\ double_pc (ccall k) (tr, tr')") apply (rule init_case[unfolded restr_trace_refine_def, rule_format], auto)[1] apply (case_tac "finite {n. \j < k. double_pc (ccall (n + j)) (tr, tr')}") apply (frule Max_in) apply simp apply (rule_tac x="0" in exI) apply (simp add: pc_ccall_less) apply (rule_tac i="Max {n. \j < k. double_pc (ccall (n + j)) (tr, tr')}" in induct_case[unfolded restr_trace_refine_def, rule_format]) apply (simp add: induct_var_def)+ apply (simp add: rel_induct unreached) apply (simp add: rel_induct unreached H_intro) apply auto[1] apply clarsimp apply (drule infinite[rotated -1], simp) apply (simp add: H_def) apply (simp add: trace_refine2_def trace_refine_def) done qed lemma restrs_condition_unique: "restrs_condition tr (restrs_list ((n, [x]) # rs)) k \ restrs_condition tr (restrs_list ((n, [y]) # rs)) k \ x = y" by (clarsimp simp: restrs_condition_def restrs_list_Cons split: if_split_asm) abbreviation(input) "split_restr split_seq restrs i \ (restrs_list ((fst split_seq, [fst (snd split_seq) + i * snd (snd split_seq)]) # restrs))" abbreviation(input) "split_pc split_seq restrs tr i \ pc' (fst split_seq) (split_restr split_seq restrs i) tr" abbreviation(input) "split_visit split_seq restrs i tr \ visit tr (NextNode (fst split_seq)) (split_restr split_seq restrs i)" definition split_visit_rel' where "split_visit_rel' rel conc_seq abs_seq restrs trs j = (split_pc conc_seq (fst restrs) (fst trs) j \ split_pc abs_seq (snd restrs) (snd trs) j \ rel j (\(x, y). the (split_visit (tuple_switch (conc_seq, abs_seq) x) (tuple_switch restrs x) (tuple_switch (0, j) y) (tuple_switch trs x))))" lemma split_visit_rel': "split_visit_rel rel (\i. (False, NextNode (fst cseq), split_restr cseq rs1 i)) (\i. (True, NextNode (fst aseq), split_restr aseq rs2 i)) j trs = split_visit_rel' rel cseq aseq (rs1, rs2) trs j" apply (simp add: split_visit_rel_def split_visit_rel'_def double_pc_def[folded pc_def] split_def double_visit_def) apply (auto elim!: rsubst[where P="rel j", OF _ ext]) apply (auto simp: tuple_switch_def switch_val_def) done theorem restr_trace_refine_Split: assumes Suc: "\i. split_pc conc_seq rs1 tr (Suc i) \ split_pc conc_seq rs1 tr i" and init: "\i. i < k \ split_pc conc_seq rs1 tr i \ split_visit_rel' rel conc_seq abs_seq (rs1, rs2) (tr, tr') i" and induct: "\i. split_pc conc_seq rs1 tr (i + k) \ (\j < k. split_visit_rel' rel conc_seq abs_seq (rs1, rs2) (tr, tr') (i + j)) \ split_visit_rel' rel conc_seq abs_seq (rs1, rs2) (tr, tr') (i + k)" and zero: "k \ 0" "snd (snd conc_seq) \ 0" "snd (snd abs_seq) \ 0" and init_case: "\ split_pc conc_seq rs1 tr k \ restr_trace_refine prec Gamma1 f1 Gamma2 f2 (restrs_list rs1) (restrs_list rs2) orel tr tr'" and induct_case: "\i. induct_var (NextNode (fst conc_seq)) i \ induct_var (NextNode (fst abs_seq)) i \ \ split_pc conc_seq rs1 tr (i + k) \ \j < k. split_visit_rel' rel conc_seq abs_seq (rs1, rs2) (tr, tr') (i + j) \ restr_trace_refine prec Gamma1 f1 Gamma2 f2 (restrs_list rs1) (restrs_list rs2) orel tr tr'" shows "restr_trace_refine prec Gamma1 f1 Gamma2 f2 (restrs_list rs1) (restrs_list rs2) orel tr tr'" apply (rule restr_trace_refine_Split_orig[where ccall="\i. (False, NextNode (fst conc_seq), split_restr conc_seq rs1 i)" and acall="\i. (True, NextNode (fst abs_seq), split_restr abs_seq rs2 i)" and k=k and rel=rel], simp_all add: Suc init induct init_case induct_case zero[simplified] double_pc_def[folded pc_def] split_visit_rel') apply (clarsimp, drule(1) restrs_condition_unique) apply (simp add: zero) apply (clarsimp, drule(1) restrs_condition_unique) apply (simp add: zero) done theorem restr_trace_refine_Split': "let cpc = split_pc conc_seq rs1 tr; rel = split_visit_rel' rel conc_seq abs_seq (rs1, rs2) (tr, tr') in (\i. cpc (Suc i) --> cpc i) --> (\i. i < k --> cpc i --> rel i) --> (\i. cpc (i + k) --> (\j < k. rel (i + j)) --> rel (i + k)) --> k \ 0 --> snd (snd conc_seq) \ 0 --> snd (snd abs_seq) \ 0 --> (~ cpc k --> restr_trace_refine prec Gamma1 f1 Gamma2 f2 (restrs_list rs1) (restrs_list rs2) orel tr tr') --> (\i. ~ cpc (i + k) --> (\j < k. rel (i + j)) --> restr_trace_refine prec Gamma1 f1 Gamma2 f2 (restrs_list rs1) (restrs_list rs2) orel tr tr') --> restr_trace_refine prec Gamma1 f1 Gamma2 f2 (restrs_list rs1) (restrs_list rs2) orel tr tr'" apply (clarsimp simp only: Let_def) apply (erule notE, rule restr_trace_refine_Split[where conc_seq=conc_seq and abs_seq=abs_seq and k=k and rel=rel], simp_all) apply blast done lemma restr_trace_refine_Restr1_offset: "induct_var (NextNode n) iv \ restr_trace_refine prec Gamma1 fn1 Gamma2 fn2 (restrs_list ((n, [iv + 1 ..< iv + j]) # rs1)) rs2 orel tr tr' \ j \ 0 \ distinct (map fst rs1) \ wf_graph_function f1 ilen olen \ Gamma1 fn1 = Some f1 \ pc' n (restrs_list ((n, [iv]) # (restrs_visit rs1 (NextNode n) f1))) tr \ \ pc' n (restrs_list ((n, [iv + (j - 1)]) # (restrs_visit rs1 (NextNode n) f1))) tr \ restr_trace_refine prec Gamma1 fn1 Gamma2 fn2 (restrs_list rs1) rs2 orel tr tr'" by (rule restr_trace_refine_Restr1[where i="iv + 1" and j="iv + j"], simp_all) lemma restr_trace_refine_Restr2_offset: "induct_var (NextNode n) iv \ restr_trace_refine prec Gamma1 fn1 Gamma2 fn2 rs1 (restrs_list ((n, [iv + 1 ..< iv + j]) # rs2)) orel tr tr' \ j \ 0 \ distinct (map fst rs2) \ wf_graph_function f2 ilen olen \ Gamma2 fn2 = Some f2 \ pc' n (restrs_list ((n, [iv]) # (restrs_visit rs2 (NextNode n) f2))) tr' \ \ pc' n (restrs_list ((n, [iv + (j - 1)]) # (restrs_visit rs2 (NextNode n) f2))) tr' \ restr_trace_refine prec Gamma1 fn1 Gamma2 fn2 rs1 (restrs_list rs2) orel tr tr'" by (rule restr_trace_refine_Restr2[where i="iv + 1" and j="iv + j"], simp_all) lemma restr_trace_refine_PCCases1: "pc nn rs1 tr \ restr_trace_refine prec Gamma1 fn1 Gamma2 fn2 rs1 rs2 orel tr tr' \ \ pc nn rs1 tr \ restr_trace_refine prec Gamma1 fn1 Gamma2 fn2 rs1 rs2 orel tr tr' \ restr_trace_refine prec Gamma1 fn1 Gamma2 fn2 rs1 rs2 orel tr tr'" by metis lemma restr_trace_refine_PCCases2: "pc nn rs2 tr' \ restr_trace_refine prec Gamma1 fn1 Gamma2 fn2 rs1 rs2 orel tr tr' \ \ pc nn rs2 tr' \ restr_trace_refine prec Gamma1 fn1 Gamma2 fn2 rs1 rs2 orel tr tr' \ restr_trace_refine prec Gamma1 fn1 Gamma2 fn2 rs1 rs2 orel tr tr'" by metis lemma restr_trace_refine_Err: "(\ pc Err restrs tr' \ restr_trace_refine prec Gamma1 fn1 Gamma2 fn2 rs1 rs2 orel tr tr') \ restr_trace_refine prec Gamma1 fn1 Gamma2 fn2 rs1 rs2 orel tr tr'" apply (clarsimp simp: restr_trace_refine_def) apply (erule impCE) apply simp apply (clarsimp simp: trace_refine2_def trace_refine_def) apply (drule(1) pc_Ret_Err_trace_end[where er=Err, simplified]) apply clarsimp apply (simp split: option.split list.split next_node.split) apply auto done definition stepwise_graph_refine :: "(string \ graph_function option) \ (string list \ string list) \ (next_node list \ next_node list) \ (state list \ state list \ bool) \ ((next_node \ state \ string) \ (next_node \ state \ string) \ bool) \ (trace \ trace) option \ nat \ bool" where "stepwise_graph_refine Gamma fnames nns irel orel trs n = (\tr tr' i j. i \ n \ j \ n \ tr i \ None \ tr' j \ None \ (map fst (the (tr i)), map fst (the (tr' j))) = nns \ (map (snd o snd) (the (tr i)), map (snd o snd) (the (tr' j))) = fnames \ irel (map (fst o snd) (the (tr i)), map (fst o snd) (the (tr' j))) \ (trs \ None \ trs = Some (tr, tr')) \ tr \ exec_trace Gamma (last (fst fnames)) \ tr' \ exec_trace Gamma (last (snd fnames)) \ (trace_end tr = None) = (trace_end tr' = None) \ (trace_end tr \ None \ orel (hd (the (trace_end tr)), hd (the (trace_end tr')))))" lemmas stepwise_graph_refineI = stepwise_graph_refine_def[THEN iffD2, rule_format] lemmas stepwise_graph_refineD = stepwise_graph_refine_def[THEN iffD1, rule_format] lemma stepwise_graph_refine_Basic: "stepwise_graph_refine Gamma (f1 # fns, f2 # fns') (nn # nns, nn' # nns') rel orel otr (Suc N) \ Gamma f1 = Some gf1 \ Gamma f2 = Some gf2 \ function_graph gf1 n = Some (Basic nn upds) \ function_graph gf2 n' = Some (Basic nn' upds') \ \x xs y ys. rel (x # xs, y # ys) \ rel (upd_vars upds x # xs, upd_vars upds' y # ys) \ stepwise_graph_refine Gamma (f1 # fns, f2 # fns') (NextNode n # nns, NextNode n' # nns') rel orel otr N" apply (rule stepwise_graph_refineI) apply (erule_tac i="Suc i" and j="Suc j" in stepwise_graph_refineD) apply clarsimp apply (frule_tac i=i and tr=tr in exec_trace_step_cases) apply (frule_tac i=j and tr=tr' in exec_trace_step_cases) apply (clarsimp simp: exec_graph_step_def exec_graph_invariant_Cons split: graph_function.split_asm) apply (simp add: K_def) done lemma stepwise_graph_refine_Cond: "stepwise_graph_refine Gamma (f1 # fns, f2 # fns') (l # nns, l' # nns') rel orel otr (Suc N) \ stepwise_graph_refine Gamma (f1 # fns, f2 # fns') (r # nns, r' # nns') rel orel otr (Suc N) \ Gamma f1 = Some gf1 \ Gamma f2 = Some gf2 \ function_graph gf1 n = Some (Cond l r cond) \ function_graph gf2 n' = Some (Cond l' r' cond') \ \x xs y ys. rel (x # xs, y # ys) \ cond x = cond' y \ stepwise_graph_refine Gamma (f1 # fns, f2 # fns') (NextNode n # nns, NextNode n' # nns') rel orel otr N" apply (rule stepwise_graph_refineI) apply (clarsimp simp del: last.simps not_None_eq) apply (frule_tac i=i and tr=tr in exec_trace_step_cases) apply (frule_tac i=j and tr=tr' in exec_trace_step_cases) apply (clarsimp simp: exec_graph_step_def exec_graph_invariant_Cons split: graph_function.split_asm split: if_split_asm simp del: last.simps not_None_eq) apply (simp_all only: exI all_simps[symmetric] simp_thms) apply (erule_tac i="Suc i" and j="Suc j" and nns="(l # xs, ys)" for xs ys in stepwise_graph_refineD, simp) apply (erule_tac i="Suc i" and j="Suc j" and nns="(r # xs, ys)" for xs ys in stepwise_graph_refineD, simp) done lemma stepwise_graph_refine_Call_same: "stepwise_graph_refine Gamma (f1 # fns, f2 # fns') (nn # nns, nn' # nns') rel orel otr (Suc N) \ Gamma f1 = Some gf1 \ Gamma f2 = Some gf2 \ function_graph gf1 n = Some (Call nn fname inputs outputs) \ function_graph gf2 n' = Some (Call nn' fname inputs' outputs') \ Gamma fname = Some gf3 \ \x xs y ys. rel (x # xs, y # ys) \ map (\i. i x) inputs = map (\i. i y) inputs' \ \x xs y ys zs st. rel (x # xs, y # ys) \ rel (return_vars zs outputs st x # xs, return_vars zs outputs' st y # ys) \ \sts. fst (fst sts) = Err \ fst (snd sts) = Err \ orel sts \ stepwise_graph_refine Gamma (f1 # fns, f2 # fns') (NextNode n # nns, NextNode n' # nns') rel orel otr N" apply (rule stepwise_graph_refineI, clarsimp) apply (simp only: all_simps[symmetric]) apply (frule(4) exec_trace_drop_n_Cons[where gf=gf1]) apply (frule(4) exec_trace_drop_n_Cons[where gf=gf2]) apply (frule_tac i=i and tr=tr in exec_trace_step_cases) apply (frule_tac i=j and tr=tr' in exec_trace_step_cases) apply (clarsimp simp: exec_graph_step_def init_vars_def split: graph_function.split_asm) apply (subgoal_tac "trace_drop_n (Suc i) (length (the (tr i))) tr = trace_drop_n (Suc j) (length (the (tr' j))) tr'") apply (frule_tac f=trace_end in arg_cong) apply (drule fun_eq_iff[THEN iffD1]) apply (case_tac "trace_end (trace_drop_n (Suc i) (length (the (tr i))) tr)") apply simp apply (drule(1) trace_end_trace_drop_n_None, simp)+ apply simp apply clarsimp apply (drule(1) exec_trace_end_SomeD)+ apply clarsimp apply (erule disjE) apply (frule_tac tr=tr and i=i and k=na in exec_trace_drop_n_rest_Cons[rule_format], (simp only: function_graph.simps)+) apply (frule_tac tr=tr' and i=j and k=na in exec_trace_drop_n_rest_Cons[rule_format], (simp only: function_graph.simps)+) apply (frule_tac i="Suc i + na" and tr=tr in exec_trace_step_cases) apply (frule_tac i="Suc j + na" and tr=tr' in exec_trace_step_cases) apply (clarsimp simp: field_simps exec_graph_step_def) apply (drule_tac tr=tr and tr'=tr' and i="i + na + 2" and j="j + na + 2" in stepwise_graph_refineD) apply (simp add: ) apply auto[1] apply auto[1] apply (drule(1) trace_end_trace_drop_n_Err[rotated], simp add: trace_end_exec_SomeI)+ apply (case_tac "hd (the (trace_end tr))") apply (case_tac "hd (the (trace_end tr'))") apply clarsimp apply (rule ext) apply (rule exec_trace_deterministic, simp+) apply (simp add: trace_drop_n_def) apply (simp only: all_simps[symmetric]) done lemma stepwise_graph_refine_nop_left: "stepwise_graph_refine Gamma (f1 # fns, fns') (nn # nns, nns') rel orel otr N \ Gamma f1 = Some gf1 \ function_graph gf1 n = Some (Basic nn []) \ stepwise_graph_refine Gamma (f1 # fns, fns') (NextNode n # nns, nns') rel orel otr N" apply (rule stepwise_graph_refineI) apply (erule_tac i="Suc i" and j="j" in stepwise_graph_refineD) apply clarsimp apply (frule_tac i=i and tr=tr in exec_trace_step_cases) apply (clarsimp simp: exec_graph_step_def exec_graph_invariant_Cons split: graph_function.split_asm) apply (simp add: K_def upd_vars_def save_vals_def) done lemma stepwise_graph_refine_flip: "stepwise_graph_refine Gamma (fns', fns) (nns', nns) (\x. rel (snd x, fst x)) (\x. orel (snd x, fst x)) (option_map (\x. (snd x, fst x)) otr) N = stepwise_graph_refine Gamma (fns, fns') (nns, nns') rel orel otr N" apply (intro iffI stepwise_graph_refineI) apply (frule_tac i=j and j=i and tr=tr' and tr'=tr in stepwise_graph_refineD) apply simp apply auto[1] apply (frule_tac i=j and j=i and tr=tr' and tr'=tr in stepwise_graph_refineD) apply simp apply auto[1] done lemma stepwise_graph_refine_nop_right: "stepwise_graph_refine Gamma (fns, f2 # fns') (nns, nn' # nns') rel orel otr N \ Gamma f2 = Some gf2 \ function_graph gf2 n' = Some (Basic nn' []) \ stepwise_graph_refine Gamma (fns, f2 # fns') (nns, NextNode n' # nns') rel orel otr N" apply (rule stepwise_graph_refineI) apply (erule_tac i="i" and j="Suc j" in stepwise_graph_refineD) apply clarsimp apply (frule_tac i=j and tr=tr' in exec_trace_step_cases) apply (clarsimp simp: exec_graph_step_def exec_graph_invariant_Cons split: graph_function.split_asm) apply (simp add: K_def upd_vars_def save_vals_def) done lemma stepwise_graph_refine_inline_left: "stepwise_graph_refine Gamma (fname # f1 # fns, f2 # fns') (NextNode (entry_point gf3) # NextNode n # nns, nn' # nns') rel' orel otr (Suc N) \ Gamma f1 = Some gf1 \ Gamma f2 = Some gf2 \ function_graph gf1 n = Some (Call nn fname inputs outputs) \ function_graph gf2 n' = Some (Basic nn' upds) \ Gamma fname = Some gf3 \ \x xs y ys. rel (x # xs, y # ys) \ rel' (init_vars (function_inputs gf3) inputs x # x # xs, upd_vars upds y # ys) \ stepwise_graph_refine Gamma (f1 # fns, f2 # fns') (NextNode n # nns, NextNode n' # nns') rel orel otr N" apply (rule stepwise_graph_refineI) apply (erule_tac i="Suc i" and j="Suc j" in stepwise_graph_refineD) apply clarsimp apply (frule_tac i=i and tr=tr in exec_trace_step_cases) apply (frule_tac i=j and tr=tr' in exec_trace_step_cases) apply (clarsimp simp: exec_graph_step_def exec_graph_invariant_Cons split: graph_function.split_asm) apply (simp add: K_def) done lemma stepwise_graph_refine_end_inline_left: "stepwise_graph_refine Gamma (f1 # fns, f2 # fns') (nn # nns, nn' # nns') rel' orel otr (Suc N) \ Gamma f1 = Some gf1 \ Gamma f2 = Some gf2 \ function_graph gf1 n = Some (Call nn fname inputs outputs) \ function_graph gf2 n' = Some (Basic nn' upds) \ Gamma fname = Some gf3 \ \x x' xs y ys. rel (x # x' # xs, y # ys) \ rel' (return_vars (function_outputs gf3) outputs x x' # xs, upd_vars upds y # ys) \ stepwise_graph_refine Gamma (fname # f1 # fns, f2 # fns') (Ret # NextNode n # nns, NextNode n' # nns') rel orel otr N" apply (rule stepwise_graph_refineI) apply (erule_tac i="Suc i" and j="Suc j" in stepwise_graph_refineD) apply clarsimp apply (frule_tac i=i and tr=tr in exec_trace_step_cases) apply (frule_tac i=j and tr=tr' in exec_trace_step_cases) apply (clarsimp simp: exec_graph_step_def exec_graph_invariant_Cons split: graph_function.split_asm) apply (simp add: K_def) done lemma stepwise_graph_refine_inline_right: "stepwise_graph_refine Gamma (f1 # fns, fname # f2 # fns') (nn # nns, NextNode (entry_point gf3) # NextNode n' # nns') rel' orel otr (Suc N) \ Gamma f1 = Some gf1 \ Gamma f2 = Some gf2 \ function_graph gf1 n = Some (Basic nn upds) \ function_graph gf2 n' = Some (Call nn' fname inputs outputs) \ Gamma fname = Some gf3 \ \x xs y ys. rel (x # xs, y # ys) \ rel' (upd_vars upds x # xs, init_vars (function_inputs gf3) inputs y # y # ys) \ stepwise_graph_refine Gamma (f1 # fns, f2 # fns') (NextNode n # nns, NextNode n' # nns') rel orel otr N" apply (rule stepwise_graph_refineI) apply (erule_tac i="Suc i" and j="Suc j" in stepwise_graph_refineD) apply clarsimp apply (frule_tac i=i and tr=tr in exec_trace_step_cases) apply (frule_tac i=j and tr=tr' in exec_trace_step_cases) apply (clarsimp simp: exec_graph_step_def exec_graph_invariant_Cons split: graph_function.split_asm) apply (simp add: K_def) done lemma stepwise_graph_refine_end_inline_right: "stepwise_graph_refine Gamma (f1 # fns, f2 # fns') (nn # nns, nn' # nns') rel' orel otr (Suc N) \ Gamma f1 = Some gf1 \ Gamma f2 = Some gf2 \ function_graph gf1 n = Some (Basic nn upds) \ function_graph gf2 n' = Some (Call nn' fname inputs outputs) \ Gamma fname = Some gf3 \ \x xs y y' ys. rel (x # xs, y # y' # ys) \ rel' (upd_vars upds x # xs, return_vars (function_outputs gf3) outputs y y' # ys) \ stepwise_graph_refine Gamma (f1 # fns, fname # f2 # fns') (NextNode n # nns, Ret # NextNode n' # nns') rel orel otr N" apply (rule stepwise_graph_refineI) apply (erule_tac i="Suc i" and j="Suc j" in stepwise_graph_refineD) apply clarsimp apply (frule_tac i=i and tr=tr in exec_trace_step_cases) apply (frule_tac i=j and tr=tr' in exec_trace_step_cases) apply (clarsimp simp: exec_graph_step_def exec_graph_invariant_Cons split: graph_function.split_asm) apply (simp add: K_def) done lemma stepwise_graph_refine_induct: assumes Suc: "\n otr. stepwise_graph_refine Gamma fns nns rel orel otr (Suc n) \ stepwise_graph_refine Gamma fns nns rel orel otr n" shows "stepwise_graph_refine Gamma fns nns rel orel otr n" proof - have induct: "\n m otr. n \ m \ stepwise_graph_refine Gamma fns nns rel orel otr n \ stepwise_graph_refine Gamma fns nns rel orel otr m" by (erule(1) inc_induct, erule Suc) show ?thesis apply (rule stepwise_graph_refineI) apply (simp add: prod_eq_iff) apply (case_tac "\N. (\i \ N. tr i = None) \ (\i \ N. tr' i = None)") apply clarsimp apply (cut_tac n="Suc N" and m="n" and otr="Some (tr, tr')" in induct) apply (rule ccontr, auto)[1] apply (auto simp add: stepwise_graph_refine_def)[1] apply (drule_tac tr=tr and tr'=tr' in stepwise_graph_refineD, (erule conjI)+) apply simp apply simp apply (case_tac "trace_end tr") apply (rule ccontr, clarsimp) apply (frule(1) exec_trace_end_SomeD, clarsimp) apply (frule(1) exec_trace_None_dom_subset) apply (drule_tac x="Suc na" in spec) apply auto[1] apply clarsimp apply (frule(1) exec_trace_end_SomeD, clarsimp) apply (frule(1) exec_trace_None_dom_subset) apply (drule_tac x="Suc na" in spec) apply auto[1] done qed end