diff --git a/thys/CAVA_LTL_Modelchecker/SM/Impl/SM_Ample_Impl.thy b/thys/CAVA_LTL_Modelchecker/SM/Impl/SM_Ample_Impl.thy --- a/thys/CAVA_LTL_Modelchecker/SM/Impl/SM_Ample_Impl.thy +++ b/thys/CAVA_LTL_Modelchecker/SM/Impl/SM_Ample_Impl.thy @@ -1,1413 +1,1413 @@ theory SM_Ample_Impl imports SM_Datastructures "../Analysis/SM_POR" "HOL-Library.RBT_Mapping" begin text \ Given a sticky program, we implement a valid ample-function by picking the transitions from the first PID that qualifies as an ample-set\ context visible_prog begin definition ga_gample :: "(cmdc\cmdc) set \ pid_global_config \ global_action set" where "ga_gample sticky_E gc \ let lcs = pid_global_config.processes gc; gs = pid_global_config.state gc; as = find_min_idx_f (\lc. let c = local_config.command lc; ls = local_config.state lc; as = set (comp.succ_impl c) in if (\(a,_)\as. read_globals a = {} \ write_globals a = {}) then let as = {(a,c')\as. la_en' (ls,gs) a} in if as={} \ (\(a,c')\as. (c,c')\sticky_E) then None else Some (c,as) else None ) lcs in case as of Some (pid,c,as) \ (\(a,c'). (pid,c,a,c'))`as | None \ ga_gen gc" definition "ga_ample sticky_E \ stutter_extend_en (ga_gample sticky_E)" lemma ga_gample_subset: "ga_gample sticky_E gc \ ga_gen gc" unfolding ga_gample_def ga_gen_def by (auto simp: comp.astep_impl_def split: option.splits if_split_asm dest!: find_min_idx_f_SomeD) lemma ga_gample_empty_iff: "ga_gample sticky_E gc = {} \ ga_gen gc = {}" unfolding ga_gample_def ga_gen_def by (fastforce simp: comp.astep_impl_def split: option.splits if_split_asm dest!: find_min_idx_f_SomeD) lemma ga_ample_None_iff: "None \ ga_ample sticky_E gc \ None \ ga_en gc" by (auto simp: ga_ample_def ga_en_def stutter_extend_en_def ga_gample_empty_iff) lemma ga_ample_neq_en_eq: assumes "ga_ample sticky_E gc \ ga_en gc" shows "ga_ample sticky_E gc = Some`ga_gample sticky_E gc \ ga_en gc = Some`ga_gen gc" using assms ga_gample_subset unfolding ga_ample_def ga_en_def stutter_extend_en_def by (auto split: if_split_asm simp: ga_gample_empty_iff) lemma pid_pac_alt: "pid_pac pid = insert None (Some`{(pid',cac). pid'=pid})" apply (auto simp: pid_pac_def split: option.splits) apply (case_tac x) apply auto done end context sticky_prog begin sublocale ample_prog prog is_vis_var sticky_E "ga_ample sticky_E" proof - have aux: "Some ` A = Some ` B \ pid_pac pid \ A = {(pid', cac). (pid', cac) \ B \ pid' = pid}" for A B pid proof assume 1: "Some ` A = Some ` B \ pid_pac pid" show "A = {(pid', cac). (pid', cac) \ B \ pid' = pid}" using 1 apply (auto simp: pid_pac_alt) unfolding image_def apply auto using 1 by blast next show "Some ` A = Some ` B \ pid_pac pid" if "A = {(pid', cac). (pid', cac) \ B \ pid' = pid}" using that by (auto simp: pid_pac_alt) qed show "ample_prog prog is_vis_var sticky_E (ga_ample sticky_E)" apply unfold_locales apply (frule ga_ample_neq_en_eq, clarsimp) apply (intro conjI) apply (auto simp: ga_gample_def sticky_ga_def dest!: find_min_idx_f_SomeD split: option.splits if_split_asm) [] apply (fastforce simp: ga_gample_def sticky_ga_def dest!: find_min_idx_f_SomeD split: option.splits if_split_asm) [] unfolding aux (* TODO: Clean up! *) apply (simp add: inj_image_eq_iff pid_pac_def) apply (thin_tac "ga_ample c a = b" for a b c) apply (simp add: ga_gample_def split: option.splits prod.splits) apply (thin_tac "a \ ga_gen b" for a b) apply (drule find_min_idx_f_SomeD) apply clarsimp apply (rename_tac gc pid c as) apply (rule_tac x=pid in exI) apply (auto split: if_split_asm prod.splits simp: ga_gen_def comp.astep_impl_def) unfolding is_ample_static_def comp.astep_impl_def apply force done qed end text \ The following locale expresses a correct ample-reduction on the level of subset and stuttering-equivalences of traces \ locale hl_reduction = visible_prog + ample: sa "\ g_V = UNIV, g_E = step, g_V0 = {pid_init_gc}, sa_L = pid_interp_gc \" for step + assumes ample_accept_subset: "ample.accept w \ lv.sa.accept w" assumes ample_accept_stuttering: "lv.sa.accept w \ \w'. w\w' \ ample.accept w'" locale hl_ample_reduction = visible_prog + ample: sa "\ g_V = UNIV, g_E = rel_of_enex (ample, ga_ex), g_V0 = {pid_init_gc}, sa_L = pid_interp_gc \" for ample + assumes ample_accept_subset: "ample.accept w \ lv.sa.accept w" assumes ample_accept_stuttering: "lv.sa.accept w \ \w'. w\w' \ ample.accept w'" begin sublocale hl_reduction prog is_vis_var "rel_of_enex (ample, ga_ex)" apply unfold_locales using ample_accept_subset ample_accept_stuttering by auto end context sticky_prog begin sublocale hl_ample: hl_ample_reduction prog is_vis_var "ga_ample sticky_E" apply unfold_locales unfolding ample_rel_def[symmetric] unfolding reduced_automaton_def[symmetric] apply simp apply simp using ample_accept_subset apply simp apply (erule ample_accept_stuttering) apply blast done end text \Next, we implement the selection of a set of sticky edges\ context visible_prog begin definition "cr_ample \ do { sticky_E \ find_fas_init cfgc_G vis_edges; RETURN (ga_ample (sticky_E)) }" lemma cr_ample_reduction: "cr_ample \ SPEC (hl_ample_reduction prog is_vis_var)" unfolding cr_ample_def apply (refine_vcg order_trans[OF find_fas_init_correct]) apply unfold_locales[] apply simp proof clarify fix sticky_E assume "is_fas cfgc_G sticky_E" "vis_edges \ sticky_E" then interpret sticky_prog prog is_vis_var sticky_E by unfold_locales show "hl_ample_reduction prog is_vis_var (ga_ample sticky_E)" by unfold_locales qed end text \We derive an implementation for finding the feedback arcs, and an efficient implementation for ga_gample.\ text \We replace finding the sticky edges by an efficient algorithm\ context visible_prog begin definition ga_gample_m :: "_ \ _ \ pid_global_config \ global_action set" where "ga_gample_m mem sticky_E gc \ let lcs = pid_global_config.processes gc; gs = pid_global_config.state gc; as = find_min_idx_f (\lc. let c = local_config.command lc; ls = local_config.state lc; as = set (comp.succ_impl c) in if (\(a,_)\as. read_globals a = {} \ write_globals a = {}) then let as = {(a,c')\as. la_en' (ls,gs) a} in if as={} \ (\(a,c')\as. mem (c,c') sticky_E) then None else Some (c,as) else None ) lcs in case as of Some (pid,c,as) \ (\(a,c'). (pid,c,a,c'))`as | None \ ga_gen gc" lemma ga_gample_m_refine: fixes Rs :: "('si\'s) set" shows "(ga_gample_m, ga_gample_m) \ (Id\\<^sub>rId \ Rs \ bool_rel) \ Rs \ Id \ \Id\set_rel" proof (intro fun_relI, simp) fix m :: "cmdc\cmdc \ 's \ bool" and m' :: "cmdc\cmdc \ 'si \ bool" and se se' gc assume "(m',m)\Id \ Rs \ bool_rel" and "(se',se)\Rs" hence "\c c'. (m' (c,c') se', m (c,c') se)\Id" by parametricity simp_all hence A: "\c c'. m' (c,c') se' = m (c,c') se" by simp show "ga_gample_m m' se' gc = ga_gample_m m se gc" unfolding ga_gample_m_def by (subst A) (rule refl) qed lemma ga_gample_eq_gample_m: "ga_gample = ga_gample_m (\)" unfolding ga_gample_def[abs_def] ga_gample_m_def[abs_def] by auto lemma stutter_extend_en_refine: "(stutter_extend_en, stutter_extend_en) \ (R \ \Id\set_rel) \ R \ \\Id\option_rel\set_rel" unfolding stutter_extend_en_def[abs_def] apply parametricity by auto lemma is_vis_var_refine: "(is_vis_var, is_vis_var) \ Id \ bool_rel" by simp lemma vis_edges_refine: "(\x. x\vis_edges,vis_edges)\\Id\\<^sub>rId\fun_set_rel" by (simp add: fun_set_rel_def br_def) lemma [autoref_op_pat_def]: "vis_edges \ Autoref_Tagging.OP vis_edges" "ga_gample_m \ Autoref_Tagging.OP ga_gample_m" "cfgc_G \ Autoref_Tagging.OP cfgc_G" by simp_all schematic_goal cr_ample_impl1_aux: notes [autoref_rules] = ga_gample_m_refine stutter_extend_en_refine is_vis_var_refine cfgc_G_impl_refine vis_edges_refine IdI[of prog] shows "(RETURN (?c::?'c),cr_ample)\?R" unfolding cr_ample_def ga_ample_def ga_gample_eq_gample_m using [[autoref_trace_failed_id, autoref_trace_intf_unif]] by (autoref_monadic (plain, trace)) concrete_definition cr_ample_impl1 uses cr_ample_impl1_aux lemma cr_ample_impl1_reduction: "hl_ample_reduction prog is_vis_var cr_ample_impl1" proof - note cr_ample_impl1.refine[THEN nres_relD] also note cr_ample_reduction finally show ?thesis by simp qed end text \The efficient implementation of @{const visible_prog.ga_gample} uses the @{const collecti_index} function\ lemma collecti_index_cong: assumes C: "\i. i f i (l!i) = f' i (l'!i)" assumes [simp]: "l=l'" shows "collecti_index f l = collecti_index f' l'" proof - { fix i0 s0 assume "\iix. x\set l \ f x = f' x" assumes "l=l'" shows "find_min_idx_f f l = find_min_idx_f f' l'" unfolding assms(2)[symmetric] using assms(1) apply (induction l) apply (auto split: option.split) done lemma find_min_idx_impl_collecti: "( case find_min_idx_f f l of Some (i,r) \ {i} \ s1 i r | None \ {(i,r) | i r. i r\s2 i (l!i)} ) = ( collecti_index (\i x. case f x of Some r \ (True, s1 i r) | None \ (False, s2 i x) ) l )" (is "?l=?r" is "_ = collecti_index ?fc l") proof (cases "\i None") define i where "i \ LEAST i. i < length l \ f (l ! i) \ None" have C_ALT: "(\i None) \ (\i. i < length l \ fst (?fc i (l ! i)))" by (auto split: option.split) have i_alt: "i = (LEAST i. i < length l \ fst (?fc i (l ! i)))" unfolding i_def apply (fo_rule arg_cong) by (auto split: option.split) { case True from LeastI_ex[OF True] obtain r where [simp]: "i s1 i r" by simp from collecti_index_collect[of ?fc l, folded i_alt] True C_ALT have REQ: "?r = {i} \ snd (?fc i (l!i))" by simp show "?l=?r" unfolding LEQ unfolding REQ by auto } { case False from False find_min_idx_f_LEAST_eq[of f l] have "find_min_idx_f f l = None" by simp hence LEQ: "?l = {(i,r) | i r. i r\s2 i (l!i)}" by simp from False C_ALT collecti_index_collect[of ?fc l] have REQ: "?r = {(i, x) |i x. i < length l \ x \ snd (?fc i (l ! i))}" by auto show "?l=?r" unfolding LEQ unfolding REQ using False by auto } qed lemma find_min_idx_impl_collecti_scheme: assumes "\x. s11 x = (case x of (i,r) \ {i} \ iss i r)" assumes "s12 = {(i,r) | i r. i r\isn i (l!i)}" assumes "\i. i f2 i (l!i) = (case f (l!i) of Some r \ (True, iss i r) | None \ (False, isn i (l!i)))" shows "(case find_min_idx_f f l of Some x \ s11 x | None \ s12) = (collecti_index f2 l)" apply (simp only: assms(1,2,3) cong: collecti_index_cong) apply (subst find_min_idx_impl_collecti) apply simp done schematic_goal stutter_extend_en_list_aux: "(?c, stutter_extend_en) \ (R \ \S\list_set_rel) \ R \ \\S\option_rel\list_set_rel" unfolding stutter_extend_en_def[abs_def] apply (autoref) done concrete_definition stutter_extend_en_list uses stutter_extend_en_list_aux lemma succ_of_enex_impl_aux: "{s'. \a\en s. s'=ex a s} = (\a. ex a s)`en s" by auto schematic_goal succ_of_enex_list_aux: "(?c, succ_of_enex) \ (Id \ \Id\list_set_rel) \\<^sub>r (Id \ Id \ Id) \ Id \ \Id\list_set_rel" unfolding succ_of_enex_def[abs_def] unfolding succ_of_enex_impl_aux by (autoref (trace, keep_goal)) concrete_definition succ_of_enex_list uses succ_of_enex_list_aux schematic_goal pred_of_enex_list_aux: "(?c, pred_of_enex) \ (Id \ \Id\list_set_rel) \\<^sub>r (Id \ Id \ Id) \ Id \ Id \ bool_rel" unfolding pred_of_enex_def[abs_def] by (autoref (trace, keep_goal)) concrete_definition pred_of_enex_list uses pred_of_enex_list_aux (* TODO: can this be defined using autoref? *) definition "rel_of_enex_list enex \ rel_of_pred (pred_of_enex_list enex)" lemma rel_of_enex_list_refine: "(rel_of_enex_list, rel_of_enex) \ (Id \ \Id\list_set_rel) \\<^sub>r (Id \ Id \ Id) \ \Id \\<^sub>r Id\ set_rel" proof - have 1: "rel_of_enex_list = rel_of_pred \ pred_of_enex_list" using rel_of_enex_list_def by force have 2: "rel_of_enex = rel_of_pred \ pred_of_enex" by auto have 3: "(rel_of_pred, rel_of_pred) \ (Id \ Id \ bool_rel) \ \Id \\<^sub>r Id\set_rel" by auto show ?thesis unfolding 1 2 by (parametricity add: 3 pred_of_enex_list.refine) qed context visible_prog begin definition "ga_gample_mi2 mem sticky_E (gc::pid_global_config) = ( let lcs = pid_global_config.processes gc; gs = pid_global_config.state gc in collecti_index (\_ lc. let c = local_config.command lc; ls = local_config.state lc; as = set (comp.succ_impl c); s={(a,c')\as. la_en' (ls,gs) a}; ample = s\{} \ (\(a,_)\as. read_globals a = {} \ write_globals a = {}) \ (\(a,c')\s. \ mem (c,c') sticky_E) in (ample,{c}\s) ) lcs )" lemma ga_gample_mi2_impl: "ga_gample_m mem sticky_E gc = ga_gample_mi2 mem sticky_E gc" unfolding ga_gample_m_def ga_gample_mi2_def apply simp apply (rule find_min_idx_impl_collecti_scheme[where iss="\pid (c,S). (\(a,c'). (c,a,c'))`S" and isn="\pid lc. {(c,a,c'). cfgc c a c' \ c = cmd_of_pid gc pid \ la_en' (local_config.state lc, pid_global_config.state gc) a}"] ) apply auto [] unfolding ga_gen_def apply auto [] apply (auto simp: comp.astep_impl_def) done definition "cr_ample_impl2 \ let sticky = find_fas_init_code (=) bounded_hashcode_nat (def_hashmap_size TYPE(nat)) cfgc_G_impl (\x. x \ vis_edges) in stutter_extend_en (ga_gample_mi2 (\x f. f x) sticky)" definition "ample_step_impl2 \ rel_of_enex (cr_ample_impl2,ga_ex)" lemma cr_ample_impl2_reduction: "hl_reduction prog is_vis_var ample_step_impl2" proof - from cr_ample_impl1_reduction interpret hl_ample_reduction prog is_vis_var cr_ample_impl1 . have 1: "ample_step_impl2 = rel_of_enex (cr_ample_impl1, ga_ex)" unfolding ample_step_impl2_def cr_ample_impl1_def cr_ample_impl2_def unfolding ga_gample_mi2_impl[abs_def] by simp show ?thesis apply unfold_locales unfolding 1 apply simp apply simp apply (blast intro: ample_accept_subset ample_accept_stuttering)+ done qed end text \Implementing the State\ type_synonym valuation_impl = "(ident,uint32) mapping" definition "vi_\ \ map_option Rep_uint32 \\ Mapping.rep" abbreviation "vi_rel \ br vi_\ (\_. True)" abbreviation "val_rel \ br Rep_uint32 (\_. True)" lemma vi_rel_rew: "(x = vi_\ y) \ (y,x)\vi_rel" "(vi_\ y = x) \ (y,x)\vi_rel" by (auto simp: br_def) export_code Mapping.lookup Mapping.update Mapping.empty checking SML lift_definition vi_empty :: "valuation_impl" is "Map.empty::valuation" . lemma [code]: "vi_empty = Mapping.empty" by transfer simp lemma vi_empty_correct: "vi_\ vi_empty = Map.empty" unfolding vi_\_def by transfer auto lift_definition vi_lookup :: "valuation_impl \ ident \ uint32" is "\vi::valuation. \x. vi x" . lemma [code]: "vi_lookup = Mapping.lookup" apply (intro ext) unfolding vi_lookup_def Mapping.lookup_def map_fun_def[abs_def] o_def id_apply option.map_comp Rep_uint32_inverse option.map_ident by simp lemma vi_lookup_correct: "vi_lookup v x = map_option Abs_uint32 (vi_\ v x)" unfolding vi_\_def apply transfer apply (simp add: option.map_comp o_def option.map_ident) done lift_definition vi_update :: "ident \ uint32 \ valuation_impl \ valuation_impl" is "\x v. \vi::valuation. vi(x\v)" . lemma [code]: "vi_update = Mapping.update" by transfer simp -lemma vi_update_correct: "vi_\ (vi_update k v m) = vi_\ m (k\Rep_uint32 v)" +lemma vi_update_correct: "vi_\ (vi_update k v m) = (vi_\ m) (k\Rep_uint32 v)" unfolding vi_\_def by transfer simp export_code vi_lookup vi_update vi_empty checking SML record local_state_impl = variables :: valuation_impl record global_state_impl = variables :: valuation_impl type_synonym focused_state_impl = "local_state_impl \ global_state_impl" type_synonym local_config_impl = "(cmdc,local_state_impl) Gen_Scheduler.local_config" type_synonym pid_global_config_impl = "(cmdc,local_state_impl,global_state_impl) Pid_Scheduler.pid_global_config" definition "local_state_rel \ { (\ local_state_impl.variables = vi \, \ local_state.variables = v \) | vi v. v = vi_\ vi }" definition "global_state_rel \ { (\ global_state_impl.variables = vi \, \ global_state.variables = v \) | vi v. v = vi_\ vi }" abbreviation "focused_state_rel \ local_state_rel \\<^sub>r global_state_rel" definition "local_config_rel \ { (\ local_config.command = ci, local_config.state=lsi \, \ local_config.command = c, local_config.state=ls \) | ci c lsi ls. (ci::cmdc,c)\Id \ (lsi,ls)\local_state_rel }" definition "global_config_rel \ { (\ pid_global_config.processes = psi, pid_global_config.state = gsi \, \ pid_global_config.processes = ps, pid_global_config.state = gs \) | psi ps gsi gs. (psi,ps)\\local_config_rel\list_rel \ (gsi,gs)\global_state_rel }" definition "init_valuation_impl vd \ fold (\x v. vi_update x 0 v) vd (vi_empty)" lemma init_valuation_impl: "(init_valuation_impl, init_valuation) \ \Id\list_rel \ vi_rel" proof - { fix vd m have "m ++ init_valuation vd = fold (\x v. v(x\0)) vd m" apply (rule sym) apply (induction vd arbitrary: m) unfolding init_valuation_def[abs_def] apply (auto intro!: ext simp: Map.map_add_def) done } note aux=this have E1: "\vd. init_valuation vd = fold (\x v. v(x\0)) vd Map.empty" by (subst aux[symmetric]) simp have E2: "\vdi vd. (vdi,vd)\\Id\list_rel \ (init_valuation_impl vdi, fold (\x v. v(x\0)) vd Map.empty) \ vi_rel" unfolding init_valuation_impl_def apply parametricity apply (auto simp: vi_update_correct vi_empty_correct br_def zero_uint32.rep_eq) done show ?thesis apply (intro fun_relI) apply (subst E1) apply (rule E2) by auto qed context cprog begin definition init_pc_impl :: "proc_decl \ local_config_impl" where "init_pc_impl pd \ \ local_config.command = comp.\ (proc_decl.body pd), local_config.state = \ local_state_impl.variables = init_valuation_impl (proc_decl.local_vars pd) \ \" lemma init_pc_impl: "(init_pc_impl, init_pc) \ Id \ local_config_rel" apply (rule fun_relI) unfolding init_pc_impl_def init_pc_def local_config_rel_def local_state_rel_def apply (simp del: pair_in_Id_conv, intro conjI) apply simp apply (simp only: vi_rel_rew) apply (parametricity add: init_valuation_impl) apply simp done definition pid_init_gc_impl :: "pid_global_config_impl" where "pid_init_gc_impl \ \ pid_global_config.processes = (map init_pc_impl (program.processes prog)), pid_global_config.state = \ global_state_impl.variables = init_valuation_impl (program.global_vars prog) \ \" lemma pid_init_gc_impl: "(pid_init_gc_impl, pid_init_gc) \ global_config_rel" unfolding pid_init_gc_impl_def pid_init_gc_def global_config_rel_def global_state_rel_def apply (simp del: pair_in_Id_conv, intro conjI) apply (parametricity add: init_pc_impl) apply simp apply (simp only: vi_rel_rew) apply (parametricity add: init_valuation_impl) apply simp done end lift_definition val_of_bool_impl :: "bool \ uint32" is val_of_bool . lemma [code]: "val_of_bool_impl b = (if b then 1 else 0)" by transfer auto lift_definition bool_of_val_impl :: "uint32 \ bool" is bool_of_val . lemma [code]: "bool_of_val_impl v = (v\0)" by transfer (auto simp: bool_of_val_def) definition "set_of_rel R \ {(a,b) . R b a}" definition "rel_of_set S \ \b a. (a,b)\S" lemma [simp]: "(a,b)\set_of_rel R \ R b a" "rel_of_set S b a \ (a,b)\S" unfolding rel_of_set_def set_of_rel_def by auto lemma [simp]: "set_of_rel (rel_of_set S) = S" "rel_of_set (set_of_rel R) = R" unfolding rel_of_set_def set_of_rel_def by auto lemma rel_fun_2set: "rel_fun A B = rel_of_set (set_of_rel A \ set_of_rel B)" unfolding rel_fun_def fun_rel_def unfolding rel_of_set_def set_of_rel_def by (auto) lemma [simp]: "set_of_rel cr_uint32 = val_rel" "set_of_rel (=) = Id" by (auto simp: br_def cr_uint32_def) lemma bool_of_val_impl: "(bool_of_val_impl, bool_of_val) \ val_rel \ bool_rel" using bool_of_val_impl.transfer by (simp add: rel_fun_2set) lemma smod_by_div_abs: "a smod b = a - a sdiv b * b" for a b :: \'a::len word\ by (subst (2) sdiv_smod_id [of a b, symmetric]) simp lift_definition sdiv_impl :: "uint32 \ uint32 \ uint32" is "(sdiv)" . lift_definition smod_impl :: "uint32 \ uint32 \ uint32" is "(smod)" . (* TODO: Is there a more efficient way? *) lemma [code]: "sdiv_impl x y = Abs_uint32' (Rep_uint32' x sdiv Rep_uint32' y)" apply transfer by simp lemma [code]: "smod_impl a b = a - sdiv_impl a b * b" by transfer (simp add: smod_by_div_abs) context includes bit_operations_syntax begin primrec eval_bin_op_impl_aux :: "bin_op \ uint32 \ uint32 \ uint32" where "eval_bin_op_impl_aux bo_plus v1 v2 = v1+v2" | "eval_bin_op_impl_aux bo_minus v1 v2 = v1-v2" | "eval_bin_op_impl_aux bo_mul v1 v2 = v1*v2" | "eval_bin_op_impl_aux bo_div v1 v2 = sdiv_impl v1 v2" | "eval_bin_op_impl_aux bo_mod v1 v2 = smod_impl v1 v2" | "eval_bin_op_impl_aux bo_less v1 v2 = val_of_bool_impl (v1 < v2)" | "eval_bin_op_impl_aux bo_less_eq v1 v2 = val_of_bool_impl (v1 \ v2)" | "eval_bin_op_impl_aux bo_eq v1 v2 = val_of_bool_impl (v1 = v2)" | "eval_bin_op_impl_aux bo_and v1 v2 = v1 AND v2" | "eval_bin_op_impl_aux bo_or v1 v2 = v1 OR v2" | "eval_bin_op_impl_aux bo_xor v1 v2 = v1 XOR v2" end lift_definition eval_bin_op_impl :: "bin_op \ uint32 \ uint32 \ uint32" is eval_bin_op . lemma [code]: "eval_bin_op_impl bop v1 v2 = eval_bin_op_impl_aux bop v1 v2" apply (cases bop) apply (transfer; simp)+ done primrec eval_un_op_impl_aux :: "un_op \ uint32 \ uint32" where "eval_un_op_impl_aux uo_minus v = -v" | "eval_un_op_impl_aux uo_not v = Bit_Operations.not v" lift_definition eval_un_op_impl :: "un_op \ uint32 \ uint32" is eval_un_op . lemma [code]: "eval_un_op_impl uop v = eval_un_op_impl_aux uop v" apply (cases uop) apply simp_all apply (transfer, simp)+ done fun eval_exp_impl :: "exp \ focused_state_impl \ uint32" where "eval_exp_impl (e_var x) (ls,gs) = do { let lv = local_state_impl.variables ls; let gv = global_state_impl.variables gs; case vi_lookup lv x of Some v \ Some v | None \ (case vi_lookup gv x of Some v \ Some v | None \ None) }" | "eval_exp_impl (e_localvar x) (ls,gs) = vi_lookup (local_state_impl.variables ls) x" | "eval_exp_impl (e_globalvar x) (ls,gs) = vi_lookup (global_state_impl.variables gs) x" | "eval_exp_impl (e_const n) fs = do { assert_option (n\min_signed \ n\max_signed); Some (uint32_of_int n) }" | "eval_exp_impl (e_bin bop e1 e2) fs = do { v1\eval_exp_impl e1 fs; v2\eval_exp_impl e2 fs; Some (eval_bin_op_impl bop v1 v2) }" | "eval_exp_impl (e_un uop e) fs = do { v\eval_exp_impl e fs; Some (eval_un_op_impl uop v) }" (* TODO: Move *) lemma map_option_bind: "map_option f (m\g) = m \ (map_option f o g)" by (auto split: Option.bind_split) lemma val_option_rel_as_eq: "(a,b)\\val_rel\option_rel \ map_option Rep_uint32 a = b" unfolding br_def apply (cases a) apply (cases b, simp_all) apply (cases b, auto) done lemma eval_exp_impl: "(eval_exp_impl, eval_exp) \ Id \ focused_state_rel \ \val_rel\option_rel" proof - { fix e ls gs lsi gsi assume "(lsi,ls)\local_state_rel" "(gsi,gs)\global_state_rel" hence "map_option Rep_uint32 (eval_exp_impl e (lsi,gsi)) = eval_exp e (ls,gs)" apply (induction e) apply (auto simp: local_state_rel_def global_state_rel_def br_def simp: vi_lookup_correct vi_update_correct simp: Abs_uint32_inverse[simplified] uint32_of_int.rep_eq signed_of_int_def simp: option.map_comp option.map_ident o_def map_option_bind simp: eval_bin_op_impl.rep_eq eval_un_op_impl.rep_eq split: option.splits) apply (drule sym) apply (drule sym) apply (simp split: Option.bind_split) apply (drule sym) apply (simp split: Option.bind_split) done } thus ?thesis by (auto simp: val_option_rel_as_eq) qed text \Enabledness and effects of actions\ primrec la_en_impl :: "focused_state_impl \ action \ bool" where "la_en_impl fs (AAssign e _ _) = do { v \ eval_exp_impl e fs; Some (bool_of_val_impl v)}" | "la_en_impl fs (AAssign_local e _ _) = do { v \ eval_exp_impl e fs; Some (bool_of_val_impl v)}" | "la_en_impl fs (AAssign_global e _ _) = do { v \ eval_exp_impl e fs; Some (bool_of_val_impl v)}" | "la_en_impl fs (ATest e) = do { v \ eval_exp_impl e fs; Some (bool_of_val_impl v)}" | "la_en_impl _ (ASkip) = Some True" lemma param_rec_action[param]: "(rec_action, rec_action) \ (Id \ Id \ Id \ A) \ (Id \ Id \ Id \ A) \ (Id \ Id \ Id \ A) \ (Id \ A) \ A \ Id \ A" apply (intro fun_relI) subgoal for fi1 f1 fi2 f2 fi3 f3 fi4 f4 fi5 f5 ai a apply simp apply (induction a) apply simp_all apply (parametricity, simp_all?)+ done done (* TODO: Move *) lemma param_option_bind[param]: "(Option.bind, Option.bind) \ \A\option_rel \ (A \ \B\option_rel) \ \B\option_rel" unfolding Option.bind_def by parametricity lemma la_en_impl: "(la_en_impl,la_en) \ focused_state_rel \ Id \ \Id\option_rel" unfolding la_en_impl_def la_en_def by (parametricity add: eval_exp_impl bool_of_val_impl) definition "la_en'_impl fs a \ case la_en_impl fs a of Some b \ b | None \ False" lemma la_en'_impl: "(la_en'_impl,la_en') \ focused_state_rel \ Id \ bool_rel" unfolding la_en'_impl_def[abs_def] la_en'_def[abs_def] by (parametricity add: la_en_impl) abbreviation "exists_var_impl v x \ vi_lookup v x \ None" fun la_ex_impl :: "focused_state_impl \ action \ focused_state_impl" where "la_ex_impl (ls,gs) (AAssign ge x e) = do { v \ eval_exp_impl ge (ls,gs); assert_option (bool_of_val_impl v); v \ eval_exp_impl e (ls,gs); if exists_var_impl (local_state_impl.variables ls) x then do { let ls = local_state_impl.variables_update (vi_update x v) ls; Some (ls,gs) } else do { assert_option (exists_var_impl (global_state_impl.variables gs) x); let gs = global_state_impl.variables_update (vi_update x v) gs; Some (ls,gs) } }" | "la_ex_impl (ls,gs) (AAssign_local ge x e) = do { v \ eval_exp_impl ge (ls,gs); assert_option (bool_of_val_impl v); v \ eval_exp_impl e (ls,gs); assert_option (exists_var_impl (local_state_impl.variables ls) x); let ls = local_state_impl.variables_update (vi_update x v) ls; Some (ls,gs) }" | "la_ex_impl (ls,gs) (AAssign_global ge x e) = do { v \ eval_exp_impl ge (ls,gs); assert_option (bool_of_val_impl v); v \ eval_exp_impl e (ls,gs); assert_option (exists_var_impl (global_state_impl.variables gs) x); let gs = global_state_impl.variables_update (vi_update x v) gs; Some (ls,gs) }" | "la_ex_impl fs (ATest e) = do { v \ eval_exp_impl e fs; assert_option (bool_of_val_impl v); Some fs }" | "la_ex_impl fs ASkip = Some fs" (* TODO: Move *) lemma param_assert_option[param]: "(assert_option, assert_option) \ bool_rel \ \unit_rel\option_rel" unfolding assert_option_def by parametricity lemma la_ex_impl: "(la_ex_impl, la_ex) \ focused_state_rel \ Id \ \focused_state_rel\option_rel" proof (intro fun_relI) fix fsi fs and ai a :: action assume "(fsi,fs)\focused_state_rel" "(ai,a)\Id" thus "(la_ex_impl fsi ai, la_ex fs a)\\focused_state_rel\option_rel" apply (cases fs, cases fsi, clarsimp) apply (cases a) apply simp_all apply (parametricity add: eval_exp_impl bool_of_val_impl) apply (auto simp: local_state_rel_def global_state_rel_def simp: vi_lookup_correct vi_update_correct simp: br_def intro: domI) [7] apply (parametricity add: eval_exp_impl bool_of_val_impl) apply (auto simp: local_state_rel_def global_state_rel_def simp: vi_lookup_correct vi_update_correct br_def intro: domI) [9] apply (parametricity add: eval_exp_impl bool_of_val_impl) apply (auto simp: local_state_rel_def global_state_rel_def simp: vi_lookup_correct vi_update_correct br_def intro: domI) [4] apply (parametricity add: eval_exp_impl bool_of_val_impl) apply simp done qed definition "la_ex'_impl fs a \ case la_ex_impl fs a of Some fs' \ fs' | None \ fs" lemma la_ex'_impl: "(la_ex'_impl,la_ex') \ focused_state_rel \ Id \ focused_state_rel" unfolding la_ex'_impl_def[abs_def] la_ex'_def[abs_def] by (parametricity add: la_ex_impl) lemma param_collecti_index: "(collecti_index, collecti_index) \ (nat_rel \ B \ bool_rel \\<^sub>r \Id\set_rel) \ \B\list_rel \ \nat_rel \\<^sub>r Id\set_rel" unfolding collecti_index'_def apply parametricity apply auto done export_code la_ex_impl la_en_impl checking SML context visible_prog begin definition "ga_gample_mi3 mem sticky_E (gc::pid_global_config_impl) = ( let lcs = pid_global_config.processes gc; gs = pid_global_config.state gc in collecti_index (\_ lc. let c = local_config.command lc; ls = local_config.state lc; as = set (comp.succ_impl c); s={(a,c')\as. la_en'_impl (ls,gs) a}; ample = s\{} \ (\(a,_)\as. read_globals a = {} \ write_globals a = {}) \ (\(a,c')\s. \ mem (c,c') sticky_E) in (ample,{c}\s) ) lcs )" lemma ga_gample_mi3_refine: "(ga_gample_mi3, ga_gample_mi2)\ Id \ Id \ global_config_rel \ \nat_rel \\<^sub>r Id\set_rel" proof - have [param]: "(comp.succ_impl,comp.succ_impl)\Id\\Id\\<^sub>rId\list_rel" by simp from la_en'_impl have aux1: "\fsi fs a. (fsi,fs)\focused_state_rel \ la_en'_impl fsi a = la_en' fs a" apply (rule_tac IdD) apply parametricity by auto show ?thesis using [[goals_limit=1]] apply (intro fun_relI) unfolding ga_gample_mi3_def ga_gample_mi2_def apply (parametricity add: param_collecti_index la_en'_impl) apply (auto simp add: global_config_rel_def local_config_rel_def) [] apply (auto simp add: global_config_rel_def local_config_rel_def) [] apply (auto simp add: global_config_rel_def local_config_rel_def) [] apply (auto simp add: global_config_rel_def local_config_rel_def) [] apply simp using aux1 apply (auto simp add: global_config_rel_def local_config_rel_def) [] apply simp_all done qed definition (in cprog) "is_vis_edge is_vis_var \ \(c,c'). \(a,cc')\set (comp.succ_impl c). c'=cc' \ (\v\write_globals a. is_vis_var v)" lemma is_vis_edge_correct: "is_vis_edge is_vis_var = (\x. x\vis_edges)" unfolding is_vis_edge_def vis_edges_def by (auto intro!: ext simp: comp.astep_impl_def) definition "cr_ample_impl3 \ let sticky = find_fas_init_code (=) bounded_hashcode_nat (def_hashmap_size TYPE(nat)) cfgc_G_impl (\x. x \ vis_edges) in stutter_extend_en (ga_gample_mi3 (\x f. f x) sticky)" lemma cr_ample_impl3_refine: "(cr_ample_impl3, cr_ample_impl2) \ global_config_rel \ \\nat_rel\\<^sub>rId\option_rel\set_rel" proof - have [param]: "\R S. S=Id \ (stutter_extend_en, stutter_extend_en) \ (R \ \S\set_rel) \ R \ \\S\option_rel\set_rel" apply (simp only: ) by (rule stutter_extend_en_refine) show ?thesis unfolding cr_ample_impl3_def cr_ample_impl2_def is_vis_edge_correct apply (parametricity add: stutter_extend_en_refine ga_gample_mi3_refine) apply (rule IdI) by auto qed (* TODO: The current Pid_Scheduler-locale is focused on refinement. We cannot use it to get an instantiation of ga_gex_impl.*) definition (in -) ga_gex_impl :: "global_action \ pid_global_config_impl \ pid_global_config_impl" where "ga_gex_impl ga gc \ let (pid,c,a,c') = ga; fs = fs_of_pid gc pid; (ls',gs') = la_ex'_impl fs a; gc' = \ pid_global_config.processes = (pid_global_config.processes gc) [pid := \ local_config.command = c', local_config.state = ls'\ ], pid_global_config.state = gs'\ in gc'" lemma lc_of_pid_refine: assumes V: "pid_valid gc pid" assumes [param]: "(pidi,pid)\nat_rel" assumes GCI: "(gci,gc)\global_config_rel" shows "(lc_of_pid gci pidi, lc_of_pid gc pid)\local_config_rel" apply parametricity apply fact using GCI by (auto simp: global_config_rel_def) lemma ls_of_pid_refine: assumes "pid_valid gc pid" assumes "(pidi,pid)\nat_rel" assumes "(gci,gc)\global_config_rel" shows "(ls_of_pid gci pidi, ls_of_pid gc pid)\local_state_rel" using lc_of_pid_refine[OF assms] unfolding local_config_rel_def by auto lemma ga_gex_impl: assumes V: "pid_valid gc pid" assumes [param]: "(pidi,pid)\Id" "(ci,c)\Id" "(ai,a)\Id" and CI'I: "(ci',c')\Id" and GCI[param]: "(gci,gc)\global_config_rel" shows "(ga_gex_impl (pidi,ci,ai,ci') gci, ga_gex (pid,c,a,c') gc) \ global_config_rel" unfolding ga_gex_impl_def[abs_def] ga_gex_def[abs_def] apply simp apply (parametricity add: la_ex'_impl ls_of_pid_refine[OF V]) using GCI CI'I apply (auto simp: global_config_rel_def local_config_rel_def) [] apply parametricity apply auto [] using GCI apply (auto simp: global_config_rel_def) [] done definition (in -) "ga_ex_impl \ stutter_extend_ex ga_gex_impl" fun ga_valid where "ga_valid gc None = True" | "ga_valid gc (Some (pid,c,a,c')) = pid_valid gc pid" lemma ga_ex_impl: assumes V: "ga_valid gc ga" assumes P: "(gai,ga)\Id" "(gci,gc)\global_config_rel" shows "(ga_ex_impl gai gci, ga_ex ga gc) \ global_config_rel" using assms apply (cases "(gc,ga)" rule: ga_valid.cases) using P apply (simp add: ga_ex_impl_def) apply (simp add: ga_ex_impl_def ga_ex_def) apply (parametricity add: ga_gex_impl) by auto definition "ample_step_impl3 \ rel_of_enex (cr_ample_impl3,ga_ex_impl)" definition (in -) pid_interp_gc_impl :: "(ident \ bool) \ pid_global_config_impl \ exp set" where "pid_interp_gc_impl is_vis_var gci \ sm_props ( vi_\ (global_state_impl.variables (pid_global_config.state gci)) |` Collect is_vis_var )" lemma pid_interp_gc_impl: "(pid_interp_gc_impl is_vis_var, pid_interp_gc) \ global_config_rel \ Id" apply rule unfolding pid_interp_gc_impl_def pid_interp_gc_def apply (auto simp: global_config_rel_def global_state_rel_def interp_gs_def sm_props_def br_def) done lemma cr_ample_impl3_sim_aux: assumes GCI: "(gci, gc) \ global_config_rel" assumes GA: "ga \ cr_ample_impl3 gci" shows "\gc'. ga\cr_ample_impl2 gc \ gc'=ga_ex ga gc \ (ga_ex_impl ga gci, gc')\global_config_rel" proof (clarsimp, safe) show G1: "ga\cr_ample_impl2 gc" using cr_ample_impl3_refine[param_fo, OF GCI] GA by auto { (* TODO: We are re-proving the validity of the action, due to the lack of a clean refinement argumentation *) fix mem sticky_E gc pid c assume "(pid,c)\ga_gample_m mem sticky_E gc" hence "pid_valid gc pid" unfolding ga_gample_m_def apply (auto split: option.splits) apply (auto simp: ga_gen_def) [] apply (auto dest: find_min_idx_f_SomeD) done } note aux=this from G1 have "ga_valid gc ga" unfolding cr_ample_impl2_def stutter_extend_en_def apply (simp split: if_split_asm) unfolding ga_gample_mi2_impl[symmetric] apply (auto dest: aux) done thus "(ga_ex_impl ga gci, ga_ex ga gc)\global_config_rel" by (parametricity add: ga_ex_impl GCI IdI[of ga]) qed lemma cr_ample_impl3_sim_aux2: assumes GCI: "(gci, gc) \ global_config_rel" assumes GA: "ga \ cr_ample_impl2 gc" shows "ga\cr_ample_impl3 gci \ (ga_ex_impl ga gci, ga_ex ga gc) \ global_config_rel" proof show G1: "ga\cr_ample_impl3 gci" using cr_ample_impl3_refine[param_fo, OF GCI] GA by auto from cr_ample_impl3_sim_aux[OF GCI G1] show "(ga_ex_impl ga gci, ga_ex ga gc) \ global_config_rel" by auto qed sublocale impl3: sa "\ g_V = UNIV, g_E = ample_step_impl3, g_V0 = {pid_init_gc_impl}, sa_L = pid_interp_gc_impl is_vis_var \" by unfold_locales auto sublocale impl3_sim: lbisimulation global_config_rel "\ g_V = UNIV, g_E = ample_step_impl3, g_V0 = {pid_init_gc_impl}, sa_L = pid_interp_gc_impl is_vis_var \" "\ g_V = UNIV, g_E = ample_step_impl2, g_V0 = {pid_init_gc}, sa_L = pid_interp_gc \" apply unfold_locales apply simp apply simp apply simp using pid_init_gc_impl apply (auto dest: fun_relD) [] unfolding ample_step_impl2_def ample_step_impl3_def apply (auto dest: cr_ample_impl3_sim_aux) [] using pid_interp_gc_impl apply (auto dest: fun_relD) [] apply simp using pid_init_gc_impl apply (auto dest: fun_relD) [] defer using pid_interp_gc_impl apply (auto dest: fun_relD) [] apply auto apply (auto dest: cr_ample_impl3_sim_aux2) [] done text \Correctness of impl3\ (* TODO: Locale hl_reduction seems very special, and we cannot use it here*) lemma impl3_accept_subset: "impl3.accept w \ lv.sa.accept w" using impl3_sim.accept_bisim cr_ample_impl2_reduction hl_reduction.ample_accept_subset by simp lemma impl3_accept_stuttering: "lv.sa.accept w \ \w'. w\w' \ impl3.accept w'" using impl3_sim.accept_bisim cr_ample_impl2_reduction hl_reduction.ample_accept_stuttering by simp text \Next, we go from sets of actions to lists of actions\ definition (in cprog) "ga_gample_mil4 mem sticky_Ei (gc::pid_global_config_impl) = ( let lcs = pid_global_config.processes gc; gs = pid_global_config.state gc in collecti_index_list (\_ lc. let c = local_config.command lc; ls = local_config.state lc; as = comp.succ_impl c; s = filter (la_en'_impl (ls,gs) o fst) as; ample = s\[] \ (\(a,_)\set as. read_globals a = {} \ write_globals a = {}) \ (\(a,c')\set s. \mem (c,c') sticky_Ei) in (ample,map (Pair c) s) ) lcs )" lemma ga_gample_mil4_refine: "(ga_gample_mil4, ga_gample_mi3) \ (Id\\<^sub>rId \ Id \ bool_rel) \ Id \ Id \ \nat_rel \\<^sub>r Id\list_set_rel" apply (intro fun_relI) unfolding ga_gample_mil4_def ga_gample_mi3_def apply (parametricity) apply simp apply (rule IdI) apply simp apply (rule IdI) apply (intro fun_relI) apply (rule collecti_index_list_refine[param_fo]) apply (intro fun_relI) apply (simp add: list_set_rel_def br_def distinct_map comp.succ_impl_invar) apply (auto simp: filter_empty_conv) apply force done definition (in cprog) "cr_ample_impl4 is_vis_var \ let sticky = find_fas_init_code (=) bounded_hashcode_nat (def_hashmap_size TYPE(nat)) cfgc_G_impl (is_vis_edge is_vis_var) in stutter_extend_en_list (ga_gample_mil4 (\x f. f x) sticky)" lemma cr_ample_impl4_refine: "(cr_ample_impl4 is_vis_var, cr_ample_impl3)\Id \ \\nat_rel \\<^sub>r Id\option_rel\list_set_rel" unfolding cr_ample_impl4_def cr_ample_impl3_def apply (parametricity add: stutter_extend_en_list.refine ga_gample_mil4_refine) apply (simp_all add: is_vis_edge_correct) done text \Finally, we combine the ample-implementation and the executable implementation to get a step function\ definition (in cprog) "ample_step_impl4 is_vis_var \ rel_of_enex_list (cr_ample_impl4 is_vis_var,ga_ex_impl)" lemma ample_step_impl4: "(ample_step_impl4 is_vis_var, ample_step_impl3) \ \Id \\<^sub>r Id\ set_rel" unfolding ample_step_impl4_def ample_step_impl3_def rel_of_pred_def apply (parametricity add: rel_of_enex_list_refine cr_ample_impl4_refine[simplified]) by auto sublocale impl4: sa "\ g_V = UNIV, g_E = ample_step_impl4 is_vis_var, g_V0 = {pid_init_gc_impl}, sa_L = pid_interp_gc_impl is_vis_var \" by unfold_locales auto lemma impl4_accept_subset: "impl4.accept w \ lv.sa.accept w" using impl3_accept_subset ample_step_impl4 by simp lemma impl4_accept_stuttering: "lv.sa.accept w \ \w'. w\w' \ impl4.accept w'" using impl3_accept_stuttering ample_step_impl4 by simp lemma (in -) pred_of_succ_of_enex_list: fixes enex :: "('c \ 'a list) \ ('a \ 'c \ 'c)" shows "pred_of_succ (set o succ_of_enex_list enex) = pred_of_enex_list enex" proof - { fix x and l have "glist_member (=) x l \ x\set l" by (induction l) auto } note [simp] = this { fix l :: "'a list" and g and l0 have "set (foldli l (\_. True) (\x. glist_insert (=) (g x)) l0) = set l0 \ g`set l" by (induction l arbitrary: l0) (auto simp: glist_insert_def) } note [simp] = this { fix l::"'a list" and P i have "foldli l Not (\x _. P x) i \ i \ (\x\set l. P x)" by (induction l arbitrary: i) auto } note [simp] = this show ?thesis unfolding succ_of_enex_list_def[abs_def] pred_of_enex_list_def[abs_def] by (auto simp: pred_of_succ_def gen_image_def gen_bex_def intro!: ext) qed lemma (in -) rel_of_succ_of_enex_list: fixes enex :: "('c \ 'a list) \ ('a \ 'c \ 'c)" shows "rel_of_succ (set o succ_of_enex_list enex) = rel_of_enex_list enex" unfolding rel_of_enex_list_def unfolding pred_of_succ_of_enex_list[symmetric] by simp definition (in cprog) "impl4_succ is_vis_var \ succ_of_enex_list (cr_ample_impl4 is_vis_var,ga_ex_impl)" lemma impl4_succ_pred: "rel_of_succ (set o (impl4_succ is_vis_var)) = ample_step_impl4 is_vis_var" unfolding impl4_succ_def ample_step_impl4_def by (simp add: rel_of_succ_of_enex_list) end export_code ga_ex_impl checking SML definition "ccfg_E_succ_impl mst \ remdups o map snd o ccfg_succ_impl mst" lemma (in cprog) ccfg_E_succ_impl: "ccfg_E_succ_impl (comp_info_of prog) = cfgc_E_succ" apply (intro ext) unfolding ccfg_E_succ_impl_def cfgc_E_succ_def by (simp add: ccfg_succ_impl) definition init_pc_impl_impl :: "comp_info \ proc_decl \ local_config_impl" where "init_pc_impl_impl mst pd \ \ local_config.command = comp_\_impl mst (proc_decl.body pd), local_config.state = \ local_state_impl.variables = init_valuation_impl (proc_decl.local_vars pd) \ \" lemma (in cprog) init_pc_impl_impl: "init_pc_impl_impl (comp_info_of prog) = init_pc_impl" apply (intro ext) unfolding init_pc_impl_impl_def init_pc_impl_def by (simp add: comp_\_impl) definition pid_init_gc_impl_impl :: "program \ comp_info \ pid_global_config_impl" where "pid_init_gc_impl_impl prog mst \ \ pid_global_config.processes = (map (init_pc_impl_impl mst) (program.processes prog)), pid_global_config.state = \ global_state_impl.variables = init_valuation_impl (program.global_vars prog) \ \" lemma (in cprog) pid_init_gc_impl_impl: "pid_init_gc_impl_impl prog (comp_info_of prog) = pid_init_gc_impl" unfolding pid_init_gc_impl_impl_def pid_init_gc_impl_def by (simp add: init_pc_impl_impl) definition "ccfg_V0_impl prog mst \ map (comp_\_impl mst) (cfg_V0_list prog)" lemma (in cprog) ccfg_V0_impl: "ccfg_V0_impl prog (comp_info_of prog) = cfgc_V0_list" unfolding ccfg_V0_impl_def cfgc_V0_list_def by (simp add: comp_\_impl) definition "ccfg_G_impl_impl prog mst \ \ gi_V = \ _. True, gi_E = ccfg_E_succ_impl mst, gi_V0 = ccfg_V0_impl prog mst \" lemma (in cprog) ccfg_G_impl_impl: "ccfg_G_impl_impl prog (comp_info_of prog) = cfgc_G_impl" unfolding ccfg_G_impl_impl_def cfgc_G_impl_def unfolding ccfg_E_succ_impl ccfg_V0_impl by rule definition "is_vis_edge_impl mst is_vis_var \ \(c,c'). \(a,cc')\set (ccfg_succ_impl mst c). c'=cc' \ (\v\write_globals a. is_vis_var v)" lemma (in cprog) is_vis_edge_impl: "is_vis_edge_impl (comp_info_of prog) = is_vis_edge" apply (intro ext) unfolding is_vis_edge_impl_def is_vis_edge_def by (simp add: ccfg_succ_impl) definition "ga_gample_mil4_impl mst mem sticky_Ei (gc::pid_global_config_impl) = ( let lcs = pid_global_config.processes gc; gs = pid_global_config.state gc in collecti_index_list (\_ lc. let c = local_config.command lc; ls = local_config.state lc; as = ccfg_succ_impl mst c; s = filter (la_en'_impl (ls,gs) o fst) as; ample = s\[] \ (\(a,_)\set as. read_globals a = {} \ write_globals a = {}) \ (\(a,c')\set s. \mem (c,c') sticky_Ei) in (ample,map (Pair c) s) ) lcs )" lemma (in cprog) ga_gample_mil4_impl: "ga_gample_mil4_impl (comp_info_of prog) = ga_gample_mil4" apply (intro ext) unfolding ga_gample_mil4_impl_def ga_gample_mil4_def by (simp add: ccfg_succ_impl) definition "cr_ample_impl4_impl prog mst is_vis_var \ let sticky = find_fas_init_code (=) bounded_hashcode_nat (def_hashmap_size TYPE(nat)) (ccfg_G_impl_impl prog mst) (is_vis_edge_impl mst is_vis_var) in stutter_extend_en_list (ga_gample_mil4_impl mst (\x f. f x) sticky)" lemma (in cprog) cr_ample_impl4_impl: "cr_ample_impl4_impl prog (comp_info_of prog) = cr_ample_impl4" apply (intro ext) unfolding cr_ample_impl4_impl_def cr_ample_impl4_def by (simp add: ga_gample_mil4_impl ccfg_G_impl_impl is_vis_edge_impl) definition "impl4_succ_impl prog mst is_vis_var \ succ_of_enex_list (cr_ample_impl4_impl prog mst is_vis_var,ga_ex_impl)" lemma (in cprog) impl4_succ_impl: "impl4_succ_impl prog (comp_info_of prog) = impl4_succ" apply (intro ext) unfolding impl4_succ_impl_def impl4_succ_def by (simp add: cr_ample_impl4_impl) definition "ample_step_impl4_impl prog mst is_vis_var \ rel_of_enex_list (cr_ample_impl4_impl prog mst is_vis_var,ga_ex_impl)" lemma (in cprog) ample_step_impl4_impl: "ample_step_impl4_impl prog (comp_info_of prog) = ample_step_impl4" apply (intro ext) unfolding ample_step_impl4_impl_def ample_step_impl4_def by (simp add: cr_ample_impl4_impl) export_code impl4_succ_impl pid_init_gc_impl_impl comp_info_of checking SML end diff --git a/thys/DFS_Framework/Invars/DFS_Invars_Basic.thy b/thys/DFS_Framework/Invars/DFS_Invars_Basic.thy --- a/thys/DFS_Framework/Invars/DFS_Invars_Basic.thy +++ b/thys/DFS_Framework/Invars/DFS_Invars_Basic.thy @@ -1,1657 +1,1657 @@ section \Basic Invariant Library\ theory DFS_Invars_Basic imports "../Param_DFS" begin text \We provide more basic invariants of the DFS algorithm\ subsection \Basic Timing Invariants\ (* Timing *) abbreviation "the_discovered s v \ the (discovered s v)" abbreviation "the_finished s v \ the (finished s v)" locale timing_syntax begin (* Timing related syntax shortcuts *) notation the_discovered ("\") notation the_finished ("\") end context param_DFS begin context begin interpretation timing_syntax . definition "timing_common_inv s \ \ \\\ s v < \ s v\\ (\v \ dom (finished s). \ s v < \ s v) \ \\v \ w \ \ s v \ \ s w \ \ s v \ \ s w\\ \ \Can't use \card dom = card ran\ as the maps may be infinite ...\ \ (\v \ dom (discovered s). \w \ dom (discovered s). v \ w \ \ s v \ \ s w) \ (\v \ dom (finished s). \w \ dom (finished s). v \ w \ \ s v \ \ s w) \ \\\ s v < counter \ \ s v < counter\\ \ (\v \ dom (discovered s). \ s v < counter s) \ (\v \ dom (finished s). \ s v < counter s) \ (\v \ dom (finished s). \w \ succ v. \ s w < \ s v)" lemma timing_common_inv: "is_invar timing_common_inv" proof (induction rule: is_invarI) case (finish s s') then interpret DFS_invar where s=s by simp from finish have NE: "stack s \ []" by (simp add: cond_alt) have *: "hd (stack s) \ dom (finished s)" "hd (stack s) \ dom (discovered s)" using stack_not_finished stack_discovered hd_in_set[OF NE] by blast+ from discovered_closed have "(E - pending s) `` {hd (stack s)} \ dom (discovered s)" using hd_in_set[OF NE] by (auto simp add: discovered_closed_def) hence succ_hd: "pending s `` {hd (stack s)} = {} \ succ (hd (stack s)) \ dom (discovered s)" by blast from finish show ?case apply (simp add: timing_common_inv_def) apply (intro conjI) using * apply simp using * apply simp apply (metis less_irrefl) apply (metis less_irrefl) apply (metis less_SucI) apply (metis less_SucI) apply (blast dest!: succ_hd) using * apply simp done next case (discover s) then interpret DFS_invar where s=s by simp from discover show ?case apply (simp add: timing_common_inv_def) apply (intro conjI) using finished_discovered apply fastforce apply (metis less_irrefl) apply (metis less_irrefl) apply (metis less_SucI) apply (metis less_SucI) using finished_imp_succ_discovered apply fastforce done next case (new_root s s' v0) then interpret DFS_invar where s=s by simp from new_root show ?case apply (simp add: timing_common_inv_def) apply (intro conjI) using finished_discovered apply fastforce apply (metis less_irrefl) apply (metis less_irrefl) apply (metis less_SucI) apply (metis less_SucI) using finished_imp_succ_discovered apply fastforce done qed (simp_all add: timing_common_inv_def) end end context DFS_invar begin context begin interpretation timing_syntax . lemmas s_timing_common_inv = timing_common_inv[THEN make_invar_thm] lemma timing_less_counter: "v \ dom (discovered s) \ \ s v < counter s" "v \ dom (finished s) \ \ s v < counter s" using s_timing_common_inv by (auto simp add: timing_common_inv_def) lemma disc_lt_fin: "v \ dom (finished s) \ \ s v < \ s v" using s_timing_common_inv by (auto simp add: timing_common_inv_def) lemma disc_unequal: assumes "v \ dom (discovered s)" "w \ dom (discovered s)" and "v \ w" shows "\ s v \ \ s w" using s_timing_common_inv assms by (auto simp add: timing_common_inv_def) lemma fin_unequal: assumes "v \ dom (finished s)" "w \ dom (finished s)" and "v \ w" shows "\ s v \ \ s w" using s_timing_common_inv assms by (auto simp add: timing_common_inv_def) lemma finished_succ_fin: assumes "v \ dom (finished s)" and "w \ succ v" shows "\ s w < \ s v" using assms s_timing_common_inv by (simp add: timing_common_inv_def) end end context param_DFS begin context begin interpretation timing_syntax . lemma i_prev_stack_discover_all: "is_invar (\s. \ n < length (stack s). \ v \ set (drop (Suc n) (stack s)). \ s (stack s ! n) > \ s v)" proof (induct rule: is_invarI) case (finish s) thus ?case by (cases "stack s") auto next case (discover s s' u v) hence EQ[simp]: "discovered s' = (discovered s)(v \ counter s)" "stack s' = v#stack s" by simp_all from discover interpret DFS_invar where s=s by simp from discover stack_discovered have v_ni: "v \ set (stack s)" by auto from stack_discovered timing_less_counter have "\w. w \ set (stack s) \ \ s w < counter s" by blast with v_ni have "\w. w \ set (stack s) \ \ s' w < \ s' v" by auto hence "\w. w \ set (drop (Suc 0) (stack s')) \ \ s' w < \ s' (stack s' ! 0)" by auto moreover from v_ni have "\n. \n < (length (stack s')); n > 0\ \ \ s' (stack s' ! n) = \ s (stack s' ! n)" by auto with discover(1) v_ni have "\n. \n < (length (stack s')) - 1; n > 0\ \ \ w \ set (drop (Suc n) (stack s')). \ s' (stack s' ! n) > \ s' w" by (auto dest: in_set_dropD) ultimately show ?case by (metis drop_Suc_Cons length_drop length_pos_if_in_set length_tl list.sel(3) neq0_conv nth_Cons_0 EQ(2) zero_less_diff) qed simp_all end end context DFS_invar begin context begin interpretation timing_syntax . lemmas prev_stack_discover_all = i_prev_stack_discover_all[THEN make_invar_thm] lemma prev_stack_discover: "\n < length (stack s); v \ set (drop (Suc n) (stack s)) \ \ \ s (stack s ! n) > \ s v" by (metis prev_stack_discover_all) lemma Suc_stack_discover: assumes n: "n < (length (stack s)) - 1" shows "\ s (stack s ! n) > \ s (stack s ! Suc n)" proof - from prev_stack_discover assms have "\ v. v \ set (drop (Suc n) (stack s)) \ \ s (stack s ! n) > \ s v" by fastforce moreover from n have "stack s ! Suc n \ set (drop (Suc n) (stack s))" using in_set_conv_nth by fastforce ultimately show ?thesis . qed lemma tl_lt_stack_hd_discover: assumes notempty: "stack s \ []" and "x \ set (tl (stack s))" shows "\ s x < \ s (hd (stack s))" proof - from notempty obtain y ys where "stack s = y#ys" by (metis list.exhaust) with assms show ?thesis using prev_stack_discover by (cases ys) force+ qed lemma stack_nth_order: assumes l: "i < length (stack s)" "j < length (stack s)" shows "\ s (stack s ! i) < \ s (stack s ! j) \ i > j" (is "\ s ?i < \ s ?j \ _") proof assume \: "\ s ?i < \ s ?j" from l stack_set_def have disc: "?i \ dom (discovered s)" "?j \ dom (discovered s)" by auto with disc_unequal[OF disc] \ have "i \ j" by auto moreover { assume "i < j" with l have "stack s ! j \ set (drop (Suc i) (stack s))" using in_set_drop_conv_nth[of "stack s ! j" "Suc i" "stack s"] by fastforce with prev_stack_discover l have "\ s (stack s ! j) < \ s (stack s ! i)" by simp with \ have "False" by simp } ultimately show "i > j" by force next assume "i > j" with l have "stack s ! i \ set (drop (Suc j) (stack s))" using in_set_drop_conv_nth[of "stack s ! i" "Suc j" "stack s"] by fastforce with prev_stack_discover l show "\ s ?i < \ s ?j" by simp qed end end subsection \Paranthesis Theorem\ (* Parenthesis Thm *) context param_DFS begin context begin interpretation timing_syntax . definition "parenthesis s \ \v \ dom (discovered s). \w \ dom (discovered s). \ s v < \ s w \ v \ dom (finished s) \ ( \ s v < \ s w \ \disjoint\ \ (\ s v > \ s w \ w \ dom (finished s) \ \ s w < \ s v))" lemma i_parenthesis: "is_invar parenthesis" proof (induct rule: is_invarI) case (finish s s') hence EQ[simp]: "discovered s' = discovered s" "counter s' = Suc (counter s)" - "finished s' = finished s(hd (stack s) \ counter s)" + "finished s' = (finished s)(hd (stack s) \ counter s)" by simp_all from finish interpret DFS_invar where s=s by simp from finish have NE[simp]: "stack s \ []" by (simp add: cond_alt) { fix x y assume dom: "x \ dom (discovered s')" "y \ dom (discovered s')" and \: "\ s' x < \ s' y" and f: "x \ dom (finished s')" hence neq: "x \ y" by force note assms = dom \ f EQ let ?DISJ = "\ s' x < \ s' y" let ?IN = "\ s' y < \ s' x \ y \ dom (finished s') \ \ s' y < \ s' x" have "?DISJ \ ?IN" proof (cases "x = hd (stack s)") case True note x_is_hd = this hence \x: "\ s' x = counter s" by simp from x_is_hd neq have y_not_hd: "y \ hd (stack s)" by simp have "\ s y < \ s' x \ y \ dom (finished s) \ \ s y < \ s' x" proof (cases "y \ set (stack s)") \ \y on stack is not possible: According to @{thm [display] \} it is discovered after @{text "x (= hd (stack s))"}\ case True with y_not_hd have "y \ set (tl (stack s))" by (cases "stack s") simp_all with tl_lt_stack_hd_discover[OF NE] \ x_is_hd have "\ s y < \ s x" by simp with \ have False by simp thus ?thesis .. next case False \ \y must be a successor of @{text "x (= (hd (stack s)))"}\ from dom have "y \ dom (discovered s)" by simp with False discovered_not_stack_imp_finished have *: "y \ dom (finished s)" by simp moreover with timing_less_counter \x have "\ s y < \ s' x" by simp moreover with * disc_lt_fin \x have "\ s y < \ s' x" by (metis less_trans) ultimately show ?thesis by simp qed with y_not_hd show ?thesis by simp next case False note [simp] = this show ?thesis proof (cases "y = hd (stack s)") case False with finish assms show ?thesis by (simp add: parenthesis_def) next case True with stack_not_finished have "y \ dom (finished s)" using hd_in_set[OF NE] by auto with finish assms have "\ s x < \ s y" unfolding parenthesis_def by auto hence ?DISJ by simp thus ?thesis .. qed qed } thus ?case by (simp add: parenthesis_def) next case (discover s s' u v) hence EQ[simp]: "discovered s' = (discovered s)(v \ counter s)" "finished s' = finished s" "counter s' = Suc (counter s)" by simp_all from discover interpret DFS_invar where s=s by simp from discover finished_discovered have V': "v \ dom (discovered s)" "v \ dom (finished s)" by auto { fix x y assume dom: "x \ dom (discovered s')" "y \ dom (discovered s')" and \: "\ s' x < \ s' y" and f: "x \ dom (finished s')" let ?DISJ = "\ s' x < \ s' y" let ?IN = "\ s' y < \ s' x \ y \ dom (finished s') \ \ s' y < \ s' x" from dom V' f have x: "x \ dom (discovered s)""x \ v" by auto have "?DISJ \ ?IN" proof (cases "y = v") case True hence "\ s' y = counter s" by simp moreover from timing_less_counter x f have "\ s' x < counter s" by auto ultimately have "?DISJ" by simp thus ?thesis .. next case False with dom have "y \ dom (discovered s)" by simp with discover False \ f x show ?thesis by (simp add: parenthesis_def) qed } thus ?case by (simp add: parenthesis_def) next case (new_root s s' v0) then interpret DFS_invar where s=s by simp from finished_discovered new_root have "v0 \ dom (finished s')" by auto with new_root timing_less_counter show ?case by (simp add: parenthesis_def) qed (simp_all add: parenthesis_def) end end context DFS_invar begin context begin interpretation timing_syntax . lemma parenthesis: assumes "v \ dom (finished s)" "w \ dom (discovered s)" and "\ s v < \ s w" shows "\ s v < \ s w \ \disjoint\ \ (\ s v > \ s w \ w \ dom (finished s) \ \ s w < \ s v)" using assms using i_parenthesis[THEN make_invar_thm] using finished_discovered unfolding parenthesis_def by blast lemma parenthesis_contained: assumes "v \ dom (finished s)" "w \ dom (discovered s)" and "\ s v < \ s w" "\ s v > \ s w" shows "w \ dom (finished s) \ \ s w < \ s v" using parenthesis assms by force lemma parenthesis_disjoint: assumes "v \ dom (finished s)" "w \ dom (discovered s)" and "\ s v < \ s w" "\ s w > \ s v" shows "\ s v < \ s w" using parenthesis assms by force lemma finished_succ_contained: assumes "v \ dom (finished s)" and "w \ succ v" and "\ s v < \ s w" shows "w \ dom (finished s) \ \ s w < \ s v" using finished_succ_fin finished_imp_succ_discovered parenthesis_contained using assms by metis end end subsection \Edge Types\ context param_DFS begin abbreviation "edges s \ tree_edges s \ cross_edges s \ back_edges s" (* Demo for simple invariant proof *) lemma "is_invar (\s. finite (edges s))" by (induction rule: establish_invarI) auto text \Sometimes it's useful to just chose between tree-edges and non-tree.\ lemma edgesE_CB: assumes "x \ edges s" and "x \ tree_edges s \ P" and "x \ cross_edges s \ back_edges s \ P" shows "P" using assms by auto definition "edges_basic s \ Field (back_edges s) \ dom (discovered s) \ back_edges s \ E - pending s \ Field (cross_edges s) \ dom (discovered s) \ cross_edges s \ E - pending s \ Field (tree_edges s) \ dom (discovered s) \ tree_edges s \ E - pending s \ back_edges s \ cross_edges s = {} \ back_edges s \ tree_edges s = {} \ cross_edges s \ tree_edges s = {} " lemma i_edges_basic: "is_invar edges_basic" unfolding edges_basic_def[abs_def] proof (induct rule: is_invarI_full) case (back_edge s) then interpret DFS_invar where s=s by simp from back_edge show ?case by (auto dest: pendingD) next case (cross_edge s) then interpret DFS_invar where s=s by simp from cross_edge show ?case by (auto dest: pendingD) next case (discover s) then interpret DFS_invar where s=s by simp from discover show ?case (* Speed optimized proof, using only auto takes too long *) apply (simp add: Field_def Range_def Domain_def) apply (drule pendingD) apply simp by (blast) next case (new_root s) thus ?case by (simp add: Field_def) blast qed auto lemmas (in DFS_invar) edges_basic = i_edges_basic[THEN make_invar_thm] lemma i_edges_covered: "is_invar (\s. (E \ dom (discovered s) \ UNIV) - pending s = edges s)" proof (induction rule: is_invarI_full) case (new_root s s' v0) interpret DFS_invar G param s by fact from new_root empty_stack_imp_empty_pending have [simp]: "pending s = {}" by simp from \v0 \ dom (discovered s)\ have [simp]: "E \ insert v0 (dom (discovered s)) \ UNIV - {v0} \ succ v0 = E \ dom (discovered s) \ UNIV" by auto from new_root show ?case by simp next case (cross_edge s s' u v) interpret DFS_invar G param s by fact from cross_edge stack_discovered have "u \ dom (discovered s)" by (cases "stack s") auto with cross_edge(2-) pending_ssE have "E \ dom (discovered s) \ UNIV - (pending s - {(hd (stack s), v)}) = insert (hd (stack s), v) (E \ dom (discovered s) \ UNIV - pending s)" by auto thus ?case using cross_edge by simp next case (back_edge s s' u v) interpret DFS_invar G param s by fact from back_edge stack_discovered have "u \ dom (discovered s)" by (cases "stack s") auto with back_edge(2-) pending_ssE have "E \ dom (discovered s) \ UNIV - (pending s - {(hd (stack s), v)}) = insert (hd (stack s), v) (E \ dom (discovered s) \ UNIV - pending s)" by auto thus ?case using back_edge by simp next case (discover s s' u v) interpret DFS_invar G param s by fact from discover stack_discovered have "u \ dom (discovered s)" by (cases "stack s") auto with discover(2-) pending_ssE have "E \ insert v (dom (discovered s)) \ UNIV - (pending s - {(hd (stack s), v)} \ {v} \ succ v) = insert (hd (stack s), v) (E \ dom (discovered s) \ UNIV - pending s)" by auto thus ?case using discover by simp qed simp_all end context DFS_invar begin lemmas edges_covered = i_edges_covered[THEN make_invar_thm] lemma edges_ss_reachable_edges: "edges s \ E \ reachable \ UNIV" using edges_covered discovered_reachable by (fast intro: rtrancl_image_advance_rtrancl) lemma nc_edges_covered: assumes "\cond s" "\is_break param s" shows "E \ reachable \ UNIV = edges s" proof - from assms have [simp]: "stack s = []" unfolding cond_def by (auto simp: pred_defs) hence [simp]: "pending s = {}" by (rule empty_stack_imp_empty_pending) from edges_covered nc_discovered_eq_reachable[OF assms] show ?thesis by simp qed lemma tree_edges_ssE: "tree_edges s \ E" and tree_edges_not_pending: "tree_edges s \ - pending s" and tree_edge_is_succ: "(v,w) \ tree_edges s \ w \ succ v" and tree_edges_discovered: "Field (tree_edges s) \ dom (discovered s)" and cross_edges_ssE: "cross_edges s \ E" and cross_edges_not_pending: "cross_edges s \ - pending s" and cross_edge_is_succ: "(v,w) \ cross_edges s \ w \ succ v" and cross_edges_discovered: "Field (cross_edges s) \ dom (discovered s)" and back_edges_ssE: "back_edges s \ E" and back_edges_not_pending: "back_edges s \ - pending s" and back_edge_is_succ: "(v,w) \ back_edges s \ w \ succ v" and back_edges_discovered: "Field (back_edges s) \ dom (discovered s)" using edges_basic unfolding edges_basic_def by auto lemma edges_disjoint: "back_edges s \ cross_edges s = {}" "back_edges s \ tree_edges s = {}" "cross_edges s \ tree_edges s = {}" using edges_basic unfolding edges_basic_def by auto lemma tree_edge_imp_discovered: "(v,w) \ tree_edges s \ v \ dom (discovered s)" "(v,w) \ tree_edges s \ w \ dom (discovered s)" using tree_edges_discovered by (auto simp add: Field_def) lemma back_edge_imp_discovered: "(v,w) \ back_edges s \ v \ dom (discovered s)" "(v,w) \ back_edges s \ w \ dom (discovered s)" using back_edges_discovered by (auto simp add: Field_def) lemma cross_edge_imp_discovered: "(v,w) \ cross_edges s \ v \ dom (discovered s)" "(v,w) \ cross_edges s \ w \ dom (discovered s)" using cross_edges_discovered by (auto simp add: Field_def) lemma edge_imp_discovered: "(v,w) \ edges s \ v \ dom (discovered s)" "(v,w) \ edges s \ w \ dom (discovered s)" using tree_edge_imp_discovered cross_edge_imp_discovered back_edge_imp_discovered by blast+ lemma tree_edges_finite[simp, intro!]: "finite (tree_edges s)" using finite_subset[OF tree_edges_discovered discovered_finite] by simp lemma cross_edges_finite[simp, intro!]: "finite (cross_edges s)" using finite_subset[OF cross_edges_discovered discovered_finite] by simp lemma back_edges_finite[simp, intro!]: "finite (back_edges s)" using finite_subset[OF back_edges_discovered discovered_finite] by simp lemma edges_finite: "finite (edges s)" by auto end subsubsection \Properties of the DFS Tree\ (* Tree *) context DFS_invar begin context begin interpretation timing_syntax . lemma tree_edge_disc_lt_fin: "(v,w) \ tree_edges s \ v \ dom (finished s) \ \ s w < \ s v" by (metis finished_succ_fin tree_edge_is_succ) lemma back_edge_disc_lt_fin: "(v,w) \ back_edges s \ v \ dom (finished s) \ \ s w < \ s v" by (metis finished_succ_fin back_edge_is_succ) lemma cross_edge_disc_lt_fin: "(v,w) \ cross_edges s \ v \ dom (finished s) \ \ s w < \ s v" by (metis finished_succ_fin cross_edge_is_succ) end end (* Stack & Tree *) context param_DFS begin lemma i_stack_is_tree_path: "is_invar (\s. stack s \ [] \ (\v0 \ V0. path (tree_edges s) v0 (rev (tl (stack s))) (hd (stack s))))" proof (induct rule: is_invarI) case (discover s s' u v) hence EQ[simp]: "stack s' = v # stack s" "tree_edges s' = insert (hd (stack s), v) (tree_edges s)" by simp_all from discover have NE[simp]: "stack s \ []" by simp from discover obtain v0 where "v0 \ V0" "path (tree_edges s) v0 (rev (tl (stack s))) (hd (stack s))" by blast with path_mono[OF _ this(2)] EQ have "path (tree_edges s') v0 (rev (tl (stack s))) (hd (stack s))" by blast with \v0 \ V0\ show ?case by (cases "stack s") (auto simp: path_simps) next case (finish s s') hence EQ[simp]: "stack s' = tl (stack s)" "tree_edges s' = tree_edges s" by simp_all from finish obtain v0 where "v0 \ V0" "path (tree_edges s) v0 (rev (tl (stack s))) (hd (stack s))" by blast hence P: "path (tree_edges s') v0 (rev (stack s')) (hd (stack s))" by simp show ?case proof assume A: "stack s' \ []" with P have "(hd (stack s'), hd (stack s)) \ tree_edges s'" by (auto simp: neq_Nil_conv path_simps) moreover from P A have "path (tree_edges s') v0 (rev (tl (stack s')) @ [hd (stack s')]) (hd (stack s))" by (simp) moreover note \v0 \ V0\ ultimately show "\v0\V0. path (tree_edges s') v0 (rev (tl (stack s'))) (hd (stack s'))" by (auto simp add: path_append_conv) qed qed simp_all end context DFS_invar begin lemmas stack_is_tree_path = i_stack_is_tree_path[THEN make_invar_thm, rule_format] lemma stack_is_path: "stack s \ [] \ \v0\V0. path E v0 (rev (tl (stack s))) (hd (stack s))" using stack_is_tree_path path_mono[OF tree_edges_ssE] by blast lemma hd_succ_stack_is_path: assumes ne: "stack s \ []" and succ: "v \ succ (hd (stack s))" shows "\v0\V0. path E v0 (rev (stack s)) v" proof - from stack_is_path[OF ne] succ obtain v0 where "v0 \ V0" "path E v0 (rev (tl (stack s)) @ [hd (stack s)]) v" by (auto simp add: path_append_conv) thus ?thesis using ne by (cases "stack s") auto qed lemma tl_stack_hd_tree_path: assumes "stack s \ []" and "v \ set (tl (stack s))" shows "(v, hd (stack s)) \ (tree_edges s)\<^sup>+" proof - from stack_is_tree_path assms obtain v0 where "path (tree_edges s) v0 (rev (tl (stack s))) (hd (stack s))" by auto from assms path_member_reach_end[OF this] show ?thesis by simp qed end context param_DFS begin definition "tree_discovered_inv s \ (tree_edges s = {} \ dom (discovered s) \ V0 \ (stack s = [] \ (\v0\V0. stack s = [v0]))) \ (tree_edges s \ {} \ (tree_edges s)\<^sup>+ `` V0 \ V0 = dom (discovered s) \ V0)" lemma i_tree_discovered_inv: "is_invar tree_discovered_inv" proof (induct rule: is_invarI) case (discover s s' u v) hence EQ[simp]: "stack s' = v # stack s" "tree_edges s' = insert (hd (stack s), v) (tree_edges s)" "discovered s' = (discovered s)(v \ counter s)" by simp_all from discover interpret DFS_invar where s=s by simp from discover have NE[simp]: "stack s \ []" by simp note TDI = \tree_discovered_inv s\[unfolded tree_discovered_inv_def] have "tree_edges s' = {} \ dom (discovered s') \ V0 \ (stack s' = [] \ (\v0\V0. stack s' = [v0]))" by simp \ \@{text "tree_edges s' \ {}"}\ moreover { fix x assume A: "x \ (tree_edges s')\<^sup>+ `` V0 \ V0" "x \ V0" then obtain y where y: "(y,x) \ (tree_edges s')\<^sup>+" "y \ V0" by auto have "x \ dom (discovered s') \ V0" proof (cases "tree_edges s = {}") case True with discover A have "(tree_edges s')\<^sup>+ = {(hd (stack s), v)}" by (simp add: trancl_single) with A show ?thesis by auto next case False note t_ne = this show ?thesis proof (cases "x = v") case True thus ?thesis by simp next case False with y have "(y,x) \ (tree_edges s)\<^sup>+" proof (induct rule: trancl_induct) case (step a b) hence "(a,b) \ tree_edges s" by simp with tree_edge_imp_discovered have "a \ dom (discovered s)" by simp with discover have "a \ v" by blast with step show ?case by auto qed simp with \y \ V0\ have "x \ (tree_edges s)\<^sup>+ `` V0" by auto with t_ne TDI show ?thesis by auto qed qed } note t_d = this { fix x assume "x \ dom (discovered s') \ V0" "x \ V0" hence A: "x \ dom (discovered s')" by simp have "x \ (tree_edges s')\<^sup>+ `` V0 \ V0" proof (cases "tree_edges s = {}") case True with trancl_single have "(tree_edges s')\<^sup>+ = {(hd (stack s), v)}" by simp moreover from True TDI have "hd (stack s) \ V0" "dom (discovered s) \ V0" by auto ultimately show ?thesis using A \x\V0\ by auto next case False note t_ne = this show ?thesis proof (cases "x=v") case False with A have "x \ dom (discovered s)" by simp with TDI t_ne \x \ V0\ have "x \ (tree_edges s)\<^sup>+ `` V0" by auto with trancl_sub_insert_trancl show ?thesis by simp blast next case True from t_ne TDI have "dom (discovered s) \ V0 = (tree_edges s)\<^sup>+ `` V0 \ V0" by simp moreover from stack_is_tree_path[OF NE] obtain v0 where "v0 \ V0" and "(v0, hd (stack s)) \ (tree_edges s)\<^sup>*" by (blast intro!: path_is_rtrancl) with EQ have "(v0, hd (stack s)) \ (tree_edges s')\<^sup>*" by (auto intro: rtrancl_mono_mp) ultimately show ?thesis using \v0 \ V0\ True by (auto elim: rtrancl_into_trancl1) qed qed } with t_d have "(tree_edges s')\<^sup>+ `` V0 \ V0 = dom (discovered s') \ V0" by blast ultimately show ?case by (simp add: tree_discovered_inv_def) qed (auto simp add: tree_discovered_inv_def) lemmas (in DFS_invar) tree_discovered_inv = i_tree_discovered_inv[THEN make_invar_thm] lemma (in DFS_invar) discovered_iff_tree_path: "v \ V0 \ v \ dom (discovered s) \ (\v0\V0. (v0,v) \ (tree_edges s)\<^sup>+)" using tree_discovered_inv by (auto simp add: tree_discovered_inv_def) lemma i_tree_one_predecessor: "is_invar (\s. \(v,v') \ tree_edges s. \y. y \ v \ (y,v') \ tree_edges s)" proof (induct rule: is_invarI) case (discover s s' u v) hence EQ[simp]: "tree_edges s' = insert (hd (stack s),v) (tree_edges s)" by simp from discover interpret DFS_invar where s=s by simp from discover have NE[simp]: "stack s \ []" by (simp add: cond_alt) { fix w w' y assume *: "(w,w') \ tree_edges s'" and "y \ w" from discover stack_discovered have v_hd: "hd (stack s) \ v" using hd_in_set[OF NE] by blast from discover tree_edges_discovered have v_notin_tree: "\(x,x') \ tree_edges s. x \ v \ x' \ v" by (blast intro!: Field_not_elem) have "(y,w') \ tree_edges s'" proof (cases "w = hd (stack s)") case True have "(y,v) \ tree_edges s'" proof (rule notI) assume "(y,v) \ tree_edges s'" with True \y\w\ have "(y,v) \ tree_edges s" by simp with v_notin_tree show False by auto qed with True * \y\w\ v_hd show ?thesis apply (cases "w = v") apply simp using discover apply simp apply blast done next case False with v_notin_tree * \y\w\ v_hd show ?thesis apply (cases "w' = v") apply simp apply blast using discover apply simp apply blast done qed } thus ?case by blast qed simp_all lemma (in DFS_invar) tree_one_predecessor: assumes "(v,w) \ tree_edges s" and "a \ v" shows "(a,w) \ tree_edges s" using assms make_invar_thm[OF i_tree_one_predecessor] by blast lemma (in DFS_invar) tree_eq_rule: "\(v,w) \ tree_edges s; (u,w) \ tree_edges s\ \ v=u" using tree_one_predecessor by blast context begin interpretation timing_syntax . lemma i_tree_edge_disc: "is_invar (\s. \(v,v') \ tree_edges s. \ s v < \ s v')" proof (induct rule: is_invarI) case (discover s s' u v) hence EQ[simp]: "tree_edges s' = insert (hd (stack s), v) (tree_edges s)" "discovered s' = (discovered s)(v \ counter s)" by simp_all from discover interpret DFS_invar where s=s by simp from discover have NE[simp]: "stack s \ []" by (simp add: cond_alt) from discover tree_edges_discovered have v_notin_tree: "\(x,x') \ tree_edges s. x \ v \ x' \ v" by (blast intro!: Field_not_elem) from discover stack_discovered have v_hd: "hd (stack s) \ v" using hd_in_set[OF NE] by blast { fix a b assume T: "(a,b) \ tree_edges s'" have "\ s' a < \ s' b" proof (cases "b = v") case True with T v_notin_tree have [simp]: "a = hd (stack s)" by auto with stack_discovered have "a \ dom (discovered s)" by (metis hd_in_set NE subsetD) with v_hd True timing_less_counter show ?thesis by simp next case False with v_notin_tree T have "(a,b) \ tree_edges s" "a \ v" by auto with discover have "\ s a < \ s b" by auto with False \a\v\ show ?thesis by simp qed } thus ?case by blast next case (new_root s s' v0) then interpret DFS_invar where s=s by simp from new_root have "tree_edges s' = tree_edges s" by simp moreover from tree_edge_imp_discovered new_root have "\(v,v') \ tree_edges s. v \ v0 \ v' \ v0" by blast ultimately show ?case using new_root by auto qed simp_all end end context DFS_invar begin context begin interpretation timing_syntax . lemma tree_edge_disc: "(v,w) \ tree_edges s \ \ s v < \ s w" using i_tree_edge_disc[THEN make_invar_thm] by blast lemma tree_path_disc: "(v,w) \ (tree_edges s)\<^sup>+ \ \ s v < \ s w" by (auto elim!: trancl_induct dest: tree_edge_disc) lemma no_loop_in_tree: "(v,v) \ (tree_edges s)\<^sup>+" using tree_path_disc by auto lemma tree_acyclic: "acyclic (tree_edges s)" by (metis acyclicI no_loop_in_tree) lemma no_self_loop_in_tree: "(v,v) \ tree_edges s" using tree_edge_disc by auto lemma tree_edge_unequal: "(v,w) \ tree_edges s \ v \ w" by (metis no_self_loop_in_tree) lemma tree_path_unequal: "(v,w) \ (tree_edges s)\<^sup>+ \ v \ w" by (metis no_loop_in_tree) lemma tree_subpath': assumes x: "(x,v) \ (tree_edges s)\<^sup>+" and y: "(y,v) \ (tree_edges s)\<^sup>+" and "x \ y" shows "(x,y) \ (tree_edges s)\<^sup>+ \ (y,x) \ (tree_edges s)\<^sup>+" proof - from x obtain px where px: "path (tree_edges s) x px v" and "px \ []" using trancl_is_path by metis from y obtain py where py: "path (tree_edges s) y py v" and "py \ []" using trancl_is_path by metis from \px \ []\ \py \ []\ px py show ?thesis proof (induction arbitrary: v rule: rev_nonempty_induct2') case (single) hence "(x,v) \ tree_edges s" "(y,v) \ tree_edges s" by (simp_all add: path_simps) with tree_eq_rule have "x=y" by simp with \x\y\ show ?case by contradiction next case (snocl a as) hence "(y,v) \ tree_edges s" by (simp add: path_simps) moreover from snocl have "path (tree_edges s) x as a" "(a,v) \ tree_edges s" by (simp_all add: path_simps) ultimately have "path (tree_edges s) x as y" using tree_eq_rule by auto with path_is_trancl \as \ []\ show ?case by metis next case (snocr _ a as) hence "(x,v) \ tree_edges s" by (simp add: path_simps) moreover from snocr have "path (tree_edges s) y as a" "(a,v) \ tree_edges s" by (simp_all add: path_simps) ultimately have "path (tree_edges s) y as x" using tree_eq_rule by auto with path_is_trancl \as \ []\ show ?case by metis next case (snoclr a as b bs) hence "path (tree_edges s) x as a" "(a,v) \ tree_edges s" "path (tree_edges s) y bs b" "(b,v) \ tree_edges s" by (simp_all add: path_simps) moreover hence "a=b" using tree_eq_rule by simp ultimately show ?thesis using snoclr.IH by metis qed qed lemma tree_subpath: assumes "(x,v) \ (tree_edges s)\<^sup>+" and "(y,v) \ (tree_edges s)\<^sup>+" and \: "\ s x < \ s y" shows "(x,y) \ (tree_edges s)\<^sup>+" proof - from \ have "x \ y" by auto with assms tree_subpath' have "(x,y) \ (tree_edges s)\<^sup>+ \ (y,x) \ (tree_edges s)\<^sup>+" by simp moreover from \ tree_path_disc have "(y,x) \ (tree_edges s)\<^sup>+" by force ultimately show ?thesis by simp qed lemma on_stack_is_tree_path: assumes x: "x \ set (stack s)" and y: "y \ set (stack s)" and \: "\ s x < \ s y" shows "(x,y) \ (tree_edges s)\<^sup>+" proof - from x obtain i where i: "stack s ! i = x" "i < length (stack s)" by (metis in_set_conv_nth) from y obtain j where j: "stack s ! j = y" "j < length (stack s)" by (metis in_set_conv_nth) with i \ stack_nth_order have "j < i" by force from x have ne[simp]: "stack s \ []" by auto from \j have "x \ set (tl (stack s))" using nth_mem nth_tl[OF ne, of "i - 1"] i by auto with tl_stack_hd_tree_path have x_path: "(x, hd (stack s)) \ (tree_edges s)\<^sup>+" by simp then show ?thesis proof (cases "j=0") case True with j have "hd (stack s) = y" by (metis hd_conv_nth ne) with x_path show ?thesis by simp next case False hence "y \ set (tl (stack s))" using nth_mem nth_tl[OF ne, of "j - 1"] j by auto with tl_stack_hd_tree_path have "(y, hd (stack s)) \ (tree_edges s)\<^sup>+" by simp with x_path \ show ?thesis using tree_subpath by metis qed qed lemma hd_stack_tree_path_finished: assumes "stack s \ []" assumes "(hd (stack s), v) \ (tree_edges s)\<^sup>+" shows "v \ dom (finished s)" proof (cases "v \ set (stack s)") case True from assms no_loop_in_tree have "hd (stack s) \ v" by auto with True have "v \ set (tl (stack s))" by (cases "stack s") auto with tl_stack_hd_tree_path assms have "(hd (stack s), hd (stack s)) \ (tree_edges s)\<^sup>+" by (metis trancl_trans) with no_loop_in_tree show ?thesis by contradiction next case False from assms obtain x where "(x,v) \ tree_edges s" by (metis tranclE) with tree_edge_imp_discovered have "v \ dom (discovered s)" by blast with False show ?thesis by (simp add: stack_set_def) qed lemma tree_edge_impl_parenthesis: assumes t: "(v,w) \ tree_edges s" and f: "v \ dom (finished s)" shows "w \ dom (finished s) \ \ s v < \ s w \ \ s w < \ s v " proof - from tree_edge_disc_lt_fin assms have "\ s w < \ s v" by simp with f tree_edge_imp_discovered[OF t] tree_edge_disc[OF t] show ?thesis using parenthesis_contained by metis qed lemma tree_path_impl_parenthesis: assumes "(v,w) \ (tree_edges s)\<^sup>+" and "v \ dom (finished s)" shows "w \ dom (finished s) \ \ s v < \ s w \ \ s w < \ s v " using assms by (auto elim!: trancl_induct dest: tree_edge_impl_parenthesis) lemma nc_reachable_v0_parenthesis: assumes C: "\ cond s" "\ is_break param s" and v: "v \ reachable" "v \ V0" obtains v0 where "v0 \ V0" and "\ s v0 < \ s v \ \ s v < \ s v0 " proof - from nc_discovered_eq_reachable[OF C] discovered_iff_tree_path v obtain v0 where "v0 \ V0" and "(v0,v) \ (tree_edges s)\<^sup>+" by auto moreover with nc_V0_finished[OF C] have "v0 \ dom (finished s)" by auto ultimately show ?thesis using tree_path_impl_parenthesis that[OF \v0 \ V0\] by simp qed end end context param_DFS begin context begin interpretation timing_syntax . definition paren_imp_tree_reach where "paren_imp_tree_reach s \ \v \ dom (discovered s). \w \ dom (finished s). \ s v < \ s w \ (v \ dom (finished s) \ \ s v > \ s w) \ (v,w) \ (tree_edges s)\<^sup>+" lemma paren_imp_tree_reach: "is_invar paren_imp_tree_reach" unfolding paren_imp_tree_reach_def[abs_def] proof (induct rule: is_invarI) case (discover s s' u v) hence EQ[simp]: "tree_edges s' = insert (hd (stack s), v) (tree_edges s)" "finished s' = finished s" "discovered s' = (discovered s)(v \ counter s)" by simp_all from discover interpret DFS_invar where s=s by simp from discover have NE[simp]: "stack s \ []" by (simp add: cond_alt) show ?case proof (intro ballI impI) fix a b assume F:"a \ dom (discovered s')" "b \ dom (finished s')" and D: "\ s' a < \ s' b \ (a \ dom (finished s') \ \ s' a > \ s' b)" from F finished_discovered discover have "b \ v" by auto show "(a,b) \ (tree_edges s')\<^sup>+" proof (cases "a = v") case True with D \b\v\ have "counter s < \ s b" by simp also from F have "b \ dom (discovered s)" using finished_discovered by auto with timing_less_counter have "\ s b < counter s" by simp finally have False . thus ?thesis .. next case False with \b\v\ F D discover have "(a,b) \ (tree_edges s)\<^sup>+" by simp thus ?thesis by (auto intro: trancl_mono_mp) qed qed next case (finish s s' u) hence EQ[simp]: "tree_edges s' = tree_edges s" "finished s' = (finished s)(hd (stack s) \ counter s)" "discovered s' = discovered s" "stack s' = tl (stack s)" by simp_all from finish interpret DFS_invar where s=s by simp from finish have NE[simp]: "stack s \ []" by (simp add: cond_alt) show ?case proof (intro ballI impI) fix a b assume F: "a \ dom (discovered s')" "b \ dom (finished s')" and paren: "\ s' a < \ s' b \ (a \ dom (finished s') \ \ s' a > \ s' b)" hence "a \ b" by auto show "(a,b) \ (tree_edges s')\<^sup>+" proof (cases "b = hd (stack s)") case True hence \b: "\ s' b = counter s" by simp have "a \ set (stack s)" unfolding stack_set_def proof from F show "a \ dom (discovered s)" by simp from True \a\b\ \b paren have "a \ dom (finished s) \ \ s a > counter s" by simp with timing_less_counter show "a \ dom (finished s)" by force qed with paren True on_stack_is_tree_path have "(a,b) \ (tree_edges s)\<^sup>+" by auto thus ?thesis by (auto intro: trancl_mono_mp) next case False note b_not_hd = this show ?thesis proof (cases "a = hd (stack s)") case False with b_not_hd F paren finish show ?thesis by simp next case True with paren b_not_hd F have "a \ dom (discovered s)" "b \ dom (finished s)" "\ s a < \ s b" by simp_all moreover from True stack_not_finished have "a \ dom (finished s)" by simp ultimately show ?thesis by (simp add: finish) qed qed qed next case (new_root s s' v0) then interpret DFS_invar where s=s by simp from new_root finished_discovered have "v0 \ dom (finished s)" by auto moreover note timing_less_counter finished_discovered ultimately show ?case using new_root by clarsimp force qed simp_all end end context DFS_invar begin context begin interpretation timing_syntax . lemmas s_paren_imp_tree_reach = paren_imp_tree_reach[THEN make_invar_thm] lemma parenthesis_impl_tree_path_not_finished: assumes "v \ dom (discovered s)" and "w \ dom (finished s)" and "\ s v < \ s w" and "v \ dom (finished s)" shows "(v,w) \ (tree_edges s)\<^sup>+" using s_paren_imp_tree_reach assms by (auto simp add: paren_imp_tree_reach_def) lemma parenthesis_impl_tree_path: assumes "v \ dom (finished s)" "w \ dom (finished s)" and "\ s v < \ s w" "\ s v > \ s w" shows "(v,w) \ (tree_edges s)\<^sup>+" proof - from assms(1) have "v \ dom (discovered s)" using finished_discovered by blast with assms show ?thesis using s_paren_imp_tree_reach assms by (auto simp add: paren_imp_tree_reach_def) qed lemma tree_path_iff_parenthesis: assumes "v \ dom (finished s)" "w \ dom (finished s)" shows "(v,w) \ (tree_edges s)\<^sup>+ \ \ s v < \ s w \ \ s v > \ s w" using assms by (metis parenthesis_impl_tree_path tree_path_impl_parenthesis) lemma no_pending_succ_impl_path_in_tree: assumes v: "v \ dom (discovered s)" "pending s `` {v} = {}" and w: "w \ succ v" and \: "\ s v < \ s w" shows "(v,w) \ (tree_edges s)\<^sup>+" proof (cases "v \ dom (finished s)") case True with assms assms have "\ s w < \ s v" "w \ dom (discovered s)" using finished_succ_fin finished_imp_succ_discovered by simp_all with True \ show ?thesis using parenthesis_contained parenthesis_impl_tree_path by blast next case False show ?thesis proof (cases "w \ dom (finished s)") case True with False v \ show ?thesis by (simp add: parenthesis_impl_tree_path_not_finished) next case False with \v \ dom (finished s)\ no_pending_imp_succ_discovered v w have "v \ set (stack s)" "w \ set (stack s)" by (simp_all add: stack_set_def) with on_stack_is_tree_path \ show ?thesis by simp qed qed lemma finished_succ_impl_path_in_tree: assumes f: "v \ dom (finished s)" and s: "w \ succ v" and \: "\ s v < \ s w" shows "(v,w) \ (tree_edges s)\<^sup>+" using no_pending_succ_impl_path_in_tree finished_no_pending finished_discovered using assms by blast end end subsubsection \Properties of Cross Edges\ (* Cross Edges *) context param_DFS begin context begin interpretation timing_syntax . lemma i_cross_edges_finished: "is_invar (\s. \(u,v)\cross_edges s. v \ dom (finished s) \ (u \ dom (finished s) \ \ s v < \ s u))" proof (induction rule: is_invarI_full) case (finish s s' u e) interpret DFS_invar G param s by fact from finish stack_not_finished have "u \ dom (finished s)" by auto with finish show ?case by (auto intro: timing_less_counter) next case (cross_edge s s' u v e) interpret DFS_invar G param s by fact from cross_edge stack_not_finished have "u \ dom (finished s)" by auto with cross_edge show ?case by (auto intro: timing_less_counter) qed simp_all end end context DFS_invar begin context begin interpretation timing_syntax . lemmas cross_edges_finished = i_cross_edges_finished[THEN make_invar_thm] lemma cross_edges_target_finished: "(u,v)\cross_edges s \ v \ dom (finished s)" using cross_edges_finished by auto lemma cross_edges_finished_decr: "\(u,v)\cross_edges s; u\dom (finished s)\ \ \ s v < \ s u" using cross_edges_finished by auto lemma cross_edge_unequal: assumes cross: "(v,w) \ cross_edges s" shows "v \ w" proof - from cross_edges_target_finished[OF cross] have w_fin: "w \ dom (finished s)" . show ?thesis proof (cases "v \ dom (finished s)") case True with cross_edges_finished_decr[OF cross] show ?thesis by force next case False with w_fin show ?thesis by force qed qed end end subsubsection \Properties of Back Edges\ (* Back Edges *) context param_DFS begin context begin interpretation timing_syntax . lemma i_back_edge_impl_tree_path: "is_invar (\s. \(v,w) \ back_edges s. (w,v) \ (tree_edges s)\<^sup>+ \ w = v)" proof (induct rule: is_invarI_full) case (back_edge s s' u v) then interpret DFS_invar where s=s by simp from back_edge have st: "v \ set (stack s)" "u \ set (stack s)" using stack_set_def by auto have "(v,u) \ (tree_edges s)\<^sup>+ \ u = v" proof (rule disjCI) assume "u \ v" with st back_edge have "v \ set (tl (stack s))" by (metis not_hd_in_tl) with tl_lt_stack_hd_discover st back_edge have "\ s v < \ s u" by simp with on_stack_is_tree_path st show "(v,u) \ (tree_edges s)\<^sup>+" by simp qed with back_edge show ?case by auto next case discover thus ?case using trancl_sub_insert_trancl by force qed simp_all end end context DFS_invar begin context begin interpretation timing_syntax . lemma back_edge_impl_tree_path: "\(v,w) \ back_edges s; v \ w\ \ (w,v) \ (tree_edges s)\<^sup>+" using i_back_edge_impl_tree_path[THEN make_invar_thm] by blast lemma back_edge_disc: assumes "(v,w) \ back_edges s" shows "\ s w \ \ s v" proof cases assume "v\w" with assms back_edge_impl_tree_path have "(w,v) \ (tree_edges s)\<^sup>+" by simp with tree_path_disc show ?thesis by force qed simp lemma back_edges_tree_disjoint: "back_edges s \ tree_edges s = {}" using back_edge_disc tree_edge_disc by force lemma back_edges_tree_pathes_disjoint: "back_edges s \ (tree_edges s)\<^sup>+ = {}" using back_edge_disc tree_path_disc by force lemma back_edge_finished: assumes "(v,w) \ back_edges s" and "w \ dom (finished s)" shows "v \ dom (finished s) \ \ s v \ \ s w" proof (cases "v=w") case True with assms show ?thesis by simp next case False with back_edge_impl_tree_path assms have "(w,v) \ (tree_edges s)\<^sup>+" by simp with tree_path_impl_parenthesis assms show ?thesis by fastforce qed end end context param_DFS begin context begin interpretation timing_syntax . (* The lemma should probably replaced by: is_invar (\s. \(v,w) \ (tree_edges s)\<^sup>+. v \ succ w \ (w,v) \ back_edges s \ (w,v) \ pending s) *) lemma i_disc_imp_back_edge_or_pending: "is_invar (\s. \(v,w) \ E. v \ dom (discovered s) \ w \ dom (discovered s) \ \ s v \ \ s w \ (w \ dom (finished s) \ v \ dom (finished s) \ \ s w \ \ s v) \ (v,w) \ back_edges s \ (v,w) \ pending s)" proof (induct rule: is_invarI_full) case (cross_edge s s' u v) then interpret DFS_invar where s=s by simp from cross_edge stack_not_finished[of u] have "u \ dom (finished s)" using hd_in_set by (auto simp add: cond_alt) with cross_edge show ?case by auto next case (finish s s' u v) then interpret DFS_invar where s=s by simp from finish have IH: "\v w. \w \ succ v; v \ dom (discovered s); w \ dom (discovered s); \ s w \ \ s v; (w \ dom (finished s) \ v \ dom (finished s) \ \ s v \ \ s w)\ \ (v, w) \ back_edges s \ (v, w) \ pending s" by blast from finish have ne[simp]: "stack s \ []" and p[simp]: "pending s `` {hd (stack s)} = {}" by (simp_all) from hd_in_set[OF ne] have disc: "hd (stack s) \ dom (discovered s)" and not_fin: "hd (stack s) \ dom (finished s)" using stack_discovered stack_not_finished by blast+ { fix w assume w: "w \ succ (hd (stack s))" "w \ hd (stack s)" "w \ dom (discovered s)" and f: "w \ dom (finished s) \ counter s \ \ s w" and \: "\ s w \ \ s (hd (stack s))" with timing_less_counter have "w \ dom (finished s)" by force with finish w \ disc have "(hd (stack s), w) \ back_edges s" by blast } moreover { fix w assume "hd (stack s) \ succ w" "w \ hd (stack s)" and "w \ dom (finished s)" "\ s (hd (stack s)) \ \ s w" with IH[of "hd (stack s)" w] disc not_fin have "(w, hd (stack s)) \ back_edges s" using finished_discovered finished_no_pending[of w] by blast } ultimately show ?case using finish by clarsimp auto next case (discover s s' u v) then interpret DFS_invar where s=s by simp from discover show ?case using timing_less_counter by clarsimp fastforce next case (new_root s s' v0) then interpret DFS_invar where s=s by simp from new_root empty_stack_imp_empty_pending have "pending s = {}" by simp with new_root show ?case using timing_less_counter by clarsimp fastforce qed auto end end context DFS_invar begin context begin interpretation timing_syntax . lemma disc_imp_back_edge_or_pending: "\w \ succ v; v \ dom (discovered s); w \ dom (discovered s); \ s w \ \ s v; (w \ dom (finished s) \ v \ dom (finished s) \ \ s v \ \ s w)\ \ (v, w) \ back_edges s \ (v, w) \ pending s" using i_disc_imp_back_edge_or_pending[THEN make_invar_thm] by blast lemma finished_imp_back_edge: "\w \ succ v; v \ dom (finished s); w \ dom (finished s); \ s w \ \ s v; \ s v \ \ s w\ \ (v, w) \ back_edges s" using disc_imp_back_edge_or_pending finished_discovered finished_no_pending by fast lemma finished_not_finished_imp_back_edge: "\w \ succ v; v \ dom (finished s); w \ dom (discovered s); w \ dom (finished s); \ s w \ \ s v\ \ (v, w) \ back_edges s" using disc_imp_back_edge_or_pending finished_discovered finished_no_pending by fast lemma finished_self_loop_in_back_edges: assumes "v \ dom (finished s)" and "(v,v) \ E" shows "(v,v) \ back_edges s" using assms using finished_imp_back_edge by blast end end (* Back edges and Cycles *) context DFS_invar begin context begin interpretation timing_syntax . (* Cross and tree_edges edges are always acyclic *) lemma tree_cross_acyclic: "acyclic (tree_edges s \ cross_edges s)" (is "acyclic ?E") proof (rule ccontr) { fix u v assume *: "u \ dom (finished s)" and "(u,v) \ ?E\<^sup>+" from this(2) have "\ s v < \ s u \ v \ dom (finished s)" proof induct case base thus ?case by (metis Un_iff * cross_edges_finished_decr cross_edges_target_finished tree_edge_impl_parenthesis) next case (step v w) hence "\ s w < \ s v \ w \ dom (finished s)" by (metis Un_iff cross_edges_finished_decr cross_edges_target_finished tree_edge_impl_parenthesis) with step show ?case by auto qed } note aux = this assume "\ acyclic ?E" then obtain u where path: "(u,u) \ ?E\<^sup>+" by (auto simp add: acyclic_def) show False proof cases assume "u \ dom (finished s)" with aux path show False by blast next assume *: "u \ dom (finished s)" moreover from no_loop_in_tree have "(u,u) \ (tree_edges s)\<^sup>+" . with trancl_union_outside[OF path] obtain x y where "(u,x) \ ?E\<^sup>*" "(x,y) \ cross_edges s" "(y,u) \ ?E\<^sup>*" by auto with cross_edges_target_finished have "y \ dom (finished s)" by simp moreover with * \(y,u) \ ?E\<^sup>*\ have "(y,u) \ ?E\<^sup>+" by (auto simp add: rtrancl_eq_or_trancl) ultimately show False by (metis aux) qed qed end lemma cycle_contains_back_edge: assumes cycle: "(u,u) \ (edges s)\<^sup>+" shows "\v w. (u,v) \ (edges s)\<^sup>* \ (v,w) \ back_edges s \ (w,u) \ (edges s)\<^sup>*" proof - from tree_cross_acyclic have "(u,u) \ (tree_edges s \ cross_edges s)\<^sup>+" by (simp add: acyclic_def) with trancl_union_outside[OF cycle] show ?thesis . qed lemma cycle_needs_back_edge: assumes "back_edges s = {}" shows "acyclic (edges s)" proof (rule ccontr) assume "\ acyclic (edges s)" then obtain u where "(u,u) \ (edges s)\<^sup>+" by (auto simp: acyclic_def) with assms have "(u,u) \ (tree_edges s \ cross_edges s)\<^sup>+" by auto with tree_cross_acyclic show False by (simp add: acyclic_def) qed lemma back_edge_closes_cycle: assumes "back_edges s \ {}" shows "\ acyclic (edges s)" proof - from assms obtain v w where be: "(v,w) \ back_edges s" by auto hence "(w,w) \ (edges s)\<^sup>+" proof (cases "v=w") case False with be back_edge_impl_tree_path have "(w,v) \ (tree_edges s)\<^sup>+" by simp hence "(w,v) \ (edges s)\<^sup>+" by (blast intro: trancl_mono_mp) also from be have "(v,w) \ edges s" by simp finally show ?thesis . qed auto thus ?thesis by (auto simp add: acyclic_def) qed lemma back_edge_closes_reachable_cycle: "back_edges s \ {} \ \ acyclic (E \ reachable \ UNIV)" by (metis back_edge_closes_cycle edges_ss_reachable_edges cyclic_subset) lemma cycle_iff_back_edges: "acyclic (edges s) \ back_edges s = {}" by (metis back_edge_closes_cycle cycle_needs_back_edge) end subsection \White Path Theorem\ context DFS begin context begin interpretation timing_syntax . definition white_path where "white_path s x y \ x\y \ (\p. path E x p y \ (\ s x < \ s y \ (\ v \ set (tl p). \ s x < \ s v)))" lemma white_path: "it_dfs \ SPEC(\s. \x \ reachable. \y \ reachable. \ is_break param s \ white_path s x y \ (x,y) \ (tree_edges s)\<^sup>*)" proof (rule it_dfs_SPEC, intro ballI impI) fix s x y assume DI: "DFS_invar G param s" and C: "\ cond s" "\ is_break param s" and reach: "x \ reachable" "y \ reachable" from DI interpret DFS_invar where s=s . note fin_eq_reach = nc_finished_eq_reachable[OF C] show "white_path s x y \ (x,y) \ (tree_edges s)\<^sup>*" proof (cases "x=y") case True thus ?thesis by (simp add: white_path_def) next case False show ?thesis proof assume "(x,y) \ (tree_edges s)\<^sup>*" with \x\y\ have T: "(x,y) \ (tree_edges s)\<^sup>+" by (metis rtranclD) then obtain p where P: "path (tree_edges s) x p y" by (metis trancl_is_path) with tree_edges_ssE have "path E x p y" using path_mono[where E="tree_edges s"] by simp moreover from P have "\ s x < \ s y \ (\ v \ set (tl p). \ s x < \ s v)" using \x\y\ proof (induct rule: path_tl_induct) case (single u) thus ?case by (fact tree_edge_disc) next case (step u v) note \\ s x < \ s u\ also from step have "\ s u < \ s v" by (metis tree_edge_disc) finally show ?case . qed ultimately show "white_path s x y" by (auto simp add: \x\y\ white_path_def) next assume "white_path s x y" with \x\y\ obtain p where P:"path E x p y" and white: "\ s x < \ s y \ (\ v \ set (tl p). \ s x < \ s v)" unfolding white_path_def by blast hence "p \ []" by auto thus "(x,y) \ (tree_edges s)\<^sup>*" using P white reach(2) proof (induction p arbitrary: y rule: rev_nonempty_induct) case single hence "y \ succ x" by (simp add: path_cons_conv) with reach single show ?case using fin_eq_reach finished_succ_impl_path_in_tree[of x y] by simp next case (snoc u us) hence "path E x us u" by (simp add: path_append_conv) moreover hence "(x,u) \ E\<^sup>*" by (simp add: path_is_rtrancl) with reach have ureach: "u \ reachable" by (metis rtrancl_image_advance_rtrancl) moreover from snoc have "\ s x < \ s u" "(\v\set (tl us). \ s x < \ s v)" by simp_all ultimately have x_u: "(x,u) \ (tree_edges s)\<^sup>*" by (metis snoc.IH) from snoc have "y \ succ u" by (simp add: path_append_conv) from snoc(5) fin_eq_reach finished_discovered have y_f_d: "y \ dom (finished s)" "y \ dom (discovered s)" by auto from \y \ succ u\ ureach fin_eq_reach have "\ s y < \ s u" using finished_succ_fin by simp also from \\ s x < \ s u\ have "x \ u" by auto with x_u have "(x,u) \ (tree_edges s)\<^sup>+" by (metis rtrancl_eq_or_trancl) with fin_eq_reach reach have "\ s u < \ s x" using tree_path_impl_parenthesis by simp finally have "\ s y < \ s x" using reach fin_eq_reach y_f_d snoc using parenthesis_contained by blast hence "(x,y) \ (tree_edges s)\<^sup>+" using reach fin_eq_reach y_f_d snoc using parenthesis_impl_tree_path by blast thus ?case by auto qed qed qed qed end end end diff --git a/thys/Eval_FO/Ailamazyan.thy b/thys/Eval_FO/Ailamazyan.thy --- a/thys/Eval_FO/Ailamazyan.thy +++ b/thys/Eval_FO/Ailamazyan.thy @@ -1,5597 +1,5597 @@ theory Ailamazyan imports Eval_FO Cluster Mapping_Code begin fun SP :: "('a, 'b) fo_fmla \ nat set" where "SP (Eqa (Var n) (Var n')) = (if n \ n' then {n, n'} else {})" | "SP (Neg \) = SP \" | "SP (Conj \ \) = SP \ \ SP \" | "SP (Disj \ \) = SP \ \ SP \" | "SP (Exists n \) = SP \ - {n}" | "SP (Forall n \) = SP \ - {n}" | "SP _ = {}" lemma SP_fv: "SP \ \ fv_fo_fmla \" by (induction \ rule: SP.induct) auto lemma finite_SP: "finite (SP \)" using SP_fv finite_fv_fo_fmla finite_subset by fastforce fun SP_list_rec :: "('a, 'b) fo_fmla \ nat list" where "SP_list_rec (Eqa (Var n) (Var n')) = (if n \ n' then [n, n'] else [])" | "SP_list_rec (Neg \) = SP_list_rec \" | "SP_list_rec (Conj \ \) = SP_list_rec \ @ SP_list_rec \" | "SP_list_rec (Disj \ \) = SP_list_rec \ @ SP_list_rec \" | "SP_list_rec (Exists n \) = filter (\m. n \ m) (SP_list_rec \)" | "SP_list_rec (Forall n \) = filter (\m. n \ m) (SP_list_rec \)" | "SP_list_rec _ = []" definition SP_list :: "('a, 'b) fo_fmla \ nat list" where "SP_list \ = remdups_adj (sort (SP_list_rec \))" lemma SP_list_set: "set (SP_list \) = SP \" unfolding SP_list_def by (induction \ rule: SP.induct) (auto simp: fv_fo_terms_set_list) lemma sorted_distinct_SP_list: "sorted_distinct (SP_list \)" unfolding SP_list_def by (auto intro: distinct_remdups_adj_sort) fun d :: "('a, 'b) fo_fmla \ nat" where "d (Eqa (Var n) (Var n')) = (if n \ n' then 2 else 1)" | "d (Neg \) = d \" | "d (Conj \ \) = max (d \) (max (d \) (card (SP (Conj \ \))))" | "d (Disj \ \) = max (d \) (max (d \) (card (SP (Disj \ \))))" | "d (Exists n \) = d \" | "d (Forall n \) = d \" | "d _ = 1" lemma d_pos: "1 \ d \" by (induction \ rule: d.induct) auto lemma card_SP_d: "card (SP \) \ d \" using dual_order.trans by (induction \ rule: SP.induct) (fastforce simp: card_Diff1_le finite_SP)+ fun eval_eterm :: "('a + 'c) val \ 'a fo_term \ 'a + 'c" (infix "\e" 60) where "eval_eterm \ (Const c) = Inl c" | "eval_eterm \ (Var n) = \ n" definition eval_eterms :: "('a + 'c) val \ ('a fo_term) list \ ('a + 'c) list" (infix "\e" 60) where "eval_eterms \ ts = map (eval_eterm \) ts" lemma eval_eterm_cong: "(\n. n \ fv_fo_term_set t \ \ n = \' n) \ eval_eterm \ t = eval_eterm \' t" by (cases t) auto lemma eval_eterms_fv_fo_terms_set: "\ \e ts = \' \e ts \ n \ fv_fo_terms_set ts \ \ n = \' n" proof (induction ts) case (Cons t ts) then show ?case by (cases t) (auto simp: eval_eterms_def fv_fo_terms_set_def) qed (auto simp: eval_eterms_def fv_fo_terms_set_def) lemma eval_eterms_cong: "(\n. n \ fv_fo_terms_set ts \ \ n = \' n) \ eval_eterms \ ts = eval_eterms \' ts" by (auto simp: eval_eterms_def fv_fo_terms_set_def intro: eval_eterm_cong) lemma eval_terms_eterms: "map Inl (\ \ ts) = (Inl \ \) \e ts" proof (induction ts) case (Cons t ts) then show ?case by (cases t) (auto simp: eval_terms_def eval_eterms_def) qed (auto simp: eval_terms_def eval_eterms_def) fun ad_equiv_pair :: "'a set \ ('a + 'c) \ ('a + 'c) \ bool" where "ad_equiv_pair X (a, a') \ (a \ Inl ` X \ a = a') \ (a' \ Inl ` X \ a = a')" fun sp_equiv_pair :: "'a \ 'b \ 'a \ 'b \ bool" where "sp_equiv_pair (a, b) (a', b') \ (a = a' \ b = b')" definition ad_equiv_list :: "'a set \ ('a + 'c) list \ ('a + 'c) list \ bool" where "ad_equiv_list X xs ys \ length xs = length ys \ (\x \ set (zip xs ys). ad_equiv_pair X x)" definition sp_equiv_list :: "('a + 'c) list \ ('a + 'c) list \ bool" where "sp_equiv_list xs ys \ length xs = length ys \ pairwise sp_equiv_pair (set (zip xs ys))" definition ad_agr_list :: "'a set \ ('a + 'c) list \ ('a + 'c) list \ bool" where "ad_agr_list X xs ys \ length xs = length ys \ ad_equiv_list X xs ys \ sp_equiv_list xs ys" lemma ad_equiv_pair_refl[simp]: "ad_equiv_pair X (a, a)" by auto declare ad_equiv_pair.simps[simp del] lemma ad_equiv_pair_comm: "ad_equiv_pair X (a, a') \ ad_equiv_pair X (a', a)" by (auto simp: ad_equiv_pair.simps) lemma ad_equiv_pair_mono: "X \ Y \ ad_equiv_pair Y (a, a') \ ad_equiv_pair X (a, a')" unfolding ad_equiv_pair.simps by fastforce lemma sp_equiv_pair_comm: "sp_equiv_pair x y \ sp_equiv_pair y x" by (cases x; cases y) auto definition sp_equiv :: "('a + 'c) val \ ('a + 'c) val \ nat set \ bool" where "sp_equiv \ \ I \ pairwise sp_equiv_pair ((\n. (\ n, \ n)) ` I)" lemma sp_equiv_mono: "I \ J \ sp_equiv \ \ J \ sp_equiv \ \ I" by (auto simp: sp_equiv_def pairwise_def) definition ad_agr_sets :: "nat set \ nat set \ 'a set \ ('a + 'c) val \ ('a + 'c) val \ bool" where "ad_agr_sets FV S X \ \ \ (\i \ FV. ad_equiv_pair X (\ i, \ i)) \ sp_equiv \ \ S" lemma ad_agr_sets_comm: "ad_agr_sets FV S X \ \ \ ad_agr_sets FV S X \ \" unfolding ad_agr_sets_def sp_equiv_def pairwise_def by (subst ad_equiv_pair_comm) auto lemma ad_agr_sets_mono: "X \ Y \ ad_agr_sets FV S Y \ \ \ ad_agr_sets FV S X \ \" using ad_equiv_pair_mono by (fastforce simp: ad_agr_sets_def) lemma ad_agr_sets_mono': "S \ S' \ ad_agr_sets FV S' X \ \ \ ad_agr_sets FV S X \ \" by (auto simp: ad_agr_sets_def sp_equiv_def pairwise_def) lemma ad_equiv_list_comm: "ad_equiv_list X xs ys \ ad_equiv_list X ys xs" by (auto simp: ad_equiv_list_def) (smt (verit, del_insts) ad_equiv_pair_comm in_set_zip prod.sel(1) prod.sel(2)) lemma ad_equiv_list_mono: "X \ Y \ ad_equiv_list Y xs ys \ ad_equiv_list X xs ys" using ad_equiv_pair_mono by (fastforce simp: ad_equiv_list_def) lemma ad_equiv_list_trans: assumes "ad_equiv_list X xs ys" "ad_equiv_list X ys zs" shows "ad_equiv_list X xs zs" proof - have lens: "length xs = length ys" "length xs = length zs" "length ys = length zs" using assms by (auto simp: ad_equiv_list_def) have "\x z. (x, z) \ set (zip xs zs) \ ad_equiv_pair X (x, z)" proof - fix x z assume "(x, z) \ set (zip xs zs)" then obtain i where i_def: "i < length xs" "xs ! i = x" "zs ! i = z" by (auto simp: set_zip) define y where "y = ys ! i" have "ad_equiv_pair X (x, y)" "ad_equiv_pair X (y, z)" using assms lens i_def by (fastforce simp: set_zip y_def ad_equiv_list_def)+ then show "ad_equiv_pair X (x, z)" unfolding ad_equiv_pair.simps by blast qed then show ?thesis using assms by (auto simp: ad_equiv_list_def) qed lemma ad_equiv_list_link: "(\i \ set ns. ad_equiv_pair X (\ i, \ i)) \ ad_equiv_list X (map \ ns) (map \ ns)" by (auto simp: ad_equiv_list_def set_zip) (metis in_set_conv_nth nth_map) lemma set_zip_comm: "(x, y) \ set (zip xs ys) \ (y, x) \ set (zip ys xs)" by (metis in_set_zip prod.sel(1) prod.sel(2)) lemma set_zip_map: "set (zip (map \ ns) (map \ ns)) = (\n. (\ n, \ n)) ` set ns" by (induction ns) auto lemma sp_equiv_list_comm: "sp_equiv_list xs ys \ sp_equiv_list ys xs" unfolding sp_equiv_list_def using set_zip_comm by (auto simp: pairwise_def) force+ lemma sp_equiv_list_trans: assumes "sp_equiv_list xs ys" "sp_equiv_list ys zs" shows "sp_equiv_list xs zs" proof - have lens: "length xs = length ys" "length xs = length zs" "length ys = length zs" using assms by (auto simp: sp_equiv_list_def) have "pairwise sp_equiv_pair (set (zip xs zs))" proof (rule pairwiseI) fix xz xz' assume "xz \ set (zip xs zs)" "xz' \ set (zip xs zs)" then obtain x z i x' z' i' where xz_def: "i < length xs" "xs ! i = x" "zs ! i = z" "xz = (x, z)" "i' < length xs" "xs ! i' = x'" "zs ! i' = z'" "xz' = (x', z')" by (auto simp: set_zip) define y where "y = ys ! i" define y' where "y' = ys ! i'" have "sp_equiv_pair (x, y) (x', y')" "sp_equiv_pair (y, z) (y', z')" using assms lens xz_def by (auto simp: sp_equiv_list_def pairwise_def y_def y'_def set_zip) metis+ then show "sp_equiv_pair xz xz'" by (auto simp: xz_def) qed then show ?thesis using assms by (auto simp: sp_equiv_list_def) qed lemma sp_equiv_list_link: "sp_equiv_list (map \ ns) (map \ ns) \ sp_equiv \ \ (set ns)" apply (auto simp: sp_equiv_list_def sp_equiv_def pairwise_def set_zip in_set_conv_nth) apply (metis nth_map) apply (metis nth_map) apply fastforce+ done lemma ad_agr_list_comm: "ad_agr_list X xs ys \ ad_agr_list X ys xs" using ad_equiv_list_comm sp_equiv_list_comm by (fastforce simp: ad_agr_list_def) lemma ad_agr_list_mono: "X \ Y \ ad_agr_list Y ys xs \ ad_agr_list X ys xs" using ad_equiv_list_mono by (force simp: ad_agr_list_def) lemma ad_agr_list_rev_mono: assumes "Y \ X" "ad_agr_list Y ys xs" "Inl -` set xs \ Y" "Inl -` set ys \ Y" shows "ad_agr_list X ys xs" proof - have "(a, b) \ set (zip ys xs) \ ad_equiv_pair Y (a, b) \ ad_equiv_pair X (a, b)" for a b using assms apply (cases a; cases b) apply (auto simp: ad_agr_list_def ad_equiv_list_def vimage_def set_zip) unfolding ad_equiv_pair.simps apply (metis Collect_mem_eq Collect_mono_iff imageI nth_mem) apply (metis Collect_mem_eq Collect_mono_iff imageI nth_mem) apply (metis Collect_mem_eq Collect_mono_iff imageI nth_mem) apply (metis Inl_Inr_False image_iff) done then show ?thesis using assms by (fastforce simp: ad_agr_list_def ad_equiv_list_def) qed lemma ad_agr_list_trans: "ad_agr_list X xs ys \ ad_agr_list X ys zs \ ad_agr_list X xs zs" using ad_equiv_list_trans sp_equiv_list_trans by (force simp: ad_agr_list_def) lemma ad_agr_list_refl: "ad_agr_list X xs xs" by (auto simp: ad_agr_list_def ad_equiv_list_def set_zip ad_equiv_pair.simps sp_equiv_list_def pairwise_def) lemma ad_agr_list_set: "ad_agr_list X xs ys \ y \ X \ Inl y \ set ys \ Inl y \ set xs" by (auto simp: ad_agr_list_def ad_equiv_list_def set_zip in_set_conv_nth) (metis ad_equiv_pair.simps image_eqI) lemma ad_agr_list_length: "ad_agr_list X xs ys \ length xs = length ys" by (auto simp: ad_agr_list_def) lemma ad_agr_list_eq: "set ys \ AD \ ad_agr_list AD (map Inl xs) (map Inl ys) \ xs = ys" by (fastforce simp: ad_agr_list_def ad_equiv_list_def set_zip ad_equiv_pair.simps intro!: nth_equalityI) lemma sp_equiv_list_subset: assumes "set ms \ set ns" "sp_equiv_list (map \ ns) (map \' ns)" shows "sp_equiv_list (map \ ms) (map \' ms)" unfolding sp_equiv_list_def length_map pairwise_def proof (rule conjI, rule refl, (rule ballI)+, rule impI) fix x y assume "x \ set (zip (map \ ms) (map \' ms))" "y \ set (zip (map \ ms) (map \' ms))" "x \ y" then have "x \ set (zip (map \ ns) (map \' ns))" "y \ set (zip (map \ ns) (map \' ns))" "x \ y" using assms(1) by (auto simp: set_zip) (metis in_set_conv_nth nth_map subset_iff)+ then show "sp_equiv_pair x y" using assms(2) by (auto simp: sp_equiv_list_def pairwise_def) qed lemma ad_agr_list_subset: "set ms \ set ns \ ad_agr_list X (map \ ns) (map \' ns) \ ad_agr_list X (map \ ms) (map \' ms)" by (auto simp: ad_agr_list_def ad_equiv_list_def sp_equiv_list_subset set_zip) (metis (no_types, lifting) in_set_conv_nth nth_map subset_iff) lemma ad_agr_list_link: "ad_agr_sets (set ns) (set ns) AD \ \ \ ad_agr_list AD (map \ ns) (map \ ns)" unfolding ad_agr_sets_def ad_agr_list_def using ad_equiv_list_link sp_equiv_list_link by fastforce definition ad_agr :: "('a, 'b) fo_fmla \ 'a set \ ('a + 'c) val \ ('a + 'c) val \ bool" where "ad_agr \ X \ \ \ ad_agr_sets (fv_fo_fmla \) (SP \) X \ \" lemma ad_agr_sets_restrict: "ad_agr_sets (set (fv_fo_fmla_list \)) (set (fv_fo_fmla_list \)) AD \ \ \ ad_agr \ AD \ \" using sp_equiv_mono SP_fv unfolding fv_fo_fmla_list_set by (auto simp: ad_agr_sets_def ad_agr_def) blast lemma finite_Inl: "finite X \ finite (Inl -` X)" using finite_vimageI[of X Inl] by (auto simp: vimage_def) lemma ex_out: assumes "finite X" shows "\k. k \ X \ k < Suc (card X)" using card_mono[OF assms, of "{..: assumes "ad_agr_sets (FV - {n}) (S - {n}) X \ \" "S \ FV" "finite S" "\ ` (FV - {n}) \ Z" "Inl ` X \ Inr ` {.. ` (S - {n})) + (if n \ S then 1 else 0))} \ Z" shows "\k \ Z. ad_agr_sets FV S X (\(n := x)) (\(n := k))" proof (cases "n \ S") case True note n_in_S = True show ?thesis proof (cases "x \ Inl ` X") case True show ?thesis using assms n_in_S True apply (auto simp: ad_agr_sets_def sp_equiv_def pairwise_def intro!: bexI[of _ "x"]) unfolding ad_equiv_pair.simps apply (metis True insert_Diff insert_iff subsetD)+ done next case False note \_n_not_Inl = False show ?thesis proof (cases "\m \ S - {n}. x = \ m") case True obtain m where m_def: "m \ S - {n}" "x = \ m" using True by auto have \_m_in: "\ m \ Z" using assms m_def by auto show ?thesis using assms n_in_S \_n_not_Inl True m_def by (auto simp: ad_agr_sets_def sp_equiv_def pairwise_def intro!: bexI[of _ "\ m"]) next case False have out: "x \ \ ` (S - {n})" using False by auto have fin: "finite (Inr -` \ ` (S - {n}))" using assms(3) by (simp add: finite_vimageI) obtain k where k_def: "Inr k \ \ ` (S - {n})" "k < Suc (card (Inr -` \ ` (S - {n})))" using ex_out[OF fin] True by auto show ?thesis using assms n_in_S \_n_not_Inl out k_def assms(5) apply (auto simp: ad_agr_sets_def sp_equiv_def pairwise_def intro!: bexI[of _ "Inr k"]) unfolding ad_equiv_pair.simps apply fastforce apply (metis image_eqI insertE insert_Diff) done qed qed next case False show ?thesis proof (cases "x \ Inl ` X") case x_in: True then show ?thesis using assms False by (auto simp: ad_agr_sets_def sp_equiv_def pairwise_def intro!: bexI[of _ "x"]) next case x_out: False then show ?thesis using assms False apply (auto simp: ad_agr_sets_def sp_equiv_def pairwise_def intro!: bexI[of _ "Inr 0"]) unfolding ad_equiv_pair.simps apply fastforce done qed qed lemma esat_Pred: assumes "ad_agr_sets FV S (\(set ` X)) \ \" "fv_fo_terms_set ts \ FV" "\ \e ts \ map Inl ` X" "t \ set ts" shows "\ \e t = \ \e t" proof (cases t) case (Var n) obtain vs where vs_def: "\ \e ts = map Inl vs" "vs \ X" using assms(3) by auto have "\ n \ set (\ \e ts)" using assms(4) by (force simp: eval_eterms_def Var) then have "\ n \ Inl ` \ (set ` X)" using vs_def(2) unfolding vs_def(1) by auto moreover have "n \ FV" using assms(2,4) by (fastforce simp: Var fv_fo_terms_set_def) ultimately show ?thesis using assms(1) unfolding ad_equiv_pair.simps ad_agr_sets_def Var by fastforce qed auto lemma sp_equiv_list_fv: assumes "(\i. i \ fv_fo_terms_set ts \ ad_equiv_pair X (\ i, \ i))" "\(set_fo_term ` set ts) \ X" "sp_equiv \ \ (fv_fo_terms_set ts)" shows "sp_equiv_list (map ((\e) \) ts) (map ((\e) \) ts)" using assms proof (induction ts) case (Cons t ts) have ind: "sp_equiv_list (map ((\e) \) ts) (map ((\e) \) ts)" using Cons by (auto simp: fv_fo_terms_set_def sp_equiv_def pairwise_def) show ?case proof (cases t) case (Const c) have c_X: "c \ X" using Cons(3) by (auto simp: Const) have fv_t: "fv_fo_term_set t = {}" by (auto simp: Const) have "t' \ set ts \ sp_equiv_pair (\ \e t, \ \e t) (\ \e t', \ \e t')" for t' using c_X Const Cons(2) apply (cases t') apply (auto simp: fv_fo_terms_set_def) unfolding ad_equiv_pair.simps by (metis Cons(2) ad_equiv_pair.simps fv_fo_terms_setI image_insert insert_iff list.set(2) mk_disjoint_insert)+ then show "sp_equiv_list (map ((\e) \) (t # ts)) (map ((\e) \) (t # ts))" using ind pairwise_insert[of sp_equiv_pair "(\ \e t, \ \e t)"] unfolding sp_equiv_list_def set_zip_map by (auto simp: sp_equiv_pair_comm fv_fo_terms_set_def fv_t) next case (Var n) have ad_n: "ad_equiv_pair X (\ n, \ n)" using Cons(2) by (auto simp: fv_fo_terms_set_def Var) have sp_equiv_Var: "\n'. Var n' \ set ts \ sp_equiv_pair (\ n, \ n) (\ n', \ n')" using Cons(4) by (auto simp: sp_equiv_def pairwise_def fv_fo_terms_set_def Var) have "t' \ set ts \ sp_equiv_pair (\ \e t, \ \e t) (\ \e t', \ \e t')" for t' using Cons(2,3) sp_equiv_Var apply (cases t') apply (auto simp: Var) apply (metis SUP_le_iff ad_equiv_pair.simps ad_n fo_term.set_intros imageI subset_eq) apply (metis SUP_le_iff ad_equiv_pair.simps ad_n fo_term.set_intros imageI subset_eq) done then show ?thesis using ind pairwise_insert[of sp_equiv_pair "(\ \e t, \ \e t)" "(\n. (\ \e n, \ \e n)) ` set ts"] unfolding sp_equiv_list_def set_zip_map by (auto simp: sp_equiv_pair_comm) qed qed (auto simp: sp_equiv_def sp_equiv_list_def fv_fo_terms_set_def) lemma esat_Pred_inf: assumes "fv_fo_terms_set ts \ FV" "fv_fo_terms_set ts \ S" "ad_agr_sets FV S AD \ \" "ad_agr_list AD (\ \e ts) vs" "\(set_fo_term ` set ts) \ AD" shows "ad_agr_list AD (\ \e ts) vs" proof - have sp: "sp_equiv \ \ (fv_fo_terms_set ts)" using assms(2,3) sp_equiv_mono unfolding ad_agr_sets_def by auto have "(\i. i \ fv_fo_terms_set ts \ ad_equiv_pair AD (\ i, \ i))" using assms(1,3) by (auto simp: ad_agr_sets_def) then have "sp_equiv_list (map ((\e) \) ts) (map ((\e) \) ts)" using sp_equiv_list_fv[OF _ assms(5) sp] by auto moreover have "t \ set ts \ \i\fv_fo_terms_set ts. ad_equiv_pair AD (\ i, \ i) \ sp_equiv \ \ S \ ad_equiv_pair AD (\ \e t, \ \e t)" for t by (cases t) (auto simp: ad_equiv_pair.simps intro!: fv_fo_terms_setI) ultimately have ad_agr_list: "ad_agr_list AD (\ \e ts) (\ \e ts)" unfolding eval_eterms_def ad_agr_list_def ad_equiv_list_link[symmetric] using assms(1,3) by (auto simp: ad_agr_sets_def) show ?thesis by (rule ad_agr_list_comm[OF ad_agr_list_trans[OF ad_agr_list_comm[OF assms(4)] ad_agr_list]]) qed type_synonym ('a, 'c) fo_t = "'a set \ nat \ ('a + 'c) table" fun esat :: "('a, 'b) fo_fmla \ ('a table, 'b) fo_intp \ ('a + nat) val \ ('a + nat) set \ bool" where "esat (Pred r ts) I \ X \ \ \e ts \ map Inl ` I (r, length ts)" | "esat (Bool b) I \ X \ b" | "esat (Eqa t t') I \ X \ \ \e t = \ \e t'" | "esat (Neg \) I \ X \ \esat \ I \ X" | "esat (Conj \ \) I \ X \ esat \ I \ X \ esat \ I \ X" | "esat (Disj \ \) I \ X \ esat \ I \ X \ esat \ I \ X" | "esat (Exists n \) I \ X \ (\x \ X. esat \ I (\(n := x)) X)" | "esat (Forall n \) I \ X \ (\x \ X. esat \ I (\(n := x)) X)" fun sz_fmla :: "('a, 'b) fo_fmla \ nat" where "sz_fmla (Neg \) = Suc (sz_fmla \)" | "sz_fmla (Conj \ \) = Suc (sz_fmla \ + sz_fmla \)" | "sz_fmla (Disj \ \) = Suc (sz_fmla \ + sz_fmla \)" | "sz_fmla (Exists n \) = Suc (sz_fmla \)" | "sz_fmla (Forall n \) = Suc (Suc (Suc (Suc (sz_fmla \))))" | "sz_fmla _ = 0" lemma sz_fmla_induct[case_names Pred Bool Eqa Neg Conj Disj Exists Forall]: "(\r ts. P (Pred r ts)) \ (\b. P (Bool b)) \ (\t t'. P (Eqa t t')) \ (\\. P \ \ P (Neg \)) \ (\\ \. P \ \ P \ \ P (Conj \ \)) \ (\\ \. P \ \ P \ \ P (Disj \ \)) \ (\n \. P \ \ P (Exists n \)) \ (\n \. P (Exists n (Neg \)) \ P (Forall n \)) \ P \" proof (induction "sz_fmla \" arbitrary: \ rule: nat_less_induct) case 1 have IH: "\\. sz_fmla \ < sz_fmla \ \ P \" using 1 by auto then show ?case using 1(2,3,4,5,6,7,8,9) by (cases \) auto qed lemma esat_fv_cong: "(\n. n \ fv_fo_fmla \ \ \ n = \' n) \ esat \ I \ X \ esat \ I \' X" proof (induction \ arbitrary: \ \' rule: sz_fmla_induct) case (Pred r ts) then show ?case by (auto simp: eval_eterms_def fv_fo_terms_set_def) (smt comp_apply eval_eterm_cong fv_fo_term_set_cong image_insert insertCI map_eq_conv mk_disjoint_insert)+ next case (Eqa t t') then show ?case by (cases t; cases t') auto next case (Neg \) show ?case using Neg(1)[of \ \'] Neg(2) by auto next case (Conj \1 \2) show ?case using Conj(1,2)[of \ \'] Conj(3) by auto next case (Disj \1 \2) show ?case using Disj(1,2)[of \ \'] Disj(3) by auto next case (Exists n \) show ?case proof (rule iffI) assume "esat (Exists n \) I \ X" then obtain x where x_def: "x \ X" "esat \ I (\(n := x)) X" by auto from x_def(2) have "esat \ I (\'(n := x)) X" using Exists(1)[of "\(n := x)" "\'(n := x)"] Exists(2) by fastforce with x_def(1) show "esat (Exists n \) I \' X" by auto next assume "esat (Exists n \) I \' X" then obtain x where x_def: "x \ X" "esat \ I (\'(n := x)) X" by auto from x_def(2) have "esat \ I (\(n := x)) X" using Exists(1)[of "\(n := x)" "\'(n := x)"] Exists(2) by fastforce with x_def(1) show "esat (Exists n \) I \ X" by auto qed next case (Forall n \) then show ?case by auto qed auto fun ad_terms :: "('a fo_term) list \ 'a set" where "ad_terms ts = \(set (map set_fo_term ts))" fun act_edom :: "('a, 'b) fo_fmla \ ('a table, 'b) fo_intp \ 'a set" where "act_edom (Pred r ts) I = ad_terms ts \ \(set ` I (r, length ts))" | "act_edom (Bool b) I = {}" | "act_edom (Eqa t t') I = set_fo_term t \ set_fo_term t'" | "act_edom (Neg \) I = act_edom \ I" | "act_edom (Conj \ \) I = act_edom \ I \ act_edom \ I" | "act_edom (Disj \ \) I = act_edom \ I \ act_edom \ I" | "act_edom (Exists n \) I = act_edom \ I" | "act_edom (Forall n \) I = act_edom \ I" lemma finite_act_edom: "wf_fo_intp \ I \ finite (act_edom \ I)" using finite_Inl by (induction \ I rule: wf_fo_intp.induct) (auto simp: finite_set_fo_term vimage_def) fun fo_adom :: "('a, 'c) fo_t \ 'a set" where "fo_adom (AD, n, X) = AD" theorem main: "ad_agr \ AD \ \ \ act_edom \ I \ AD \ Inl ` AD \ Inr ` {..} \ X \ \ ` fv_fo_fmla \ \ X \ esat \ I \ UNIV \ esat \ I \ X" proof (induction \ arbitrary: \ \ rule: sz_fmla_induct) case (Pred r ts) have fv_sub: "fv_fo_terms_set ts \ fv_fo_fmla (Pred r ts)" by auto have sub_AD: "\(set ` I (r, length ts)) \ AD" using Pred(2) by auto show ?case unfolding esat.simps proof (rule iffI) assume assm: "\ \e ts \ map Inl ` I (r, length ts)" have "\ \e ts = \ \e ts" using esat_Pred[OF ad_agr_sets_mono[OF sub_AD Pred(1)[unfolded ad_agr_def]] fv_sub assm] by (auto simp: eval_eterms_def) with assm show "\ \e ts \ map Inl ` I (r, length ts)" by auto next assume assm: "\ \e ts \ map Inl ` I (r, length ts)" have "\ \e ts = \ \e ts" using esat_Pred[OF ad_agr_sets_comm[OF ad_agr_sets_mono[OF sub_AD Pred(1)[unfolded ad_agr_def]]] fv_sub assm] by (auto simp: eval_eterms_def) with assm show "\ \e ts \ map Inl ` I (r, length ts)" by auto qed next case (Eqa x1 x2) show ?case proof (cases x1; cases x2) fix c c' assume "x1 = Const c" "x2 = Const c'" with Eqa show ?thesis by auto next fix c m' assume assms: "x1 = Const c" "x2 = Var m'" with Eqa(1,2) have "\ m' = Inl c \ \ m' = Inl c" apply (auto simp: ad_agr_def ad_agr_sets_def) unfolding ad_equiv_pair.simps by fastforce+ with assms show ?thesis by fastforce next fix m c' assume assms: "x1 = Var m" "x2 = Const c'" with Eqa(1,2) have "\ m = Inl c' \ \ m = Inl c'" apply (auto simp: ad_agr_def ad_agr_sets_def) unfolding ad_equiv_pair.simps by fastforce+ with assms show ?thesis by auto next fix m m' assume assms: "x1 = Var m" "x2 = Var m'" with Eqa(1,2) have "\ m = \ m' \ \ m = \ m'" by (auto simp: ad_agr_def ad_agr_sets_def sp_equiv_def pairwise_def split: if_splits) with assms show ?thesis by auto qed next case (Neg \) from Neg(2) have "ad_agr \ AD \ \" by (auto simp: ad_agr_def) with Neg show ?case by auto next case (Conj \1 \2) have aux: "ad_agr \1 AD \ \" "ad_agr \2 AD \ \" "Inl ` AD \ Inr ` {..1} \ X" "Inl ` AD \ Inr ` {..2} \ X" "\ ` fv_fo_fmla \1 \ X" "\ ` fv_fo_fmla \2 \ X" using Conj(3,5,6) by (auto simp: ad_agr_def ad_agr_sets_def sp_equiv_def pairwise_def) show ?case using Conj(1)[OF aux(1) _ aux(3) aux(5)] Conj(2)[OF aux(2) _ aux(4) aux(6)] Conj(4) by auto next case (Disj \1 \2) have aux: "ad_agr \1 AD \ \" "ad_agr \2 AD \ \" "Inl ` AD \ Inr ` {..1} \ X" "Inl ` AD \ Inr ` {..2} \ X" "\ ` fv_fo_fmla \1 \ X" "\ ` fv_fo_fmla \2 \ X" using Disj(3,5,6) by (auto simp: ad_agr_def ad_agr_sets_def sp_equiv_def pairwise_def) show ?case using Disj(1)[OF aux(1) _ aux(3) aux(5)] Disj(2)[OF aux(2) _ aux(4) aux(6)] Disj(4) by auto next case (Exists m \) show ?case proof (rule iffI) assume "esat (Exists m \) I \ UNIV" then obtain x where assm: "esat \ I (\(m := x)) UNIV" by auto have "m \ SP \ \ Suc (card (Inr -` \ ` (SP \ - {m}))) \ card (SP \)" by (metis Diff_insert_absorb card_image card_le_Suc_iff finite_Diff finite_SP image_vimage_subset inj_Inr mk_disjoint_insert surj_card_le) moreover have "card (Inr -` \ ` SP \) \ card (SP \)" by (metis card_image finite_SP image_vimage_subset inj_Inr surj_card_le) ultimately have "max 1 (card (Inr -` \ ` (SP \ - {m})) + (if m \ SP \ then 1 else 0)) \ d \" using d_pos card_SP_d[of \] by auto then have "\x' \ X. ad_agr \ AD (\(m := x)) (\(m := x'))" using extend_\[OF Exists(2)[unfolded ad_agr_def fv_fo_fmla.simps SP.simps] SP_fv[of \] finite_SP Exists(5)[unfolded fv_fo_fmla.simps]] Exists(4) by (force simp: ad_agr_def) then obtain x' where x'_def: "x' \ X" "ad_agr \ AD (\(m := x)) (\(m := x'))" by auto from Exists(5) have "\(m := x') ` fv_fo_fmla \ \ X" using x'_def(1) by fastforce then have "esat \ I (\(m := x')) X" using Exists x'_def(1,2) assm by fastforce with x'_def show "esat (Exists m \) I \ X" by auto next assume "esat (Exists m \) I \ X" then obtain z where assm: "z \ X" "esat \ I (\(m := z)) X" by auto have ad_agr: "ad_agr_sets (fv_fo_fmla \ - {m}) (SP \ - {m}) AD \ \" using Exists(2)[unfolded ad_agr_def fv_fo_fmla.simps SP.simps] by (rule ad_agr_sets_comm) have "\x. ad_agr \ AD (\(m := x)) (\(m := z))" using extend_\[OF ad_agr SP_fv[of \] finite_SP subset_UNIV subset_UNIV] ad_agr_sets_comm unfolding ad_agr_def by fastforce then obtain x where x_def: "ad_agr \ AD (\(m := x)) (\(m := z))" by auto have "\(m := z) ` fv_fo_fmla (Exists m \) \ X" using Exists by fastforce with x_def have "esat \ I (\(m := x)) UNIV" using Exists assm by fastforce then show "esat (Exists m \) I \ UNIV" by auto qed next case (Forall n \) have unfold: "act_edom (Forall n \) I = act_edom (Exists n (Neg \)) I" "Inl ` AD \ Inr ` {..)} = Inl ` AD \ Inr ` {..))}" "fv_fo_fmla (Forall n \) = fv_fo_fmla (Exists n (Neg \))" by auto have pred: "ad_agr (Exists n (Neg \)) AD \ \" using Forall(2) by (auto simp: ad_agr_def) show ?case using Forall(1)[OF pred Forall(3,4,5)[unfolded unfold]] by auto qed auto lemma main_cor_inf: assumes "ad_agr \ AD \ \" "act_edom \ I \ AD" "d \ \ n" "\ ` fv_fo_fmla \ \ Inl ` AD \ Inr ` {.. I \ UNIV \ esat \ I \ (Inl ` AD \ Inr ` {.. :: "nat \ 'a + nat" assumes "ad_agr \ AD \ \" "act_edom \ I \ AD" shows "esat \ I \ UNIV \ esat \ I \ UNIV" proof - show ?thesis using main[OF assms(1,2) subset_UNIV subset_UNIV] by auto qed lemma esat_UNIV_ad_agr_list: fixes \ :: "nat \ 'a + nat" assumes "ad_agr_list AD (map \ (fv_fo_fmla_list \)) (map \ (fv_fo_fmla_list \))" "act_edom \ I \ AD" shows "esat \ I \ UNIV \ esat \ I \ UNIV" using esat_UNIV_cong[OF iffD2[OF ad_agr_def, OF ad_agr_sets_mono'[OF SP_fv], OF iffD2[OF ad_agr_list_link, OF assms(1), unfolded fv_fo_fmla_list_set]] assms(2)] . fun fo_rep :: "('a, 'c) fo_t \ 'a table" where "fo_rep (AD, n, X) = {ts. \ts' \ X. ad_agr_list AD (map Inl ts) ts'}" lemma sat_esat_conv: fixes \ :: "('a :: infinite, 'b) fo_fmla" assumes fin: "wf_fo_intp \ I" shows "sat \ I \ \ esat \ I (Inl \ \ :: nat \ 'a + nat) UNIV" using assms proof (induction \ arbitrary: I \ rule: sz_fmla_induct) case (Pred r ts) show ?case unfolding sat.simps esat.simps comp_def[symmetric] eval_terms_eterms[symmetric] by auto next case (Eqa t t') show ?case by (cases t; cases t') auto next case (Exists n \) show ?case proof (rule iffI) assume "sat (Exists n \) I \" then obtain x where x_def: "esat \ I (Inl \ \(n := x)) UNIV" using Exists by fastforce have Inl_unfold: "Inl \ \(n := x) = (Inl \ \)(n := Inl x)" by auto show "esat (Exists n \) I (Inl \ \) UNIV" using x_def unfolding Inl_unfold by auto next assume "esat (Exists n \) I (Inl \ \) UNIV" then obtain x where x_def: "esat \ I ((Inl \ \)(n := x)) UNIV" by auto show "sat (Exists n \) I \" proof (cases x) case (Inl a) have Inl_unfold: "(Inl \ \)(n := x) = Inl \ \(n := a)" by (auto simp: Inl) show ?thesis using x_def[unfolded Inl_unfold] Exists by fastforce next case (Inr b) obtain c where c_def: "c \ act_edom \ I \ \ ` fv_fo_fmla \" using arb_element finite_act_edom[OF Exists(2), simplified] finite_fv_fo_fmla by (metis finite_Un finite_imageI) have wf_local: "wf_fo_intp \ I" using Exists(2) by auto have "(a, a') \ set (zip (map (\x. if x = n then Inr b else (Inl \ \) x) (fv_fo_fmla_list \)) (map (\a. Inl (if a = n then c else \ a)) (fv_fo_fmla_list \))) \ ad_equiv_pair (act_edom \ I) (a, a')" for a a' using c_def by (cases a; cases a') (auto simp: set_zip ad_equiv_pair.simps split: if_splits) then have "sat \ I (\(n := c))" using c_def[folded fv_fo_fmla_list_set] by (auto simp: ad_agr_list_def ad_equiv_list_def fun_upd_def sp_equiv_list_def pairwise_def set_zip split: if_splits intro!: Exists(1)[OF wf_local, THEN iffD2, OF esat_UNIV_ad_agr_list[OF _ subset_refl, THEN iffD1, OF _ x_def[unfolded Inr]]]) then show ?thesis by auto qed qed next case (Forall n \) show ?case using Forall(1)[of I \] Forall(2) by auto qed auto lemma sat_ad_agr_list: fixes \ :: "('a :: infinite, 'b) fo_fmla" and J :: "(('a, nat) fo_t, 'b) fo_intp" assumes "wf_fo_intp \ I" "ad_agr_list AD (map (Inl \ \ :: nat \ 'a + nat) (fv_fo_fmla_list \)) (map (Inl \ \) (fv_fo_fmla_list \))" "act_edom \ I \ AD" shows "sat \ I \ \ sat \ I \" using esat_UNIV_ad_agr_list[OF assms(2,3)] sat_esat_conv[OF assms(1)] by auto definition nfv :: "('a, 'b) fo_fmla \ nat" where "nfv \ = length (fv_fo_fmla_list \)" lemma nfv_card: "nfv \ = card (fv_fo_fmla \)" proof - have "distinct (fv_fo_fmla_list \)" using sorted_distinct_fv_list by auto then have "length (fv_fo_fmla_list \) = card (set (fv_fo_fmla_list \))" using distinct_card by fastforce then show ?thesis unfolding fv_fo_fmla_list_set by (auto simp: nfv_def) qed fun rremdups :: "'a list \ 'a list" where "rremdups [] = []" | "rremdups (x # xs) = x # rremdups (filter ((\) x) xs)" lemma filter_rremdups_filter: "filter P (rremdups (filter Q xs)) = rremdups (filter (\x. P x \ Q x) xs)" apply (induction xs arbitrary: Q) apply auto by metis lemma filter_rremdups: "filter P (rremdups xs) = rremdups (filter P xs)" using filter_rremdups_filter[where Q="\_. True"] by auto lemma filter_take: "\j. filter P (take i xs) = take j (filter P xs)" apply (induction xs arbitrary: i) apply (auto) apply (metis filter.simps(1) filter.simps(2) take_Cons' take_Suc_Cons) apply (metis filter.simps(2) take0 take_Cons') done lemma rremdups_take: "\j. rremdups (take i xs) = take j (rremdups xs)" proof (induction xs arbitrary: i) case (Cons x xs) show ?case proof (cases i) case (Suc n) obtain j where j_def: "rremdups (take n xs) = take j (rremdups xs)" using Cons by auto obtain j' where j'_def: "filter ((\) x) (take j (rremdups xs)) = take j' (filter ((\) x) (rremdups xs))" using filter_take by blast show ?thesis by (auto simp: Suc filter_rremdups[symmetric] j_def j'_def intro: exI[of _ "Suc j'"]) qed (auto simp add: take_Cons') qed auto lemma rremdups_app: "rremdups (xs @ [x]) = rremdups xs @ (if x \ set xs then [] else [x])" apply (induction xs) apply auto apply (smt filter.simps(1) filter.simps(2) filter_append filter_rremdups)+ done lemma rremdups_set: "set (rremdups xs) = set xs" by (induction xs) (auto simp: filter_rremdups[symmetric]) lemma distinct_rremdups: "distinct (rremdups xs)" proof (induction "length xs" arbitrary: xs rule: nat_less_induct) case 1 then have IH: "\m ys. length (ys :: 'a list) < length xs \ distinct (rremdups ys)" by auto show ?case proof (cases xs) case (Cons z zs) show ?thesis using IH by (auto simp: Cons rremdups_set le_imp_less_Suc) qed auto qed lemma length_rremdups: "length (rremdups xs) = card (set xs)" using distinct_card[OF distinct_rremdups] by (subst eq_commute) (auto simp: rremdups_set) lemma set_map_filter_sum: "set (List.map_filter (case_sum Map.empty Some) xs) = Inr -` set xs" by (induction xs) (auto simp: List.map_filter_simps split: sum.splits) definition nats :: "nat list \ bool" where "nats ns = (ns = [0.. ('a + nat) list \ bool" where "fo_nmlzd AD xs \ Inl -` set xs \ AD \ (let ns = List.map_filter (case_sum Map.empty Some) xs in nats (rremdups ns))" lemma fo_nmlzd_all_AD: assumes "set xs \ Inl ` AD" shows "fo_nmlzd AD xs" proof - have "List.map_filter (case_sum Map.empty Some) xs = []" using assms by (induction xs) (auto simp: List.map_filter_simps) then show ?thesis using assms by (auto simp: fo_nmlzd_def nats_def Let_def) qed lemma card_Inr_vimage_le_length: "card (Inr -` set xs) \ length xs" proof - have "card (Inr -` set xs) \ card (set xs)" by (meson List.finite_set card_inj_on_le image_vimage_subset inj_Inr) moreover have "\ \ length xs" by (rule card_length) finally show ?thesis . qed lemma fo_nmlzd_set: assumes "fo_nmlzd AD xs" shows "set xs = set xs \ Inl ` AD \ Inr ` {.. AD" using assms by (auto simp: fo_nmlzd_def) moreover have "Inr -` set xs = {.. Inl ` AD \ Inr ` {..j. List.map_filter f (take i xs) = take j (List.map_filter f xs)" apply (induction xs arbitrary: i) apply (auto simp: List.map_filter_simps split: option.splits) apply (metis map_filter_simps(1) option.case(1) take0 take_Cons') apply (metis map_filter_simps(1) map_filter_simps(2) option.case(2) take_Cons' take_Suc_Cons) done lemma fo_nmlzd_take: assumes "fo_nmlzd AD xs" shows "fo_nmlzd AD (take i xs)" proof - have aux: "rremdups zs = [0.. rremdups (take j zs) = [0.. [y] | _ \ [])" by (induction xs) (auto simp: List.map_filter_simps split: option.splits) lemma fo_nmlzd_app_Inr: "Inr n \ set xs \ Inr n' \ set xs \ fo_nmlzd AD (xs @ [Inr n]) \ fo_nmlzd AD (xs @ [Inr n']) \ n = n'" by (auto simp: List.map_filter_simps fo_nmlzd_def nats_def Let_def map_filter_app rremdups_app set_map_filter_sum) fun all_tuples :: "'c set \ nat \ 'c table" where "all_tuples xs 0 = {[]}" | "all_tuples xs (Suc n) = \((\as. (\x. x # as) ` xs) ` (all_tuples xs n))" definition nall_tuples :: "'a set \ nat \ ('a + nat) table" where "nall_tuples AD n = {zs \ all_tuples (Inl ` AD \ Inr ` {.. finite (all_tuples xs n)" by (induction xs n rule: all_tuples.induct) auto lemma nall_tuples_finite: "finite AD \ finite (nall_tuples AD n)" by (auto simp: nall_tuples_def all_tuples_finite) lemma all_tuplesI: "length vs = n \ set vs \ xs \ vs \ all_tuples xs n" proof (induction xs n arbitrary: vs rule: all_tuples.induct) case (2 xs n) then obtain w ws where "vs = w # ws" "length ws = n" "set ws \ xs" "w \ xs" by (metis Suc_length_conv contra_subsetD list.set_intros(1) order_trans set_subset_Cons) with 2(1) show ?case by auto qed auto lemma nall_tuplesI: "length vs = n \ fo_nmlzd AD vs \ vs \ nall_tuples AD n" using fo_nmlzd_set[of AD vs] by (auto simp: nall_tuples_def intro!: all_tuplesI) lemma all_tuplesD: "vs \ all_tuples xs n \ length vs = n \ set vs \ xs" by (induction xs n arbitrary: vs rule: all_tuples.induct) auto+ lemma all_tuples_setD: "vs \ all_tuples xs n \ set vs \ xs" by (auto dest: all_tuplesD) lemma nall_tuplesD: "vs \ nall_tuples AD n \ length vs = n \ set vs \ Inl ` AD \ Inr ` {.. fo_nmlzd AD vs" by (auto simp: nall_tuples_def dest: all_tuplesD) lemma all_tuples_set: "all_tuples xs n = {ys. length ys = n \ set ys \ xs}" proof (induction xs n rule: all_tuples.induct) case (2 xs n) show ?case proof (rule subset_antisym; rule subsetI) fix ys assume "ys \ all_tuples xs (Suc n)" then show "ys \ {ys. length ys = Suc n \ set ys \ xs}" using 2 by auto next fix ys assume "ys \ {ys. length ys = Suc n \ set ys \ xs}" then have assm: "length ys = Suc n" "set ys \ xs" by auto then obtain z zs where zs_def: "ys = z # zs" "z \ xs" "length zs = n" "set zs \ xs" by (cases ys) auto with 2 have "zs \ all_tuples xs n" by auto with zs_def(1,2) show "ys \ all_tuples xs (Suc n)" by auto qed qed auto lemma nall_tuples_set: "nall_tuples AD n = {ys. length ys = n \ fo_nmlzd AD ys}" using fo_nmlzd_set[of AD] card_Inr_vimage_le_length by (auto simp: nall_tuples_def all_tuples_set) (smt UnE nall_tuplesD nall_tuplesI subsetD) fun pos :: "'a \ 'a list \ nat option" where "pos a [] = None" | "pos a (x # xs) = (if a = x then Some 0 else (case pos a xs of Some n \ Some (Suc n) | _ \ None))" lemma pos_set: "pos a xs = Some i \ a \ set xs" by (induction a xs arbitrary: i rule: pos.induct) (auto split: if_splits option.splits) lemma pos_length: "pos a xs = Some i \ i < length xs" by (induction a xs arbitrary: i rule: pos.induct) (auto split: if_splits option.splits) lemma pos_sound: "pos a xs = Some i \ i < length xs \ xs ! i = a" by (induction a xs arbitrary: i rule: pos.induct) (auto split: if_splits option.splits) lemma pos_complete: "pos a xs = None \ a \ set xs" by (induction a xs rule: pos.induct) (auto split: if_splits option.splits) fun rem_nth :: "nat \ 'a list \ 'a list" where "rem_nth _ [] = []" | "rem_nth 0 (x # xs) = xs" | "rem_nth (Suc n) (x # xs) = x # rem_nth n xs" lemma rem_nth_length: "i < length xs \ length (rem_nth i xs) = length xs - 1" by (induction i xs rule: rem_nth.induct) auto lemma rem_nth_take_drop: "i < length xs \ rem_nth i xs = take i xs @ drop (Suc i) xs" by (induction i xs rule: rem_nth.induct) auto lemma rem_nth_sound: "distinct xs \ pos n xs = Some i \ rem_nth i (map \ xs) = map \ (filter ((\) n) xs)" apply (induction xs arbitrary: i) apply (auto simp: pos_set split: option.splits) by (metis (mono_tags, lifting) filter_True) fun add_nth :: "nat \ 'a \ 'a list \ 'a list" where "add_nth 0 a xs = a # xs" | "add_nth (Suc n) a zs = (case zs of x # xs \ x # add_nth n a xs)" lemma add_nth_length: "i \ length zs \ length (add_nth i z zs) = Suc (length zs)" by (induction i z zs rule: add_nth.induct) (auto split: list.splits) lemma add_nth_take_drop: "i \ length zs \ add_nth i v zs = take i zs @ v # drop i zs" by (induction i v zs rule: add_nth.induct) (auto split: list.splits) lemma add_nth_rem_nth_map: "distinct xs \ pos n xs = Some i \ add_nth i a (rem_nth i (map \ xs)) = map (\(n := a)) xs" by (induction xs arbitrary: i) (auto simp: pos_set split: option.splits) lemma add_nth_rem_nth_self: "i < length xs \ add_nth i (xs ! i) (rem_nth i xs) = xs" by (induction i xs rule: rem_nth.induct) auto lemma rem_nth_add_nth: "i \ length zs \ rem_nth i (add_nth i z zs) = zs" by (induction i z zs rule: add_nth.induct) (auto split: list.splits) fun merge :: "(nat \ 'a) list \ (nat \ 'a) list \ (nat \ 'a) list" where "merge [] mys = mys" | "merge nxs [] = nxs" | "merge ((n, x) # nxs) ((m, y) # mys) = (if n \ m then (n, x) # merge nxs ((m, y) # mys) else (m, y) # merge ((n, x) # nxs) mys)" lemma merge_Nil2[simp]: "merge nxs [] = nxs" by (cases nxs) auto lemma merge_length: "length (merge nxs mys) = length (map fst nxs @ map fst mys)" by (induction nxs mys rule: merge.induct) auto lemma insort_aux_le: "\x\set nxs. n \ fst x \ \x\set mys. m \ fst x \ n \ m \ insort n (sort (map fst nxs @ m # map fst mys)) = n # sort (map fst nxs @ m # map fst mys)" by (induction nxs) (auto simp: insort_is_Cons insort_left_comm) lemma insort_aux_gt: "\x\set nxs. n \ fst x \ \x\set mys. m \ fst x \ \ n \ m \ insort n (sort (map fst nxs @ m # map fst mys)) = m # insort n (sort (map fst nxs @ map fst mys))" apply (induction nxs) apply (auto simp: insort_is_Cons) by (metis dual_order.trans insort_key.simps(2) insort_left_comm) lemma map_fst_merge: "sorted_distinct (map fst nxs) \ sorted_distinct (map fst mys) \ map fst (merge nxs mys) = sort (map fst nxs @ map fst mys)" by (induction nxs mys rule: merge.induct) (auto simp add: sorted_sort_id insort_is_Cons insort_aux_le insort_aux_gt) lemma merge_map': "sorted_distinct (map fst nxs) \ sorted_distinct (map fst mys) \ fst ` set nxs \ fst ` set mys = {} \ map snd nxs = map \ (map fst nxs) \ map snd mys = map \ (map fst mys) \ map snd (merge nxs mys) = map \ (sort (map fst nxs @ map fst mys))" by (induction nxs mys rule: merge.induct) (auto simp: sorted_sort_id insort_is_Cons insort_aux_le insort_aux_gt) lemma merge_map: "sorted_distinct ns \ sorted_distinct ms \ set ns \ set ms = {} \ map snd (merge (zip ns (map \ ns)) (zip ms (map \ ms))) = map \ (sort (ns @ ms))" using merge_map'[of "zip ns (map \ ns)" "zip ms (map \ ms)" \] by auto (metis length_map list.set_map map_fst_zip) fun fo_nmlz_rec :: "nat \ ('a + nat \ nat) \ 'a set \ ('a + nat) list \ ('a + nat) list" where "fo_nmlz_rec i m AD [] = []" | "fo_nmlz_rec i m AD (Inl x # xs) = (if x \ AD then Inl x # fo_nmlz_rec i m AD xs else (case m (Inl x) of None \ Inr i # fo_nmlz_rec (Suc i) (m(Inl x \ i)) AD xs | Some j \ Inr j # fo_nmlz_rec i m AD xs))" | "fo_nmlz_rec i m AD (Inr n # xs) = (case m (Inr n) of None \ Inr i # fo_nmlz_rec (Suc i) (m(Inr n \ i)) AD xs | Some j \ Inr j # fo_nmlz_rec i m AD xs)" lemma fo_nmlz_rec_sound: "ran m \ {.. filter ((\) i) (rremdups (List.map_filter (case_sum Map.empty Some) (fo_nmlz_rec i m AD xs))) = ns \ ns = [i.. AD") case False show ?thesis proof (cases "m (Inl x)") case None have pred: "ran (m(Inl x \ i)) \ {..) (Suc i)) (rremdups (List.map_filter (case_sum Map.empty Some) (fo_nmlz_rec (Suc i) (m(Inl x \ i)) AD xs)))" using 2(5) False None by (auto simp: List.map_filter_simps filter_rremdups) (metis Suc_leD antisym not_less_eq_eq) then show ?thesis by (auto simp: 2(2)[OF False None pred, OF refl]) (smt Suc_le_eq Suc_pred le_add1 le_zero_eq less_add_same_cancel1 not_less_eq_eq upt_Suc_append upt_rec) next case (Some j) then have j_lt_i: "j < i" using 2(4) by (auto simp: ran_def) have ns_def: "ns = filter ((\) i) (rremdups (List.map_filter (case_sum Map.empty Some) (fo_nmlz_rec i m AD xs)))" using 2(5) False Some j_lt_i by (auto simp: List.map_filter_simps filter_rremdups) (metis leD) show ?thesis by (rule 2(3)[OF False Some 2(4) ns_def[symmetric]]) qed qed (auto simp: List.map_filter_simps split: option.splits) next case (3 i m AD n xs) show ?case proof (cases "m (Inr n)") case None have pred: "ran (m(Inr n \ i)) \ {..) (Suc i)) (rremdups (List.map_filter (case_sum Map.empty Some) (fo_nmlz_rec (Suc i) (m(Inr n \ i)) AD xs)))" using 3(4) None by (auto simp: List.map_filter_simps filter_rremdups) (metis Suc_leD antisym not_less_eq_eq) then show ?thesis by (auto simp add: 3(1)[OF None pred, OF refl]) (smt Suc_le_eq Suc_pred le_add1 le_zero_eq less_add_same_cancel1 not_less_eq_eq upt_Suc_append upt_rec) next case (Some j) then have j_lt_i: "j < i" using 3(3) by (auto simp: ran_def) have ns_def: "ns = filter ((\) i) (rremdups (List.map_filter (case_sum Map.empty Some) (fo_nmlz_rec i m AD xs)))" using 3(4) Some j_lt_i by (auto simp: List.map_filter_simps filter_rremdups) (metis leD) show ?thesis by (rule 3(2)[OF Some 3(3) ns_def[symmetric]]) qed qed (auto simp: List.map_filter_simps) definition id_map :: "nat \ ('a + nat \ nat)" where "id_map n = (\x. case x of Inl x \ None | Inr x \ if x < n then Some x else None)" lemma fo_nmlz_rec_idem: "Inl -` set ys \ AD \ rremdups (List.map_filter (case_sum Map.empty Some) ys) = ns \ set (filter (\n. n < i) ns) \ {.. filter ((\) i) ns = [i.. fo_nmlz_rec i (id_map i) AD ys = ys" proof (induction ys arbitrary: i k ns) case (Cons y ys) show ?case proof (cases y) case (Inl a) show ?thesis using Cons(1)[OF _ _ Cons(4,5)] Cons(2,3) by (auto simp: Inl List.map_filter_simps) next case (Inr j) show ?thesis proof (cases "j < i") case False have j_i: "j = i" using False Cons(3,5) by (auto simp: Inr List.map_filter_simps filter_rremdups in_mono split: if_splits) (metis (no_types, lifting) upt_eq_Cons_conv) obtain kk where k_def: "k = Suc kk" using Cons(3,5) by (cases k) (auto simp: Inr List.map_filter_simps j_i) define ns' where "ns' = rremdups (List.map_filter (case_sum Map.empty Some) ys)" have id_map_None: "id_map i (Inr i) = None" by (auto simp: id_map_def) - have id_map_upd: "id_map i(Inr i \ i) = id_map (Suc i)" + have id_map_upd: "(id_map i)(Inr i \ i) = id_map (Suc i)" by (auto simp: id_map_def split: sum.splits) have "set (filter (\n. n < Suc i) ns') \ {..) (Suc i)) ns' = [Suc i.. AD" using Cons(2) by (auto simp: vimage_def) ultimately have "fo_nmlz_rec (Suc i) ((id_map i)(Inr i \ i)) AD ys = ys" using Cons(1)[OF _ ns'_def[symmetric], of "Suc i" kk] by (auto simp: ns'_def k_def id_map_upd split: if_splits) then show ?thesis by (auto simp: Inr j_i id_map_None) next case True define ns' where "ns' = rremdups (List.map_filter (case_sum Map.empty Some) ys)" have "set (filter (\y. y < i) ns') \ set (filter (\y. y < i) ns)" "filter ((\) i) ns' = filter ((\) i) ns" using Cons(3) True by (auto simp: Inr List.map_filter_simps filter_rremdups[symmetric] ns'_def[symmetric]) (smt filter_cong leD) then have "fo_nmlz_rec i (id_map i) AD ys = ys" using Cons(1)[OF _ ns'_def[symmetric]] Cons(3,5) Cons(2) by (auto simp: vimage_def) then show ?thesis using True by (auto simp: Inr id_map_def) qed qed qed (auto simp: List.map_filter_simps intro!: exI[of _ "[]"]) lemma fo_nmlz_rec_length: "length (fo_nmlz_rec i m AD xs) = length xs" by (induction i m AD xs rule: fo_nmlz_rec.induct) (auto simp: fun_upd_def split: option.splits) lemma insert_Inr: "\X. insert (Inr i) (X \ Inr ` {.. Inr ` {.. {.. set (fo_nmlz_rec i m AD xs) \ Inr ` {.. Inl ` AD \ Inr ` {.. AD") case True have "card (set (Inl x # xs) - Inl ` AD - dom m) = card (set xs - Inl ` AD - dom m)" using True by auto then show ?thesis using 2(1)[OF True 2(4)] True by auto next case False show ?thesis proof (cases "m (Inl x)") case None have pred: "ran (m(Inl x \ i)) \ {.. (set xs - Inl ` AD - dom (m(Inl x \ i)))" using None False by (auto simp: dom_def) then have Suc: "Suc i + card (set xs - Inl ` AD - dom (m(Inl x \ i))) = i + card (set (Inl x # xs) - Inl ` AD - dom m)" using None by auto show ?thesis using 2(2)[OF False None pred] False None unfolding Suc by (auto simp: fun_upd_def[symmetric] insert_Inr) next case (Some j) then have j_lt_i: "j < i" using 2(4) by (auto simp: ran_def) have "card (set (Inl x # xs) - Inl ` AD - dom m) = card (set xs - Inl ` AD - dom m)" by (auto simp: Some intro: arg_cong[of _ _ card]) then show ?thesis using 2(3)[OF False Some 2(4)] False Some j_lt_i by auto qed qed next case (3 i m AD k xs) then show ?case proof (cases "m (Inr k)") case None have preds: "ran (m(Inr k \ i)) \ {.. (set xs - Inl ` AD - dom (m(Inr k \ i)))" using None by (auto simp: dom_def) then have Suc: "Suc i + card (set xs - Inl ` AD - dom (m(Inr k \ i))) = i + card (set (Inr k # xs) - Inl ` AD - dom m)" using None by auto show ?thesis using None 3(1)[OF None preds] unfolding Suc by (auto simp: fun_upd_def[symmetric] insert_Inr) next case (Some j) have fin: "finite (set (Inr k # xs) - Inl ` AD - dom m)" by auto have card_eq: "card (set xs - Inl ` AD - dom m) = card (set (Inr k # xs) - Inl ` AD - dom m)" by (auto simp: Some intro!: arg_cong[of _ _ card]) have j_lt_i: "j < i" using 3(3) Some by (auto simp: ran_def) show ?thesis using 3(2)[OF Some 3(3)] j_lt_i unfolding card_eq by (auto simp: ran_def insert_Inr Some) qed qed auto lemma fo_nmlz_rec_set_rev: "set (fo_nmlz_rec i m AD xs) \ Inl ` AD \ set xs \ Inl ` AD" by (induction i m AD xs rule: fo_nmlz_rec.induct) (auto split: if_splits option.splits) lemma fo_nmlz_rec_map: "inj_on m (dom m) \ ran m \ {.. \m'. inj_on m' (dom m') \ (\n. m n \ None \ m' n = m n) \ (\(x, y) \ set (zip xs (fo_nmlz_rec i m AD xs)). (case x of Inl x' \ if x' \ AD then x = y else \j. m' (Inl x') = Some j \ y = Inr j | Inr n \ \j. m' (Inr n) = Some j \ y = Inr j))" proof (induction i m AD xs rule: fo_nmlz_rec.induct) case (2 i m AD x xs) show ?case using 2(1)[OF _ 2(4,5)] proof (cases "x \ AD") case False show ?thesis proof (cases "m (Inl x)") case None have preds: "inj_on (m(Inl x \ i)) (dom (m(Inl x \ i)))" "ran (m(Inl x \ i)) \ {.. i)) (dom (m(Inr n \ i)))" "ran (m(Inr n \ i)) \ {..x y. (x, y) \ set (zip xs ys) \ (case x of Inl x' \ if x' \ AD then x = y else m x = Some y \ (case y of Inl z \ z \ AD | Inr _ \ True) | Inr n \ m x = Some y \ (case y of Inl z \ z \ AD | Inr _ \ True))" shows "ad_agr_list AD xs ys" proof - have "ad_equiv_pair AD (a, b)" if "(a, b) \ set (zip xs ys)" for a b unfolding ad_equiv_pair.simps using assms(3)[OF that] by (auto split: sum.splits if_splits) moreover have "False" if "(a, c) \ set (zip xs ys)" "(b, c) \ set (zip xs ys)" "a \ b" for a b c using assms(3)[OF that(1)] assms(3)[OF that(2)] assms(2) that(3) by (auto split: sum.splits if_splits) (metis domI inj_onD that(3))+ moreover have "False" if "(a, b) \ set (zip xs ys)" "(a, c) \ set (zip xs ys)" "b \ c" for a b c using assms(3)[OF that(1)] assms(3)[OF that(2)] assms(2) that(3) by (auto split: sum.splits if_splits) ultimately show ?thesis using assms by (fastforce simp: ad_agr_list_def ad_equiv_list_def sp_equiv_list_def pairwise_def) qed lemma fo_nmlz_rec_take: "take n (fo_nmlz_rec i m AD xs) = fo_nmlz_rec i m AD (take n xs)" by (induction i m AD xs arbitrary: n rule: fo_nmlz_rec.induct) (auto simp: take_Cons' split: option.splits) definition fo_nmlz :: "'a set \ ('a + nat) list \ ('a + nat) list" where "fo_nmlz = fo_nmlz_rec 0 Map.empty" lemma fo_nmlz_Nil[simp]: "fo_nmlz AD [] = []" by (auto simp: fo_nmlz_def) lemma fo_nmlz_Cons: "fo_nmlz AD [x] = (case x of Inl x \ if x \ AD then [Inl x] else [Inr 0] | _ \ [Inr 0])" by (auto simp: fo_nmlz_def split: sum.splits) lemma fo_nmlz_Cons_Cons: "fo_nmlz AD [x, x] = (case x of Inl x \ if x \ AD then [Inl x, Inl x] else [Inr 0, Inr 0] | _ \ [Inr 0, Inr 0])" by (auto simp: fo_nmlz_def split: sum.splits) lemma fo_nmlz_sound: "fo_nmlzd AD (fo_nmlz AD xs)" using fo_nmlz_rec_sound[of Map.empty 0] fo_nmlz_rec_set[of Map.empty 0 AD xs] by (auto simp: fo_nmlzd_def fo_nmlz_def nats_def Let_def) lemma fo_nmlz_length: "length (fo_nmlz AD xs) = length xs" using fo_nmlz_rec_length by (auto simp: fo_nmlz_def) lemma fo_nmlz_map: "\\. fo_nmlz AD (map \ ns) = map \ ns" proof - obtain m' where m'_def: "\(x, y)\set (zip (map \ ns) (fo_nmlz AD (map \ ns))). case x of Inl x' \ if x' \ AD then x = y else \j. m' (Inl x') = Some j \ y = Inr j | Inr n \ \j. m' (Inr n) = Some j \ y = Inr j" using fo_nmlz_rec_map[of Map.empty 0, of "map \ ns"] by (auto simp: fo_nmlz_def) define \ where "\ \ (\n. case \ n of Inl x \ if x \ AD then Inl x else Inr (the (m' (Inl x))) | Inr j \ Inr (the (m' (Inr j))))" have "fo_nmlz AD (map \ ns) = map \ ns" proof (rule nth_equalityI) show "length (fo_nmlz AD (map \ ns)) = length (map \ ns)" using fo_nmlz_length[of AD "map \ ns"] by auto fix i assume "i < length (fo_nmlz AD (map \ ns))" then show "fo_nmlz AD (map \ ns) ! i = map \ ns ! i" using m'_def fo_nmlz_length[of AD "map \ ns"] apply (auto simp: set_zip \_def split: sum.splits) apply (metis nth_map) apply (metis nth_map option.sel)+ done qed then show ?thesis by auto qed lemma card_set_minus: "card (set xs - X) \ length xs" by (meson Diff_subset List.finite_set card_length card_mono order_trans) lemma fo_nmlz_set: "set (fo_nmlz AD xs) = set xs \ Inl ` AD \ Inr ` {.. Inl ` AD \ set xs \ Inl ` AD" using fo_nmlz_rec_set_rev[of 0 Map.empty AD xs] by (auto simp: fo_nmlz_def) lemma inj_on_empty: "inj_on Map.empty (dom Map.empty)" and ran_empty_upto: "ran Map.empty \ {..<0}" by auto lemma fo_nmlz_ad_agr: "ad_agr_list AD xs (fo_nmlz AD xs)" using fo_nmlz_rec_map[OF inj_on_empty ran_empty_upto, of xs AD] unfolding fo_nmlz_def apply safe subgoal for m' by (fastforce simp: inj_on_def dom_def split: sum.splits if_splits intro!: ad_agr_map[OF fo_nmlz_rec_length[symmetric], of "map_option Inr \ m'"]) done lemma fo_nmlzd_mono: "Inl -` set xs \ AD \ fo_nmlzd AD' xs \ fo_nmlzd AD xs" by (auto simp: fo_nmlzd_def) lemma fo_nmlz_idem: "fo_nmlzd AD ys \ fo_nmlz AD ys = ys" using fo_nmlz_rec_idem[where ?i=0] by (auto simp: fo_nmlzd_def fo_nmlz_def id_map_def nats_def Let_def) lemma fo_nmlz_take: "take n (fo_nmlz AD xs) = fo_nmlz AD (take n xs)" using fo_nmlz_rec_take by (auto simp: fo_nmlz_def) fun nall_tuples_rec :: "'a set \ nat \ nat \ ('a + nat) table" where "nall_tuples_rec AD i 0 = {[]}" | "nall_tuples_rec AD i (Suc n) = \((\as. (\x. x # as) ` (Inl ` AD \ Inr ` {.. (\as. Inr i # as) ` nall_tuples_rec AD (Suc i) n" lemma nall_tuples_rec_Inl: "vs \ nall_tuples_rec AD i n \ Inl -` set vs \ AD" by (induction AD i n arbitrary: vs rule: nall_tuples_rec.induct) (fastforce simp: vimage_def)+ lemma nall_tuples_rec_length: "xs \ nall_tuples_rec AD i n \ length xs = n" by (induction AD i n arbitrary: xs rule: nall_tuples_rec.induct) auto -lemma fun_upd_id_map: "id_map i(Inr i \ i) = id_map (Suc i)" +lemma fun_upd_id_map: "(id_map i)(Inr i \ i) = id_map (Suc i)" by (rule ext) (auto simp: id_map_def split: sum.splits) lemma id_mapD: "id_map j (Inr i) = None \ j \ i" "id_map j (Inr i) = Some x \ i < j \ i = x" by (auto simp: id_map_def split: if_splits) lemma nall_tuples_rec_fo_nmlz_rec_sound: "i \ j \ xs \ nall_tuples_rec AD i n \ fo_nmlz_rec j (id_map j) AD xs = xs" apply (induction n arbitrary: i j xs) apply (auto simp: fun_upd_id_map dest!: id_mapD split: option.splits) apply (meson dual_order.strict_trans2 id_mapD(1) not_Some_eq sup.strict_order_iff) using Suc_leI apply blast+ done lemma nall_tuples_rec_fo_nmlz_rec_complete: assumes "fo_nmlz_rec j (id_map j) AD xs = xs" shows "xs \ nall_tuples_rec AD j (length xs)" using assms proof (induction xs arbitrary: j) case (Cons x xs) show ?case proof (cases x) case (Inl a) have a_AD: "a \ AD" using Cons(2) by (auto simp: Inl split: if_splits option.splits) show ?thesis using Cons a_AD by (auto simp: Inl) next case (Inr b) have b_j: "b \ j" using Cons(2) by (auto simp: Inr split: option.splits dest: id_mapD) show ?thesis proof (cases "b = j") case True have preds: "fo_nmlz_rec (Suc j) (id_map (Suc j)) AD xs = xs" using Cons(2) by (auto simp: Inr True fun_upd_id_map dest: id_mapD split: option.splits) show ?thesis using Cons(1)[OF preds] by (auto simp: Inr True) next case False have b_lt_j: "b < j" using b_j False by auto have id_map: "id_map j (Inr b) = Some b" using b_lt_j by (auto simp: id_map_def) have preds: "fo_nmlz_rec j (id_map j) AD xs = xs" using Cons(2) by (auto simp: Inr id_map) show ?thesis using Cons(1)[OF preds] b_lt_j by (auto simp: Inr) qed qed qed auto lemma nall_tuples_rec_fo_nmlz: "xs \ nall_tuples_rec AD 0 (length xs) \ fo_nmlz AD xs = xs" using nall_tuples_rec_fo_nmlz_rec_sound[of 0 0 xs AD "length xs"] nall_tuples_rec_fo_nmlz_rec_complete[of 0 AD xs] by (auto simp: fo_nmlz_def id_map_def) lemma fo_nmlzd_code[code]: "fo_nmlzd AD xs \ fo_nmlz AD xs = xs" using fo_nmlz_idem fo_nmlz_sound by metis lemma nall_tuples_code[code]: "nall_tuples AD n = nall_tuples_rec AD 0 n" unfolding nall_tuples_set using nall_tuples_rec_length trans[OF nall_tuples_rec_fo_nmlz fo_nmlzd_code[symmetric]] by fastforce lemma exists_map: "length xs = length ys \ distinct xs \ \f. ys = map f xs" proof (induction xs ys rule: list_induct2) case (Cons x xs y ys) then obtain f where f_def: "ys = map f xs" by auto with Cons(3) have "y # ys = map (f(x := y)) (x # xs)" by auto then show ?case by metis qed auto lemma exists_fo_nmlzd: assumes "length xs = length ys" "distinct xs" "fo_nmlzd AD ys" shows "\f. ys = fo_nmlz AD (map f xs)" using fo_nmlz_idem[OF assms(3)] exists_map[OF _ assms(2)] assms(1) by metis lemma list_induct2_rev[consumes 1]: "length xs = length ys \ (P [] []) \ (\x y xs ys. P xs ys \ P (xs @ [x]) (ys @ [y])) \ P xs ys" proof (induction "length xs" arbitrary: xs ys) case (Suc n) then show ?case by (cases xs rule: rev_cases; cases ys rule: rev_cases) auto qed auto lemma ad_agr_list_fo_nmlzd: assumes "ad_agr_list AD vs vs'" "fo_nmlzd AD vs" "fo_nmlzd AD vs'" shows "vs = vs'" using ad_agr_list_length[OF assms(1)] assms proof (induction vs vs' rule: list_induct2_rev) case (2 x y xs ys) have norms: "fo_nmlzd AD xs" "fo_nmlzd AD ys" using 2(3,4) by (auto simp: fo_nmlzd_def nats_def Let_def map_filter_app rremdups_app split: sum.splits if_splits) have ad_agr: "ad_agr_list AD xs ys" using 2(2) by (auto simp: ad_agr_list_def ad_equiv_list_def sp_equiv_list_def pairwise_def) note xs_ys = 2(1)[OF ad_agr norms] have "x = y" proof (cases "isl x \ isl y") case True then have "isl x \ projl x \ AD" "isl y \ projl y \ AD" using 2(3,4) by (auto simp: fo_nmlzd_def) then show ?thesis using 2(2) True apply (auto simp: ad_agr_list_def ad_equiv_list_def isl_def) unfolding ad_equiv_pair.simps by blast+ next case False then obtain x' y' where inr: "x = Inr x'" "y = Inr y'" by (cases x; cases y) auto show ?thesis using 2(2) xs_ys proof (cases "x \ set xs \ y \ set ys") case False then show ?thesis using fo_nmlzd_app_Inr 2(3,4) unfolding inr xs_ys by auto qed (auto simp: ad_agr_list_def sp_equiv_list_def pairwise_def set_zip in_set_conv_nth) qed then show ?case using xs_ys by auto qed auto lemma fo_nmlz_eqI: assumes "ad_agr_list AD vs vs'" shows "fo_nmlz AD vs = fo_nmlz AD vs'" using ad_agr_list_fo_nmlzd[OF ad_agr_list_trans[OF ad_agr_list_trans[OF ad_agr_list_comm[OF fo_nmlz_ad_agr[of AD vs]] assms] fo_nmlz_ad_agr[of AD vs']] fo_nmlz_sound fo_nmlz_sound] . lemma fo_nmlz_eqD: assumes "fo_nmlz AD vs = fo_nmlz AD vs'" shows "ad_agr_list AD vs vs'" using ad_agr_list_trans[OF fo_nmlz_ad_agr[of AD vs, unfolded assms] ad_agr_list_comm[OF fo_nmlz_ad_agr[of AD vs']]] . lemma fo_nmlz_eq: "fo_nmlz AD vs = fo_nmlz AD vs' \ ad_agr_list AD vs vs'" using fo_nmlz_eqI[where ?AD=AD] fo_nmlz_eqD[where ?AD=AD] by blast lemma fo_nmlz_mono: assumes "AD \ AD'" "Inl -` set xs \ AD" shows "fo_nmlz AD' xs = fo_nmlz AD xs" proof - have "fo_nmlz AD (fo_nmlz AD' xs) = fo_nmlz AD' xs" apply (rule fo_nmlz_idem[OF fo_nmlzd_mono[OF _ fo_nmlz_sound]]) using assms by (auto simp: fo_nmlz_set) moreover have "fo_nmlz AD xs = fo_nmlz AD (fo_nmlz AD' xs)" apply (rule fo_nmlz_eqI) apply (rule ad_agr_list_mono[OF assms(1)]) apply (rule fo_nmlz_ad_agr) done ultimately show ?thesis by auto qed definition proj_vals :: "'c val set \ nat list \ 'c table" where "proj_vals R ns = (\\. map \ ns) ` R" definition proj_fmla :: "('a, 'b) fo_fmla \ 'c val set \ 'c table" where "proj_fmla \ R = proj_vals R (fv_fo_fmla_list \)" lemmas proj_fmla_map = proj_fmla_def[unfolded proj_vals_def] definition "extends_subst \ \ = (\x. \ x \ None \ \ x = \ x)" definition ext_tuple :: "'a set \ nat list \ nat list \ ('a + nat) list \ ('a + nat) list set" where "ext_tuple AD fv_sub fv_sub_comp as = (if fv_sub_comp = [] then {as} else (\fs. map snd (merge (zip fv_sub as) (zip fv_sub_comp fs))) ` (nall_tuples_rec AD (card (Inr -` set as)) (length fv_sub_comp)))" lemma ext_tuple_eq: "length fv_sub = length as \ ext_tuple AD fv_sub fv_sub_comp as = (\fs. map snd (merge (zip fv_sub as) (zip fv_sub_comp fs))) ` (nall_tuples_rec AD (card (Inr -` set as)) (length fv_sub_comp))" using fo_nmlz_idem[of AD as] by (auto simp: ext_tuple_def) lemma map_map_of: "length xs = length ys \ distinct xs \ ys = map (the \ (map_of (zip xs ys))) xs" by (induction xs ys rule: list_induct2) (auto simp: fun_upd_comp) lemma id_map_empty: "id_map 0 = Map.empty" by (rule ext) (auto simp: id_map_def split: sum.splits) lemma fo_nmlz_rec_shift: fixes xs :: "('a + nat) list" shows "fo_nmlz_rec i (id_map i) AD xs = xs \ i' = card (Inr -` (Inr ` {.. set (take n xs))) \ n \ length xs \ fo_nmlz_rec i' (id_map i') AD (drop n xs) = drop n xs" proof (induction i "id_map i :: 'a + nat \ nat" AD xs arbitrary: n rule: fo_nmlz_rec.induct) case (2 i AD x xs) have preds: "x \ AD" "fo_nmlz_rec i (id_map i) AD xs = xs" using 2(4) by (auto split: if_splits option.splits) show ?case using 2(4,5) proof (cases n) case (Suc k) have k_le: "k \ length xs" using 2(6) by (auto simp: Suc) have i'_def: "i' = card (Inr -` (Inr ` {.. set (take k xs)))" using 2(5) by (auto simp: Suc vimage_def) show ?thesis using 2(1)[OF preds i'_def k_le] by (auto simp: Suc) qed (auto simp: inj_vimage_image_eq) next case (3 i AD j xs) show ?case using 3(3,4) proof (cases n) case (Suc k) have k_le: "k \ length xs" using 3(5) by (auto simp: Suc) have j_le_i: "j \ i" using 3(3) by (auto split: option.splits dest: id_mapD) show ?thesis proof (cases "j = i") case True - have id_map: "id_map i (Inr j) = None" "id_map i(Inr j \ i) = id_map (Suc i)" + have id_map: "id_map i (Inr j) = None" "(id_map i)(Inr j \ i) = id_map (Suc i)" unfolding True fun_upd_id_map by (auto simp: id_map_def) have norm_xs: "fo_nmlz_rec (Suc i) (id_map (Suc i)) AD xs = xs" using 3(3) by (auto simp: id_map split: option.splits dest: id_mapD) have i'_def: "i' = card (Inr -` (Inr ` {.. set (take k xs)))" using 3(4) by (auto simp: Suc True inj_vimage_image_eq) (metis Un_insert_left image_insert inj_Inr inj_vimage_image_eq lessThan_Suc vimage_Un) show ?thesis using 3(1)[OF id_map norm_xs i'_def k_le] by (auto simp: Suc) next case False have id_map: "id_map i (Inr j) = Some j" using j_le_i False by (auto simp: id_map_def) have norm_xs: "fo_nmlz_rec i (id_map i) AD xs = xs" using 3(3) by (auto simp: id_map) have i'_def: "i' = card (Inr -` (Inr ` {.. set (take k xs)))" using 3(4) j_le_i False by (auto simp: Suc inj_vimage_image_eq insert_absorb) show ?thesis using 3(2)[OF id_map norm_xs i'_def k_le] by (auto simp: Suc) qed qed (auto simp: inj_vimage_image_eq) qed auto fun proj_tuple :: "nat list \ (nat \ ('a + nat)) list \ ('a + nat) list" where "proj_tuple [] mys = []" | "proj_tuple ns [] = []" | "proj_tuple (n # ns) ((m, y) # mys) = (if m < n then proj_tuple (n # ns) mys else if m = n then y # proj_tuple ns mys else proj_tuple ns ((m, y) # mys))" lemma proj_tuple_idle: "proj_tuple (map fst nxs) nxs = map snd nxs" by (induction nxs) auto lemma proj_tuple_merge: "sorted_distinct (map fst nxs) \ sorted_distinct (map fst mys) \ set (map fst nxs) \ set (map fst mys) = {} \ proj_tuple (map fst nxs) (merge nxs mys) = map snd nxs" using proj_tuple_idle by (induction nxs mys rule: merge.induct) auto+ lemma proj_tuple_map: assumes "sorted_distinct ns" "sorted_distinct ms" "set ns \ set ms" shows "proj_tuple ns (zip ms (map \ ms)) = map \ ns" proof - define ns' where "ns' = filter (\n. n \ set ns) ms" have sd_ns': "sorted_distinct ns'" using assms(2) sorted_filter[of id] by (auto simp: ns'_def) have disj: "set ns \ set ns' = {}" by (auto simp: ns'_def) have ms_def: "ms = sort (ns @ ns')" apply (rule sorted_distinct_set_unique) using assms by (auto simp: ns'_def) have zip: "zip ms (map \ ms) = merge (zip ns (map \ ns)) (zip ns' (map \ ns'))" unfolding merge_map[OF assms(1) sd_ns' disj, folded ms_def, symmetric] using map_fst_merge assms(1) by (auto simp: ms_def) (smt length_map map_fst_merge map_fst_zip sd_ns' zip_map_fst_snd) show ?thesis unfolding zip using proj_tuple_merge by (smt assms(1) disj length_map map_fst_zip map_snd_zip sd_ns') qed lemma proj_tuple_length: assumes "sorted_distinct ns" "sorted_distinct ms" "set ns \ set ms" "length ms = length xs" shows "length (proj_tuple ns (zip ms xs)) = length ns" proof - obtain \ where \: "xs = map \ ms" using exists_map[OF assms(4)] assms(2) by auto show ?thesis unfolding \ by (auto simp: proj_tuple_map[OF assms(1-3)]) qed lemma ext_tuple_sound: assumes "sorted_distinct fv_sub" "sorted_distinct fv_sub_comp" "sorted_distinct fv_all" "set fv_sub \ set fv_sub_comp = {}" "set fv_sub \ set fv_sub_comp = set fv_all" "ass = fo_nmlz AD ` proj_vals R fv_sub" "\\ \. ad_agr_sets (set fv_sub) (set fv_sub) AD \ \ \ \ \ R \ \ \ R" "xs \ fo_nmlz AD ` \(ext_tuple AD fv_sub fv_sub_comp ` ass)" shows "fo_nmlz AD (proj_tuple fv_sub (zip fv_all xs)) \ ass" "xs \ fo_nmlz AD ` proj_vals R fv_all" proof - have fv_all_sort: "fv_all = sort (fv_sub @ fv_sub_comp)" using assms(1,2,3,4,5) by (simp add: sorted_distinct_set_unique) have len_in_ass: "\xs. xs \ ass \ xs = fo_nmlz AD xs \ length xs = length fv_sub" by (auto simp: assms(6) proj_vals_def fo_nmlz_length fo_nmlz_idem fo_nmlz_sound) obtain as fs where as_fs_def: "as \ ass" "fs \ nall_tuples_rec AD (card (Inr -` set as)) (length fv_sub_comp)" "xs = fo_nmlz AD (map snd (merge (zip fv_sub as) (zip fv_sub_comp fs)))" using fo_nmlz_sound len_in_ass assms(8) by (auto simp: ext_tuple_def split: if_splits) then have vs_norm: "fo_nmlzd AD xs" using fo_nmlz_sound by auto obtain \ where \_def: "\ \ R" "as = fo_nmlz AD (map \ fv_sub)" using as_fs_def(1) assms(6) by (auto simp: proj_vals_def) then obtain \ where \_def: "as = map \ fv_sub" "ad_agr_list AD (map \ fv_sub) (map \ fv_sub)" using fo_nmlz_map fo_nmlz_ad_agr by metis have \_R: "\ \ R" using assms(7) ad_agr_list_link \_def(1) \_def(2) by fastforce define \' where "\' \ \n. if n \ set fv_sub_comp then the (map_of (zip fv_sub_comp fs) n) else \ n" then have "\n \ set fv_sub. \ n = \' n" using assms(4) by auto then have \'_S: "\' \ R" using assms(7) \_R by (fastforce simp: ad_agr_sets_def sp_equiv_def pairwise_def ad_equiv_pair.simps) have length_as: "length as = length fv_sub" using as_fs_def(1) assms(6) by (auto simp: proj_vals_def fo_nmlz_length) have length_fs: "length fs = length fv_sub_comp" using as_fs_def(2) by (auto simp: nall_tuples_rec_length) have map_fv_sub: "map \' fv_sub = map \ fv_sub" using assms(4) \_def(2) by (auto simp: \'_def) have fs_map_map_of: "fs = map (the \ (map_of (zip fv_sub_comp fs))) fv_sub_comp" using map_map_of length_fs assms(2) by metis have fs_map: "fs = map \' fv_sub_comp" using \'_def length_fs by (subst fs_map_map_of) simp have vs_map_fv_all: "xs = fo_nmlz AD (map \' fv_all)" unfolding as_fs_def(3) \_def(1) map_fv_sub[symmetric] fs_map fv_all_sort using merge_map[OF assms(1,2,4)] by metis show "xs \ fo_nmlz AD ` proj_vals R fv_all" using \'_S vs_map_fv_all by (auto simp: proj_vals_def) obtain \'' where \''_def: "xs = map \'' fv_all" using exists_map[of fv_all xs] fo_nmlz_map vs_map_fv_all by blast have proj: "proj_tuple fv_sub (zip fv_all xs) = map \'' fv_sub" using proj_tuple_map assms(1,3,5) unfolding \''_def by blast have \''_\': "fo_nmlz AD (map \'' fv_sub) = as" using \''_def vs_map_fv_all \_def(2) by (metis \_def(2) ad_agr_list_subset assms(5) fo_nmlz_ad_agr fo_nmlz_eqI map_fv_sub sup_ge1) show "fo_nmlz AD (proj_tuple fv_sub (zip fv_all xs)) \ ass" unfolding proj \''_\' map_fv_sub by (rule as_fs_def(1)) qed lemma ext_tuple_complete: assumes "sorted_distinct fv_sub" "sorted_distinct fv_sub_comp" "sorted_distinct fv_all" "set fv_sub \ set fv_sub_comp = {}" "set fv_sub \ set fv_sub_comp = set fv_all" "ass = fo_nmlz AD ` proj_vals R fv_sub" "\\ \. ad_agr_sets (set fv_sub) (set fv_sub) AD \ \ \ \ \ R \ \ \ R" "xs = fo_nmlz AD (map \ fv_all)" "\ \ R" shows "xs \ fo_nmlz AD ` \(ext_tuple AD fv_sub fv_sub_comp ` ass)" proof - have fv_all_sort: "fv_all = sort (fv_sub @ fv_sub_comp)" using assms(1,2,3,4,5) by (simp add: sorted_distinct_set_unique) note \_def = assms(9,8) have vs_norm: "fo_nmlzd AD xs" using \_def(2) fo_nmlz_sound by auto define fs where "fs = map \ fv_sub_comp" define as where "as = map \ fv_sub" define nos where "nos = fo_nmlz AD (as @ fs)" define as' where "as' = take (length fv_sub) nos" define fs' where "fs' = drop (length fv_sub) nos" have length_as': "length as' = length fv_sub" by (auto simp: as'_def nos_def as_def fo_nmlz_length) have length_fs': "length fs' = length fv_sub_comp" by (auto simp: fs'_def nos_def as_def fs_def fo_nmlz_length) have len_fv_sub_nos: "length fv_sub \ length nos" by (auto simp: nos_def fo_nmlz_length as_def) have norm_as': "fo_nmlzd AD as'" using fo_nmlzd_take[OF fo_nmlz_sound] by (auto simp: as'_def nos_def) have as'_norm_as: "as' = fo_nmlz AD as" by (auto simp: as'_def nos_def as_def fo_nmlz_take) have ad_agr_as': "ad_agr_list AD as as'" using fo_nmlz_ad_agr unfolding as'_norm_as . have nos_as'_fs': "nos = as' @ fs'" using length_as' length_fs' by (auto simp: as'_def fs'_def) obtain \ where \_def: "as' = map \ fv_sub" "fs' = map \ fv_sub_comp" using exists_map[of "fv_sub @ fv_sub_comp" "as' @ fs'"] assms(1,2,4) length_as' length_fs' by auto have "length fv_sub + length fv_sub_comp \ length fv_all" using assms(1,2,3,4,5) by (metis distinct_append distinct_card eq_iff length_append set_append) then have nos_sub: "set nos \ Inl ` AD \ Inr ` {.. nall_tuples_rec AD (card (Inr -` set as')) (length fv_sub_comp)" unfolding len_fs'[symmetric] by (rule nall_tuples_rec_fo_nmlz_rec_complete) (rule fo_nmlz_rec_shift[OF norm_nos_idem, simplified, OF refl len_fv_sub_nos, folded as'_def fs'_def]) have "as' \ nall_tuples AD (length fv_sub)" using length_as' apply (rule nall_tuplesI) using norm_as' . then have as'_ass: "as' \ ass" using as'_norm_as \_def(1) as_def unfolding assms(6) by (auto simp: proj_vals_def) have vs_norm: "xs = fo_nmlz AD (map snd (merge (zip fv_sub as) (zip fv_sub_comp fs)))" using assms(1,2,4) \_def(2) by (auto simp: merge_map as_def fs_def fv_all_sort) have set_sort': "set (sort (fv_sub @ fv_sub_comp)) = set (fv_sub @ fv_sub_comp)" by auto have "xs = fo_nmlz AD (map snd (merge (zip fv_sub as') (zip fv_sub_comp fs')))" unfolding vs_norm as_def fs_def \_def merge_map[OF assms(1,2,4)] apply (rule fo_nmlz_eqI) apply (rule ad_agr_list_subset[OF equalityD1, OF set_sort']) using fo_nmlz_ad_agr[of AD "as @ fs", folded nos_def, unfolded nos_as'_fs'] unfolding as_def fs_def \_def map_append[symmetric] . then show ?thesis using as'_ass fs'_all by (auto simp: ext_tuple_def length_as') qed definition "ext_tuple_set AD ns ns' X = (if ns' = [] then X else fo_nmlz AD ` \(ext_tuple AD ns ns' ` X))" lemma ext_tuple_set_eq: "Ball X (fo_nmlzd AD) \ ext_tuple_set AD ns ns' X = fo_nmlz AD ` \(ext_tuple AD ns ns' ` X)" by (auto simp: ext_tuple_set_def ext_tuple_def fo_nmlzd_code) lemma ext_tuple_set_mono: "A \ B \ ext_tuple_set AD ns ns' A \ ext_tuple_set AD ns ns' B" by (auto simp: ext_tuple_set_def) lemma ext_tuple_correct: assumes "sorted_distinct fv_sub" "sorted_distinct fv_sub_comp" "sorted_distinct fv_all" "set fv_sub \ set fv_sub_comp = {}" "set fv_sub \ set fv_sub_comp = set fv_all" "ass = fo_nmlz AD ` proj_vals R fv_sub" "\\ \. ad_agr_sets (set fv_sub) (set fv_sub) AD \ \ \ \ \ R \ \ \ R" shows "ext_tuple_set AD fv_sub fv_sub_comp ass = fo_nmlz AD ` proj_vals R fv_all" proof (rule set_eqI, rule iffI) fix xs assume xs_in: "xs \ ext_tuple_set AD fv_sub fv_sub_comp ass" show "xs \ fo_nmlz AD ` proj_vals R fv_all" using ext_tuple_sound(2)[OF assms] xs_in by (auto simp: ext_tuple_set_def ext_tuple_def assms(6) fo_nmlz_idem[OF fo_nmlz_sound] image_iff split: if_splits) next fix xs assume "xs \ fo_nmlz AD ` proj_vals R fv_all" then obtain \ where \_def: "xs = fo_nmlz AD (map \ fv_all)" "\ \ R" by (auto simp: proj_vals_def) show "xs \ ext_tuple_set AD fv_sub fv_sub_comp ass" using ext_tuple_complete[OF assms \_def] by (auto simp: ext_tuple_set_def ext_tuple_def assms(6) fo_nmlz_idem[OF fo_nmlz_sound] image_iff split: if_splits) qed lemma proj_tuple_sound: assumes "sorted_distinct fv_sub" "sorted_distinct fv_sub_comp" "sorted_distinct fv_all" "set fv_sub \ set fv_sub_comp = {}" "set fv_sub \ set fv_sub_comp = set fv_all" "ass = fo_nmlz AD ` proj_vals R fv_sub" "\\ \. ad_agr_sets (set fv_sub) (set fv_sub) AD \ \ \ \ \ R \ \ \ R" "fo_nmlz AD xs = xs" "length xs = length fv_all" "fo_nmlz AD (proj_tuple fv_sub (zip fv_all xs)) \ ass" shows "xs \ fo_nmlz AD ` \(ext_tuple AD fv_sub fv_sub_comp ` ass)" proof - have fv_all_sort: "fv_all = sort (fv_sub @ fv_sub_comp)" using assms(1,2,3,4,5) by (simp add: sorted_distinct_set_unique) obtain \ where \_def: "xs = map \ fv_all" using exists_map[of fv_all xs] assms(3,9) by auto have xs_norm: "xs = fo_nmlz AD (map \ fv_all)" using assms(8) by (auto simp: \_def) have proj: "proj_tuple fv_sub (zip fv_all xs) = map \ fv_sub" unfolding \_def apply (rule proj_tuple_map[OF assms(1,3)]) using assms(5) by blast obtain \ where \_def: "fo_nmlz AD (map \ fv_sub) = fo_nmlz AD (map \ fv_sub)" "\ \ R" using assms(10) by (auto simp: assms(6) proj proj_vals_def) have \_R: "\ \ R" using assms(7) fo_nmlz_eqD[OF \_def(1)] \_def(2) unfolding ad_agr_list_link[symmetric] by auto show ?thesis by (rule ext_tuple_complete[OF assms(1,2,3,4,5,6,7) xs_norm \_R]) assumption qed lemma proj_tuple_correct: assumes "sorted_distinct fv_sub" "sorted_distinct fv_sub_comp" "sorted_distinct fv_all" "set fv_sub \ set fv_sub_comp = {}" "set fv_sub \ set fv_sub_comp = set fv_all" "ass = fo_nmlz AD ` proj_vals R fv_sub" "\\ \. ad_agr_sets (set fv_sub) (set fv_sub) AD \ \ \ \ \ R \ \ \ R" "fo_nmlz AD xs = xs" "length xs = length fv_all" shows "xs \ fo_nmlz AD ` \(ext_tuple AD fv_sub fv_sub_comp ` ass) \ fo_nmlz AD (proj_tuple fv_sub (zip fv_all xs)) \ ass" using ext_tuple_sound(1)[OF assms(1,2,3,4,5,6,7)] proj_tuple_sound[OF assms] by blast fun unify_vals_terms :: "('a + 'c) list \ ('a fo_term) list \ (nat \ ('a + 'c)) \ (nat \ ('a + 'c)) option" where "unify_vals_terms [] [] \ = Some \" | "unify_vals_terms (v # vs) ((Const c') # ts) \ = (if v = Inl c' then unify_vals_terms vs ts \ else None)" | "unify_vals_terms (v # vs) ((Var n) # ts) \ = (case \ n of Some x \ (if v = x then unify_vals_terms vs ts \ else None) | None \ unify_vals_terms vs ts (\(n := Some v)))" | "unify_vals_terms _ _ _ = None" lemma unify_vals_terms_extends: "unify_vals_terms vs ts \ = Some \' \ extends_subst \ \'" unfolding extends_subst_def by (induction vs ts \ arbitrary: \' rule: unify_vals_terms.induct) (force split: if_splits option.splits)+ lemma unify_vals_terms_sound: "unify_vals_terms vs ts \ = Some \' \ (the \ \') \e ts = vs" using unify_vals_terms_extends by (induction vs ts \ arbitrary: \' rule: unify_vals_terms.induct) (force simp: eval_eterms_def extends_subst_def fv_fo_terms_set_def split: if_splits option.splits)+ lemma unify_vals_terms_complete: "\'' \e ts = vs \ (\n. \ n \ None \ \ n = Some (\'' n)) \ \\'. unify_vals_terms vs ts \ = Some \'" by (induction vs ts \ rule: unify_vals_terms.induct) (force simp: eval_eterms_def extends_subst_def split: if_splits option.splits)+ definition eval_table :: "'a fo_term list \ ('a + 'c) table \ ('a + 'c) table" where "eval_table ts X = (let fvs = fv_fo_terms_list ts in \((\vs. case unify_vals_terms vs ts Map.empty of Some \ \ {map (the \ \) fvs} | _ \ {}) ` X))" lemma eval_table: fixes X :: "('a + 'c) table" shows "eval_table ts X = proj_vals {\. \ \e ts \ X} (fv_fo_terms_list ts)" proof (rule set_eqI, rule iffI) fix vs assume "vs \ eval_table ts X" then obtain as \ where as_def: "as \ X" "unify_vals_terms as ts Map.empty = Some \" "vs = map (the \ \) (fv_fo_terms_list ts)" by (auto simp: eval_table_def split: option.splits) have "(the \ \) \e ts \ X" using unify_vals_terms_sound[OF as_def(2)] as_def(1) by auto with as_def(3) show "vs \ proj_vals {\. \ \e ts \ X} (fv_fo_terms_list ts)" by (fastforce simp: proj_vals_def) next fix vs :: "('a + 'c) list" assume "vs \ proj_vals {\. \ \e ts \ X} (fv_fo_terms_list ts)" then obtain \ where \_def: "vs = map \ (fv_fo_terms_list ts)" "\ \e ts \ X" by (auto simp: proj_vals_def) obtain \' where \'_def: "unify_vals_terms (\ \e ts) ts Map.empty = Some \'" using unify_vals_terms_complete[OF refl, of Map.empty \ ts] by auto have "(the \ \') \e ts = (\ \e ts)" using unify_vals_terms_sound[OF \'_def(1)] by auto then have "vs = map (the \ \') (fv_fo_terms_list ts)" using fv_fo_terms_set_list eval_eterms_fv_fo_terms_set unfolding \_def(1) by fastforce then show "vs \ eval_table ts X" using \_def(2) \'_def by (force simp: eval_table_def) qed fun ad_agr_close_rec :: "nat \ (nat \ 'a + nat) \ 'a set \ ('a + nat) list \ ('a + nat) list set" where "ad_agr_close_rec i m AD [] = {[]}" | "ad_agr_close_rec i m AD (Inl x # xs) = (\xs. Inl x # xs) ` ad_agr_close_rec i m AD xs" | "ad_agr_close_rec i m AD (Inr n # xs) = (case m n of None \ \((\x. (\xs. Inl x # xs) ` ad_agr_close_rec i (m(n := Some (Inl x))) (AD - {x}) xs) ` AD) \ (\xs. Inr i # xs) ` ad_agr_close_rec (Suc i) (m(n := Some (Inr i))) AD xs | Some v \ (\xs. v # xs) ` ad_agr_close_rec i m AD xs)" lemma ad_agr_close_rec_length: "ys \ ad_agr_close_rec i m AD xs \ length xs = length ys" by (induction i m AD xs arbitrary: ys rule: ad_agr_close_rec.induct) (auto split: option.splits) lemma ad_agr_close_rec_sound: "ys \ ad_agr_close_rec i m AD xs \ fo_nmlz_rec j (id_map j) X xs = xs \ X \ AD = {} \ X \ Y = {} \ Y \ AD = {} \ inj_on m (dom m) \ dom m = {.. ran m \ Inl ` Y \ Inr ` {.. i \ j \ fo_nmlz_rec i (id_map i) (X \ Y \ AD) ys = ys \ (\m'. inj_on m' (dom m') \ (\n v. m n = Some v \ m' (Inr n) = Some v) \ (\(x, y) \ set (zip xs ys). case x of Inl x' \ if x' \ X then x = y else m' x = Some y \ (case y of Inl z \ z \ X | Inr x \ True) | Inr n \ m' x = Some y \ (case y of Inl z \ z \ X | Inr x \ True)))" proof (induction i m AD xs arbitrary: Y j ys rule: ad_agr_close_rec.induct) case (1 i m AD) then show ?case by (auto simp: ad_agr_list_def ad_equiv_list_def sp_equiv_list_def inj_on_def dom_def split: sum.splits intro!: exI[of _ "case_sum Map.empty m"]) next case (2 i m AD x xs) obtain zs where ys_def: "ys = Inl x # zs" "zs \ ad_agr_close_rec i m AD xs" using 2(2) by auto have preds: "fo_nmlz_rec j (id_map j) X xs = xs" "x \ X" using 2(3) by (auto split: if_splits option.splits) show ?case using 2(1)[OF ys_def(2) preds(1) 2(4,5,6,7,8,9,10)] preds(2) by (auto simp: ys_def(1)) next case (3 i m AD n xs) show ?case proof (cases "m n") case None obtain v zs where ys_def: "ys = v # zs" using 3(4) by (auto simp: None) have n_ge_j: "j \ n" using 3(9,10) None by (metis domIff leI lessThan_iff) show ?thesis proof (cases v) case (Inl x) have zs_def: "zs \ ad_agr_close_rec i (m(n \ Inl x)) (AD - {x}) xs" "x \ AD" using 3(4) by (auto simp: None ys_def Inl) have preds: "fo_nmlz_rec (Suc j) (id_map (Suc j)) X xs = xs" "X \ (AD - {x}) = {}" "X \ (Y \ {x}) = {}" "(Y \ {x}) \ (AD - {x}) = {}" "dom (m(n \ Inl x)) = {.. Inl x)) \ Inl ` (Y \ {x}) \ Inr ` {.. Suc j" "n = j" using 3(5,6,7,8,10,11,12) n_ge_j zs_def(2) by (auto simp: fun_upd_id_map ran_def dest: id_mapD split: option.splits) have inj: "inj_on (m(n \ Inl x)) (dom (m(n \ Inl x)))" using 3(8,9,10,11,12) preds(8) zs_def(2) by (fastforce simp: inj_on_def dom_def ran_def) have sets_unfold: "X \ (Y \ {x}) \ (AD - {x}) = X \ Y \ AD" using zs_def(2) by auto note IH = 3(1)[OF None zs_def(2,1) preds(1,2,3,4) inj preds(5,6,7), unfolded sets_unfold] have norm_ys: "fo_nmlz_rec i (id_map i) (X \ Y \ AD) ys = ys" using conjunct1[OF IH] zs_def(2) by (auto simp: ys_def(1) Inl split: option.splits) show ?thesis using norm_ys conjunct2[OF IH] None zs_def(2) 3(6) unfolding ys_def(1) apply safe subgoal for m' apply (auto simp: Inl dom_def intro!: exI[of _ m'] split: if_splits) apply (metis option.distinct(1)) apply (fastforce split: prod.splits sum.splits) done done next case (Inr k) have zs_def: "zs \ ad_agr_close_rec (Suc i) (m(n \ Inr i)) AD xs" "i = k" using 3(4) by (auto simp: None ys_def Inr) have preds: "fo_nmlz_rec (Suc n) (id_map (Suc n)) X xs = xs" "dom (m(n \ Inr i)) = {.. Inr i)) \ Inl ` Y \ Inr ` {.. Suc n" using 3(5,10,11,12) n_ge_j by (auto simp: fun_upd_id_map ran_def dest: id_mapD split: option.splits) have inj: "inj_on (m(n \ Inr i)) (dom (m(n \ Inr i)))" using 3(9,11) by (auto simp: inj_on_def dom_def ran_def) note IH = 3(2)[OF None zs_def(1) preds(1) 3(6,7,8) inj preds(2,3,4)] have norm_ys: "fo_nmlz_rec i (id_map i) (X \ Y \ AD) ys = ys" using conjunct1[OF IH] zs_def(2) by (auto simp: ys_def Inr fun_upd_id_map dest: id_mapD split: option.splits) show ?thesis using norm_ys conjunct2[OF IH] None unfolding ys_def(1) zs_def(2) apply safe subgoal for m' apply (auto simp: Inr dom_def intro!: exI[of _ m'] split: if_splits) apply (metis option.distinct(1)) apply (fastforce split: prod.splits sum.splits) done done qed next case (Some v) obtain zs where ys_def: "ys = v # zs" "zs \ ad_agr_close_rec i m AD xs" using 3(4) by (auto simp: Some) have preds: "fo_nmlz_rec j (id_map j) X xs = xs" "n < j" using 3(5,8,10) Some by (auto simp: dom_def split: option.splits) note IH = 3(3)[OF Some ys_def(2) preds(1) 3(6,7,8,9,10,11,12)] have norm_ys: "fo_nmlz_rec i (id_map i) (X \ Y \ AD) ys = ys" using conjunct1[OF IH] 3(11) Some by (auto simp: ys_def(1) ran_def id_map_def) have "case v of Inl z \ z \ X | Inr x \ True" using 3(7,11) Some by (auto simp: ran_def split: sum.splits) then show ?thesis using norm_ys conjunct2[OF IH] Some unfolding ys_def(1) apply safe subgoal for m' by (auto intro!: exI[of _ m'] split: sum.splits) done qed qed lemma ad_agr_close_rec_complete: fixes xs :: "('a + nat) list" shows "fo_nmlz_rec j (id_map j) X xs = xs \ X \ AD = {} \ X \ Y = {} \ Y \ AD = {} \ inj_on m (dom m) \ dom m = {.. ran m = Inl ` Y \ Inr ` {.. i \ j \ (\n b. (Inr n, b) \ set (zip xs ys) \ case m n of Some v \ v = b | None \ b \ ran m) \ fo_nmlz_rec i (id_map i) (X \ Y \ AD) ys = ys \ ad_agr_list X xs ys \ ys \ ad_agr_close_rec i m AD xs" proof (induction j "id_map j :: 'a + nat \ nat option" X xs arbitrary: m i ys AD Y rule: fo_nmlz_rec.induct) case (2 j X x xs) have x_X: "x \ X" "fo_nmlz_rec j (id_map j) X xs = xs" using 2(4) by (auto split: if_splits option.splits) obtain z zs where ys_def: "ys = Inl z # zs" "z = x" using 2(14) x_X(1) by (cases ys) (auto simp: ad_agr_list_def ad_equiv_list_def ad_equiv_pair.simps) have norm_zs: "fo_nmlz_rec i (id_map i) (X \ Y \ AD) zs = zs" using 2(13) ys_def(2) x_X(1) by (auto simp: ys_def(1)) have ad_agr: "ad_agr_list X xs zs" using 2(14) by (auto simp: ys_def ad_agr_list_def ad_equiv_list_def sp_equiv_list_def pairwise_def) show ?case using 2(1)[OF x_X 2(5,6,7,8,9,10,11) _ norm_zs ad_agr] 2(12) by (auto simp: ys_def) next case (3 j X n xs) obtain z zs where ys_def: "ys = z # zs" using 3(13) apply (cases ys) apply (auto simp: ad_agr_list_def) done show ?case proof (cases "j \ n") case True then have n_j: "n = j" using 3(3) by (auto split: option.splits dest: id_mapD) - have id_map: "id_map j (Inr n) = None" "id_map j(Inr n \ j) = id_map (Suc j)" + have id_map: "id_map j (Inr n) = None" "(id_map j)(Inr n \ j) = id_map (Suc j)" unfolding n_j fun_upd_id_map by (auto simp: id_map_def) have norm_xs: "fo_nmlz_rec (Suc j) (id_map (Suc j)) X xs = xs" using 3(3) by (auto simp: ys_def fun_upd_id_map id_map(1) split: option.splits) have None: "m n = None" using 3(8) by (auto simp: dom_def n_j) have z_out: "z \ Inl ` Y \ Inr ` {.. AD" using 3(12,13) z_out by (auto simp: ys_def Inl ad_agr_list_def ad_equiv_list_def ad_equiv_pair.simps split: if_splits option.splits) have norm_zs: "fo_nmlz_rec i (id_map i) (X \ Y \ AD) zs = zs" using 3(12) a_in by (auto simp: ys_def Inl) have preds: "X \ (AD - {a}) = {}" "X \ (Y \ {a}) = {}" "(Y \ {a}) \ (AD - {a}) = {}" using 3(4,5,6) a_in by auto have inj: "inj_on (m(n := Some (Inl a))) (dom (m(n := Some (Inl a))))" using 3(6,7,9) None a_in by (auto simp: inj_on_def dom_def ran_def) blast+ have preds': "dom (m(n \ Inl a)) = {.. Inl a)) = Inl ` (Y \ {a}) \ Inr ` {.. Suc j" using 3(6,8,9,10) None less_Suc_eq a_in apply (auto simp: n_j dom_def ran_def) apply (smt Un_iff image_eqI mem_Collect_eq option.simps(3)) apply (smt 3(8) domIff image_subset_iff lessThan_iff mem_Collect_eq sup_ge2) done have a_unfold: "X \ (Y \ {a}) \ (AD - {a}) = X \ Y \ AD" "Y \ {a} \ (AD - {a}) = Y \ AD" using a_in by auto have ad_agr: "ad_agr_list X xs zs" using 3(13) by (auto simp: ys_def Inl ad_agr_list_def ad_equiv_list_def sp_equiv_list_def pairwise_def) have "zs \ ad_agr_close_rec i (m(n \ Inl a)) (AD - {a}) xs" apply (rule 3(1)[OF id_map norm_xs preds inj preds' _ _ ad_agr]) using 3(11,13) norm_zs unfolding 3(9) preds'(2) a_unfold apply (auto simp: None Inl ys_def ad_agr_list_def sp_equiv_list_def pairwise_def split: option.splits) apply (metis Un_iff image_eqI option.simps(4)) apply (metis image_subset_iff lessThan_iff option.simps(4) sup_ge2) apply fastforce done then show ?thesis using a_in by (auto simp: ys_def Inl None) next case (Inr b) have i_b: "i = b" using 3(12) z_out by (auto simp: ys_def Inr split: option.splits dest: id_mapD) have norm_zs: "fo_nmlz_rec (Suc i) (id_map (Suc i)) (X \ Y \ AD) zs = zs" using 3(12) by (auto simp: ys_def Inr i_b fun_upd_id_map split: option.splits dest: id_mapD) have ad_agr: "ad_agr_list X xs zs" using 3(13) by (auto simp: ys_def ad_agr_list_def ad_equiv_list_def sp_equiv_list_def pairwise_def) define m' where "m' \ m(n := Some (Inr i))" have preds: "inj_on m' (dom m')" "dom m' = {.. Suc j" using 3(7,8,9,10) by (auto simp: m'_def n_j inj_on_def dom_def ran_def image_iff) (metis 3(8) domI lessThan_iff less_SucI) have ran: "ran m' = Inl ` Y \ Inr ` {.. ad_agr_close_rec (Suc i) m' AD xs" apply (rule 3(1)[OF id_map norm_xs 3(4,5,6) preds(1,2) ran preds(3) _ norm_zs ad_agr]) using 3(11,13) unfolding 3(9) ys_def Inr i_b m'_def unfolding ran[unfolded m'_def i_b] apply (auto simp: ad_agr_list_def sp_equiv_list_def pairwise_def split: option.splits) apply (metis Un_upper1 image_subset_iff option.simps(4)) apply (metis UnI1 image_eqI insert_iff lessThan_Suc lessThan_iff option.simps(4) sp_equiv_pair.simps sum.inject(2) sup_commute) apply fastforce done then show ?thesis by (auto simp: ys_def Inr None m'_def i_b) qed next case False have id_map: "id_map j (Inr n) = Some n" using False by (auto simp: id_map_def) have norm_xs: "fo_nmlz_rec j (id_map j) X xs = xs" using 3(3) by (auto simp: id_map) have Some: "m n = Some z" using False 3(11)[unfolded ys_def] by (metis (mono_tags) 3(8) domD insert_iff leI lessThan_iff list.simps(15) option.simps(5) zip_Cons_Cons) have z_in: "z \ Inl ` Y \ Inr ` {.. Y \ AD" using 3(12,13) by (auto simp: ys_def Inl ad_agr_list_def ad_equiv_list_def ad_equiv_pair.simps split: if_splits option.splits) have norm_zs: "fo_nmlz_rec i (id_map i) (X \ Y \ AD) zs = zs" using 3(12) a_in by (auto simp: ys_def Inl) show ?thesis using 3(2)[OF id_map norm_xs 3(4,5,6,7,8,9,10) _ norm_zs ad_agr] 3(11) a_in by (auto simp: ys_def Inl Some split: option.splits) next case (Inr b) have b_lt: "b < i" using z_in by (auto simp: Inr) have norm_zs: "fo_nmlz_rec i (id_map i) (X \ Y \ AD) zs = zs" using 3(12) b_lt by (auto simp: ys_def Inr split: option.splits) show ?thesis using 3(2)[OF id_map norm_xs 3(4,5,6,7,8,9,10) _ norm_zs ad_agr] 3(11) by (auto simp: ys_def Inr Some) qed qed qed (auto simp: ad_agr_list_def) definition ad_agr_close :: "'a set \ ('a + nat) list \ ('a + nat) list set" where "ad_agr_close AD xs = ad_agr_close_rec 0 Map.empty AD xs" lemma ad_agr_close_sound: assumes "ys \ ad_agr_close Y xs" "fo_nmlzd X xs" "X \ Y = {}" shows "fo_nmlzd (X \ Y) ys \ ad_agr_list X xs ys" using ad_agr_close_rec_sound[OF assms(1)[unfolded ad_agr_close_def] fo_nmlz_idem[OF assms(2), unfolded fo_nmlz_def, folded id_map_empty] assms(3) Int_empty_right Int_empty_left] ad_agr_map[OF ad_agr_close_rec_length[OF assms(1)[unfolded ad_agr_close_def]], of _ X] fo_nmlzd_code[unfolded fo_nmlz_def, folded id_map_empty, of "X \ Y" ys] by (auto simp: fo_nmlz_def) lemma ad_agr_close_complete: assumes "X \ Y = {}" "fo_nmlzd X xs" "fo_nmlzd (X \ Y) ys" "ad_agr_list X xs ys" shows "ys \ ad_agr_close Y xs" using ad_agr_close_rec_complete[OF fo_nmlz_idem[OF assms(2), unfolded fo_nmlz_def, folded id_map_empty] assms(1) Int_empty_right Int_empty_left _ _ _ order.refl _ _ assms(4), of Map.empty] fo_nmlzd_code[unfolded fo_nmlz_def, folded id_map_empty, of "X \ Y" ys] assms(3) unfolding ad_agr_close_def by (auto simp: fo_nmlz_def) lemma ad_agr_close_empty: "fo_nmlzd X xs \ ad_agr_close {} xs = {xs}" using ad_agr_close_complete[where ?X=X and ?Y="{}" and ?xs=xs and ?ys=xs] ad_agr_close_sound[where ?X=X and ?Y="{}" and ?xs=xs] ad_agr_list_refl ad_agr_list_fo_nmlzd by fastforce lemma ad_agr_close_set_correct: assumes "AD' \ AD" "sorted_distinct ns" "\\ \. ad_agr_sets (set ns) (set ns) AD' \ \ \ \ \ R \ \ \ R" shows "\(ad_agr_close (AD - AD') ` fo_nmlz AD' ` proj_vals R ns) = fo_nmlz AD ` proj_vals R ns" proof (rule set_eqI, rule iffI) fix vs assume "vs \ \(ad_agr_close (AD - AD') ` fo_nmlz AD' ` proj_vals R ns)" then obtain \ where \_def: "vs \ ad_agr_close (AD - AD') (fo_nmlz AD' (map \ ns))" "\ \ R" by (auto simp: proj_vals_def) have vs: "fo_nmlzd AD vs" "ad_agr_list AD' (fo_nmlz AD' (map \ ns)) vs" using ad_agr_close_sound[OF \_def(1) fo_nmlz_sound] assms(1) Diff_partition by fastforce+ obtain \ where \_def: "vs = map \ ns" using exists_map[of ns vs] assms(2) vs(2) by (auto simp: ad_agr_list_def fo_nmlz_length) show "vs \ fo_nmlz AD ` proj_vals R ns" apply (subst fo_nmlz_idem[OF vs(1), symmetric]) using iffD1[OF assms(3) \_def(2), OF iffD2[OF ad_agr_list_link ad_agr_list_trans[OF fo_nmlz_ad_agr[of AD' "map \ ns"] vs(2), unfolded \_def]]] unfolding \_def by (auto simp: proj_vals_def) next fix vs assume "vs \ fo_nmlz AD ` proj_vals R ns" then obtain \ where \_def: "vs = fo_nmlz AD (map \ ns)" "\ \ R" by (auto simp: proj_vals_def) define xs where "xs = fo_nmlz AD' vs" have preds: "AD' \ (AD - AD') = {}" "fo_nmlzd AD' xs" "fo_nmlzd (AD' \ (AD - AD')) vs" using assms(1) fo_nmlz_sound Diff_partition by (fastforce simp: \_def(1) xs_def)+ obtain \ where \_def: "vs = map \ ns" using exists_map[of "ns" vs] assms(2) \_def(1) by (auto simp: fo_nmlz_length) have "vs \ ad_agr_close (AD - AD') xs" using ad_agr_close_complete[OF preds] ad_agr_list_comm[OF fo_nmlz_ad_agr] by (auto simp: xs_def) then show "vs \ \(ad_agr_close (AD - AD') ` fo_nmlz AD' ` proj_vals R ns)" unfolding xs_def \_def using iffD1[OF assms(3) \_def(2), OF ad_agr_sets_mono[OF assms(1) iffD2[OF ad_agr_list_link fo_nmlz_ad_agr[of AD "map \ ns", folded \_def(1), unfolded \_def]]]] by (auto simp: proj_vals_def) qed lemma ad_agr_close_correct: assumes "AD' \ AD" "\\ \. ad_agr_sets (set (fv_fo_fmla_list \)) (set (fv_fo_fmla_list \)) AD' \ \ \ \ \ R \ \ \ R" shows "\(ad_agr_close (AD - AD') ` fo_nmlz AD' ` proj_fmla \ R) = fo_nmlz AD ` proj_fmla \ R" using ad_agr_close_set_correct[OF _ sorted_distinct_fv_list, OF assms] by (auto simp: proj_fmla_def) definition "ad_agr_close_set AD X = (if Set.is_empty AD then X else \(ad_agr_close AD ` X))" lemma ad_agr_close_set_eq: "Ball X (fo_nmlzd AD') \ ad_agr_close_set AD X = \(ad_agr_close AD ` X)" by (force simp: ad_agr_close_set_def Set.is_empty_def ad_agr_close_empty) lemma Ball_fo_nmlzd: "Ball (fo_nmlz AD ` X) (fo_nmlzd AD)" by (auto simp: fo_nmlz_sound) lemmas ad_agr_close_set_nmlz_eq = ad_agr_close_set_eq[OF Ball_fo_nmlzd] definition eval_pred :: "('a fo_term) list \ 'a table \ ('a, 'c) fo_t" where "eval_pred ts X = (let AD = \(set (map set_fo_term ts)) \ \(set ` X) in (AD, length (fv_fo_terms_list ts), eval_table ts (map Inl ` X)))" definition eval_bool :: "bool \ ('a, 'c) fo_t" where "eval_bool b = (if b then ({}, 0, {[]}) else ({}, 0, {}))" definition eval_eq :: "'a fo_term \ 'a fo_term \ ('a, nat) fo_t" where "eval_eq t t' = (case t of Var n \ (case t' of Var n' \ if n = n' then ({}, 1, {[Inr 0]}) else ({}, 2, {[Inr 0, Inr 0]}) | Const c' \ ({c'}, 1, {[Inl c']})) | Const c \ (case t' of Var n' \ ({c}, 1, {[Inl c]}) | Const c' \ if c = c' then ({c}, 0, {[]}) else ({c, c'}, 0, {})))" fun eval_neg :: "nat list \ ('a, nat) fo_t \ ('a, nat) fo_t" where "eval_neg ns (AD, _, X) = (AD, length ns, nall_tuples AD (length ns) - X)" definition "eval_conj_tuple AD ns\ ns\ xs ys = (let cxs = filter (\(n, x). n \ set ns\ \ isl x) (zip ns\ xs); nxs = map fst (filter (\(n, x). n \ set ns\ \ \isl x) (zip ns\ xs)); cys = filter (\(n, y). n \ set ns\ \ isl y) (zip ns\ ys); nys = map fst (filter (\(n, y). n \ set ns\ \ \isl y) (zip ns\ ys)) in fo_nmlz AD ` ext_tuple {} (sort (ns\ @ map fst cys)) nys (map snd (merge (zip ns\ xs) cys)) \ fo_nmlz AD ` ext_tuple {} (sort (ns\ @ map fst cxs)) nxs (map snd (merge (zip ns\ ys) cxs)))" definition "eval_conj_set AD ns\ X\ ns\ X\ = \((\xs. \(eval_conj_tuple AD ns\ ns\ xs ` X\)) ` X\)" definition "idx_join AD ns ns\ X\ ns\ X\ = (let idx\' = cluster (Some \ (\xs. fo_nmlz AD (proj_tuple ns (zip ns\ xs)))) X\; idx\' = cluster (Some \ (\ys. fo_nmlz AD (proj_tuple ns (zip ns\ ys)))) X\ in set_of_idx (mapping_join (\X\'' X\''. eval_conj_set AD ns\ X\'' ns\ X\'') idx\' idx\'))" fun eval_conj :: "nat list \ ('a, nat) fo_t \ nat list \ ('a, nat) fo_t \ ('a, nat) fo_t" where "eval_conj ns\ (AD\, _, X\) ns\ (AD\, _, X\) = (let AD = AD\ \ AD\; AD\\ = AD - AD\; AD\\ = AD - AD\; ns = filter (\n. n \ set ns\) ns\ in (AD, card (set ns\ \ set ns\), idx_join AD ns ns\ (ad_agr_close_set AD\\ X\) ns\ (ad_agr_close_set AD\\ X\)))" fun eval_ajoin :: "nat list \ ('a, nat) fo_t \ nat list \ ('a, nat) fo_t \ ('a, nat) fo_t" where "eval_ajoin ns\ (AD\, _, X\) ns\ (AD\, _, X\) = (let AD = AD\ \ AD\; AD\\ = AD - AD\; AD\\ = AD - AD\; ns = filter (\n. n \ set ns\) ns\; ns\' = filter (\n. n \ set ns\) ns\; idx\ = cluster (Some \ (\xs. fo_nmlz AD\ (proj_tuple ns (zip ns\ xs)))) (ad_agr_close_set AD\\ X\); idx\ = cluster (Some \ (\ys. fo_nmlz AD\ (proj_tuple ns (zip ns\ ys)))) X\ in (AD, card (set ns\ \ set ns\), set_of_idx (Mapping.map_values (\xs X. case Mapping.lookup idx\ xs of Some Y \ idx_join AD ns ns\ X ns\ (ad_agr_close_set AD\\ (ext_tuple_set AD\ ns ns\' {xs} - Y)) | _ \ ext_tuple_set AD ns\ ns\' X) idx\)))" fun eval_disj :: "nat list \ ('a, nat) fo_t \ nat list \ ('a, nat) fo_t \ ('a, nat) fo_t" where "eval_disj ns\ (AD\, _, X\) ns\ (AD\, _, X\) = (let AD = AD\ \ AD\; ns\' = filter (\n. n \ set ns\) ns\; ns\' = filter (\n. n \ set ns\) ns\; AD\\ = AD - AD\; AD\\ = AD - AD\ in (AD, card (set ns\ \ set ns\), ext_tuple_set AD ns\ ns\' (ad_agr_close_set AD\\ X\) \ ext_tuple_set AD ns\ ns\' (ad_agr_close_set AD\\ X\)))" fun eval_exists :: "nat \ nat list \ ('a, nat) fo_t \ ('a, nat) fo_t" where "eval_exists i ns (AD, _, X) = (case pos i ns of Some j \ (AD, length ns - 1, fo_nmlz AD ` rem_nth j ` X) | None \ (AD, length ns, X))" fun eval_forall :: "nat \ nat list \ ('a, nat) fo_t \ ('a, nat) fo_t" where "eval_forall i ns (AD, _, X) = (case pos i ns of Some j \ let n = card AD in (AD, length ns - 1, Mapping.keys (Mapping.filter (\t Z. n + card (Inr -` set t) + 1 \ card Z) (cluster (Some \ (\ts. fo_nmlz AD (rem_nth j ts))) X))) | None \ (AD, length ns, X))" lemma combine_map2: assumes "length ys = length xs" "length ys' = length xs'" "distinct xs" "distinct xs'" "set xs \ set xs' = {}" shows "\f. ys = map f xs \ ys' = map f xs'" proof - obtain f g where fg_def: "ys = map f xs" "ys' = map g xs'" using assms exists_map by metis show ?thesis using assms by (auto simp: fg_def intro!: exI[of _ "\x. if x \ set xs then f x else g x"]) qed lemma combine_map3: assumes "length ys = length xs" "length ys' = length xs'" "length ys'' = length xs''" "distinct xs" "distinct xs'" "distinct xs''" "set xs \ set xs' = {}" "set xs \ set xs'' = {}" "set xs' \ set xs'' = {}" shows "\f. ys = map f xs \ ys' = map f xs' \ ys'' = map f xs''" proof - obtain f g h where fgh_def: "ys = map f xs" "ys' = map g xs'" "ys'' = map h xs''" using assms exists_map by metis show ?thesis using assms by (auto simp: fgh_def intro!: exI[of _ "\x. if x \ set xs then f x else if x \ set xs' then g x else h x"]) qed lemma distinct_set_zip: "length nsx = length xs \ distinct nsx \ (a, b) \ set (zip nsx xs) \ (a, ba) \ set (zip nsx xs) \ b = ba" by (induction nsx xs rule: list_induct2) (auto dest: set_zip_leftD) lemma fo_nmlz_idem_isl: assumes "\x. x \ set xs \ (case x of Inl z \ z \ X | _ \ False)" shows "fo_nmlz X xs = xs" proof - have F1: "Inl x \ set xs \ x \ X" for x using assms[of "Inl x"] by auto have F2: "List.map_filter (case_sum Map.empty Some) xs = []" using assms by (induction xs) (fastforce simp: List.map_filter_def split: sum.splits)+ show ?thesis by (rule fo_nmlz_idem) (auto simp: fo_nmlzd_def nats_def F2 intro: F1) qed lemma set_zip_mapI: "x \ set xs \ (f x, g x) \ set (zip (map f xs) (map g xs))" by (induction xs) auto lemma ad_agr_list_fo_nmlzd_isl: assumes "ad_agr_list X (map f xs) (map g xs)" "fo_nmlzd X (map f xs)" "x \ set xs" "isl (f x)" shows "f x = g x" proof - have AD: "ad_equiv_pair X (f x, g x)" using assms(1) set_zip_mapI[OF assms(3)] by (auto simp: ad_agr_list_def ad_equiv_list_def split: sum.splits) then show ?thesis using assms(2-) by (auto simp: fo_nmlzd_def) (metis AD ad_equiv_pair.simps ad_equiv_pair_mono image_eqI sum.collapse(1) vimageI) qed lemma eval_conj_tuple_close_empty2: assumes "fo_nmlzd X xs" "fo_nmlzd Y ys" "length nsx = length xs" "length nsy = length ys" "sorted_distinct nsx" "sorted_distinct nsy" "sorted_distinct ns" "set ns \ set nsx \ set nsy" "fo_nmlz (X \ Y) (proj_tuple ns (zip nsx xs)) \ fo_nmlz (X \ Y) (proj_tuple ns (zip nsy ys)) \ (proj_tuple ns (zip nsx xs) \ proj_tuple ns (zip nsy ys) \ (\x \ set (proj_tuple ns (zip nsx xs)). isl x) \ (\y \ set (proj_tuple ns (zip nsy ys)). isl y))" "xs' \ ad_agr_close ((X \ Y) - X) xs" "ys' \ ad_agr_close ((X \ Y) - Y) ys" shows "eval_conj_tuple (X \ Y) nsx nsy xs' ys' = {}" proof - define cxs where "cxs = filter (\(n, x). n \ set nsy \ isl x) (zip nsx xs')" define nxs where "nxs = map fst (filter (\(n, x). n \ set nsy \ \isl x) (zip nsx xs'))" define cys where "cys = filter (\(n, y). n \ set nsx \ isl y) (zip nsy ys')" define nys where "nys = map fst (filter (\(n, y). n \ set nsx \ \isl y) (zip nsy ys'))" define both where "both = sorted_list_of_set (set nsx \ set nsy)" have close: "fo_nmlzd (X \ Y) xs'" "ad_agr_list X xs xs'" "fo_nmlzd (X \ Y) ys'" "ad_agr_list Y ys ys'" using ad_agr_close_sound[OF assms(10) assms(1)] ad_agr_close_sound[OF assms(11) assms(2)] by (auto simp add: sup_left_commute) have close': "length xs' = length xs" "length ys' = length ys" using close by (auto simp: ad_agr_list_length) have len_sort: "length (sort (nsx @ map fst cys)) = length (map snd (merge (zip nsx xs') cys))" "length (sort (nsy @ map fst cxs)) = length (map snd (merge (zip nsy ys') cxs))" by (auto simp: merge_length assms(3,4) close') { fix zs assume "zs \ fo_nmlz (X \ Y) ` (\fs. map snd (merge (zip (sort (nsx @ map fst cys)) (map snd (merge (zip nsx xs') cys))) (zip nys fs))) ` nall_tuples_rec {} (card (Inr -` set (map snd (merge (zip nsx xs') cys)))) (length nys)" "zs \ fo_nmlz (X \ Y) ` (\fs. map snd (merge (zip (sort (nsy @ map fst cxs)) (map snd (merge (zip nsy ys') cxs))) (zip nxs fs))) ` nall_tuples_rec {} (card (Inr -` set (map snd (merge (zip nsy ys') cxs)))) (length nxs)" then obtain zxs zys where nall: "zxs \ nall_tuples_rec {} (card (Inr -` set (map snd (merge (zip nsx xs') cys)))) (length nys)" "zs = fo_nmlz (X \ Y) (map snd (merge (zip (sort (nsx @ map fst cys)) (map snd (merge (zip nsx xs') cys))) (zip nys zxs)))" "zys \ nall_tuples_rec {} (card (Inr -` set (map snd (merge (zip nsy ys') cxs)))) (length nxs)" "zs = fo_nmlz (X \ Y) (map snd (merge (zip (sort (nsy @ map fst cxs)) (map snd (merge (zip nsy ys') cxs))) (zip nxs zys)))" by auto have len_zs: "length zxs = length nys" "length zys = length nxs" using nall(1,3) by (auto dest: nall_tuples_rec_length) have aux: "sorted_distinct (map fst cxs)" "sorted_distinct nxs" "sorted_distinct nsy" "sorted_distinct (map fst cys)" "sorted_distinct nys" "sorted_distinct nsx" "set (map fst cxs) \ set nsy = {}" "set (map fst cxs) \ set nxs = {}" "set nsy \ set nxs = {}" "set (map fst cys) \ set nsx = {}" "set (map fst cys) \ set nys = {}" "set nsx \ set nys = {}" using assms(3,4,5,6) close' distinct_set_zip by (auto simp: cxs_def nxs_def cys_def nys_def sorted_filter distinct_map_fst_filter) (smt (z3) distinct_set_zip)+ obtain xf where xf_def: "map snd cxs = map xf (map fst cxs)" "ys' = map xf nsy" "zys = map xf nxs" using combine_map3[where ?ys="map snd cxs" and ?xs="map fst cxs" and ?ys'=ys' and ?xs'=nsy and ?ys''=zys and ?xs''=nxs] assms(4) aux close' by (auto simp: len_zs) obtain ysf where ysf_def: "ys = map ysf nsy" using assms(4,6) exists_map by auto obtain xg where xg_def: "map snd cys = map xg (map fst cys)" "xs' = map xg nsx" "zxs = map xg nys" using combine_map3[where ?ys="map snd cys" and ?xs="map fst cys" and ?ys'=xs' and ?xs'=nsx and ?ys''=zxs and ?xs''=nys] assms(3) aux close' by (auto simp: len_zs) obtain xsf where xsf_def: "xs = map xsf nsx" using assms(3,5) exists_map by auto have set_cxs_nxs: "set (map fst cxs @ nxs) = set nsx - set nsy" using assms(3) unfolding cxs_def nxs_def close'[symmetric] by (induction nsx xs' rule: list_induct2) auto have set_cys_nys: "set (map fst cys @ nys) = set nsy - set nsx" using assms(4) unfolding cys_def nys_def close'[symmetric] by (induction nsy ys' rule: list_induct2) auto have sort_sort_both_xs: "sort (sort (nsy @ map fst cxs) @ nxs) = both" apply (rule sorted_distinct_set_unique) using assms(3,5,6) close' set_cxs_nxs by (auto simp: both_def nxs_def cxs_def intro: distinct_map_fst_filter) (metis (no_types, lifting) distinct_set_zip) have sort_sort_both_ys: "sort (sort (nsx @ map fst cys) @ nys) = both" apply (rule sorted_distinct_set_unique) using assms(4,5,6) close' set_cys_nys by (auto simp: both_def nys_def cys_def intro: distinct_map_fst_filter) (metis (no_types, lifting) distinct_set_zip) have "map snd (merge (zip nsy ys') cxs) = map xf (sort (nsy @ map fst cxs))" using merge_map[where ?\=xf and ?ns=nsy and ?ms="map fst cxs"] assms(6) aux unfolding xf_def(1)[symmetric] xf_def(2) by (auto simp: zip_map_fst_snd) then have zs_xf: "zs = fo_nmlz (X \ Y) (map xf both)" using merge_map[where \=xf and ?ns="sort (nsy @ map fst cxs)" and ?ms=nxs] aux by (fastforce simp: nall(4) xf_def(3) sort_sort_both_xs) have "map snd (merge (zip nsx xs') cys) = map xg (sort (nsx @ map fst cys))" using merge_map[where ?\=xg and ?ns=nsx and ?ms="map fst cys"] assms(5) aux unfolding xg_def(1)[symmetric] xg_def(2) by (fastforce simp: zip_map_fst_snd) then have zs_xg: "zs = fo_nmlz (X \ Y) (map xg both)" using merge_map[where \=xg and ?ns="sort (nsx @ map fst cys)" and ?ms=nys] aux by (fastforce simp: nall(2) xg_def(3) sort_sort_both_ys) have proj_map: "proj_tuple ns (zip nsx xs') = map xg ns" "proj_tuple ns (zip nsy ys') = map xf ns" "proj_tuple ns (zip nsx xs) = map xsf ns" "proj_tuple ns (zip nsy ys) = map ysf ns" unfolding xf_def(2) xg_def(2) xsf_def ysf_def using assms(5,6,7,8) proj_tuple_map by auto have "ad_agr_list (X \ Y) (map xg both) (map xf both)" using zs_xg zs_xf by (fastforce dest: fo_nmlz_eqD) then have "ad_agr_list (X \ Y) (proj_tuple ns (zip nsx xs')) (proj_tuple ns (zip nsy ys'))" using assms(8) unfolding proj_map by (fastforce simp: both_def intro: ad_agr_list_subset[rotated]) then have fo_nmlz_Un: "fo_nmlz (X \ Y) (proj_tuple ns (zip nsx xs')) = fo_nmlz (X \ Y) (proj_tuple ns (zip nsy ys'))" by (auto intro: fo_nmlz_eqI) have "False" using assms(9) proof (rule disjE) assume c: "fo_nmlz (X \ Y) (proj_tuple ns (zip nsx xs)) \ fo_nmlz (X \ Y) (proj_tuple ns (zip nsy ys))" have fo_nmlz_Int: "fo_nmlz (X \ Y) (proj_tuple ns (zip nsx xs')) = fo_nmlz (X \ Y) (proj_tuple ns (zip nsy ys'))" using fo_nmlz_Un by (rule fo_nmlz_eqI[OF ad_agr_list_mono, rotated, OF fo_nmlz_eqD]) auto have proj_xs: "fo_nmlz (X \ Y) (proj_tuple ns (zip nsx xs)) = fo_nmlz (X \ Y) (proj_tuple ns (zip nsx xs'))" unfolding proj_map apply (rule fo_nmlz_eqI) apply (rule ad_agr_list_mono[OF Int_lower1]) apply (rule ad_agr_list_subset[OF _ close(2)[unfolded xsf_def xg_def(2)]]) using assms(8) apply (auto) done have proj_ys: "fo_nmlz (X \ Y) (proj_tuple ns (zip nsy ys)) = fo_nmlz (X \ Y) (proj_tuple ns (zip nsy ys'))" unfolding proj_map apply (rule fo_nmlz_eqI) apply (rule ad_agr_list_mono[OF Int_lower2]) apply (rule ad_agr_list_subset[OF _ close(4)[unfolded ysf_def xf_def(2)]]) using assms(8) apply (auto) done show "False" using c fo_nmlz_Int proj_xs proj_ys by auto next assume c: "proj_tuple ns (zip nsx xs) \ proj_tuple ns (zip nsy ys) \ (\x\set (proj_tuple ns (zip nsx xs)). isl x) \ (\y\set (proj_tuple ns (zip nsy ys)). isl y)" have "case x of Inl z \ z \ X \ Y | Inr b \ False" if "x \ set (proj_tuple ns (zip nsx xs'))" for x using close(2) assms(1,8) c that ad_agr_list_fo_nmlzd_isl[where ?X=X and ?f=xsf and ?g=xg and ?xs=nsx] unfolding proj_map unfolding xsf_def xg_def(2) apply (auto simp: fo_nmlzd_def split: sum.splits) apply (metis image_eqI subsetD vimageI) apply (metis subsetD sum.disc(2)) done then have E1: "fo_nmlz (X \ Y) (proj_tuple ns (zip nsx xs')) = proj_tuple ns (zip nsx xs')" by (rule fo_nmlz_idem_isl) have "case y of Inl z \ z \ X \ Y | Inr b \ False" if "y \ set (proj_tuple ns (zip nsy ys'))" for y using close(4) assms(2,8) c that ad_agr_list_fo_nmlzd_isl[where ?X=Y and ?f=ysf and ?g=xf and ?xs=nsy] unfolding proj_map unfolding ysf_def xf_def(2) apply (auto simp: fo_nmlzd_def split: sum.splits) apply (metis image_eqI subsetD vimageI) apply (metis subsetD sum.disc(2)) done then have E2: "fo_nmlz (X \ Y) (proj_tuple ns (zip nsy ys')) = proj_tuple ns (zip nsy ys')" by (rule fo_nmlz_idem_isl) have ad: "ad_agr_list X (map xsf ns) (map xg ns)" using assms(8) close(2)[unfolded xsf_def xg_def(2)] ad_agr_list_subset by blast have "\x\set (proj_tuple ns (zip nsx xs)). isl x" using c by auto then have E3: "proj_tuple ns (zip nsx xs) = proj_tuple ns (zip nsx xs')" using assms(8) unfolding proj_map apply (induction ns) using ad_agr_list_fo_nmlzd_isl[OF close(2)[unfolded xsf_def xg_def(2)] assms(1)[unfolded xsf_def]] by auto have "\x\set (proj_tuple ns (zip nsy ys)). isl x" using c by auto then have E4: "proj_tuple ns (zip nsy ys) = proj_tuple ns (zip nsy ys')" using assms(8) unfolding proj_map apply (induction ns) using ad_agr_list_fo_nmlzd_isl[OF close(4)[unfolded ysf_def xf_def(2)] assms(2)[unfolded ysf_def]] by auto show "False" using c fo_nmlz_Un unfolding E1 E2 E3 E4 by auto qed } then show ?thesis by (auto simp: eval_conj_tuple_def Let_def cxs_def[symmetric] nxs_def[symmetric] cys_def[symmetric] nys_def[symmetric] ext_tuple_eq[OF len_sort(1)] ext_tuple_eq[OF len_sort(2)]) qed lemma eval_conj_tuple_close_empty: assumes "fo_nmlzd X xs" "fo_nmlzd Y ys" "length nsx = length xs" "length nsy = length ys" "sorted_distinct nsx" "sorted_distinct nsy" "ns = filter (\n. n \ set nsy) nsx" "fo_nmlz (X \ Y) (proj_tuple ns (zip nsx xs)) \ fo_nmlz (X \ Y) (proj_tuple ns (zip nsy ys))" "xs' \ ad_agr_close ((X \ Y) - X) xs" "ys' \ ad_agr_close ((X \ Y) - Y) ys" shows "eval_conj_tuple (X \ Y) nsx nsy xs' ys' = {}" proof - have aux: "sorted_distinct ns" "set ns \ set nsx \ set nsy" using assms(5) sorted_filter[of id] by (auto simp: assms(7)) show ?thesis using eval_conj_tuple_close_empty2[OF assms(1-6) aux] assms(8-) by auto qed lemma eval_conj_tuple_empty2: assumes "fo_nmlzd Z xs" "fo_nmlzd Z ys" "length nsx = length xs" "length nsy = length ys" "sorted_distinct nsx" "sorted_distinct nsy" "sorted_distinct ns" "set ns \ set nsx \ set nsy" "fo_nmlz Z (proj_tuple ns (zip nsx xs)) \ fo_nmlz Z (proj_tuple ns (zip nsy ys)) \ (proj_tuple ns (zip nsx xs) \ proj_tuple ns (zip nsy ys) \ (\x \ set (proj_tuple ns (zip nsx xs)). isl x) \ (\y \ set (proj_tuple ns (zip nsy ys)). isl y))" shows "eval_conj_tuple Z nsx nsy xs ys = {}" using eval_conj_tuple_close_empty2[OF assms(1-8)] assms(9) ad_agr_close_empty assms(1-2) by fastforce lemma eval_conj_tuple_empty: assumes "fo_nmlzd Z xs" "fo_nmlzd Z ys" "length nsx = length xs" "length nsy = length ys" "sorted_distinct nsx" "sorted_distinct nsy" "ns = filter (\n. n \ set nsy) nsx" "fo_nmlz Z (proj_tuple ns (zip nsx xs)) \ fo_nmlz Z (proj_tuple ns (zip nsy ys))" shows "eval_conj_tuple Z nsx nsy xs ys = {}" proof - have aux: "sorted_distinct ns" "set ns \ set nsx \ set nsy" using assms(5) sorted_filter[of id] by (auto simp: assms(7)) show ?thesis using eval_conj_tuple_empty2[OF assms(1-6) aux] assms(8-) by auto qed lemma nall_tuples_rec_filter: assumes "xs \ nall_tuples_rec AD n (length xs)" "ys = filter (\x. \isl x) xs" shows "ys \ nall_tuples_rec {} n (length ys)" using assms proof (induction xs arbitrary: n ys) case (Cons x xs) then show ?case proof (cases x) case (Inr b) have b_le_i: "b \ n" using Cons(2) by (auto simp: Inr) obtain zs where ys_def: "ys = Inr b # zs" "zs = filter (\x. \ isl x) xs" using Cons(3) by (auto simp: Inr) show ?thesis proof (cases "b < n") case True then show ?thesis using Cons(1)[OF _ ys_def(2), of n] Cons(2) by (auto simp: Inr ys_def(1)) next case False then show ?thesis using Cons(1)[OF _ ys_def(2), of "Suc n"] Cons(2) by (auto simp: Inr ys_def(1)) qed qed auto qed auto lemma nall_tuples_rec_filter_rev: assumes "ys \ nall_tuples_rec {} n (length ys)" "ys = filter (\x. \isl x) xs" "Inl -` set xs \ AD" shows "xs \ nall_tuples_rec AD n (length xs)" using assms proof (induction xs arbitrary: n ys) case (Cons x xs) show ?case proof (cases x) case (Inl a) have a_AD: "a \ AD" using Cons(4) by (auto simp: Inl) show ?thesis using Cons(1)[OF Cons(2)] Cons(3,4) a_AD by (auto simp: Inl) next case (Inr b) obtain zs where ys_def: "ys = Inr b # zs" "zs = filter (\x. \ isl x) xs" using Cons(3) by (auto simp: Inr) show ?thesis using Cons(1)[OF _ ys_def(2)] Cons(2,4) by (fastforce simp: ys_def(1) Inr) qed qed auto lemma eval_conj_set_aux: fixes AD :: "'a set" assumes ns\'_def: "ns\' = filter (\n. n \ set ns\) ns\" and ns\'_def: "ns\' = filter (\n. n \ set ns\) ns\" and X\_def: "X\ = fo_nmlz AD ` proj_vals R\ ns\" and X\_def: "X\ = fo_nmlz AD ` proj_vals R\ ns\" and distinct: "sorted_distinct ns\" "sorted_distinct ns\" and cxs_def: "cxs = filter (\(n, x). n \ set ns\ \ isl x) (zip ns\ xs)" and nxs_def: "nxs = map fst (filter (\(n, x). n \ set ns\ \ \isl x) (zip ns\ xs))" and cys_def: "cys = filter (\(n, y). n \ set ns\ \ isl y) (zip ns\ ys)" and nys_def: "nys = map fst (filter (\(n, y). n \ set ns\ \ \isl y) (zip ns\ ys))" and xs_ys_def: "xs \ X\" "ys \ X\" and \xs_def: "xs = map \xs ns\" "fs\ = map \xs ns\'" and \ys_def: "ys = map \ys ns\" "fs\ = map \ys ns\'" and fs\_def: "fs\ \ nall_tuples_rec AD (card (Inr -` set xs)) (length ns\')" and fs\_def: "fs\ \ nall_tuples_rec AD (card (Inr -` set ys)) (length ns\')" and ad_agr: "ad_agr_list AD (map \ys (sort (ns\ @ ns\'))) (map \xs (sort (ns\ @ ns\')))" shows "map snd (merge (zip ns\ xs) (zip ns\' fs\)) = map snd (merge (zip (sort (ns\ @ map fst cys)) (map \xs (sort (ns\ @ map fst cys)))) (zip nys (map \xs nys)))" and "map snd (merge (zip ns\ xs) cys) = map \xs (sort (ns\ @ map fst cys))" and "map \xs nys \ nall_tuples_rec {} (card (Inr -` set (map \xs (sort (ns\ @ map fst cys))))) (length nys)" proof - have len_xs_ys: "length xs = length ns\" "length ys = length ns\" using xs_ys_def by (auto simp: X\_def X\_def proj_vals_def fo_nmlz_length) have len_fs\: "length fs\ = length ns\'" using \xs_def(2) by auto have set_ns\': "set ns\' = set (map fst cys) \ set nys" using len_xs_ys(2) by (auto simp: ns\'_def cys_def nys_def dest: set_zip_leftD) (metis (no_types, lifting) image_eqI in_set_impl_in_set_zip1 mem_Collect_eq prod.sel(1) split_conv) have "\x. Inl x \ set xs \ set fs\ \ x \ AD" "\y. Inl y \ set ys \ set fs\ \ y \ AD" using xs_ys_def fo_nmlz_set[of AD] nall_tuples_rec_Inl[OF fs\_def] nall_tuples_rec_Inl[OF fs\_def] by (auto simp: X\_def X\_def) then have Inl_xs_ys: "\n. n \ set ns\ \ set ns\ \ isl (\xs n) \ (\x. \xs n = Inl x \ x \ AD)" "\n. n \ set ns\ \ set ns\ \ isl (\ys n) \ (\y. \ys n = Inl y \ y \ AD)" unfolding \xs_def \ys_def ns\'_def ns\'_def by (auto simp: isl_def) (smt imageI mem_Collect_eq)+ have sort_sort: "sort (ns\ @ ns\') = sort (ns\ @ ns\')" apply (rule sorted_distinct_set_unique) using distinct by (auto simp: ns\'_def ns\'_def) have isl_iff: "\n. n \ set ns\' \ set ns\' \ isl (\xs n) \ isl (\ys n) \ \xs n = \ys n" using ad_agr Inl_xs_ys unfolding sort_sort[symmetric] ad_agr_list_link[symmetric] unfolding ns\'_def ns\'_def apply (auto simp: ad_agr_sets_def) unfolding ad_equiv_pair.simps apply (metis (no_types, lifting) UnI2 image_eqI mem_Collect_eq) apply (metis (no_types, lifting) UnI2 image_eqI mem_Collect_eq) apply (metis (no_types, lifting) UnI1 image_eqI)+ done have "\n. n \ set (map fst cys) \ isl (\xs n)" "\n. n \ set (map fst cxs) \ isl (\ys n)" using isl_iff by (auto simp: cys_def ns\'_def \ys_def(1) cxs_def ns\'_def \xs_def(1) set_zip) (metis nth_mem)+ then have Inr_sort: "Inr -` set (map \xs (sort (ns\ @ map fst cys))) = Inr -` set xs" unfolding \xs_def(1) \ys_def(1) by (auto simp: zip_map_fst_snd dest: set_zip_leftD) (metis fst_conv image_iff sum.disc(2))+ have map_nys: "map \xs nys = filter (\x. \isl x) fs\" using isl_iff[unfolded ns\'_def] unfolding nys_def \ys_def(1) \xs_def(2) ns\'_def filter_map by (induction ns\) force+ have map_nys_in_nall: "map \xs nys \ nall_tuples_rec {} (card (Inr -` set xs)) (length nys)" using nall_tuples_rec_filter[OF fs\_def[folded len_fs\] map_nys] by auto have map_cys: "map snd cys = map \xs (map fst cys)" using isl_iff by (auto simp: cys_def set_zip ns\'_def \ys_def(1)) (metis nth_mem) show merge_xs_cys: "map snd (merge (zip ns\ xs) cys) = map \xs (sort (ns\ @ map fst cys))" apply (subst zip_map_fst_snd[of cys, symmetric]) unfolding \xs_def(1) map_cys apply (rule merge_map) using distinct by (auto simp: cys_def \ys_def sorted_filter distinct_map_filter map_fst_zip_take) have merge_nys_prems: "sorted_distinct (sort (ns\ @ map fst cys))" "sorted_distinct nys" "set (sort (ns\ @ map fst cys)) \ set nys = {}" using distinct len_xs_ys(2) by (auto simp: cys_def nys_def distinct_map_filter sorted_filter) (metis eq_key_imp_eq_value map_fst_zip) have map_snd_merge_nys: "map \xs (sort (sort (ns\ @ map fst cys) @ nys)) = map snd (merge (zip (sort (ns\ @ map fst cys)) (map \xs (sort (ns\ @ map fst cys)))) (zip nys (map \xs nys)))" by (rule merge_map[OF merge_nys_prems, symmetric]) have sort_sort_nys: "sort (sort (ns\ @ map fst cys) @ nys) = sort (ns\ @ ns\')" apply (rule sorted_distinct_set_unique) using distinct merge_nys_prems set_ns\' by (auto simp: cys_def nys_def ns\'_def dest: set_zip_leftD) have map_merge_fs\: "map snd (merge (zip ns\ xs) (zip ns\' fs\)) = map \xs (sort (ns\ @ ns\'))" unfolding \xs_def apply (rule merge_map) using distinct sorted_filter[of id] by (auto simp: ns\'_def) show "map snd (merge (zip ns\ xs) (zip ns\' fs\)) = map snd (merge (zip (sort (ns\ @ map fst cys)) (map \xs (sort (ns\ @ map fst cys)))) (zip nys (map \xs nys)))" unfolding map_merge_fs\ map_snd_merge_nys[unfolded sort_sort_nys] by auto show "map \xs nys \ nall_tuples_rec {} (card (Inr -` set (map \xs (sort (ns\ @ map fst cys))))) (length nys)" using map_nys_in_nall unfolding Inr_sort[symmetric] by auto qed lemma eval_conj_set_aux': fixes AD :: "'a set" assumes ns\'_def: "ns\' = filter (\n. n \ set ns\) ns\" and ns\'_def: "ns\' = filter (\n. n \ set ns\) ns\" and X\_def: "X\ = fo_nmlz AD ` proj_vals R\ ns\" and X\_def: "X\ = fo_nmlz AD ` proj_vals R\ ns\" and distinct: "sorted_distinct ns\" "sorted_distinct ns\" and cxs_def: "cxs = filter (\(n, x). n \ set ns\ \ isl x) (zip ns\ xs)" and nxs_def: "nxs = map fst (filter (\(n, x). n \ set ns\ \ \isl x) (zip ns\ xs))" and cys_def: "cys = filter (\(n, y). n \ set ns\ \ isl y) (zip ns\ ys)" and nys_def: "nys = map fst (filter (\(n, y). n \ set ns\ \ \isl y) (zip ns\ ys))" and xs_ys_def: "xs \ X\" "ys \ X\" and \xs_def: "xs = map \xs ns\" "map snd cys = map \xs (map fst cys)" "ys\ = map \xs nys" and \ys_def: "ys = map \ys ns\" "map snd cxs = map \ys (map fst cxs)" "xs\ = map \ys nxs" and fs\_def: "fs\ = map \xs ns\'" and fs\_def: "fs\ = map \ys ns\'" and ys\_def: "map \xs nys \ nall_tuples_rec {} (card (Inr -` set (map \xs (sort (ns\ @ map fst cys))))) (length nys)" and Inl_set_AD: "Inl -` (set (map snd cxs) \ set xs\) \ AD" "Inl -` (set (map snd cys) \ set ys\) \ AD" and ad_agr: "ad_agr_list AD (map \ys (sort (ns\ @ ns\'))) (map \xs (sort (ns\ @ ns\')))" shows "map snd (merge (zip ns\ xs) (zip ns\' fs\)) = map snd (merge (zip (sort (ns\ @ map fst cys)) (map \xs (sort (ns\ @ map fst cys)))) (zip nys (map \xs nys)))" and "map snd (merge (zip ns\ xs) cys) = map \xs (sort (ns\ @ map fst cys))" "fs\ \ nall_tuples_rec AD (card (Inr -` set xs)) (length ns\')" proof - have len_xs_ys: "length xs = length ns\" "length ys = length ns\" using xs_ys_def by (auto simp: X\_def X\_def proj_vals_def fo_nmlz_length) have len_fs\: "length fs\ = length ns\'" by (auto simp: fs\_def) have set_ns: "set ns\' = set (map fst cys) \ set nys" "set ns\' = set (map fst cxs) \ set nxs" using len_xs_ys by (auto simp: ns\'_def cys_def nys_def ns\'_def cxs_def nxs_def dest: set_zip_leftD) (metis (no_types, lifting) image_eqI in_set_impl_in_set_zip1 mem_Collect_eq prod.sel(1) split_conv)+ then have set_\_ns: "\xs ` set ns\' \ \xs ` set ns\' \ set xs \ set (map snd cys) \ set ys\" "\ys ` set ns\' \ \ys ` set ns\' \ set ys \ set (map snd cxs) \ set xs\" by (auto simp: \xs_def \ys_def ns\'_def ns\'_def) have Inl_sub_AD: "\x. Inl x \ set xs \ set (map snd cys) \ set ys\ \ x \ AD" "\y. Inl y \ set ys \ set (map snd cxs) \ set xs\ \ y \ AD" using xs_ys_def fo_nmlz_set[of AD] Inl_set_AD by (auto simp: X\_def X\_def) (metis in_set_zipE set_map subset_eq vimageI zip_map_fst_snd)+ then have Inl_xs_ys: "\n. n \ set ns\' \ set ns\' \ isl (\xs n) \ (\x. \xs n = Inl x \ x \ AD)" "\n. n \ set ns\' \ set ns\' \ isl (\ys n) \ (\y. \ys n = Inl y \ y \ AD)" using set_\_ns by (auto simp: isl_def rev_image_eqI) have sort_sort: "sort (ns\ @ ns\') = sort (ns\ @ ns\')" apply (rule sorted_distinct_set_unique) using distinct by (auto simp: ns\'_def ns\'_def) have isl_iff: "\n. n \ set ns\' \ set ns\' \ isl (\xs n) \ isl (\ys n) \ \xs n = \ys n" using ad_agr Inl_xs_ys unfolding sort_sort[symmetric] ad_agr_list_link[symmetric] unfolding ns\'_def ns\'_def apply (auto simp: ad_agr_sets_def) unfolding ad_equiv_pair.simps apply (metis (no_types, lifting) UnI2 image_eqI mem_Collect_eq) apply (metis (no_types, lifting) UnI2 image_eqI mem_Collect_eq) apply (metis (no_types, lifting) UnI1 image_eqI)+ done have "\n. n \ set (map fst cys) \ isl (\xs n)" "\n. n \ set (map fst cxs) \ isl (\ys n)" using isl_iff by (auto simp: cys_def ns\'_def \ys_def(1) cxs_def ns\'_def \xs_def(1) set_zip) (metis nth_mem)+ then have Inr_sort: "Inr -` set (map \xs (sort (ns\ @ map fst cys))) = Inr -` set xs" unfolding \xs_def(1) \ys_def(1) by (auto simp: zip_map_fst_snd dest: set_zip_leftD) (metis fst_conv image_iff sum.disc(2))+ have map_nys: "map \xs nys = filter (\x. \isl x) fs\" using isl_iff[unfolded ns\'_def] unfolding nys_def \ys_def(1) fs\_def ns\'_def by (induction ns\) force+ have map_cys: "map snd cys = map \xs (map fst cys)" using isl_iff by (auto simp: cys_def set_zip ns\'_def \ys_def(1)) (metis nth_mem) show merge_xs_cys: "map snd (merge (zip ns\ xs) cys) = map \xs (sort (ns\ @ map fst cys))" apply (subst zip_map_fst_snd[of cys, symmetric]) unfolding \xs_def(1) map_cys apply (rule merge_map) using distinct by (auto simp: cys_def \ys_def sorted_filter distinct_map_filter map_fst_zip_take) have merge_nys_prems: "sorted_distinct (sort (ns\ @ map fst cys))" "sorted_distinct nys" "set (sort (ns\ @ map fst cys)) \ set nys = {}" using distinct len_xs_ys(2) by (auto simp: cys_def nys_def distinct_map_filter sorted_filter) (metis eq_key_imp_eq_value map_fst_zip) have map_snd_merge_nys: "map \xs (sort (sort (ns\ @ map fst cys) @ nys)) = map snd (merge (zip (sort (ns\ @ map fst cys)) (map \xs (sort (ns\ @ map fst cys)))) (zip nys (map \xs nys)))" by (rule merge_map[OF merge_nys_prems, symmetric]) have sort_sort_nys: "sort (sort (ns\ @ map fst cys) @ nys) = sort (ns\ @ ns\')" apply (rule sorted_distinct_set_unique) using distinct merge_nys_prems set_ns by (auto simp: cys_def nys_def ns\'_def dest: set_zip_leftD) have map_merge_fs\: "map snd (merge (zip ns\ xs) (zip ns\' fs\)) = map \xs (sort (ns\ @ ns\'))" unfolding \xs_def fs\_def apply (rule merge_map) using distinct sorted_filter[of id] by (auto simp: ns\'_def) show "map snd (merge (zip ns\ xs) (zip ns\' fs\)) = map snd (merge (zip (sort (ns\ @ map fst cys)) (map \xs (sort (ns\ @ map fst cys)))) (zip nys (map \xs nys)))" unfolding map_merge_fs\ map_snd_merge_nys[unfolded sort_sort_nys] by auto have "Inl -` set fs\ \ AD" using Inl_sub_AD(1) set_\_ns by (force simp: fs\_def) then show "fs\ \ nall_tuples_rec AD (card (Inr -` set xs)) (length ns\')" unfolding len_fs\[symmetric] using nall_tuples_rec_filter_rev[OF _ map_nys] ys\_def[unfolded Inr_sort] by auto qed lemma eval_conj_set_correct: assumes ns\'_def: "ns\' = filter (\n. n \ set ns\) ns\" and ns\'_def: "ns\' = filter (\n. n \ set ns\) ns\" and X\_def: "X\ = fo_nmlz AD ` proj_vals R\ ns\" and X\_def: "X\ = fo_nmlz AD ` proj_vals R\ ns\" and distinct: "sorted_distinct ns\" "sorted_distinct ns\" shows "eval_conj_set AD ns\ X\ ns\ X\ = ext_tuple_set AD ns\ ns\' X\ \ ext_tuple_set AD ns\ ns\' X\" proof - have aux: "ext_tuple_set AD ns\ ns\' X\ = fo_nmlz AD ` \(ext_tuple AD ns\ ns\' ` X\)" "ext_tuple_set AD ns\ ns\' X\ = fo_nmlz AD ` \(ext_tuple AD ns\ ns\' ` X\)" by (auto simp: ext_tuple_set_def ext_tuple_def X\_def X\_def image_iff fo_nmlz_idem[OF fo_nmlz_sound]) show ?thesis unfolding aux proof (rule set_eqI, rule iffI) fix vs assume "vs \ fo_nmlz AD ` \(ext_tuple AD ns\ ns\' ` X\) \ fo_nmlz AD ` \(ext_tuple AD ns\ ns\' ` X\)" then obtain xs ys where xs_ys_def: "xs \ X\" "vs \ fo_nmlz AD ` ext_tuple AD ns\ ns\' xs" "ys \ X\" "vs \ fo_nmlz AD ` ext_tuple AD ns\ ns\' ys" by auto have len_xs_ys: "length xs = length ns\" "length ys = length ns\" using xs_ys_def(1,3) by (auto simp: X\_def X\_def proj_vals_def fo_nmlz_length) obtain fs\ where fs\_def: "vs = fo_nmlz AD (map snd (merge (zip ns\ xs) (zip ns\' fs\)))" "fs\ \ nall_tuples_rec AD (card (Inr -` set xs)) (length ns\')" using xs_ys_def(1,2) by (auto simp: X\_def proj_vals_def ext_tuple_def split: if_splits) (metis fo_nmlz_map length_map map_snd_zip) obtain fs\ where fs\_def: "vs = fo_nmlz AD (map snd (merge (zip ns\ ys) (zip ns\' fs\)))" "fs\ \ nall_tuples_rec AD (card (Inr -` set ys)) (length ns\')" using xs_ys_def(3,4) by (auto simp: X\_def proj_vals_def ext_tuple_def split: if_splits) (metis fo_nmlz_map length_map map_snd_zip) note len_fs\ = nall_tuples_rec_length[OF fs\_def(2)] note len_fs\ = nall_tuples_rec_length[OF fs\_def(2)] obtain \xs where \xs_def: "xs = map \xs ns\" "fs\ = map \xs ns\'" using exists_map[of "ns\ @ ns\'" "xs @ fs\"] len_xs_ys(1) len_fs\ distinct by (auto simp: ns\'_def) obtain \ys where \ys_def: "ys = map \ys ns\" "fs\ = map \ys ns\'" using exists_map[of "ns\ @ ns\'" "ys @ fs\"] len_xs_ys(2) len_fs\ distinct by (auto simp: ns\'_def) have map_merge_fs\: "map snd (merge (zip ns\ xs) (zip ns\' fs\)) = map \xs (sort (ns\ @ ns\'))" unfolding \xs_def apply (rule merge_map) using distinct sorted_filter[of id] by (auto simp: ns\'_def) have map_merge_fs\: "map snd (merge (zip ns\ ys) (zip ns\' fs\)) = map \ys (sort (ns\ @ ns\'))" unfolding \ys_def apply (rule merge_map) using distinct sorted_filter[of id] by (auto simp: ns\'_def) define cxs where "cxs = filter (\(n, x). n \ set ns\ \ isl x) (zip ns\ xs)" define nxs where "nxs = map fst (filter (\(n, x). n \ set ns\ \ \isl x) (zip ns\ xs))" define cys where "cys = filter (\(n, y). n \ set ns\ \ isl y) (zip ns\ ys)" define nys where "nys = map fst (filter (\(n, y). n \ set ns\ \ \isl y) (zip ns\ ys))" note ad_agr1 = fo_nmlz_eqD[OF trans[OF fs\_def(1)[symmetric] fs\_def(1)], unfolded map_merge_fs\ map_merge_fs\] note ad_agr2 = ad_agr_list_comm[OF ad_agr1] obtain \xs where aux1: "map snd (merge (zip ns\ xs) (zip ns\' fs\)) = map snd (merge (zip (sort (ns\ @ map fst cys)) (map \xs (sort (ns\ @ map fst cys)))) (zip nys (map \xs nys)))" "map snd (merge (zip ns\ xs) cys) = map \xs (sort (ns\ @ map fst cys))" "map \xs nys \ nall_tuples_rec {} (card (Inr -` set (map \xs (sort (ns\ @ map fst cys))))) (length nys)" using eval_conj_set_aux[OF ns\'_def ns\'_def X\_def X\_def distinct cxs_def nxs_def cys_def nys_def xs_ys_def(1,3) \xs_def \ys_def fs\_def(2) fs\_def(2) ad_agr2] by blast obtain \ys where aux2: "map snd (merge (zip ns\ ys) (zip ns\' fs\)) = map snd (merge (zip (sort (ns\ @ map fst cxs)) (map \ys (sort (ns\ @ map fst cxs)))) (zip nxs (map \ys nxs)))" "map snd (merge (zip ns\ ys) cxs) = map \ys (sort (ns\ @ map fst cxs))" "map \ys nxs \ nall_tuples_rec {} (card (Inr -` set (map \ys (sort (ns\ @ map fst cxs))))) (length nxs)" using eval_conj_set_aux[OF ns\'_def ns\'_def X\_def X\_def distinct(2,1) cys_def nys_def cxs_def nxs_def xs_ys_def(3,1) \ys_def \xs_def fs\_def(2) fs\_def(2) ad_agr1] by blast have vs_ext_nys: "vs \ fo_nmlz AD ` ext_tuple {} (sort (ns\ @ map fst cys)) nys (map snd (merge (zip ns\ xs) cys))" using aux1(3) unfolding fs\_def(1) aux1(1) by (simp add: ext_tuple_eq[OF length_map[symmetric]] aux1(2)) have vs_ext_nxs: "vs \ fo_nmlz AD ` ext_tuple {} (sort (ns\ @ map fst cxs)) nxs (map snd (merge (zip ns\ ys) cxs))" using aux2(3) unfolding fs\_def(1) aux2(1) by (simp add: ext_tuple_eq[OF length_map[symmetric]] aux2(2)) show "vs \ eval_conj_set AD ns\ X\ ns\ X\" using vs_ext_nys vs_ext_nxs xs_ys_def(1,3) by (auto simp: eval_conj_set_def eval_conj_tuple_def nys_def cys_def nxs_def cxs_def Let_def) next fix vs assume "vs \ eval_conj_set AD ns\ X\ ns\ X\" then obtain xs ys cxs nxs cys nys where cxs_def: "cxs = filter (\(n, x). n \ set ns\ \ isl x) (zip ns\ xs)" and nxs_def: "nxs = map fst (filter (\(n, x). n \ set ns\ \ \isl x) (zip ns\ xs))" and cys_def: "cys = filter (\(n, y). n \ set ns\ \ isl y) (zip ns\ ys)" and nys_def: "nys = map fst (filter (\(n, y). n \ set ns\ \ \isl y) (zip ns\ ys))" and xs_def: "xs \ X\" "vs \ fo_nmlz AD ` ext_tuple {} (sort (ns\ @ map fst cys)) nys (map snd (merge (zip ns\ xs) cys))" and ys_def: "ys \ X\" "vs \ fo_nmlz AD ` ext_tuple {} (sort (ns\ @ map fst cxs)) nxs (map snd (merge (zip ns\ ys) cxs))" by (auto simp: eval_conj_set_def eval_conj_tuple_def Let_def) (metis (no_types, lifting) image_eqI) have len_xs_ys: "length xs = length ns\" "length ys = length ns\" using xs_def(1) ys_def(1) by (auto simp: X\_def X\_def proj_vals_def fo_nmlz_length) have len_merge_cys: "length (map snd (merge (zip ns\ xs) cys)) = length (sort (ns\ @ map fst cys))" using merge_length[of "zip ns\ xs" cys] len_xs_ys by auto obtain ys\ where ys\_def: "vs = fo_nmlz AD (map snd (merge (zip (sort (ns\ @ map fst cys)) (map snd (merge (zip ns\ xs) cys))) (zip nys ys\)))" "ys\ \ nall_tuples_rec {} (card (Inr -` set (map snd (merge (zip ns\ xs) cys)))) (length nys)" using xs_def(2) unfolding ext_tuple_eq[OF len_merge_cys[symmetric]] by auto have distinct_nys: "distinct (ns\ @ map fst cys @ nys)" using distinct len_xs_ys by (auto simp: cys_def nys_def sorted_filter distinct_map_filter) (metis eq_key_imp_eq_value map_fst_zip) obtain \xs where \xs_def: "xs = map \xs ns\" "map snd cys = map \xs (map fst cys)" "ys\ = map \xs nys" using exists_map[OF _ distinct_nys, of "xs @ map snd cys @ ys\"] len_xs_ys(1) nall_tuples_rec_length[OF ys\_def(2)] by (auto simp: ns\'_def) have len_merge_cxs: "length (map snd (merge (zip ns\ ys) cxs)) = length (sort (ns\ @ map fst cxs))" using merge_length[of "zip ns\ ys"] len_xs_ys by auto obtain xs\ where xs\_def: "vs = fo_nmlz AD (map snd (merge (zip (sort (ns\ @ map fst cxs)) (map snd (merge (zip ns\ ys) cxs))) (zip nxs xs\)))" "xs\ \ nall_tuples_rec {} (card (Inr -` set (map snd (merge (zip ns\ ys) cxs)))) (length nxs)" using ys_def(2) unfolding ext_tuple_eq[OF len_merge_cxs[symmetric]] by auto have distinct_nxs: "distinct (ns\ @ map fst cxs @ nxs)" using distinct len_xs_ys(1) by (auto simp: cxs_def nxs_def sorted_filter distinct_map_filter) (metis eq_key_imp_eq_value map_fst_zip) obtain \ys where \ys_def: "ys = map \ys ns\" "map snd cxs = map \ys (map fst cxs)" "xs\ = map \ys nxs" using exists_map[OF _ distinct_nxs, of "ys @ map snd cxs @ xs\"] len_xs_ys(2) nall_tuples_rec_length[OF xs\_def(2)] by (auto simp: ns\'_def) have sd_cs_ns: "sorted_distinct (map fst cxs)" "sorted_distinct nxs" "sorted_distinct (map fst cys)" "sorted_distinct nys" "sorted_distinct (sort (ns\ @ map fst cxs))" "sorted_distinct (sort (ns\ @ map fst cys))" using distinct len_xs_ys by (auto simp: cxs_def nxs_def cys_def nys_def sorted_filter distinct_map_filter) have set_cs_ns_disj: "set (map fst cxs) \ set nxs = {}" "set (map fst cys) \ set nys = {}" "set (sort (ns\ @ map fst cys)) \ set nys = {}" "set (sort (ns\ @ map fst cxs)) \ set nxs = {}" using distinct nth_eq_iff_index_eq by (auto simp: cxs_def nxs_def cys_def nys_def set_zip) blast+ have merge_sort_cxs: "map snd (merge (zip ns\ ys) cxs) = map \ys (sort (ns\ @ map fst cxs))" unfolding \ys_def(1) apply (subst zip_map_fst_snd[of cxs, symmetric]) unfolding \ys_def(2) apply (rule merge_map) using distinct(2) sd_cs_ns by (auto simp: cxs_def) have merge_sort_cys: "map snd (merge (zip ns\ xs) cys) = map \xs (sort (ns\ @ map fst cys))" unfolding \xs_def(1) apply (subst zip_map_fst_snd[of cys, symmetric]) unfolding \xs_def(2) apply (rule merge_map) using distinct(1) sd_cs_ns by (auto simp: cys_def) have set_ns\': "set ns\' = set (map fst cys) \ set nys" using len_xs_ys(2) by (auto simp: ns\'_def cys_def nys_def dest: set_zip_leftD) (metis (no_types, lifting) image_eqI in_set_impl_in_set_zip1 mem_Collect_eq prod.sel(1) split_conv) have sort_sort_nys: "sort (sort (ns\ @ map fst cys) @ nys) = sort (ns\ @ ns\')" apply (rule sorted_distinct_set_unique) using distinct sd_cs_ns set_cs_ns_disj set_ns\' by (auto simp: cys_def nys_def ns\'_def dest: set_zip_leftD) have set_ns\': "set ns\' = set (map fst cxs) \ set nxs" using len_xs_ys(1) by (auto simp: ns\'_def cxs_def nxs_def dest: set_zip_leftD) (metis (no_types, lifting) image_eqI in_set_impl_in_set_zip1 mem_Collect_eq prod.sel(1) split_conv) have sort_sort_nxs: "sort (sort (ns\ @ map fst cxs) @ nxs) = sort (ns\ @ ns\')" apply (rule sorted_distinct_set_unique) using distinct sd_cs_ns set_cs_ns_disj set_ns\' by (auto simp: cxs_def nxs_def ns\'_def dest: set_zip_leftD) have ad_agr1: "ad_agr_list AD (map \ys (sort (ns\ @ ns\'))) (map \xs (sort (ns\ @ ns\')))" using fo_nmlz_eqD[OF trans[OF xs\_def(1)[symmetric] ys\_def(1)]] unfolding \xs_def(3) \ys_def(3) merge_sort_cxs merge_sort_cys unfolding merge_map[OF sd_cs_ns(5) sd_cs_ns(2) set_cs_ns_disj(4)] unfolding merge_map[OF sd_cs_ns(6) sd_cs_ns(4) set_cs_ns_disj(3)] unfolding sort_sort_nxs sort_sort_nys . note ad_agr2 = ad_agr_list_comm[OF ad_agr1] have Inl_set_AD: "Inl -` (set (map snd cxs) \ set xs\) \ AD" "Inl -` (set (map snd cys) \ set ys\) \ AD" using xs_def(1) nall_tuples_rec_Inl[OF xs\_def(2)] ys_def(1) nall_tuples_rec_Inl[OF ys\_def(2)] fo_nmlz_set[of AD] by (fastforce simp: cxs_def X\_def cys_def X\_def dest!: set_zip_rightD)+ note aux1 = eval_conj_set_aux'[OF ns\'_def ns\'_def X\_def X\_def distinct cxs_def nxs_def cys_def nys_def xs_def(1) ys_def(1) \xs_def \ys_def refl refl ys\_def(2)[unfolded \xs_def(3) merge_sort_cys] Inl_set_AD ad_agr1] note aux2 = eval_conj_set_aux'[OF ns\'_def ns\'_def X\_def X\_def distinct(2,1) cys_def nys_def cxs_def nxs_def ys_def(1) xs_def(1) \ys_def \xs_def refl refl xs\_def(2)[unfolded \ys_def(3) merge_sort_cxs] Inl_set_AD(2,1) ad_agr2] show "vs \ fo_nmlz AD ` \(ext_tuple AD ns\ ns\' ` X\) \ fo_nmlz AD ` \(ext_tuple AD ns\ ns\' ` X\)" using xs_def(1) ys_def(1) ys\_def(1) xs\_def(1) aux1(3) aux2(3) ext_tuple_eq[OF len_xs_ys(1)[symmetric], of AD ns\'] ext_tuple_eq[OF len_xs_ys(2)[symmetric], of AD ns\'] unfolding aux1(2) aux2(2) \ys_def(3) \xs_def(3) aux1(1)[symmetric] aux2(1)[symmetric] by blast qed qed lemma esat_exists_not_fv: "n \ fv_fo_fmla \ \ X \ {} \ esat (Exists n \) I \ X \ esat \ I \ X" proof (rule iffI) assume assms: "n \ fv_fo_fmla \" "esat (Exists n \) I \ X" then obtain x where "esat \ I (\(n := x)) X" by auto with assms(1) show "esat \ I \ X" using esat_fv_cong[of \ \ "\(n := x)"] by fastforce next assume assms: "n \ fv_fo_fmla \" "X \ {}" "esat \ I \ X" from assms(2) obtain x where x_def: "x \ X" by auto with assms(1,3) have "esat \ I (\(n := x)) X" using esat_fv_cong[of \ \ "\(n := x)"] by fastforce with x_def show "esat (Exists n \) I \ X" by auto qed lemma esat_forall_not_fv: "n \ fv_fo_fmla \ \ X \ {} \ esat (Forall n \) I \ X \ esat \ I \ X" using esat_exists_not_fv[of n "Neg \" X I \] by auto lemma proj_sat_vals: "proj_sat \ I = proj_vals {\. sat \ I \} (fv_fo_fmla_list \)" by (auto simp: proj_sat_def proj_vals_def) lemma fv_fo_fmla_list_Pred: "remdups_adj (sort (fv_fo_terms_list ts)) = fv_fo_terms_list ts" unfolding fv_fo_terms_list_def by (simp add: distinct_remdups_adj_sort remdups_adj_distinct sorted_sort_id) lemma ad_agr_list_fv_list': "\(set (map set_fo_term ts)) \ X \ ad_agr_list X (map \ (fv_fo_terms_list ts)) (map \ (fv_fo_terms_list ts)) \ ad_agr_list X (\ \e ts) (\ \e ts)" proof (induction ts) case (Cons t ts) have IH: "ad_agr_list X (\ \e ts) (\ \e ts)" using Cons by (auto simp: ad_agr_list_def ad_equiv_list_link[symmetric] fv_fo_terms_set_list fv_fo_terms_set_def sp_equiv_list_link sp_equiv_def pairwise_def) blast+ have ad_equiv: "\i. i \ fv_fo_term_set t \ \(fv_fo_term_set ` set ts) \ ad_equiv_pair X (\ i, \ i)" using Cons(3) by (auto simp: ad_agr_list_def ad_equiv_list_link[symmetric] fv_fo_terms_set_list fv_fo_terms_set_def) have sp_equiv: "\i j. i \ fv_fo_term_set t \ \(fv_fo_term_set ` set ts) \ j \ fv_fo_term_set t \ \(fv_fo_term_set ` set ts) \ sp_equiv_pair (\ i, \ i) (\ j, \ j)" using Cons(3) by (auto simp: ad_agr_list_def sp_equiv_list_link fv_fo_terms_set_list fv_fo_terms_set_def sp_equiv_def pairwise_def) show ?case proof (cases t) case (Const c) show ?thesis using IH Cons(2) apply (auto simp: ad_agr_list_def eval_eterms_def ad_equiv_list_def Const sp_equiv_list_def pairwise_def set_zip) unfolding ad_equiv_pair.simps apply (metis nth_map rev_image_eqI)+ done next case (Var n) note t_def = Var have ad: "ad_equiv_pair X (\ n, \ n)" using ad_equiv by (auto simp: Var) have "\y. y \ set (zip (map ((\e) \) ts) (map ((\e) \) ts)) \ y \ (\ n, \ n) \ sp_equiv_pair (\ n, \ n) y \ sp_equiv_pair y (\ n, \ n)" proof - fix y assume "y \ set (zip (map ((\e) \) ts) (map ((\e) \) ts))" then obtain t' where y_def: "t' \ set ts" "y = (\ \e t', \ \e t')" using nth_mem by (auto simp: set_zip) blast show "sp_equiv_pair (\ n, \ n) y \ sp_equiv_pair y (\ n, \ n)" proof (cases t') case (Const c') have c'_X: "c' \ X" using Cons(2) y_def(1) by (auto simp: Const) (meson SUP_le_iff fo_term.set_intros subsetD) then show ?thesis using ad_equiv[of n] y_def(1) unfolding y_def apply (auto simp: Const t_def) unfolding ad_equiv_pair.simps apply fastforce+ apply force apply (metis rev_image_eqI) done next case (Var n') show ?thesis using sp_equiv[of n n'] y_def(1) unfolding y_def by (fastforce simp: t_def Var) qed qed then show ?thesis using IH Cons(3) by (auto simp: ad_agr_list_def eval_eterms_def ad_equiv_list_def Var ad sp_equiv_list_def pairwise_insert) qed qed (auto simp: eval_eterms_def ad_agr_list_def ad_equiv_list_def sp_equiv_list_def) lemma ext_tuple_ad_agr_close: assumes S\_def: "S\ \ {\. esat \ I \ UNIV}" and AD_sub: "act_edom \ I \ AD\" "AD\ \ AD" and X\_def: "X\ = fo_nmlz AD\ ` proj_vals S\ (fv_fo_fmla_list \)" and ns\'_def: "ns\' = filter (\n. n \ fv_fo_fmla \) ns\" and sd_ns\: "sorted_distinct ns\" and fv_Un: "fv_fo_fmla \ = fv_fo_fmla \ \ set ns\" shows "ext_tuple_set AD (fv_fo_fmla_list \) ns\' (ad_agr_close_set (AD - AD\) X\) = fo_nmlz AD ` proj_vals S\ (fv_fo_fmla_list \)" "ad_agr_close_set (AD - AD\) X\ = fo_nmlz AD ` proj_vals S\ (fv_fo_fmla_list \)" proof - have ad_agr_\: "\\ \. ad_agr_sets (set (fv_fo_fmla_list \)) (set (fv_fo_fmla_list \)) AD\ \ \ \ \ \ S\ \ \ \ S\" using esat_UNIV_cong[OF ad_agr_sets_restrict, OF _ subset_refl] ad_agr_sets_mono AD_sub unfolding S\_def by blast show ad_close_alt: "ad_agr_close_set (AD - AD\) X\ = fo_nmlz AD ` proj_vals S\ (fv_fo_fmla_list \)" using ad_agr_close_correct[OF AD_sub(2) ad_agr_\] AD_sub(2) unfolding X\_def S\_def[symmetric] proj_fmla_def by (auto simp: ad_agr_close_set_def Set.is_empty_def) have fv_\: "set (fv_fo_fmla_list \) \ set (fv_fo_fmla_list \)" using fv_Un by (auto simp: fv_fo_fmla_list_set) have sd_ns\': "sorted_distinct ns\'" using sd_ns\ sorted_filter[of id] by (auto simp: ns\'_def) show "ext_tuple_set AD (fv_fo_fmla_list \) ns\' (ad_agr_close_set (AD - AD\) X\) = fo_nmlz AD ` proj_vals S\ (fv_fo_fmla_list \)" apply (rule ext_tuple_correct) using sorted_distinct_fv_list ad_close_alt ad_agr_\ ad_agr_sets_mono[OF AD_sub(2)] fv_Un sd_ns\' by (fastforce simp: ns\'_def fv_fo_fmla_list_set)+ qed lemma proj_ext_tuple: assumes S\_def: "S\ \ {\. esat \ I \ UNIV}" and AD_sub: "act_edom \ I \ AD" and X\_def: "X\ = fo_nmlz AD ` proj_vals S\ (fv_fo_fmla_list \)" and ns\'_def: "ns\' = filter (\n. n \ fv_fo_fmla \) ns\" and sd_ns\: "sorted_distinct ns\" and fv_Un: "fv_fo_fmla \ = fv_fo_fmla \ \ set ns\" and Z_props: "\xs. xs \ Z \ fo_nmlz AD xs = xs \ length xs = length (fv_fo_fmla_list \)" shows "Z \ ext_tuple_set AD (fv_fo_fmla_list \) ns\' X\ = {xs \ Z. fo_nmlz AD (proj_tuple (fv_fo_fmla_list \) (zip (fv_fo_fmla_list \) xs)) \ X\}" "Z - ext_tuple_set AD (fv_fo_fmla_list \) ns\' X\ = {xs \ Z. fo_nmlz AD (proj_tuple (fv_fo_fmla_list \) (zip (fv_fo_fmla_list \) xs)) \ X\}" proof - have ad_agr_\: "\\ \. ad_agr_sets (set (fv_fo_fmla_list \)) (set (fv_fo_fmla_list \)) AD \ \ \ \ \ S\ \ \ \ S\" using esat_UNIV_cong[OF ad_agr_sets_restrict, OF _ subset_refl] ad_agr_sets_mono AD_sub unfolding S\_def by blast have sd_ns\': "sorted_distinct ns\'" using sd_ns\ sorted_filter[of id] by (auto simp: ns\'_def) have disj: "set (fv_fo_fmla_list \) \ set ns\' = {}" by (auto simp: ns\'_def fv_fo_fmla_list_set) have Un: "set (fv_fo_fmla_list \) \ set ns\' = set (fv_fo_fmla_list \)" using fv_Un by (auto simp: ns\'_def fv_fo_fmla_list_set) note proj = proj_tuple_correct[OF sorted_distinct_fv_list sd_ns\' sorted_distinct_fv_list disj Un X\_def ad_agr_\, simplified] have "fo_nmlz AD ` X\ = X\" using fo_nmlz_idem[OF fo_nmlz_sound] by (auto simp: X\_def image_iff) then have aux: "ext_tuple_set AD (fv_fo_fmla_list \) ns\' X\ = fo_nmlz AD ` \(ext_tuple AD (fv_fo_fmla_list \) ns\' ` X\)" by (auto simp: ext_tuple_set_def ext_tuple_def) show "Z \ ext_tuple_set AD (fv_fo_fmla_list \) ns\' X\ = {xs \ Z. fo_nmlz AD (proj_tuple (fv_fo_fmla_list \) (zip (fv_fo_fmla_list \) xs)) \ X\}" using Z_props proj by (auto simp: aux) show "Z - ext_tuple_set AD (fv_fo_fmla_list \) ns\' X\ = {xs \ Z. fo_nmlz AD (proj_tuple (fv_fo_fmla_list \) (zip (fv_fo_fmla_list \) xs)) \ X\}" using Z_props proj by (auto simp: aux) qed lemma fo_nmlz_proj_sub: "fo_nmlz AD ` proj_fmla \ R \ nall_tuples AD (nfv \)" by (auto simp: proj_fmla_map fo_nmlz_length fo_nmlz_sound nfv_def intro: nall_tuplesI) lemma fin_ad_agr_list_iff: fixes AD :: "('a :: infinite) set" assumes "finite AD" "\vs. vs \ Z \ length vs = n" "Z = {ts. \ts' \ X. ad_agr_list AD (map Inl ts) ts'}" shows "finite Z \ \(set ` Z) \ AD" proof (rule iffI, rule ccontr) assume fin: "finite Z" assume "\\(set ` Z) \ AD" then obtain \ i vs where \_def: "map \ [0.. Z" "i < n" "\ i \ AD" "vs \ X" "ad_agr_list AD (map (Inl \ \) [0.. AD \ \ ` {0..y. y \ Y \ map ((\z. if z = \ i then y else z) \ \) [0.. Z" using \_def(3) by (auto simp: assms(3) intro!: bexI[OF _ \_def(4)] ad_agr_list_trans[OF _ \_def(5)]) (auto simp: ad_agr_list_def ad_equiv_list_def set_zip Y_def ad_equiv_pair.simps sp_equiv_list_def pairwise_def split: if_splits) then have "(\x'. map ((\z. if z = \ i then x' else z) \ \) [0.. Z" by auto moreover have "inj (\x'. map ((\z. if z = \ i then x' else z) \ \) [0.._def(2) by (auto simp: inj_def) ultimately show "False" using inf_UNIV_Y fin by (meson inj_on_diff inj_on_finite) next assume "\(set ` Z) \ AD" then have "Z \ all_tuples AD n" using assms(2) by (auto intro: all_tuplesI) then show "finite Z" using all_tuples_finite[OF assms(1)] finite_subset by auto qed lemma proj_out_list: fixes AD :: "('a :: infinite) set" and \ :: "nat \ 'a + nat" and ns :: "nat list" assumes "finite AD" shows "\\. ad_agr_list AD (map \ ns) (map (Inl \ \) ns) \ (\j x. j \ set ns \ \ j = Inl x \ \ j = x)" proof - have fin: "finite (AD \ Inl -` set (map \ ns))" using assms(1) finite_Inl[OF finite_set] by blast obtain f where f_def: "inj (f :: nat \ 'a)" "range f \ UNIV - (AD \ Inl -` set (map \ ns))" using arb_countable_map[OF fin] by auto define \ where "\ = case_sum id f \ \" have f_out: "\i x. i < length ns \ \ (ns ! i) = Inl (f x) \ False" using f_def(2) by (auto simp: vimage_def) (metis (no_types, lifting) DiffE UNIV_I UnCI imageI image_subset_iff mem_Collect_eq nth_mem) have "(a, b) \ set (zip (map \ ns) (map (Inl \ \) ns)) \ ad_equiv_pair AD (a, b)" for a b using f_def(2) by (auto simp: set_zip \_def ad_equiv_pair.simps split: sum.splits)+ moreover have "sp_equiv_list (map \ ns) (map (Inl \ \) ns)" using f_def(1) f_out by (auto simp: sp_equiv_list_def pairwise_def set_zip \_def inj_def split: sum.splits)+ ultimately have "ad_agr_list AD (map \ ns) (map (Inl \ \) ns)" by (auto simp: ad_agr_list_def ad_equiv_list_def) then show ?thesis by (auto simp: \_def intro!: exI[of _ \]) qed lemma proj_out: fixes \ :: "('a :: infinite, 'b) fo_fmla" and J :: "(('a, nat) fo_t, 'b) fo_intp" assumes "wf_fo_intp \ I" "esat \ I \ UNIV" shows "\\. esat \ I (Inl \ \) UNIV \ (\i x. i \ fv_fo_fmla \ \ \ i = Inl x \ \ i = x) \ ad_agr_list (act_edom \ I) (map \ (fv_fo_fmla_list \)) (map (Inl \ \) (fv_fo_fmla_list \))" using proj_out_list[OF finite_act_edom[OF assms(1)], of \ "fv_fo_fmla_list \"] esat_UNIV_ad_agr_list[OF _ subset_refl] assms(2) unfolding fv_fo_fmla_list_set by fastforce lemma proj_fmla_esat_sat: fixes \ :: "('a :: infinite, 'b) fo_fmla" and J :: "(('a, nat) fo_t, 'b) fo_intp" assumes wf: "wf_fo_intp \ I" shows "proj_fmla \ {\. esat \ I \ UNIV} \ map Inl ` UNIV = map Inl ` proj_fmla \ {\. sat \ I \}" unfolding sat_esat_conv[OF wf] proof (rule set_eqI, rule iffI) fix vs assume "vs \ proj_fmla \ {\. esat \ I \ UNIV} \ map Inl ` UNIV" then obtain \ where \_def: "vs = map \ (fv_fo_fmla_list \)" "esat \ I \ UNIV" "set vs \ range Inl" by (auto simp: proj_fmla_map) (metis image_subset_iff list.set_map range_eqI) obtain \ where \_def: "esat \ I (Inl \ \) UNIV" "\i x. i \ fv_fo_fmla \ \ \ i = Inl x \ \ i = x" using proj_out[OF assms \_def(2)] by fastforce have "vs = map (Inl \ \) (fv_fo_fmla_list \)" using \_def(1,3) \_def(2) by (auto simp: fv_fo_fmla_list_set) then show "vs \ map Inl ` proj_fmla \ {\. esat \ I (Inl \ \) UNIV}" using \_def(1) by (force simp: proj_fmla_map) qed (auto simp: proj_fmla_map) lemma norm_proj_fmla_esat_sat: fixes \ :: "('a :: infinite, 'b) fo_fmla" assumes "wf_fo_intp \ I" shows "fo_nmlz (act_edom \ I) ` proj_fmla \ {\. esat \ I \ UNIV} = fo_nmlz (act_edom \ I) ` map Inl ` proj_fmla \ {\. sat \ I \}" proof - have "fo_nmlz (act_edom \ I) (map \ (fv_fo_fmla_list \)) = fo_nmlz (act_edom \ I) x" "x \ (\\. map \ (fv_fo_fmla_list \)) ` {\. esat \ I \ UNIV} \ range (map Inl)" if "esat \ I \ UNIV" "esat \ I (Inl \ \) UNIV" "x = map (Inl \ \) (fv_fo_fmla_list \)" "ad_agr_list (act_edom \ I) (map \ (fv_fo_fmla_list \)) (map (Inl \ \) (fv_fo_fmla_list \))" for \ \ x using that by (auto intro!: fo_nmlz_eqI) (metis map_map range_eqI) then show ?thesis unfolding proj_fmla_esat_sat[OF assms, symmetric] using proj_out[OF assms] by (fastforce simp: image_iff proj_fmla_map) qed lemma proj_sat_fmla: "proj_sat \ I = proj_fmla \ {\. sat \ I \}" by (auto simp: proj_sat_def proj_fmla_map) fun fo_wf :: "('a, 'b) fo_fmla \ ('b \ nat \ 'a list set) \ ('a, nat) fo_t \ bool" where "fo_wf \ I (AD, n, X) \ finite AD \ finite X \ n = nfv \ \ wf_fo_intp \ I \ AD = act_edom \ I \ fo_rep (AD, n, X) = proj_sat \ I \ Inl -` \(set ` X) \ AD \ (\vs \ X. fo_nmlzd AD vs \ length vs = n)" fun fo_fin :: "('a, nat) fo_t \ bool" where "fo_fin (AD, n, X) \ (\x \ \(set ` X). isl x)" lemma fo_rep_fin: assumes "fo_wf \ I (AD, n, X)" "fo_fin (AD, n, X)" shows "fo_rep (AD, n, X) = map projl ` X" proof (rule set_eqI, rule iffI) fix vs assume "vs \ fo_rep (AD, n, X)" then obtain xs where xs_def: "xs \ X" "ad_agr_list AD (map Inl vs) xs" by auto obtain zs where zs_def: "xs = map Inl zs" using xs_def(1) assms by auto (meson ex_map_conv isl_def) have "set zs \ AD" using assms(1) xs_def(1) zs_def by (force simp: vimage_def) then have vs_zs: "vs = zs" using xs_def(2) unfolding zs_def by (fastforce simp: ad_agr_list_def ad_equiv_list_def set_zip ad_equiv_pair.simps intro!: nth_equalityI) show "vs \ map projl ` X" using xs_def(1) zs_def by (auto simp: image_iff comp_def vs_zs intro!: bexI[of _ "map Inl zs"]) next fix vs assume "vs \ map projl ` X" then obtain xs where xs_def: "xs \ X" "vs = map projl xs" by auto have xs_map_Inl: "xs = map Inl vs" using assms xs_def by (auto simp: map_idI) show "vs \ fo_rep (AD, n, X)" using xs_def(1) by (auto simp: xs_map_Inl intro!: bexI[of _ xs] ad_agr_list_refl) qed definition eval_abs :: "('a, 'b) fo_fmla \ ('a table, 'b) fo_intp \ ('a, nat) fo_t" where "eval_abs \ I = (act_edom \ I, nfv \, fo_nmlz (act_edom \ I) ` proj_fmla \ {\. esat \ I \ UNIV})" lemma map_projl_Inl: "map projl (map Inl xs) = xs" by (metis (mono_tags, lifting) length_map nth_equalityI nth_map sum.sel(1)) lemma fo_rep_eval_abs: fixes \ :: "('a :: infinite, 'b) fo_fmla" assumes "wf_fo_intp \ I" shows "fo_rep (eval_abs \ I) = proj_sat \ I" proof - obtain AD n X where AD_X_def: "eval_abs \ I = (AD, n, X)" "AD = act_edom \ I" "n = nfv \" "X = fo_nmlz (act_edom \ I) ` proj_fmla \ {\. esat \ I \ UNIV}" by (cases "eval_abs \ I") (auto simp: eval_abs_def) have AD_sub: "act_edom \ I \ AD" by (auto simp: AD_X_def) have X_def: "X = fo_nmlz AD ` map Inl ` proj_fmla \ {\. sat \ I \}" using AD_X_def norm_proj_fmla_esat_sat[OF assms] by auto have "{ts. \ts' \ X. ad_agr_list AD (map Inl ts) ts'} = proj_fmla \ {\. sat \ I \}" proof (rule set_eqI, rule iffI) fix vs assume "vs \ {ts. \ts' \ X. ad_agr_list AD (map Inl ts) ts'}" then obtain vs' where vs'_def: "vs' \ proj_fmla \ {\. sat \ I \}" "ad_agr_list AD (map Inl vs) (fo_nmlz AD (map Inl vs'))" using X_def by auto have "length vs = length (fv_fo_fmla_list \)" using vs'_def by (auto simp: proj_fmla_map ad_agr_list_def fo_nmlz_length) then obtain \ where \_def: "vs = map \ (fv_fo_fmla_list \)" using exists_map[of "fv_fo_fmla_list \" vs] sorted_distinct_fv_list by fastforce obtain \ where \_def: "fo_nmlz AD (map Inl vs') = map \ (fv_fo_fmla_list \)" using vs'_def fo_nmlz_map by (fastforce simp: proj_fmla_map) have ad_agr: "ad_agr_list AD (map (Inl \ \) (fv_fo_fmla_list \)) (map \ (fv_fo_fmla_list \))" by (metis \_def \_def map_map vs'_def(2)) obtain \' where \'_def: "map Inl vs' = map (Inl \ \') (fv_fo_fmla_list \)" "sat \ I \'" using vs'_def(1) by (fastforce simp: proj_fmla_map) have ad_agr': "ad_agr_list AD (map \ (fv_fo_fmla_list \)) (map (Inl \ \') (fv_fo_fmla_list \))" by (rule ad_agr_list_comm) (metis fo_nmlz_ad_agr \'_def(1) \_def map_map map_projl_Inl) have esat: "esat \ I \ UNIV" using esat_UNIV_ad_agr_list[OF ad_agr' AD_sub, folded sat_esat_conv[OF assms]] \'_def(2) by auto show "vs \ proj_fmla \ {\. sat \ I \}" using esat_UNIV_ad_agr_list[OF ad_agr AD_sub, folded sat_esat_conv[OF assms]] esat unfolding \_def by (auto simp: proj_fmla_map) next fix vs assume "vs \ proj_fmla \ {\. sat \ I \}" then have vs_X: "fo_nmlz AD (map Inl vs) \ X" using X_def by auto then show "vs \ {ts. \ts' \ X. ad_agr_list AD (map Inl ts) ts'}" using fo_nmlz_ad_agr by auto qed then show ?thesis by (auto simp: AD_X_def proj_sat_fmla) qed lemma fo_wf_eval_abs: fixes \ :: "('a :: infinite, 'b) fo_fmla" assumes "wf_fo_intp \ I" shows "fo_wf \ I (eval_abs \ I)" using fo_nmlz_set[of "act_edom \ I"] finite_act_edom[OF assms(1)] finite_subset[OF fo_nmlz_proj_sub, OF nall_tuples_finite] fo_rep_eval_abs[OF assms] assms by (auto simp: eval_abs_def fo_nmlz_sound fo_nmlz_length nfv_def proj_sat_def proj_fmla_map) blast lemma fo_fin: fixes t :: "('a :: infinite, nat) fo_t" assumes "fo_wf \ I t" shows "fo_fin t = finite (fo_rep t)" proof - obtain AD n X where t_def: "t = (AD, n, X)" using assms by (cases t) auto have fin: "finite AD" "finite X" using assms by (auto simp: t_def) have len_in_X: "\vs. vs \ X \ length vs = n" using assms by (auto simp: t_def) have Inl_X_AD: "\x. Inl x \ \(set ` X) \ x \ AD" using assms by (fastforce simp: t_def) define Z where "Z = {ts. \ts' \ X. ad_agr_list AD (map Inl ts) ts'}" have fin_Z_iff: "finite Z = (\(set ` Z) \ AD)" using assms fin_ad_agr_list_iff[OF fin(1) _ Z_def, of n] by (auto simp: Z_def t_def ad_agr_list_def) moreover have "(\(set ` Z) \ AD) \ (\x \ \(set ` X). isl x)" proof (rule iffI, rule ccontr) fix x assume Z_sub_AD: "\(set ` Z) \ AD" assume "\(\x \ \(set ` X). isl x)" then obtain vs i m where vs_def: "vs \ X" "i < n" "vs ! i = Inr m" using len_in_X by (auto simp: in_set_conv_nth) (metis sum.collapse(2)) obtain \ where \_def: "vs = map \ [0.. where \_def: "ad_agr_list AD vs (map Inl (map \ [0.. "[0.._def) have map_\_in_Z: "map \ [0.. Z" using vs_def(1) ad_agr_list_comm[OF \_def] by (auto simp: Z_def) moreover have "\ i \ AD" using \_def vs_def(2,3) apply (auto simp: ad_agr_list_def ad_equiv_list_def set_zip comp_def \_def) unfolding ad_equiv_pair.simps by (metis (no_types, lifting) Inl_Inr_False diff_zero image_iff length_upt nth_map nth_upt plus_nat.add_0) ultimately show "False" using vs_def(2) Z_sub_AD by fastforce next assume "\x \ \(set ` X). isl x" then show "\(set ` Z) \ AD" using Inl_X_AD apply (auto simp: Z_def ad_agr_list_def ad_equiv_list_def set_zip in_set_conv_nth) unfolding ad_equiv_pair.simps by (metis image_eqI isl_def nth_map nth_mem) qed ultimately show ?thesis by (auto simp: t_def Z_def[symmetric]) qed lemma eval_pred: fixes I :: "'b \ nat \ 'a :: infinite list set" assumes "finite (I (r, length ts))" shows "fo_wf (Pred r ts) I (eval_pred ts (I (r, length ts)))" proof - define \ where "\ = Pred r ts" have nfv_len: "nfv \ = length (fv_fo_terms_list ts)" by (auto simp: \_def nfv_def fv_fo_fmla_list_def fv_fo_fmla_list_Pred) have vimage_unfold: "Inl -` (\x\I (r, length ts). Inl ` set x) = \(set ` I (r, length ts))" by auto have "eval_table ts (map Inl ` I (r, length ts)) \ nall_tuples (act_edom \ I) (nfv \)" by (auto simp: \_def proj_vals_def eval_table nfv_len[unfolded \_def] fo_nmlz_length fo_nmlz_sound eval_eterms_def fv_fo_terms_set_list fv_fo_terms_set_def vimage_unfold intro!: nall_tuplesI fo_nmlzd_all_AD dest!: fv_fo_term_setD) (smt UN_I Un_iff eval_eterm.simps(2) imageE image_eqI list.set_map) then have eval: "eval_pred ts (I (r, length ts)) = eval_abs \ I" by (force simp: eval_abs_def \_def proj_fmla_def eval_pred_def eval_table fv_fo_fmla_list_def fv_fo_fmla_list_Pred nall_tuples_set fo_nmlz_idem nfv_len[unfolded \_def]) have fin: "wf_fo_intp (Pred r ts) I" using assms by auto show ?thesis using fo_wf_eval_abs[OF fin] by (auto simp: eval \_def) qed lemma ad_agr_list_eval: "\(set (map set_fo_term ts)) \ AD \ ad_agr_list AD (\ \e ts) zs \ \\. zs = \ \e ts" proof (induction ts arbitrary: zs) case (Cons t ts) obtain w ws where zs_split: "zs = w # ws" using Cons(3) by (cases zs) (auto simp: ad_agr_list_def eval_eterms_def) obtain \ where \_def: "ws = \ \e ts" using Cons by (fastforce simp: zs_split ad_agr_list_def ad_equiv_list_def sp_equiv_list_def pairwise_def eval_eterms_def) show ?case proof (cases t) case (Const c) then show ?thesis using Cons(3)[unfolded zs_split] Cons(2) unfolding Const apply (auto simp: zs_split eval_eterms_def \_def ad_agr_list_def ad_equiv_list_def) unfolding ad_equiv_pair.simps by blast next case (Var n) show ?thesis proof (cases "n \ fv_fo_terms_set ts") case True obtain i where i_def: "i < length ts" "ts ! i = Var n" using True by (auto simp: fv_fo_terms_set_def in_set_conv_nth dest!: fv_fo_term_setD) have "w = \ n" using Cons(3)[unfolded zs_split \_def] i_def using pairwiseD[of sp_equiv_pair _ "(\ n, w)" "(\ \e (ts ! i), \ \e (ts ! i))"] by (force simp: Var eval_eterms_def ad_agr_list_def sp_equiv_list_def set_zip) then show ?thesis by (auto simp: Var zs_split eval_eterms_def \_def) next case False then have "ws = (\(n := w)) \e ts" using eval_eterms_cong[of ts \ "\(n := w)"] \_def by fastforce then show ?thesis by (auto simp: zs_split eval_eterms_def Var fun_upd_def intro: exI[of _ "\(n := w)"]) qed qed qed (auto simp: ad_agr_list_def eval_eterms_def) lemma sp_equiv_list_fv_list: assumes "sp_equiv_list (\ \e ts) (\ \e ts)" shows "sp_equiv_list (map \ (fv_fo_terms_list ts)) (map \ (fv_fo_terms_list ts))" proof - have "sp_equiv_list (\ \e (map Var (fv_fo_terms_list ts))) (\ \e (map Var (fv_fo_terms_list ts)))" unfolding eval_eterms_def by (rule sp_equiv_list_subset[OF _ assms[unfolded eval_eterms_def]]) (auto simp: fv_fo_terms_set_list dest: fv_fo_terms_setD) then show ?thesis by (auto simp: eval_eterms_def comp_def) qed lemma ad_agr_list_fv_list: "ad_agr_list X (\ \e ts) (\ \e ts) \ ad_agr_list X (map \ (fv_fo_terms_list ts)) (map \ (fv_fo_terms_list ts))" using sp_equiv_list_fv_list by (auto simp: eval_eterms_def ad_agr_list_def ad_equiv_list_def sp_equiv_list_def set_zip) (metis (no_types, opaque_lifting) eval_eterm.simps(2) fv_fo_terms_setD fv_fo_terms_set_list in_set_conv_nth nth_map) lemma eval_bool: "fo_wf (Bool b) I (eval_bool b)" by (auto simp: eval_bool_def fo_nmlzd_def nats_def Let_def List.map_filter_simps proj_sat_def fv_fo_fmla_list_def ad_agr_list_def ad_equiv_list_def sp_equiv_list_def nfv_def) lemma eval_eq: fixes I :: "'b \ nat \ 'a :: infinite list set" shows "fo_wf (Eqa t t') I (eval_eq t t')" proof - define \ :: "('a, 'b) fo_fmla" where "\ = Eqa t t'" obtain AD n X where AD_X_def: "eval_eq t t' = (AD, n, X)" by (cases "eval_eq t t'") auto have AD_def: "AD = act_edom \ I" using AD_X_def by (auto simp: eval_eq_def \_def split: fo_term.splits if_splits) have n_def: "n = nfv \" using AD_X_def by (cases t; cases t') (auto simp: \_def fv_fo_fmla_list_def eval_eq_def nfv_def split: if_splits) have fo_nmlz_empty_x_x: "fo_nmlz {} [x, x] = [Inr 0, Inr 0]" for x :: "'a + nat" by (cases x) (auto simp: fo_nmlz_def) have Inr_0_in_fo_nmlz_empty: "[Inr 0, Inr 0] \ fo_nmlz {} ` (\x. [x n', x n']) ` {\ :: nat \ 'a + nat. \ n = \ n'}" for n n' by (auto simp: image_def fo_nmlz_empty_x_x intro!: exI[of _ "[Inr 0, Inr 0]"]) have X_def: "X = fo_nmlz AD ` proj_fmla \ {\. esat \ I \ UNIV}" proof (rule set_eqI, rule iffI) fix vs assume assm: "vs \ X" define pes where "pes = proj_fmla \ {\. esat \ I \ UNIV}" have "\c c'. t = Const c \ t' = Const c' \ fo_nmlz AD ` pes = (if c = c' then {[]} else {})" by (auto simp: \_def pes_def proj_fmla_map fo_nmlz_def fv_fo_fmla_list_def) moreover have "\c n. (t = Const c \ t' = Var n) \ (t' = Const c \ t = Var n) \ fo_nmlz AD ` pes = {[Inl c]}" by (auto simp: \_def AD_def pes_def proj_fmla_map fo_nmlz_Cons fv_fo_fmla_list_def image_def split: sum.splits) (auto simp: fo_nmlz_def) moreover have "\n. t = Var n \ t' = Var n \ fo_nmlz AD ` pes = {[Inr 0]}" by (auto simp: \_def AD_def pes_def proj_fmla_map fo_nmlz_Cons fv_fo_fmla_list_def image_def split: sum.splits) moreover have "\n n'. t = Var n \ t' = Var n' \ n \ n' \ fo_nmlz AD ` pes = {[Inr 0, Inr 0]}" using Inr_0_in_fo_nmlz_empty by (auto simp: \_def AD_def pes_def proj_fmla_map fo_nmlz_Cons fv_fo_fmla_list_def fo_nmlz_empty_x_x split: sum.splits) ultimately show "vs \ fo_nmlz AD ` pes" using assm AD_X_def by (cases t; cases t') (auto simp: eval_eq_def split: if_splits) next fix vs assume assm: "vs \ fo_nmlz AD ` proj_fmla \ {\. esat \ I \ UNIV}" obtain \ where \_def: "vs = fo_nmlz AD (map \ (fv_fo_fmla_list \))" "esat (Eqa t t') I \ UNIV" using assm by (auto simp: \_def fv_fo_fmla_list_def proj_fmla_map) show "vs \ X" using \_def AD_X_def by (cases t; cases t') (auto simp: \_def eval_eq_def fv_fo_fmla_list_def fo_nmlz_Cons fo_nmlz_Cons_Cons split: sum.splits) qed have eval: "eval_eq t t' = eval_abs \ I" using X_def[unfolded AD_def] by (auto simp: eval_abs_def AD_X_def AD_def n_def) have fin: "wf_fo_intp \ I" by (auto simp: \_def) show ?thesis using fo_wf_eval_abs[OF fin] by (auto simp: eval \_def) qed lemma fv_fo_terms_list_Var: "fv_fo_terms_list_rec (map Var ns) = ns" by (induction ns) auto lemma eval_eterms_map_Var: "\ \e map Var ns = map \ ns" by (auto simp: eval_eterms_def) lemma fo_wf_eval_table: fixes AD :: "'a set" assumes "fo_wf \ I (AD, n, X)" shows "X = fo_nmlz AD ` eval_table (map Var [0..(set ` X) \ AD" using assms by fastforce have fvs: "fv_fo_terms_list (map Var [0..vs. vs \ X \ length vs = n" using assms by auto then have X_map: "\vs. vs \ X \ \\. vs = map \ [0... \ \e map Var [0.. X} [0.. I (AD, n, X)" shows "X = fo_nmlz AD ` map Inl ` fo_rep (AD, n, X)" proof (rule set_eqI, rule iffI) fix vs assume vs_in: "vs \ X" have fin_AD: "finite AD" using assms(1) by auto have len_vs: "length vs = n" using vs_in assms(1) by auto obtain \ where \_def: "ad_agr_list AD vs (map Inl (map \ [0.._in: "map \ [0.. fo_rep (AD, n, X)" using vs_in ad_agr_list_comm[OF \_def] by auto have "vs = fo_nmlz AD (map Inl (map \ [0.._def] fo_nmlz_idem vs_in assms(1) by fastforce then show "vs \ fo_nmlz AD ` map Inl ` fo_rep (AD, n, X)" using map_\_in by blast next fix vs assume "vs \ fo_nmlz AD ` map Inl ` fo_rep (AD, n, X)" then obtain xs xs' where vs_def: "xs' \ X" "ad_agr_list AD (map Inl xs) xs'" "vs = fo_nmlz AD (map Inl xs)" by auto then have "vs = fo_nmlz AD xs'" using fo_nmlz_eqI[OF vs_def(2)] by auto then have "vs = xs'" using vs_def(1) assms(1) fo_nmlz_idem by fastforce then show "vs \ X" using vs_def(1) by auto qed lemma fo_wf_X: fixes \ :: "('a :: infinite, 'b) fo_fmla" assumes wf: "fo_wf \ I (AD, n, X)" shows "X = fo_nmlz AD ` proj_fmla \ {\. esat \ I \ UNIV}" proof - have fin: "wf_fo_intp \ I" using wf by auto have AD_def: "AD = act_edom \ I" using wf by auto have fo_wf: "fo_wf \ I (AD, n, X)" using wf by auto have fo_rep: "fo_rep (AD, n, X) = proj_fmla \ {\. sat \ I \}" using wf by (auto simp: proj_sat_def proj_fmla_map) show ?thesis using fo_rep_norm[OF fo_wf] norm_proj_fmla_esat_sat[OF fin] unfolding fo_rep AD_def[symmetric] by auto qed lemma eval_neg: fixes \ :: "('a :: infinite, 'b) fo_fmla" assumes wf: "fo_wf \ I t" shows "fo_wf (Neg \) I (eval_neg (fv_fo_fmla_list \) t)" proof - obtain AD n X where t_def: "t = (AD, n, X)" by (cases t) auto have eval_neg: "eval_neg (fv_fo_fmla_list \) t = (AD, nfv \, nall_tuples AD (nfv \) - X)" by (auto simp: t_def nfv_def) have fv_unfold: "fv_fo_fmla_list (Neg \) = fv_fo_fmla_list \" by (auto simp: fv_fo_fmla_list_def) then have nfv_unfold: "nfv (Neg \) = nfv \" by (auto simp: nfv_def) have AD_def: "AD = act_edom (Neg \) I" using wf by (auto simp: t_def) note X_def = fo_wf_X[OF wf[unfolded t_def]] have esat_iff: "\vs. vs \ nall_tuples AD (nfv \) \ vs \ fo_nmlz AD ` proj_fmla \ {\. esat \ I \ UNIV} \ vs \ fo_nmlz AD ` proj_fmla \ {\. esat (Neg \) I \ UNIV}" proof (rule iffI; rule ccontr) fix vs assume "vs \ fo_nmlz AD ` proj_fmla \ {\. esat \ I \ UNIV}" then obtain \ where \_def: "vs = fo_nmlz AD (map \ (fv_fo_fmla_list \))" "esat \ I \ UNIV" by (auto simp: proj_fmla_map) assume "\vs \ fo_nmlz AD ` proj_fmla \ {\. esat (Neg \) I \ UNIV}" then obtain \' where \'_def: "vs = fo_nmlz AD (map \' (fv_fo_fmla_list \))" "esat (Neg \) I \' UNIV" by (auto simp: proj_fmla_map) have "esat \ I \ UNIV = esat \ I \' UNIV" using esat_UNIV_cong[OF ad_agr_sets_restrict[OF iffD2[OF ad_agr_list_link], OF fo_nmlz_eqD[OF trans[OF \_def(1)[symmetric] \'_def(1)]]]] by (auto simp: AD_def) then show "False" using \_def(2) \'_def(2) by simp next fix vs assume assms: "vs \ fo_nmlz AD ` proj_fmla \ {\. esat (Neg \) I \ UNIV}" "vs \ fo_nmlz AD ` proj_fmla \ {\. esat \ I \ UNIV}" assume "vs \ nall_tuples AD (nfv \)" then have l_vs: "length vs = length (fv_fo_fmla_list \)" "fo_nmlzd AD vs" by (auto simp: nfv_def dest: nall_tuplesD) obtain \ where "vs = fo_nmlz AD (map \ (fv_fo_fmla_list \))" using l_vs sorted_distinct_fv_list exists_fo_nmlzd by metis with assms show "False" by (auto simp: proj_fmla_map) qed moreover have "\R. fo_nmlz AD ` proj_fmla \ R \ nall_tuples AD (nfv \)" by (auto simp: proj_fmla_map nfv_def nall_tuplesI fo_nmlz_length fo_nmlz_sound) ultimately have eval: "eval_neg (fv_fo_fmla_list \) t = eval_abs (Neg \) I" unfolding eval_neg eval_abs_def AD_def[symmetric] by (auto simp: X_def proj_fmla_def fv_unfold nfv_unfold image_subset_iff) have wf_neg: "wf_fo_intp (Neg \) I" using wf by (auto simp: t_def) show ?thesis using fo_wf_eval_abs[OF wf_neg] by (auto simp: eval) qed definition "cross_with f t t' = \((\xs. \(f xs ` t')) ` t)" lemma mapping_join_cross_with: assumes "\x x'. x \ t \ x' \ t' \ h x \ h' x' \ f x x' = {}" shows "set_of_idx (mapping_join (cross_with f) (cluster (Some \ h) t) (cluster (Some \ h') t')) = cross_with f t t'" proof - have sub: "cross_with f {y \ t. h y = h x} {y \ t'. h' y = h x} \ cross_with f t t'" for t t' x by (auto simp: cross_with_def) have "\a. a \ h ` t \ a \ h' ` t' \ z \ cross_with f {y \ t. h y = a} {y \ t'. h' y = a}" if z: "z \ cross_with f t t'" for z proof - obtain xs ys where wit: "xs \ t" "ys \ t'" "z \ f xs ys" using z by (auto simp: cross_with_def) have h: "h xs = h' ys" using assms(1)[OF wit(1-2)] wit(3) by auto have hys: "h' ys \ h ` t" using wit(1) by (auto simp: h[symmetric]) show ?thesis apply (rule exI[of _ "h xs"]) using wit hys h by (auto simp: cross_with_def) qed then show ?thesis using sub apply (transfer fixing: f h h') apply (auto simp: ran_def) apply fastforce+ done qed lemma fo_nmlzd_mono_sub: "X \ X' \ fo_nmlzd X xs \ fo_nmlzd X' xs" by (meson fo_nmlzd_def order_trans) lemma idx_join: assumes X\_props: "\vs. vs \ X\ \ fo_nmlzd AD vs \ length vs = length ns\" assumes X\_props: "\vs. vs \ X\ \ fo_nmlzd AD vs \ length vs = length ns\" assumes sd_ns: "sorted_distinct ns\" "sorted_distinct ns\" assumes ns_def: "ns = filter (\n. n \ set ns\) ns\" shows "idx_join AD ns ns\ X\ ns\ X\ = eval_conj_set AD ns\ X\ ns\ X\" proof - have ect_empty: "x \ X\ \ x' \ X\ \ fo_nmlz AD (proj_tuple ns (zip ns\ x)) \ fo_nmlz AD (proj_tuple ns (zip ns\ x')) \ eval_conj_tuple AD ns\ ns\ x x' = {}" if "X\' \ X\" "X\' \ X\" for X\' X\' and x x' apply (rule eval_conj_tuple_empty[where ?ns="filter (\n. n \ set ns\) ns\"]) using X\_props X\_props that sd_ns by (auto simp: ns_def ad_agr_close_set_def split: if_splits) have cross_eval_conj_tuple: "(\X\''. eval_conj_set AD ns\ X\'' ns\) = cross_with (eval_conj_tuple AD ns\ ns\)" for AD :: "'a set" and ns\ ns\ by (rule ext)+ (auto simp: eval_conj_set_def cross_with_def) have "idx_join AD ns ns\ X\ ns\ X\ = cross_with (eval_conj_tuple AD ns\ ns\) X\ X\" unfolding idx_join_def Let_def cross_eval_conj_tuple by (rule mapping_join_cross_with[OF ect_empty]) auto moreover have "\ = eval_conj_set AD ns\ X\ ns\ X\" by (auto simp: cross_with_def eval_conj_set_def) finally show ?thesis . qed lemma proj_fmla_conj_sub: assumes AD_sub: "act_edom \ I \ AD" shows "fo_nmlz AD ` proj_fmla (Conj \ \) {\. esat \ I \ UNIV} \ fo_nmlz AD ` proj_fmla (Conj \ \) {\. esat \ I \ UNIV} \ fo_nmlz AD ` proj_fmla (Conj \ \) {\. esat (Conj \ \) I \ UNIV}" proof (rule subsetI) fix vs assume "vs \ fo_nmlz AD ` proj_fmla (Conj \ \) {\. esat \ I \ UNIV} \ fo_nmlz AD ` proj_fmla (Conj \ \) {\. esat \ I \ UNIV}" then obtain \ \' where \_def: "\ \ {\. esat \ I \ UNIV}" "vs = fo_nmlz AD (map \ (fv_fo_fmla_list (Conj \ \)))" "\' \ {\. esat \ I \ UNIV}" "vs = fo_nmlz AD (map \' (fv_fo_fmla_list (Conj \ \)))" unfolding proj_fmla_map by blast have ad_sub: "act_edom \ I \ AD" using assms(1) by auto have ad_agr: "ad_agr_list AD (map \ (fv_fo_fmla_list \)) (map \' (fv_fo_fmla_list \))" by (rule ad_agr_list_subset[OF _ fo_nmlz_eqD[OF trans[OF \_def(2)[symmetric] \_def(4)]]]) (auto simp: fv_fo_fmla_list_set) have "\ \ {\. esat \ I \ UNIV}" using esat_UNIV_cong[OF ad_agr_sets_restrict[OF iffD2[OF ad_agr_list_link]], OF ad_agr ad_sub] \_def(3) by blast then show "vs \ fo_nmlz AD ` proj_fmla (Conj \ \) {\. esat (Conj \ \) I \ UNIV}" using \_def(1,2) by (auto simp: proj_fmla_map) qed lemma eval_conj: fixes \ :: "('a :: infinite, 'b) fo_fmla" assumes wf: "fo_wf \ I t\" "fo_wf \ I t\" shows "fo_wf (Conj \ \) I (eval_conj (fv_fo_fmla_list \) t\ (fv_fo_fmla_list \) t\)" proof - obtain AD\ n\ X\ AD\ n\ X\ where ts_def: "t\ = (AD\, n\, X\)" "t\ = (AD\, n\, X\)" "AD\ = act_edom \ I" "AD\ = act_edom \ I" using assms by (cases t\, cases t\) auto have AD_sub: "act_edom \ I \ AD\" "act_edom \ I \ AD\" by (auto simp: ts_def(3,4)) obtain AD n X where AD_X_def: "eval_conj (fv_fo_fmla_list \) t\ (fv_fo_fmla_list \) t\ = (AD, n, X)" by (cases "eval_conj (fv_fo_fmla_list \) t\ (fv_fo_fmla_list \) t\") auto have AD_def: "AD = act_edom (Conj \ \) I" "act_edom (Conj \ \) I \ AD" "AD\ \ AD" "AD\ \ AD" "AD = AD\ \ AD\" using AD_X_def by (auto simp: ts_def Let_def) have n_def: "n = nfv (Conj \ \)" using AD_X_def by (auto simp: ts_def Let_def nfv_card fv_fo_fmla_list_set) define S\ where "S\ \ {\. esat \ I \ UNIV}" define S\ where "S\ \ {\. esat \ I \ UNIV}" define AD\\ where "AD\\ = AD - AD\" define AD\\ where "AD\\ = AD - AD\" define ns\ where "ns\ = fv_fo_fmla_list \" define ns\ where "ns\ = fv_fo_fmla_list \" define ns where "ns = filter (\n. n \ fv_fo_fmla \) (fv_fo_fmla_list \)" define ns\' where "ns\' = filter (\n. n \ fv_fo_fmla \) (fv_fo_fmla_list \)" define ns\' where "ns\' = filter (\n. n \ fv_fo_fmla \) (fv_fo_fmla_list \)" note X\_def = fo_wf_X[OF wf(1)[unfolded ts_def(1)], unfolded proj_fmla_def, folded S\_def] note X\_def = fo_wf_X[OF wf(2)[unfolded ts_def(2)], unfolded proj_fmla_def, folded S\_def] have sd_ns: "sorted_distinct ns\" "sorted_distinct ns\" by (auto simp: ns\_def ns\_def sorted_distinct_fv_list) have ad_agr_X\: "ad_agr_close_set AD\\ X\ = fo_nmlz AD ` proj_vals S\ ns\" unfolding X\_def ad_agr_close_set_nmlz_eq ns\_def[symmetric] AD\\_def apply (rule ad_agr_close_set_correct[OF AD_def(3) sd_ns(1)]) using AD_sub(1) esat_UNIV_ad_agr_list by (fastforce simp: ad_agr_list_link S\_def ns\_def) have ad_agr_X\: "ad_agr_close_set AD\\ X\ = fo_nmlz AD ` proj_vals S\ ns\" unfolding X\_def ad_agr_close_set_nmlz_eq ns\_def[symmetric] AD\\_def apply (rule ad_agr_close_set_correct[OF AD_def(4) sd_ns(2)]) using AD_sub(2) esat_UNIV_ad_agr_list by (fastforce simp: ad_agr_list_link S\_def ns\_def) have idx_join_eval_conj: "idx_join AD (filter (\n. n \ set ns\) ns\) ns\ (ad_agr_close_set AD\\ X\) ns\ (ad_agr_close_set AD\\ X\) = eval_conj_set AD ns\ (ad_agr_close_set AD\\ X\) ns\ (ad_agr_close_set AD\\ X\)" apply (rule idx_join[OF _ _ sd_ns]) unfolding ad_agr_X\ ad_agr_X\ by (auto simp: fo_nmlz_sound fo_nmlz_length proj_vals_def) have fv_sub: "fv_fo_fmla (Conj \ \) = fv_fo_fmla \ \ set (fv_fo_fmla_list \)" "fv_fo_fmla (Conj \ \) = fv_fo_fmla \ \ set (fv_fo_fmla_list \)" by (auto simp: fv_fo_fmla_list_set) note res_left_alt = ext_tuple_ad_agr_close[OF S\_def AD_sub(1) AD_def(3) X\_def(1)[folded S\_def] ns\'_def sorted_distinct_fv_list fv_sub(1)] note res_right_alt = ext_tuple_ad_agr_close[OF S\_def AD_sub(2) AD_def(4) X\_def(1)[folded S\_def] ns\'_def sorted_distinct_fv_list fv_sub(2)] note eval_conj_set = eval_conj_set_correct[OF ns\'_def[folded fv_fo_fmla_list_set] ns\'_def[folded fv_fo_fmla_list_set] res_left_alt(2) res_right_alt(2) sorted_distinct_fv_list sorted_distinct_fv_list] have "X = fo_nmlz AD ` proj_fmla (Conj \ \) {\. esat \ I \ UNIV} \ fo_nmlz AD ` proj_fmla (Conj \ \) {\. esat \ I \ UNIV}" using AD_X_def apply (simp add: ts_def(1,2) Let_def ts_def(3,4)[symmetric] AD_def(5)[symmetric] idx_join_eval_conj[unfolded ns\_def ns\_def AD\\_def AD\\_def]) unfolding eval_conj_set proj_fmla_def unfolding res_left_alt(1) res_right_alt(1) S\_def S\_def by auto then have eval: "eval_conj (fv_fo_fmla_list \) t\ (fv_fo_fmla_list \) t\ = eval_abs (Conj \ \) I" using proj_fmla_conj_sub[OF AD_def(4)[unfolded ts_def(4)], of \] unfolding AD_X_def AD_def(1)[symmetric] n_def eval_abs_def by (auto simp: proj_fmla_map) have wf_conj: "wf_fo_intp (Conj \ \) I" using wf by (auto simp: ts_def) show ?thesis using fo_wf_eval_abs[OF wf_conj] by (auto simp: eval) qed lemma map_values_cluster: "(\w z Z. Z \ X \ z \ Z \ w \ f (h z) {z} \ w \ f (h z) Z) \ (\w z Z. Z \ X \ z \ Z \ w \ f (h z) Z \ (\z'\Z. w \ f (h z) {z'})) \ set_of_idx (Mapping.map_values f (cluster (Some \ h) X)) = \((\x. f (h x) {x}) ` X)" apply transfer apply (auto simp: ran_def) apply (smt (verit, del_insts) mem_Collect_eq subset_eq) apply (smt (z3) imageI mem_Collect_eq subset_iff) done lemma fo_nmlz_twice: assumes "sorted_distinct ns" "sorted_distinct ns'" "set ns \ set ns'" shows "fo_nmlz AD (proj_tuple ns (zip ns' (fo_nmlz AD (map \ ns')))) = fo_nmlz AD (map \ ns)" proof - obtain \' where \': "fo_nmlz AD (map \ ns') = map \' ns'" using exists_map[where ?ys="fo_nmlz AD (map \ ns')" and ?xs=ns'] assms by (auto simp: fo_nmlz_length) have proj: "proj_tuple ns (zip ns' (map \' ns')) = map \' ns" by (rule proj_tuple_map[OF assms]) show ?thesis unfolding \' proj apply (rule fo_nmlz_eqI) using \' by (metis ad_agr_list_comm ad_agr_list_subset assms(3) fo_nmlz_ad_agr) qed lemma map_values_cong: assumes "\x y. Mapping.lookup t x = Some y \ f x y = f' x y" shows "Mapping.map_values f t = Mapping.map_values f' t" proof - have "map_option (f x) (Mapping.lookup t x) = map_option (f' x) (Mapping.lookup t x)" for x using assms by (cases "Mapping.lookup t x") auto then show ?thesis by (auto simp: lookup_map_values intro!: mapping_eqI) qed lemma ad_agr_close_set_length: "z \ ad_agr_close_set AD X \ (\x. x \ X \ length x = n) \ length z = n" by (auto simp: ad_agr_close_set_def ad_agr_close_def split: if_splits dest: ad_agr_close_rec_length) lemma ad_agr_close_set_sound: "z \ ad_agr_close_set (AD - AD') X \ (\x. x \ X \ fo_nmlzd AD' x) \ AD' \ AD \ fo_nmlzd AD z" using ad_agr_close_sound[where ?X=AD' and ?Y="AD - AD'"] by (auto simp: ad_agr_close_set_def Set.is_empty_def split: if_splits) (metis Diff_partition Un_Diff_cancel) lemma ext_tuple_set_length: "z \ ext_tuple_set AD ns ns' X \ (\x. x \ X \ length x = length ns) \ length z = length ns + length ns'" by (auto simp: ext_tuple_set_def ext_tuple_def fo_nmlz_length merge_length dest: nall_tuples_rec_length split: if_splits) lemma eval_ajoin: fixes \ :: "('a :: infinite, 'b) fo_fmla" assumes wf: "fo_wf \ I t\" "fo_wf \ I t\" shows "fo_wf (Conj \ (Neg \)) I (eval_ajoin (fv_fo_fmla_list \) t\ (fv_fo_fmla_list \) t\)" proof - obtain AD\ n\ X\ AD\ n\ X\ where ts_def: "t\ = (AD\, n\, X\)" "t\ = (AD\, n\, X\)" "AD\ = act_edom \ I" "AD\ = act_edom \ I" using assms by (cases t\, cases t\) auto have AD_sub: "act_edom \ I \ AD\" "act_edom \ I \ AD\" by (auto simp: ts_def(3,4)) obtain AD n X where AD_X_def: "eval_ajoin (fv_fo_fmla_list \) t\ (fv_fo_fmla_list \) t\ = (AD, n, X)" by (cases "eval_ajoin (fv_fo_fmla_list \) t\ (fv_fo_fmla_list \) t\") auto have AD_def: "AD = act_edom (Conj \ (Neg \)) I" "act_edom (Conj \ (Neg \)) I \ AD" "AD\ \ AD" "AD\ \ AD" "AD = AD\ \ AD\" using AD_X_def by (auto simp: ts_def Let_def) have n_def: "n = nfv (Conj \ (Neg \))" using AD_X_def by (auto simp: ts_def Let_def nfv_card fv_fo_fmla_list_set) define S\ where "S\ \ {\. esat \ I \ UNIV}" define S\ where "S\ \ {\. esat \ I \ UNIV}" define both where "both = remdups_adj (sort (fv_fo_fmla_list \ @ fv_fo_fmla_list \))" define ns\' where "ns\' = filter (\n. n \ fv_fo_fmla \) (fv_fo_fmla_list \)" define ns\' where "ns\' = filter (\n. n \ fv_fo_fmla \) (fv_fo_fmla_list \)" define AD\\ where "AD\\ = AD - AD\" define AD\\ where "AD\\ = AD - AD\" define ns\ where "ns\ = fv_fo_fmla_list \" define ns\ where "ns\ = fv_fo_fmla_list \" define ns where "ns = filter (\n. n \ set ns\) ns\" define X\' where "X\' = ext_tuple_set AD ns\ ns\' (ad_agr_close_set AD\\ X\)" define idx\ where "idx\ = cluster (Some \ (\xs. fo_nmlz AD\ (proj_tuple ns (zip ns\ xs)))) (ad_agr_close_set AD\\ X\)" define idx\ where "idx\ = cluster (Some \ (\ys. fo_nmlz AD\ (proj_tuple ns (zip ns\ ys)))) X\" define res where "res = Mapping.map_values (\xs X. case Mapping.lookup idx\ xs of Some Y \ eval_conj_set AD ns\ X ns\ (ad_agr_close_set AD\\ (ext_tuple_set AD\ ns ns\' {xs} - Y)) | _ \ ext_tuple_set AD ns\ ns\' X) idx\" note X\_def = fo_wf_X[OF wf(1)[unfolded ts_def(1)], unfolded proj_fmla_def, folded S\_def] note X\_def = fo_wf_X[OF wf(2)[unfolded ts_def(2)], unfolded proj_fmla_def, folded S\_def] have fv_sub: "fv_fo_fmla (Conj \ (Neg \)) = fv_fo_fmla \ \ set (fv_fo_fmla_list \)" by (auto simp: fv_fo_fmla_list_set) have fv_sort: "fv_fo_fmla_list (Conj \ (Neg \)) = both" unfolding both_def apply (rule sorted_distinct_set_unique) using sorted_distinct_fv_list by (auto simp: fv_fo_fmla_list_def distinct_remdups_adj_sort) have AD_disj: "AD\ \ AD\\ = {}" "AD\ \ AD\\ = {}" by (auto simp: AD\\_def AD\\_def) have AD_delta: "AD = AD\ \ AD\\" "AD = AD\ \ AD\\" by (auto simp: AD\\_def AD\\_def AD_def ts_def) have fo_nmlzd_X: "Ball X\ (fo_nmlzd AD\)" "Ball X\ (fo_nmlzd AD\)" using wf by (auto simp: ts_def) have Ball_ad_agr: "Ball (ad_agr_close_set AD\\ X\) (fo_nmlzd AD)" using ad_agr_close_sound[where ?X="AD\" and ?Y="AD\\"] fo_nmlzd_X(1) by (auto simp: ad_agr_close_set_eq[OF fo_nmlzd_X(1)] AD_disj AD_delta) have ad_agr_\: "\\ \. ad_agr_sets (set (fv_fo_fmla_list \)) (set (fv_fo_fmla_list \)) AD\ \ \ \ \ \ S\ \ \ \ S\" "\\ \. ad_agr_sets (set (fv_fo_fmla_list \)) (set (fv_fo_fmla_list \)) AD \ \ \ \ \ S\ \ \ \ S\" using esat_UNIV_cong[OF ad_agr_sets_restrict, OF _ subset_refl] ad_agr_sets_mono AD_sub(1) subset_trans[OF AD_sub(1) AD_def(3)] unfolding S\_def by blast+ have ad_agr_S\: "\' \ S\ \ ad_agr_list AD\ (map \' ns\) (map \'' ns\) \ \'' \ S\" for \' \'' using ad_agr_\ by (auto simp: ad_agr_list_link ns\_def) have ad_agr_\: "\\ \. ad_agr_sets (set (fv_fo_fmla_list \)) (set (fv_fo_fmla_list \)) AD\ \ \ \ \ \ S\ \ \ \ S\" using esat_UNIV_cong[OF ad_agr_sets_restrict, OF _ subset_refl] ad_agr_sets_mono[OF AD_sub(2)] unfolding S\_def by blast+ have ad_agr_S\: "\' \ S\ \ ad_agr_list AD\ (map \' ns\) (map \'' ns\) \ \'' \ S\" for \' \'' using ad_agr_\ by (auto simp: ad_agr_list_link ns\_def) have aux: "sorted_distinct ns\" "sorted_distinct ns\'" "sorted_distinct both" "set ns\ \ set ns\' = {}" "set ns\ \ set ns\' = set both" by (auto simp: ns\_def ns\'_def fv_sort[symmetric] fv_fo_fmla_list_set sorted_distinct_fv_list intro: sorted_filter[where ?f=id, simplified]) have aux2: "ns\' = filter (\n. n \ set ns\) ns\'" "ns\ = filter (\n. n \ set ns\') ns\" by (auto simp: ns\_def ns\'_def ns\_def ns\'_def fv_fo_fmla_list_set) have aux3: "set ns\' \ set ns = {}" "set ns\' \ set ns = set ns\" by (auto simp: ns\_def ns\'_def ns\_def ns_def fv_fo_fmla_list_set) have aux4: "set ns \ set ns\' = {}" "set ns \ set ns\' = set ns\" by (auto simp: ns\_def ns\'_def ns\_def ns_def fv_fo_fmla_list_set) have aux5: "ns\' = filter (\n. n \ set ns\) ns\" "ns\' = filter (\n. n \ set ns\) ns\" by (auto simp: ns\_def ns\'_def ns\_def ns\'_def fv_fo_fmla_list_set) have aux6: "set ns\ \ set ns\' = {}" "set ns\ \ set ns\' = set both" by (auto simp: ns\_def ns\'_def ns\_def ns\'_def both_def fv_fo_fmla_list_set) have ns_sd: "sorted_distinct ns" "sorted_distinct ns\" "sorted_distinct ns\" "set ns \ set ns\" "set ns \ set ns\" "set ns \ set both" "set ns\' \ set ns\" "set ns\ \ set both" by (auto simp: ns_def ns\_def ns\'_def ns\_def both_def sorted_distinct_fv_list intro: sorted_filter[where ?f=id, simplified]) have ns_sd': "sorted_distinct ns\'" by (auto simp: ns\'_def sorted_distinct_fv_list intro: sorted_filter[where ?f=id, simplified]) have ns: "ns = filter (\n. n \ fv_fo_fmla \) (fv_fo_fmla_list \)" by (rule sorted_distinct_set_unique) (auto simp: ns_def ns\_def ns\_def fv_fo_fmla_list_set sorted_distinct_fv_list intro: sorted_filter[where ?f=id, simplified]) have len_ns\: "length ns + length ns\' = length ns\" using sum_length_filter_compl[where ?P="\n. n \ fv_fo_fmla \" and ?xs="fv_fo_fmla_list \"] by (auto simp: ns ns\_def ns\'_def ns\_def fv_fo_fmla_list_set) have res_eq: "res = Mapping.map_values (\xs X. case Mapping.lookup idx\ xs of Some Y \ idx_join AD ns ns\ X ns\ (ad_agr_close_set AD\\ (ext_tuple_set AD\ ns ns\' {xs} - Y)) | _ \ ext_tuple_set AD ns\ ns\' X) idx\" proof - have ad_agr_X\: "ad_agr_close_set AD\\ X\ = fo_nmlz AD ` proj_vals S\ ns\" unfolding X\_def ad_agr_close_set_nmlz_eq ns\_def[symmetric] apply (rule ad_agr_close_set_correct[OF AD_def(3) aux(1), folded AD\\_def]) using ad_agr_S\ ad_agr_list_comm by (fastforce simp: ad_agr_list_link) have idx_eval: "idx_join AD ns ns\ y ns\ (ad_agr_close_set AD\\ (ext_tuple_set AD\ ns ns\' {x} - x2)) = eval_conj_set AD ns\ y ns\ (ad_agr_close_set AD\\ (ext_tuple_set AD\ ns ns\' {x} - x2))" if lup: "Mapping.lookup idx\ x = Some y" "Mapping.lookup idx\ x = Some x2" for x y x2 proof - have "vs \ y \ fo_nmlzd AD vs \ length vs = length ns\" for vs using lup(1) by (auto simp: idx\_def lookup_cluster' ad_agr_X\ fo_nmlz_sound fo_nmlz_length proj_vals_def split: if_splits) moreover have "vs \ ad_agr_close_set AD\\ (ext_tuple_set AD\ ns ns\' {x} - x2) \ fo_nmlzd AD vs" for vs apply (rule ad_agr_close_set_sound[OF _ _ AD_def(4), folded AD\\_def, where ?X="ext_tuple_set AD\ ns ns\' {x} - x2"]) using lup(1) by (auto simp: idx\_def lookup_cluster' ext_tuple_set_def fo_nmlz_sound split: if_splits) moreover have "vs \ ad_agr_close_set AD\\ (ext_tuple_set AD\ ns ns\' {x} - x2) \ length vs = length ns\" for vs apply (erule ad_agr_close_set_length) apply (rule ext_tuple_set_length[where ?AD=AD\ and ?ns=ns and ?ns'=ns\' and ?X="{x}", unfolded len_ns\]) using lup(1) ns_sd(1,2,4) by (auto simp: idx\_def lookup_cluster' fo_nmlz_length ad_agr_X\ proj_vals_def intro!: proj_tuple_length split: if_splits) ultimately show ?thesis by (auto intro!: idx_join[OF _ _ ns_sd(2-3) ns_def]) qed show ?thesis unfolding res_def by (rule map_values_cong) (auto simp: idx_eval split: option.splits) qed have eval_conj: "eval_conj_set AD ns\ {x} ns\ (ad_agr_close_set AD\\ (ext_tuple_set AD\ ns ns\' {fo_nmlz AD\ (proj_tuple ns (zip ns\ x))} - Y)) = ext_tuple_set AD ns\ ns\' {x} \ ext_tuple_set AD ns\ ns\' (fo_nmlz AD ` proj_vals {\ \ - S\. ad_agr_list AD\ (map \ ns) (map \' ns)} ns\)" if x_ns: "proj_tuple ns (zip ns\ x) = map \' ns" and x_proj_singleton: "{x} = fo_nmlz AD ` proj_vals {\} ns\" and Some: "Mapping.lookup idx\ (fo_nmlz AD\ (proj_tuple ns (zip ns\ x))) = Some Y" for x Y \ \' proof - have "Y = {ys \ fo_nmlz AD\ ` proj_vals S\ ns\. fo_nmlz AD\ (proj_tuple ns (zip ns\ ys)) = fo_nmlz AD\ (map \' ns)}" using Some apply (auto simp: X\_def idx\_def ns\_def x_ns lookup_cluster' split: if_splits) done moreover have "\ = fo_nmlz AD\ ` proj_vals {\ \ S\. fo_nmlz AD\ (map \ ns) = fo_nmlz AD\ (map \' ns)} ns\" by (auto simp: proj_vals_def fo_nmlz_twice[OF ns_sd(1,3,5)])+ moreover have "\ = fo_nmlz AD\ ` proj_vals {\ \ S\. ad_agr_list AD\ (map \ ns) (map \' ns)} ns\" by (auto simp: fo_nmlz_eq) ultimately have Y_def: "Y = fo_nmlz AD\ ` proj_vals {\ \ S\. ad_agr_list AD\ (map \ ns) (map \' ns)} ns\" by auto have R_def: "{fo_nmlz AD\ (map \' ns)} = fo_nmlz AD\ ` proj_vals {\. ad_agr_list AD\ (map \ ns) (map \' ns)} ns" using ad_agr_list_refl by (auto simp: proj_vals_def intro: fo_nmlz_eqI) have "ext_tuple_set AD\ ns ns\' {fo_nmlz AD\ (map \' ns)} = fo_nmlz AD\ ` proj_vals {\. ad_agr_list AD\ (map \ ns) (map \' ns)} ns\" apply (rule ext_tuple_correct[OF ns_sd(1) aux(2) ns_sd(3) aux4 R_def]) using ad_agr_list_trans ad_agr_list_comm apply (auto simp: ad_agr_list_link) by fast then have "ext_tuple_set AD\ ns ns\' {fo_nmlz AD\ (map \' ns)} - Y = fo_nmlz AD\ ` proj_vals {\ \ -S\. ad_agr_list AD\ (map \ ns) (map \' ns)} ns\" apply (auto simp: Y_def proj_vals_def fo_nmlz_eq) using ad_agr_S\ ad_agr_list_comm by blast+ moreover have "ad_agr_close_set AD\\ (fo_nmlz AD\ ` proj_vals {\ \ -S\. ad_agr_list AD\ (map \ ns) (map \' ns)} ns\) = fo_nmlz AD ` proj_vals {\ \ -S\. ad_agr_list AD\ (map \ ns) (map \' ns)} ns\" unfolding ad_agr_close_set_eq[OF Ball_fo_nmlzd] apply (rule ad_agr_close_set_correct[OF AD_def(4) ns_sd(3), folded AD\\_def]) apply (auto simp: ad_agr_list_link) using ad_agr_S\ ad_agr_list_comm ad_agr_list_subset[OF ns_sd(5)] ad_agr_list_trans by blast+ ultimately have comp_proj: "ad_agr_close_set AD\\ (ext_tuple_set AD\ ns ns\' {fo_nmlz AD\ (map \' ns)} - Y) = fo_nmlz AD ` proj_vals {\ \ -S\. ad_agr_list AD\ (map \ ns) (map \' ns)} ns\" by simp have "ext_tuple_set AD ns\ ns\' (fo_nmlz AD ` proj_vals {\ \ - S\. ad_agr_list AD\ (map \ ns) (map \' ns)} ns\) = fo_nmlz AD ` proj_vals {\ \ - S\. ad_agr_list AD\ (map \ ns) (map \' ns)} both" apply (rule ext_tuple_correct[OF ns_sd(3) ns_sd'(1) aux(3) aux6 refl]) apply (auto simp: ad_agr_list_link) using ad_agr_S\ ad_agr_list_comm ad_agr_list_subset[OF ns_sd(5)] ad_agr_list_trans ad_agr_list_mono[OF AD_def(4)] by fast+ show "eval_conj_set AD ns\ {x} ns\ (ad_agr_close_set AD\\ (ext_tuple_set AD\ ns ns\' {fo_nmlz AD\ (proj_tuple ns (zip ns\ x))} - Y)) = ext_tuple_set AD ns\ ns\' {x} \ ext_tuple_set AD ns\ ns\' (fo_nmlz AD ` proj_vals {\ \ - S\. ad_agr_list AD\ (map \ ns) (map \' ns)} ns\)" unfolding x_ns comp_proj using eval_conj_set_correct[OF aux5 x_proj_singleton refl aux(1) ns_sd(3)] by auto qed have "X = set_of_idx res" using AD_X_def unfolding eval_ajoin.simps ts_def(1,2) Let_def AD_def(5)[symmetric] fv_fo_fmla_list_set ns\'_def[symmetric] fv_sort[symmetric] proj_fmla_def S\_def[symmetric] S\_def[symmetric] AD\\_def[symmetric] AD\\_def[symmetric] ns\_def[symmetric] ns\'_def[symmetric, folded fv_fo_fmla_list_set[of \, folded ns\_def] ns\_def] ns\_def[symmetric] ns_def[symmetric] X\'_def[symmetric] idx\_def[symmetric] idx\_def[symmetric] res_eq[symmetric] by auto moreover have "\ = (\x\ad_agr_close_set AD\\ X\. case Mapping.lookup idx\ (fo_nmlz AD\ (proj_tuple ns (zip ns\ x))) of None \ ext_tuple_set AD ns\ ns\' {x} | Some Y \ eval_conj_set AD ns\ {x} ns\ (ad_agr_close_set AD\\ (ext_tuple_set AD\ ns ns\' {fo_nmlz AD\ (proj_tuple ns (zip ns\ x))} - Y)))" unfolding res_def[unfolded idx\_def] apply (rule map_values_cluster) apply (auto simp: eval_conj_set_def split: option.splits) apply (auto simp: ext_tuple_set_def split: if_splits) done moreover have "\ = fo_nmlz AD ` proj_fmla (Conj \ (Neg \)) {\. esat \ I \ UNIV} - fo_nmlz AD ` proj_fmla (Conj \ (Neg \)) {\. esat \ I \ UNIV}" unfolding S\_def[symmetric] S\_def[symmetric] proj_fmla_def fv_sort proof (rule set_eqI, rule iffI) fix t assume "t \ (\x\ad_agr_close_set AD\\ X\. case Mapping.lookup idx\ (fo_nmlz AD\ (proj_tuple ns (zip ns\ x))) of None \ ext_tuple_set AD ns\ ns\' {x} | Some Y \ eval_conj_set AD ns\ {x} ns\ (ad_agr_close_set AD\\ (ext_tuple_set AD\ ns ns\' {fo_nmlz AD\ (proj_tuple ns (zip ns\ x))} - Y)))" then obtain x where x: "x \ ad_agr_close_set AD\\ X\" "Mapping.lookup idx\ (fo_nmlz AD\ (proj_tuple ns (zip ns\ x))) = None \ t \ ext_tuple_set AD ns\ ns\' {x}" "\Y. Mapping.lookup idx\ (fo_nmlz AD\ (proj_tuple ns (zip ns\ x))) = Some Y \ t \ eval_conj_set AD ns\ {x} ns\ (ad_agr_close_set AD\\ (ext_tuple_set AD\ ns ns\' {fo_nmlz AD\ (proj_tuple ns (zip ns\ x))} - Y))" by (fastforce split: option.splits) obtain \ where val: "\ \ S\" "x = fo_nmlz AD (map \ ns\)" using ad_agr_close_correct[OF AD_def(3) ad_agr_\(1), folded AD\\_def] X\_def[folded proj_fmla_def] ad_agr_close_set_eq[OF fo_nmlzd_X(1)] x(1) apply (auto simp: proj_fmla_def proj_vals_def ns\_def) apply fast done obtain \' where \': "x = map \' ns\" using exists_map[where ?ys=x and ?xs=ns\] aux(1) by (auto simp: val(2) fo_nmlz_length) have x_proj_singleton: "{x} = fo_nmlz AD ` proj_vals {\} ns\" by (auto simp: val(2) proj_vals_def) have x_ns: "proj_tuple ns (zip ns\ x) = map \' ns" unfolding \' by (rule proj_tuple_map[OF ns_sd(1-2,4)]) have ad_agr_\_\': "ad_agr_list AD (map \ ns\) (map \' ns\)" using \' by (auto simp: val(2)) (metis fo_nmlz_ad_agr) have x_proj_ad_agr: "{x} = fo_nmlz AD ` proj_vals {\. ad_agr_list AD (map \ ns\) (map \' ns\)} ns\" using ad_agr_\_\' ad_agr_list_comm ad_agr_list_trans by (auto simp: val(2) proj_vals_def fo_nmlz_eq) blast have "t \ fo_nmlz AD ` \ (ext_tuple AD ns\ ns\' ` {x}) \ fo_nmlz AD (proj_tuple ns\ (zip both t)) \ {x}" apply (rule ext_tuple_sound(1)[OF aux x_proj_ad_agr]) apply (auto simp: ad_agr_list_link) using ad_agr_list_comm ad_agr_list_trans by blast+ then have x_proj: "t \ ext_tuple_set AD ns\ ns\' {x} \ x = fo_nmlz AD (proj_tuple ns\ (zip both t))" using ext_tuple_set_eq[where ?AD=AD] Ball_ad_agr x(1) by (auto simp: val(2) proj_vals_def) have x_S\: "t \ ext_tuple_set AD ns\ ns\' {x} \ t \ fo_nmlz AD ` proj_vals S\ both" using ext_tuple_correct[OF aux refl ad_agr_\(2)[folded ns\_def]] ext_tuple_set_mono[of "{x}" "fo_nmlz AD ` proj_vals S\ ns\"] val(1) by (fastforce simp: val(2) proj_vals_def) show "t \ fo_nmlz AD ` proj_vals S\ both - fo_nmlz AD ` proj_vals S\ both" proof (cases "Mapping.lookup idx\ (fo_nmlz AD\ (proj_tuple ns (zip ns\ x)))") case None have "False" if t_in_S\: "t \ fo_nmlz AD ` proj_vals S\ both" proof - obtain \ where \: "\ \ S\" "t = fo_nmlz AD (map \ both)" using t_in_S\ by (auto simp: proj_vals_def) obtain \' where t_\': "t = map \' both" using aux(3) exists_map[where ?ys=t and ?xs=both] by (auto simp: \(2) fo_nmlz_length) obtain \'' where \'': "fo_nmlz AD\ (map \ ns\) = map \'' ns\" using ns_sd exists_map[where ?ys="fo_nmlz AD\ (map \ ns\)" and xs=ns\] by (auto simp: fo_nmlz_length) have proj_\'': "proj_tuple ns (zip ns\ (map \'' ns\)) = map \'' ns" apply (rule proj_tuple_map) using ns_sd by auto have "proj_tuple ns\ (zip both t) = map \' ns\" unfolding t_\' apply (rule proj_tuple_map) using aux by auto then have x_\': "x = fo_nmlz AD (map \' ns\)" by (auto simp: x_proj[OF x(2)[OF None]]) obtain \''' where \''': "x = map \''' ns\" using aux exists_map[where ?ys=x and ?xs=ns\] by (auto simp: x_\' fo_nmlz_length) have ad_\_\': "ad_agr_list AD (map \ both) (map \' both)" using t_\' by (auto simp: \) (metis fo_nmlz_ad_agr) have ad_\_\'': "ad_agr_list AD\ (map \ ns\) (map \'' ns\)" using \'' by (metis fo_nmlz_ad_agr) have ad_\'_\''': "ad_agr_list AD (map \' ns\) (map \''' ns\)" using \''' by (auto simp: x_\') (metis fo_nmlz_ad_agr) have proj_\''': "proj_tuple ns (zip ns\ (map \''' ns\)) = map \''' ns" apply (rule proj_tuple_map) using aux ns_sd by auto have "fo_nmlz AD\ (proj_tuple ns (zip ns\ x)) = fo_nmlz AD\ (proj_tuple ns (zip ns\ (fo_nmlz AD\ (map \ ns\))))" unfolding \'' proj_\'' \''' proj_\''' apply (rule fo_nmlz_eqI) using ad_agr_list_trans ad_agr_list_subset ns_sd(4-6) ad_agr_list_mono[OF AD_def(4)] ad_agr_list_comm[OF ad_\'_\'''] ad_agr_list_comm[OF ad_\_\'] ad_\_\'' by metis then show ?thesis using None \(1) by (auto simp: idx\_def lookup_cluster' X\_def ns\_def[symmetric] proj_vals_def split: if_splits) qed then show ?thesis using x_S\[OF x(2)[OF None]] by auto next case (Some Y) have t_in: "t \ ext_tuple_set AD ns\ ns\' {x}" "t \ ext_tuple_set AD ns\ ns\' (fo_nmlz AD ` proj_vals {\ \ - S\. ad_agr_list AD\ (map \ ns) (map \' ns)} ns\)" using x(3)[OF Some] eval_conj[OF x_ns x_proj_singleton Some] by auto have "ext_tuple_set AD ns\ ns\' (fo_nmlz AD ` proj_vals {\ \ - S\. ad_agr_list AD\ (map \ ns) (map \' ns)} ns\) = fo_nmlz AD ` proj_vals {\ \ - S\. ad_agr_list AD\ (map \ ns) (map \' ns)} both" apply (rule ext_tuple_correct[OF ns_sd(3) ns_sd'(1) aux(3) aux6 refl]) apply (auto simp: ad_agr_list_link) using ad_agr_S\ ad_agr_list_comm ad_agr_list_subset[OF ns_sd(5)] ad_agr_list_trans ad_agr_list_mono[OF AD_def(4)] by fast+ then have t_both: "t \ fo_nmlz AD ` proj_vals {\ \ - S\. ad_agr_list AD\ (map \ ns) (map \' ns)} both" using t_in(2) by auto { assume "t \ fo_nmlz AD ` proj_vals S\ both" then obtain \ where \: "\ \ S\" "t = fo_nmlz AD (map \ both)" by (auto simp: proj_vals_def) obtain \' where \': "\' \ S\" "t = fo_nmlz AD (map \' both)" using t_both by (auto simp: proj_vals_def) have "False" using \ \' apply (auto simp: fo_nmlz_eq) using ad_agr_S\ ad_agr_list_comm ad_agr_list_subset[OF ns_sd(8)] ad_agr_list_mono[OF AD_def(4)] by blast } then show ?thesis using x_S\[OF t_in(1)] by auto qed next fix t assume t_in_asm: "t \ fo_nmlz AD ` proj_vals S\ both - fo_nmlz AD ` proj_vals S\ both" then obtain \ where val: "\ \ S\" "t = fo_nmlz AD (map \ both)" by (auto simp: proj_vals_def) define x where "x = fo_nmlz AD (map \ ns\)" obtain \' where \': "x = map \' ns\" using exists_map[where ?ys=x and ?xs=ns\] aux(1) by (auto simp: x_def fo_nmlz_length) have x_proj_singleton: "{x} = fo_nmlz AD ` proj_vals {\} ns\" by (auto simp: x_def proj_vals_def) have x_in_ad_agr_close: "x \ ad_agr_close_set AD\\ X\" using ad_agr_close_correct[OF AD_def(3) ad_agr_\(1), folded AD\\_def] val(1) unfolding ad_agr_close_set_eq[OF fo_nmlzd_X(1)] x_def unfolding X\_def[folded proj_fmla_def] proj_fmla_map by (fastforce simp: x_def ns\_def) have ad_agr_\_\': "ad_agr_list AD (map \ ns\) (map \' ns\)" using \' by (auto simp: x_def) (metis fo_nmlz_ad_agr) have x_proj_ad_agr: "{x} = fo_nmlz AD ` proj_vals {\. ad_agr_list AD (map \ ns\) (map \' ns\)} ns\" using ad_agr_\_\' ad_agr_list_comm ad_agr_list_trans by (auto simp: x_def proj_vals_def fo_nmlz_eq) blast+ have x_ns: "proj_tuple ns (zip ns\ x) = map \' ns" unfolding \' by (rule proj_tuple_map[OF ns_sd(1-2,4)]) have "ext_tuple_set AD ns\ ns\' {x} = fo_nmlz AD ` proj_vals {\. ad_agr_list AD (map \ ns\) (map \' ns\)} both" apply (rule ext_tuple_correct[OF aux x_proj_ad_agr]) using ad_agr_list_comm ad_agr_list_trans by (auto simp: ad_agr_list_link) blast+ then have t_in_ext_x: "t \ ext_tuple_set AD ns\ ns\' {x}" using ad_agr_\_\' by (auto simp: val(2) proj_vals_def) { fix Y assume Some: "Mapping.lookup idx\ (fo_nmlz AD\ (map \' ns)) = Some Y" have tmp: "proj_tuple ns (zip ns\ x) = map \' ns" unfolding \' by (rule proj_tuple_map[OF ns_sd(1) aux(1) ns_sd(4)]) have unfold: "ext_tuple_set AD ns\ ns\' (fo_nmlz AD ` proj_vals {\ \ - S\. ad_agr_list AD\ (map \ ns) (map \' ns)} ns\) = fo_nmlz AD ` proj_vals {\ \ - S\. ad_agr_list AD\ (map \ ns) (map \' ns)} both" apply (rule ext_tuple_correct[OF ns_sd(3) ns_sd'(1) aux(3) aux6 refl]) apply (auto simp: ad_agr_list_link) using ad_agr_S\ ad_agr_list_mono[OF AD_def(4)] ad_agr_list_comm ad_agr_list_trans ad_agr_list_subset[OF ns_sd(5)] by blast+ have "\ \ S\" using t_in_asm by (auto simp: val(2) proj_vals_def) moreover have "ad_agr_list AD\ (map \ ns) (map \' ns)" using ad_agr_\_\' ad_agr_list_mono[OF AD_def(4)] ad_agr_list_subset[OF ns_sd(4)] by blast ultimately have "t \ ext_tuple_set AD ns\ ns\' (fo_nmlz AD ` proj_vals {\ \ - S\. ad_agr_list AD\ (map \ ns) (map \' ns)} ns\)" unfolding unfold val(2) by (auto simp: proj_vals_def) then have "t \ eval_conj_set AD ns\ {x} ns\ (ad_agr_close_set AD\\ (ext_tuple_set AD\ ns ns\' {fo_nmlz AD\ (map \' ns)} - Y))" using eval_conj[OF tmp x_proj_singleton Some[folded x_ns]] t_in_ext_x by (auto simp: x_ns) } then show "t \ (\x\ad_agr_close_set AD\\ X\. case Mapping.lookup idx\ (fo_nmlz AD\ (proj_tuple ns (zip ns\ x))) of None \ ext_tuple_set AD ns\ ns\' {x} | Some Y \ eval_conj_set AD ns\ {x} ns\ (ad_agr_close_set AD\\ (ext_tuple_set AD\ ns ns\' {fo_nmlz AD\ (proj_tuple ns (zip ns\ x))} - Y)))" using t_in_ext_x by (intro UN_I[OF x_in_ad_agr_close]) (auto simp: x_ns split: option.splits) qed ultimately have X_def: "X = fo_nmlz AD ` proj_fmla (Conj \ (Neg \)) {\. esat \ I \ UNIV} - fo_nmlz AD ` proj_fmla (Conj \ (Neg \)) {\. esat \ I \ UNIV}" by simp have AD_Neg_sub: "act_edom (Neg \) I \ AD" by (auto simp: AD_def(1)) have "X = fo_nmlz AD ` proj_fmla (Conj \ (Neg \)) {\. esat \ I \ UNIV} \ fo_nmlz AD ` proj_fmla (Conj \ (Neg \)) {\. esat (Neg \) I \ UNIV}" unfolding X_def by (auto simp: proj_fmla_map dest!: fo_nmlz_eqD) (metis AD_def(4) ad_agr_list_subset esat_UNIV_ad_agr_list fv_fo_fmla_list_set fv_sub sup_ge1 ts_def(4)) then have eval: "eval_ajoin (fv_fo_fmla_list \) t\ (fv_fo_fmla_list \) t\ = eval_abs (Conj \ (Neg \)) I" using proj_fmla_conj_sub[OF AD_Neg_sub, of \] unfolding AD_X_def AD_def(1)[symmetric] n_def eval_abs_def by (auto simp: proj_fmla_map) have wf_conj_neg: "wf_fo_intp (Conj \ (Neg \)) I" using wf by (auto simp: ts_def) show ?thesis using fo_wf_eval_abs[OF wf_conj_neg] by (auto simp: eval) qed lemma eval_disj: fixes \ :: "('a :: infinite, 'b) fo_fmla" assumes wf: "fo_wf \ I t\" "fo_wf \ I t\" shows "fo_wf (Disj \ \) I (eval_disj (fv_fo_fmla_list \) t\ (fv_fo_fmla_list \) t\)" proof - obtain AD\ n\ X\ AD\ n\ X\ where ts_def: "t\ = (AD\, n\, X\)" "t\ = (AD\, n\, X\)" "AD\ = act_edom \ I" "AD\ = act_edom \ I" using assms by (cases t\, cases t\) auto have AD_sub: "act_edom \ I \ AD\" "act_edom \ I \ AD\" by (auto simp: ts_def(3,4)) obtain AD n X where AD_X_def: "eval_disj (fv_fo_fmla_list \) t\ (fv_fo_fmla_list \) t\ = (AD, n, X)" by (cases "eval_disj (fv_fo_fmla_list \) t\ (fv_fo_fmla_list \) t\") auto have AD_def: "AD = act_edom (Disj \ \) I" "act_edom (Disj \ \) I \ AD" "AD\ \ AD" "AD\ \ AD" "AD = AD\ \ AD\" using AD_X_def by (auto simp: ts_def Let_def) have n_def: "n = nfv (Disj \ \)" using AD_X_def by (auto simp: ts_def Let_def nfv_card fv_fo_fmla_list_set) define S\ where "S\ \ {\. esat \ I \ UNIV}" define S\ where "S\ \ {\. esat \ I \ UNIV}" define ns\' where "ns\' = filter (\n. n \ fv_fo_fmla \) (fv_fo_fmla_list \)" define ns\' where "ns\' = filter (\n. n \ fv_fo_fmla \) (fv_fo_fmla_list \)" note X\_def = fo_wf_X[OF wf(1)[unfolded ts_def(1)], unfolded proj_fmla_def, folded S\_def] note X\_def = fo_wf_X[OF wf(2)[unfolded ts_def(2)], unfolded proj_fmla_def, folded S\_def] have fv_sub: "fv_fo_fmla (Disj \ \) = fv_fo_fmla \ \ set (fv_fo_fmla_list \)" "fv_fo_fmla (Disj \ \) = fv_fo_fmla \ \ set (fv_fo_fmla_list \)" by (auto simp: fv_fo_fmla_list_set) note res_left_alt = ext_tuple_ad_agr_close[OF S\_def AD_sub(1) AD_def(3) X\_def(1)[folded S\_def] ns\'_def sorted_distinct_fv_list fv_sub(1)] note res_right_alt = ext_tuple_ad_agr_close[OF S\_def AD_sub(2) AD_def(4) X\_def(1)[folded S\_def] ns\'_def sorted_distinct_fv_list fv_sub(2)] have "X = fo_nmlz AD ` proj_fmla (Disj \ \) {\. esat \ I \ UNIV} \ fo_nmlz AD ` proj_fmla (Disj \ \) {\. esat \ I \ UNIV}" using AD_X_def apply (simp add: ts_def(1,2) Let_def AD_def(5)[symmetric]) unfolding fv_fo_fmla_list_set proj_fmla_def ns\'_def[symmetric] ns\'_def[symmetric] S\_def[symmetric] S\_def[symmetric] using res_left_alt(1) res_right_alt(1) by auto then have eval: "eval_disj (fv_fo_fmla_list \) t\ (fv_fo_fmla_list \) t\ = eval_abs (Disj \ \) I" unfolding AD_X_def AD_def(1)[symmetric] n_def eval_abs_def by (auto simp: proj_fmla_map) have wf_disj: "wf_fo_intp (Disj \ \) I" using wf by (auto simp: ts_def) show ?thesis using fo_wf_eval_abs[OF wf_disj] by (auto simp: eval) qed lemma fv_ex_all: assumes "pos i (fv_fo_fmla_list \) = None" shows "fv_fo_fmla_list (Exists i \) = fv_fo_fmla_list \" "fv_fo_fmla_list (Forall i \) = fv_fo_fmla_list \" using pos_complete[of i "fv_fo_fmla_list \"] fv_fo_fmla_list_eq[of "Exists i \" \] fv_fo_fmla_list_eq[of "Forall i \" \] assms by (auto simp: fv_fo_fmla_list_set) lemma nfv_ex_all: assumes Some: "pos i (fv_fo_fmla_list \) = Some j" shows "nfv \ = Suc (nfv (Exists i \))" "nfv \ = Suc (nfv (Forall i \))" proof - have "i \ fv_fo_fmla \" "j < nfv \" "i \ set (fv_fo_fmla_list \)" using fv_fo_fmla_list_set pos_set[of i "fv_fo_fmla_list \"] pos_length[of i "fv_fo_fmla_list \"] Some by (fastforce simp: nfv_def)+ then show "nfv \ = Suc (nfv (Exists i \))" "nfv \ = Suc (nfv (Forall i \))" using nfv_card[of \] nfv_card[of "Exists i \"] nfv_card[of "Forall i \"] by (auto simp: finite_fv_fo_fmla) qed lemma fv_fo_fmla_list_exists: "fv_fo_fmla_list (Exists n \) = filter ((\) n) (fv_fo_fmla_list \)" by (auto simp: fv_fo_fmla_list_def) (metis (mono_tags, lifting) distinct_filter distinct_remdups_adj_sort distinct_remdups_id filter_set filter_sort remdups_adj_set sorted_list_of_set_sort_remdups sorted_remdups_adj sorted_sort sorted_sort_id) lemma eval_exists: fixes \ :: "('a :: infinite, 'b) fo_fmla" assumes wf: "fo_wf \ I t" shows "fo_wf (Exists i \) I (eval_exists i (fv_fo_fmla_list \) t)" proof - obtain AD n X where t_def: "t = (AD, n, X)" "AD = act_edom \ I" "AD = act_edom (Exists i \) I" using assms by (cases t) auto note X_def = fo_wf_X[OF wf[unfolded t_def], folded t_def(2)] have eval: "eval_exists i (fv_fo_fmla_list \) t = eval_abs (Exists i \) I" proof (cases "pos i (fv_fo_fmla_list \)") case None note fv_eq = fv_ex_all[OF None] have "X = fo_nmlz AD ` proj_fmla (Exists i \) {\. esat \ I \ UNIV}" unfolding X_def by (auto simp: proj_fmla_def fv_eq) also have "\ = fo_nmlz AD ` proj_fmla (Exists i \) {\. esat (Exists i \) I \ UNIV}" using esat_exists_not_fv[of i \ UNIV I] pos_complete[OF None] by (simp add: fv_fo_fmla_list_set) finally show ?thesis by (auto simp: t_def None eval_abs_def fv_eq nfv_def) next case (Some j) have "fo_nmlz AD ` rem_nth j ` X = fo_nmlz AD ` proj_fmla (Exists i \) {\. esat (Exists i \) I \ UNIV}" proof (rule set_eqI, rule iffI) fix vs assume "vs \ fo_nmlz AD ` rem_nth j ` X" then obtain ws where ws_def: "ws \ fo_nmlz AD ` proj_fmla \ {\. esat \ I \ UNIV}" "vs = fo_nmlz AD (rem_nth j ws)" unfolding X_def by auto then obtain \ where \_def: "esat \ I \ UNIV" "ws = fo_nmlz AD (map \ (fv_fo_fmla_list \))" by (auto simp: proj_fmla_map) obtain \ where \_def: "ws = map \ (fv_fo_fmla_list \)" using fo_nmlz_map \_def(2) by blast have esat_\: "esat (Exists i \) I \ UNIV" using esat_UNIV_ad_agr_list[OF fo_nmlz_ad_agr[of AD "map \ (fv_fo_fmla_list \)", folded \_def(2), unfolded \_def]] \_def(1) by (auto simp: t_def intro!: exI[of _ "\ i"]) have rem_nth_ws: "rem_nth j ws = map \ (fv_fo_fmla_list (Exists i \))" using rem_nth_sound[of "fv_fo_fmla_list \" i j \] sorted_distinct_fv_list Some unfolding fv_fo_fmla_list_exists \_def by auto have "vs \ fo_nmlz AD ` proj_fmla (Exists i \) {\. esat (Exists i \) I \ UNIV}" using ws_def(2) esat_\ unfolding rem_nth_ws by (auto simp: proj_fmla_map) then show "vs \ fo_nmlz AD ` proj_fmla (Exists i \) {\. esat (Exists i \) I \ UNIV}" by auto next fix vs assume assm: "vs \ fo_nmlz AD ` proj_fmla (Exists i \) {\. esat (Exists i \) I \ UNIV}" from assm obtain \ where \_def: "vs = fo_nmlz AD (map \ (fv_fo_fmla_list (Exists i \)))" "esat (Exists i \) I \ UNIV" by (auto simp: proj_fmla_map) then obtain x where x_def: "esat \ I (\(i := x)) UNIV" by auto define ws where "ws \ fo_nmlz AD (map (\(i := x)) (fv_fo_fmla_list \))" then have "length ws = nfv \" using nfv_def fo_nmlz_length by (metis length_map) then have ws_in: "ws \ fo_nmlz AD ` proj_fmla \ {\. esat \ I \ UNIV}" using x_def ws_def by (auto simp: fo_nmlz_sound proj_fmla_map) obtain \ where \_def: "ws = map \ (fv_fo_fmla_list \)" using fo_nmlz_map ws_def by blast have rem_nth_ws: "rem_nth j ws = map \ (fv_fo_fmla_list (Exists i \))" using rem_nth_sound[of "fv_fo_fmla_list \" i j] sorted_distinct_fv_list Some unfolding fv_fo_fmla_list_exists \_def by auto have "set (fv_fo_fmla_list (Exists i \)) \ set (fv_fo_fmla_list \)" by (auto simp: fv_fo_fmla_list_exists) then have ad_agr: "ad_agr_list AD (map (\(i := x)) (fv_fo_fmla_list (Exists i \))) (map \ (fv_fo_fmla_list (Exists i \)))" by (rule ad_agr_list_subset) (rule fo_nmlz_ad_agr[of AD "map (\(i := x)) (fv_fo_fmla_list \)", folded ws_def, unfolded \_def]) have map_fv_cong: "map (\(i := x)) (fv_fo_fmla_list (Exists i \)) = map \ (fv_fo_fmla_list (Exists i \))" by (auto simp: fv_fo_fmla_list_exists) have vs_rem_nth: "vs = fo_nmlz AD (rem_nth j ws)" unfolding \_def(1) rem_nth_ws apply (rule fo_nmlz_eqI) using ad_agr[unfolded map_fv_cong] . show "vs \ fo_nmlz AD ` rem_nth j ` X" using Some ws_in unfolding vs_rem_nth X_def by auto qed then show ?thesis using nfv_ex_all[OF Some] by (auto simp: t_def Some eval_abs_def nfv_def) qed have wf_ex: "wf_fo_intp (Exists i \) I" using wf by (auto simp: t_def) show ?thesis using fo_wf_eval_abs[OF wf_ex] by (auto simp: eval) qed lemma fv_fo_fmla_list_forall: "fv_fo_fmla_list (Forall n \) = filter ((\) n) (fv_fo_fmla_list \)" by (auto simp: fv_fo_fmla_list_def) (metis (mono_tags, lifting) distinct_filter distinct_remdups_adj_sort distinct_remdups_id filter_set filter_sort remdups_adj_set sorted_list_of_set_sort_remdups sorted_remdups_adj sorted_sort sorted_sort_id) lemma pairwise_take_drop: assumes "pairwise P (set (zip xs ys))" "length xs = length ys" shows "pairwise P (set (zip (take i xs @ drop (Suc i) xs) (take i ys @ drop (Suc i) ys)))" by (rule pairwise_subset[OF assms(1)]) (auto simp: set_zip assms(2)) lemma fo_nmlz_set_card: "fo_nmlz AD xs = xs \ set xs = set xs \ Inl ` AD \ Inr ` {.. ad_agr_list AD (take i xs @ drop (Suc i) xs) (take i ys @ drop (Suc i) ys)" apply (auto simp: ad_agr_list_def ad_equiv_list_def sp_equiv_list_def) apply (metis take_zip in_set_takeD) apply (metis drop_zip in_set_dropD) using pairwise_take_drop by fastforce lemma fo_nmlz_rem_nth_add_nth: assumes "fo_nmlz AD zs = zs" "i \ length zs" shows "fo_nmlz AD (rem_nth i (fo_nmlz AD (add_nth i z zs))) = zs" proof - have ad_agr: "ad_agr_list AD (add_nth i z zs) (fo_nmlz AD (add_nth i z zs))" using fo_nmlz_ad_agr by auto have i_lt_add: "i < length (add_nth i z zs)" "i < length (fo_nmlz AD (add_nth i z zs))" using add_nth_length assms(2) by (fastforce simp: fo_nmlz_length)+ show ?thesis using ad_agr_list_take_drop[OF ad_agr, of i, folded rem_nth_take_drop[OF i_lt_add(1)] rem_nth_take_drop[OF i_lt_add(2)], unfolded rem_nth_add_nth[OF assms(2)]] apply (subst eq_commute) apply (subst assms(1)[symmetric]) apply (auto intro: fo_nmlz_eqI) done qed lemma ad_agr_list_add: assumes "ad_agr_list AD xs ys" "i \ length xs" shows "\z' \ Inl ` AD \ Inr ` {.. set ys. ad_agr_list AD (take i xs @ z # drop i xs) (take i ys @ z' # drop i ys)" proof - define n where "n = length xs" have len_ys: "n = length ys" using assms(1) by (auto simp: ad_agr_list_def n_def) obtain \ where \_def: "xs = map \ [0.. where \_def: "ys = map \ [0.. n" using assms(2) by (auto simp: n_def) have set_n: "set [0.. \" using iffD2[OF ad_agr_list_link, OF assms(1)[unfolded \_def \_def]] unfolding set_n . have set_ys: "\ ` ({..n} - {n}) = set ys" by (auto simp: \_def) obtain z' where z'_def: "z' \ Inl ` AD \ Inr ` {.. set ys" "ad_agr_sets {..n} {..n} AD (\(n := z)) (\(n := z'))" using extend_\[OF ad_agr subset_refl, of "Inl ` AD \ Inr ` {.. set ys" z] by (auto simp: set_ys) have map_take: "map (\(n := z)) ([0..(n := z')) ([0.._def \_def take_map drop_map) show ?thesis using iffD1[OF ad_agr_list_link, OF z'_def(2)[unfolded set_n[symmetric]]] z'_def(1) unfolding map_take by auto qed lemma add_nth_restrict: assumes "fo_nmlz AD zs = zs" "i \ length zs" shows "\z' \ Inl ` AD \ Inr ` {.. Inl ` AD \ Inr ` {.. Inl ` AD \ Inr ` {.. length zs" shows "\z'. fo_nmlz AD (add_nth i z zs) = fo_nmlz AD (add_nth i z' (fo_nmlz AD zs))" proof - have ad_agr: "ad_agr_list AD zs (fo_nmlz AD zs)" using fo_nmlz_ad_agr by auto have i_le_fo_nmlz: "i \ length (fo_nmlz AD zs)" using assms(1) by (auto simp: fo_nmlz_length) obtain x where x_def: "ad_agr_list AD (add_nth i z zs) (add_nth i x (fo_nmlz AD zs))" using ad_agr_list_add[OF ad_agr assms(1)] by (auto simp: add_nth_take_drop[OF assms(1)] add_nth_take_drop[OF i_le_fo_nmlz]) then show ?thesis using fo_nmlz_eqI by auto qed lemma fo_nmlz_add_rem': assumes "i \ length zs" shows "\z'. fo_nmlz AD (add_nth i z (fo_nmlz AD zs)) = fo_nmlz AD (add_nth i z' zs)" proof - have ad_agr: "ad_agr_list AD (fo_nmlz AD zs) zs" using ad_agr_list_comm[OF fo_nmlz_ad_agr] by auto have i_le_fo_nmlz: "i \ length (fo_nmlz AD zs)" using assms(1) by (auto simp: fo_nmlz_length) obtain x where x_def: "ad_agr_list AD (add_nth i z (fo_nmlz AD zs)) (add_nth i x zs)" using ad_agr_list_add[OF ad_agr i_le_fo_nmlz] by (auto simp: add_nth_take_drop[OF assms(1)] add_nth_take_drop[OF i_le_fo_nmlz]) then show ?thesis using fo_nmlz_eqI by auto qed lemma fo_nmlz_add_nth_rem_nth: assumes "fo_nmlz AD xs = xs" "i < length xs" shows "\z. fo_nmlz AD (add_nth i z (fo_nmlz AD (rem_nth i xs))) = xs" using rem_nth_length[OF assms(2)] fo_nmlz_add_rem[of i "rem_nth i xs" AD "xs ! i", unfolded assms(1) add_nth_rem_nth_self[OF assms(2)]] assms(2) by (subst eq_commute) auto lemma sp_equiv_list_almost_same: "sp_equiv_list (xs @ v # ys) (xs @ w # ys) \ v \ set xs \ set ys \ w \ set xs \ set ys \ v = w" by (auto simp: sp_equiv_list_def pairwise_def) (metis UnCI sp_equiv_pair.simps zip_same)+ lemma ad_agr_list_add_nth: assumes "i \ length zs" "ad_agr_list AD (add_nth i v zs) (add_nth i w zs)" "v \ w" shows "{v, w} \ (Inl ` AD \ set zs) = {}" using assms(2)[unfolded add_nth_take_drop[OF assms(1)]] assms(1,3) sp_equiv_list_almost_same by (auto simp: ad_agr_list_def ad_equiv_list_def ad_equiv_pair.simps) (smt append_take_drop_id set_append sp_equiv_list_almost_same)+ lemma Inr_in_tuple: assumes "fo_nmlz AD zs = zs" "n < card (Inr -` set zs)" shows "Inr n \ set zs" using assms fo_nmlz_set_card[OF assms(1)] by (auto simp: fo_nmlzd_code[symmetric]) lemma card_wit_sub: assumes "finite Z" "card Z \ card {ts \ X. \z \ Z. ts = f z}" shows "f ` Z \ X" proof - have set_unfold: "{ts \ X. \z \ Z. ts = f z} = f ` Z \ X" by auto show ?thesis using assms unfolding set_unfold by (metis Int_lower1 card_image_le card_seteq finite_imageI inf.absorb_iff1 le_antisym surj_card_le) qed lemma add_nth_iff_card: assumes "(\xs. xs \ X \ fo_nmlz AD xs = xs)" "(\xs. xs \ X \ i < length xs)" "fo_nmlz AD zs = zs" "i \ length zs" "finite AD" "finite X" shows "(\z. fo_nmlz AD (add_nth i z zs) \ X) \ Suc (card AD + card (Inr -` set zs)) \ card {ts \ X. \z. ts = fo_nmlz AD (add_nth i z zs)}" proof - have inj: "inj_on (\z. fo_nmlz AD (add_nth i z zs)) (Inl ` AD \ Inr ` {.. Inr ` {..z. fo_nmlz AD (add_nth i z zs) \ X) \ (\z \ Inl ` AD \ Inr ` {.. X)" using add_nth_restrict[OF assms(3,4)] by metis have restrict_z': "{ts \ X. \z. ts = fo_nmlz AD (add_nth i z zs)} = {ts \ X. \z \ Inl ` AD \ Inr ` {..z. fo_nmlz AD (add_nth i z zs) \ X" then have image_sub: "(\z. fo_nmlz AD (add_nth i z zs)) ` (Inl ` AD \ Inr ` {.. {ts \ X. \z. ts = fo_nmlz AD (add_nth i z zs)}" by auto have "Suc (card AD + card (Inr -` set zs)) \ card {ts \ X. \z. ts = fo_nmlz AD (add_nth i z zs)}" unfolding card_Un[symmetric] using card_inj_on_le[OF inj image_sub] assms(6) by auto then have "Suc (card AD + card (Inr -` set zs)) \ card {ts \ X. \z. ts = fo_nmlz AD (add_nth i z zs)}" by (auto simp: card_image) } moreover { assume assm: "card (Inl ` AD \ Inr ` {.. card {ts \ X. \z \ Inl ` AD \ Inr ` {..z \ Inl ` AD \ Inr ` {.. X" using card_wit_sub[OF _ assm] assms(5) by auto } ultimately show ?thesis unfolding restrict_z[symmetric] restrict_z'[symmetric] card_Un by auto qed lemma set_fo_nmlz_add_nth_rem_nth: assumes "j < length xs" "\x. x \ X \ fo_nmlz AD x = x" "\x. x \ X \ j < length x" shows "{ts \ X. \z. ts = fo_nmlz AD (add_nth j z (fo_nmlz AD (rem_nth j xs)))} = {y \ X. fo_nmlz AD (rem_nth j y) = fo_nmlz AD (rem_nth j xs)}" using fo_nmlz_rem_nth_add_nth[where ?zs="fo_nmlz AD (rem_nth j xs)"] rem_nth_length[OF assms(1)] fo_nmlz_add_nth_rem_nth assms by (fastforce simp: fo_nmlz_idem[OF fo_nmlz_sound] fo_nmlz_length) lemma eval_forall: fixes \ :: "('a :: infinite, 'b) fo_fmla" assumes wf: "fo_wf \ I t" shows "fo_wf (Forall i \) I (eval_forall i (fv_fo_fmla_list \) t)" proof - obtain AD n X where t_def: "t = (AD, n, X)" "AD = act_edom \ I" "AD = act_edom (Forall i \) I" using assms by (cases t) auto have AD_sub: "act_edom \ I \ AD" by (auto simp: t_def(2)) have fin_AD: "finite AD" using finite_act_edom wf by (auto simp: t_def) have fin_X: "finite X" using wf by (auto simp: t_def) note X_def = fo_wf_X[OF wf[unfolded t_def], folded t_def(2)] have eval: "eval_forall i (fv_fo_fmla_list \) t = eval_abs (Forall i \) I" proof (cases "pos i (fv_fo_fmla_list \)") case None note fv_eq = fv_ex_all[OF None] have "X = fo_nmlz AD ` proj_fmla (Forall i \) {\. esat \ I \ UNIV}" unfolding X_def by (auto simp: proj_fmla_def fv_eq) also have "\ = fo_nmlz AD ` proj_fmla (Forall i \) {\. esat (Forall i \) I \ UNIV}" using esat_forall_not_fv[of i \ UNIV I] pos_complete[OF None] by (auto simp: fv_fo_fmla_list_set) finally show ?thesis by (auto simp: t_def None eval_abs_def fv_eq nfv_def) next case (Some j) have i_in_fv: "i \ fv_fo_fmla \" by (rule pos_set[OF Some, unfolded fv_fo_fmla_list_set]) have fo_nmlz_X: "\xs. xs \ X \ fo_nmlz AD xs = xs" by (auto simp: X_def proj_fmla_map fo_nmlz_idem[OF fo_nmlz_sound]) have j_lt_len: "\xs. xs \ X \ j < length xs" using pos_sound[OF Some] by (auto simp: X_def proj_fmla_map fo_nmlz_length) have rem_nth_j_le_len: "\xs. xs \ X \ j \ length (fo_nmlz AD (rem_nth j xs))" using rem_nth_length j_lt_len by (fastforce simp: fo_nmlz_length) have img_proj_fmla: "Mapping.keys (Mapping.filter (\t Z. Suc (card AD + card (Inr -` set t)) \ card Z) (cluster (Some \ (\ts. fo_nmlz AD (rem_nth j ts))) X)) = fo_nmlz AD ` proj_fmla (Forall i \) {\. esat (Forall i \) I \ UNIV}" proof (rule set_eqI, rule iffI) fix vs assume "vs \ Mapping.keys (Mapping.filter (\t Z. Suc (card AD + card (Inr -` set t)) \ card Z) (cluster (Some \ (\ts. fo_nmlz AD (rem_nth j ts))) X))" then obtain ws where ws_def: "ws \ X" "vs = fo_nmlz AD (rem_nth j ws)" "\a. fo_nmlz AD (add_nth j a (fo_nmlz AD (rem_nth j ws))) \ X" using add_nth_iff_card[OF fo_nmlz_X j_lt_len fo_nmlz_idem[OF fo_nmlz_sound] rem_nth_j_le_len fin_AD fin_X] set_fo_nmlz_add_nth_rem_nth[OF j_lt_len fo_nmlz_X j_lt_len] by transfer (fastforce split: option.splits if_splits) then obtain \ where \_def: "esat \ I \ UNIV" "ws = fo_nmlz AD (map \ (fv_fo_fmla_list \))" unfolding X_def by (auto simp: proj_fmla_map) obtain \ where \_def: "ws = map \ (fv_fo_fmla_list \)" using fo_nmlz_map \_def(2) by blast have fo_nmlzd_\: "fo_nmlzd AD (map \ (fv_fo_fmla_list \))" unfolding \_def[symmetric] \_def(2) by (rule fo_nmlz_sound) have rem_nth_j_ws: "rem_nth j ws = map \ (filter ((\) i) (fv_fo_fmla_list \))" using rem_nth_sound[OF _ Some] sorted_distinct_fv_list by (auto simp: \_def) have esat_\: "esat (Forall i \) I \ UNIV" unfolding esat.simps proof (rule ballI) fix x have "fo_nmlz AD (add_nth j x (rem_nth j ws)) \ X" using fo_nmlz_add_rem[of j "rem_nth j ws" AD x] rem_nth_length j_lt_len[OF ws_def(1)] ws_def(3) by fastforce then have "fo_nmlz AD (map (\(i := x)) (fv_fo_fmla_list \)) \ X" using add_nth_rem_nth_map[OF _ Some, of x] sorted_distinct_fv_list unfolding \_def by fastforce then show "esat \ I (\(i := x)) UNIV" by (auto simp: X_def proj_fmla_map esat_UNIV_ad_agr_list[OF _ AD_sub] dest!: fo_nmlz_eqD) qed have rem_nth_ws: "rem_nth j ws = map \ (fv_fo_fmla_list (Forall i \))" using rem_nth_sound[OF _ Some] sorted_distinct_fv_list by (auto simp: fv_fo_fmla_list_forall \_def) then show "vs \ fo_nmlz AD ` proj_fmla (Forall i \) {\. esat (Forall i \) I \ UNIV}" using ws_def(2) esat_\ by (auto simp: proj_fmla_map rem_nth_ws) next fix vs assume assm: "vs \ fo_nmlz AD ` proj_fmla (Forall i \) {\. esat (Forall i \) I \ UNIV}" from assm obtain \ where \_def: "vs = fo_nmlz AD (map \ (fv_fo_fmla_list (Forall i \)))" "esat (Forall i \) I \ UNIV" by (auto simp: proj_fmla_map) then have all_esat: "\x. esat \ I (\(i := x)) UNIV" by auto define ws where "ws \ fo_nmlz AD (map \ (fv_fo_fmla_list \))" then have "length ws = nfv \" using nfv_def fo_nmlz_length by (metis length_map) then have ws_in: "ws \ fo_nmlz AD ` proj_fmla \ {\. esat \ I \ UNIV}" using all_esat[of "\ i"] ws_def by (auto simp: fo_nmlz_sound proj_fmla_map) then have ws_in_X: "ws \ X" by (auto simp: X_def) obtain \ where \_def: "ws = map \ (fv_fo_fmla_list \)" using fo_nmlz_map ws_def by blast have rem_nth_ws: "rem_nth j ws = map \ (fv_fo_fmla_list (Forall i \))" using rem_nth_sound[of "fv_fo_fmla_list \" i j] sorted_distinct_fv_list Some unfolding fv_fo_fmla_list_forall \_def by auto have "set (fv_fo_fmla_list (Forall i \)) \ set (fv_fo_fmla_list \)" by (auto simp: fv_fo_fmla_list_forall) then have ad_agr: "ad_agr_list AD (map \ (fv_fo_fmla_list (Forall i \))) (map \ (fv_fo_fmla_list (Forall i \)))" apply (rule ad_agr_list_subset) using fo_nmlz_ad_agr[of AD] ws_def \_def by metis have map_fv_cong: "\x. map (\(i := x)) (fv_fo_fmla_list (Forall i \)) = map \ (fv_fo_fmla_list (Forall i \))" by (auto simp: fv_fo_fmla_list_forall) have vs_rem_nth: "vs = fo_nmlz AD (rem_nth j ws)" unfolding \_def(1) rem_nth_ws apply (rule fo_nmlz_eqI) using ad_agr[unfolded map_fv_cong] . have "\a. fo_nmlz AD (add_nth j a (fo_nmlz AD (rem_nth j ws))) \ fo_nmlz AD ` proj_fmla \ {\. esat \ I \ UNIV}" proof - fix a obtain x where add_rem: "fo_nmlz AD (add_nth j a (fo_nmlz AD (rem_nth j ws))) = fo_nmlz AD (map (\(i := x)) (fv_fo_fmla_list \))" using add_nth_rem_nth_map[OF _ Some, of _ \] sorted_distinct_fv_list fo_nmlz_add_rem'[of j "rem_nth j ws"] rem_nth_length[of j ws] j_lt_len[OF ws_in_X] by (fastforce simp: \_def) have "esat (Forall i \) I \ UNIV" apply (rule iffD1[OF esat_UNIV_ad_agr_list \_def(2), OF _ subset_refl, folded t_def]) using fo_nmlz_ad_agr[of AD "map \ (fv_fo_fmla_list \)", folded ws_def, unfolded \_def] unfolding ad_agr_list_link[symmetric] by (auto simp: fv_fo_fmla_list_set ad_agr_sets_def sp_equiv_def pairwise_def) then have "esat \ I (\(i := x)) UNIV" by auto then show "fo_nmlz AD (add_nth j a (fo_nmlz AD (rem_nth j ws))) \ fo_nmlz AD ` proj_fmla \ {\. esat \ I \ UNIV}" by (auto simp: add_rem proj_fmla_map) qed then show "vs \ Mapping.keys (Mapping.filter (\t Z. Suc (card AD + card (Inr -` set t)) \ card Z) (cluster (Some \ (\ts. fo_nmlz AD (rem_nth j ts))) X))" unfolding vs_rem_nth X_def[symmetric] using add_nth_iff_card[OF fo_nmlz_X j_lt_len fo_nmlz_idem[OF fo_nmlz_sound] rem_nth_j_le_len fin_AD fin_X] set_fo_nmlz_add_nth_rem_nth[OF j_lt_len fo_nmlz_X j_lt_len] ws_in_X by transfer (fastforce split: option.splits if_splits) qed show ?thesis using nfv_ex_all[OF Some] by (simp add: t_def Some eval_abs_def nfv_def img_proj_fmla[unfolded t_def(2)] split: option.splits) qed have wf_all: "wf_fo_intp (Forall i \) I" using wf by (auto simp: t_def) show ?thesis using fo_wf_eval_abs[OF wf_all] by (auto simp: eval) qed fun fo_res :: "('a, nat) fo_t \ 'a eval_res" where "fo_res (AD, n, X) = (if fo_fin (AD, n, X) then Fin (map projl ` X) else Infin)" lemma fo_res_fin: fixes t :: "('a :: infinite, nat) fo_t" assumes "fo_wf \ I t" "finite (fo_rep t)" shows "fo_res t = Fin (fo_rep t)" proof - obtain AD n X where t_def: "t = (AD, n, X)" using assms(1) by (cases t) auto show ?thesis using fo_fin assms by (fastforce simp only: t_def fo_res.simps fo_rep_fin split: if_splits) qed lemma fo_res_infin: fixes t :: "('a :: infinite, nat) fo_t" assumes "fo_wf \ I t" "\finite (fo_rep t)" shows "fo_res t = Infin" proof - obtain AD n X where t_def: "t = (AD, n, X)" using assms(1) by (cases t) auto show ?thesis using fo_fin assms by (fastforce simp only: t_def fo_res.simps split: if_splits) qed lemma fo_rep: "fo_wf \ I t \ fo_rep t = proj_sat \ I" by (cases t) auto global_interpretation Ailamazyan: eval_fo fo_wf eval_pred fo_rep fo_res eval_bool eval_eq eval_neg eval_conj eval_ajoin eval_disj eval_exists eval_forall defines eval_fmla = Ailamazyan.eval_fmla and eval = Ailamazyan.eval apply standard apply (rule fo_rep, assumption+) apply (rule fo_res_fin, assumption+) apply (rule fo_res_infin, assumption+) apply (rule eval_pred, assumption+) apply (rule eval_bool) apply (rule eval_eq) apply (rule eval_neg, assumption+) apply (rule eval_conj, assumption+) apply (rule eval_ajoin, assumption+) apply (rule eval_disj, assumption+) apply (rule eval_exists, assumption+) apply (rule eval_forall, assumption+) done definition esat_UNIV :: "('a :: infinite, 'b) fo_fmla \ ('a table, 'b) fo_intp \ ('a + nat) val \ bool" where "esat_UNIV \ I \ = esat \ I \ UNIV" lemma esat_UNIV_code[code]: "esat_UNIV \ I \ \ (if wf_fo_intp \ I then (case eval_fmla \ I of (AD, n, X) \ fo_nmlz (act_edom \ I) (map \ (fv_fo_fmla_list \)) \ X) else esat_UNIV \ I \)" proof - obtain AD n T where t_def: "Ailamazyan.eval_fmla \ I = (AD, n, T)" by (cases "Ailamazyan.eval_fmla \ I") auto { assume wf_fo_intp: "wf_fo_intp \ I" note fo_wf = Ailamazyan.eval_fmla_correct[OF wf_fo_intp, unfolded t_def] note T_def = fo_wf_X[OF fo_wf] have AD_def: "AD = act_edom \ I" using fo_wf by auto have "esat_UNIV \ I \ \ fo_nmlz (act_edom \ I) (map \ (fv_fo_fmla_list \)) \ T" using esat_UNIV_ad_agr_list[OF _ subset_refl] by (force simp add: esat_UNIV_def T_def AD_def proj_fmla_map dest!: fo_nmlz_eqD) } then show ?thesis by (auto simp: t_def) qed lemma sat_code[code]: fixes \ :: "('a :: infinite, 'b) fo_fmla" shows "sat \ I \ \ (if wf_fo_intp \ I then (case eval_fmla \ I of (AD, n, X) \ fo_nmlz (act_edom \ I) (map (Inl \ \) (fv_fo_fmla_list \)) \ X) else sat \ I \)" using esat_UNIV_code sat_esat_conv[folded esat_UNIV_def] by metis end diff --git a/thys/HRB-Slicing/JinjaVM_Inter/JVMCFG.thy b/thys/HRB-Slicing/JinjaVM_Inter/JVMCFG.thy --- a/thys/HRB-Slicing/JinjaVM_Inter/JVMCFG.thy +++ b/thys/HRB-Slicing/JinjaVM_Inter/JVMCFG.thy @@ -1,825 +1,825 @@ chapter \A Control Flow Graph for Jinja Byte Code\ section \Formalizing the CFG\ theory JVMCFG imports "../StaticInter/BasicDefs" Jinja.BVExample begin declare lesub_list_impl_same_size [simp del] declare nlistsE_length [simp del] subsection \Type definitions\ subsubsection \Wellformed Programs\ definition "wf_jvmprog = {(P, Phi). wf_jvm_prog\<^bsub>Phi\<^esub> P}" typedef wf_jvmprog = "wf_jvmprog" proof show "(E, Phi) \ wf_jvmprog" unfolding wf_jvmprog_def by (auto intro: wf_prog) qed hide_const Phi E abbreviation PROG :: "wf_jvmprog \ jvm_prog" where "PROG P \ fst(Rep_wf_jvmprog(P))" abbreviation TYPING :: "wf_jvmprog \ ty\<^sub>P" where "TYPING P \ snd(Rep_wf_jvmprog(P))" lemma wf_jvmprog_is_wf_typ: "wf_jvm_prog\<^bsub>TYPING P\<^esub> (PROG P)" using Rep_wf_jvmprog [of P] by (auto simp: wf_jvmprog_def split_beta) lemma wf_jvmprog_is_wf: "wf_jvm_prog (PROG P)" using wf_jvmprog_is_wf_typ unfolding wf_jvm_prog_def by blast subsubsection \Interprocedural CFG\ type_synonym jvm_method = "wf_jvmprog \ cname \ mname" datatype var = Heap | Local "nat" | Stack "nat" | Exception datatype val = Hp "heap" | Value "Value.val" type_synonym state = "var \ val" definition valid_state :: "state \ bool" where "valid_state s \ (\val. s Heap \ Some (Value val)) \ (s Exception = None \ (\addr. s Exception = Some (Value (Addr addr)))) \ (\var. var \ Heap \ var \ Exception \ (\h. s var \ Some (Hp h)))" fun the_Heap :: "val \ heap" where "the_Heap (Hp h) = h" fun the_Value :: "val \ Value.val" where "the_Value (Value v) = v" abbreviation heap_of :: "state \ heap" where "heap_of s \ the_Heap (the (s Heap))" abbreviation exc_flag :: "state \ addr option" where "exc_flag s \ case (s Exception) of None \ None | Some v \ Some (THE a. v = Value (Addr a))" abbreviation stkAt :: "state \ nat \ Value.val" where "stkAt s n \ the_Value (the (s (Stack n)))" abbreviation locAt :: "state \ nat \ Value.val" where "locAt s n \ the_Value (the (s (Local n)))" datatype nodeType = Enter | Normal | Return | Exceptional "pc option" "nodeType" type_synonym cfg_node = "cname \ mname \ pc option \ nodeType" type_synonym cfg_edge = "cfg_node \ (var, val, cname \ mname \ pc, cname \ mname) edge_kind \ cfg_node" definition ClassMain :: "wf_jvmprog \ cname" where "ClassMain P \ SOME Name. \ is_class (PROG P) Name" definition MethodMain :: "wf_jvmprog \ mname" where "MethodMain P \ SOME Name. \C D fs ms. class (PROG P) C = \(D, fs, ms)\ \ (\m \ set ms. Name \ fst m)" definition stkLength :: "jvm_method \ pc \ nat" where "stkLength m pc \ let (P, C, M) = m in ( if (C = ClassMain P) then 1 else ( length (fst(the(((TYPING P) C M) ! pc))) ))" definition locLength :: "jvm_method \ pc \ nat" where "locLength m pc \ let (P, C, M) = m in ( if (C = ClassMain P) then 1 else ( length (snd(the(((TYPING P) C M) ! pc))) ))" lemma ex_new_class_name: "\C. \ is_class P C" proof - have "\ finite (UNIV :: cname set)" by (rule infinite_UNIV_listI) hence "\C. C \ set (map fst P)" by -(rule ex_new_if_finite, auto) then obtain C where "C \ set (map fst P)" by blast have "\ is_class P C" proof assume "is_class P C" then obtain D fs ms where "class P C = \(D, fs, ms)\" by auto with \C \ set (map fst P)\ show False by (auto dest: map_of_SomeD intro!: image_eqI simp: class_def) qed thus ?thesis by blast qed lemma ClassMain_unique_in_P: assumes "is_class (PROG P) C" shows "ClassMain P \ C" proof - from ex_new_class_name [of "PROG P"] obtain D where "\ is_class (PROG P) D" by blast with \is_class (PROG P) C\ show ?thesis unfolding ClassMain_def by -(rule someI2, fastforce+) qed lemma map_of_fstD: "\ map_of xs a = \b\; \x \ set xs. fst x \ a \ \ False" by (induct xs, auto) lemma map_of_fstE: "\ map_of xs a = \b\; \x \ set xs. fst x = a \ thesis \ \ thesis" by (induct xs) (auto split: if_split_asm) lemma ex_unique_method_name: "\Name. \C D fs ms. class (PROG P) C = \(D, fs, ms)\ \ (\m\set ms. Name \ fst m)" proof - from wf_jvmprog_is_wf [of P] have "distinct_fst (PROG P)" by (simp add: wf_jvm_prog_def wf_jvm_prog_phi_def wf_prog_def) hence "{C. \D fs ms. class (PROG P) C = \(D, fs, ms)\} = fst ` set (PROG P)" by (fastforce elim: map_of_fstE simp: class_def intro: map_of_SomeI) hence "finite {C. \D fs ms. class (PROG P) C = \(D, fs, ms)\}" by auto moreover have "{ms. \C D fs. class (PROG P) C = \(D, fs, ms)\} = snd ` snd ` the ` (\C. class (PROG P) C) ` {C. \D fs ms. class (PROG P) C = \(D, fs, ms)\}" by (fastforce intro: rev_image_eqI map_of_SomeI simp: class_def) ultimately have "finite {ms. \C D fs. class (PROG P) C = \(D, fs, ms)\}" by auto moreover have "\ finite (UNIV :: mname set)" by (rule infinite_UNIV_listI) ultimately have "\Name. Name \ fst ` (\ms \ {ms. \C D fs. class (PROG P) C = \(D, fs, ms)\}. set ms)" by -(rule ex_new_if_finite, auto) thus ?thesis by fastforce qed lemma MethodMain_unique_in_P: assumes "PROG P \ D sees M:Ts\T = mb in C" shows "MethodMain P \ M" proof - from ex_unique_method_name [of P] obtain M' where "\C D fs ms. class (PROG P) C = \(D, fs, ms)\ \ (\m \ set ms. M' \ fst m)" by blast with \PROG P \ D sees M:Ts\T = mb in C\ show ?thesis unfolding MethodMain_def by -(rule someI2_ex, fastforce, fastforce dest!: visible_method_exists elim: map_of_fstE) qed lemma ClassMain_is_no_class [dest!]: "is_class (PROG P) (ClassMain P) \ False" proof (erule rev_notE) from ex_new_class_name [of "PROG P"] obtain C where "\ is_class (PROG P) C" by blast thus "\ is_class (PROG P) (ClassMain P)" unfolding ClassMain_def by (rule someI) qed lemma MethodMain_not_seen [dest!]: "PROG P \ C sees (MethodMain P):Ts\T = mb in D \ False" by (fastforce dest: MethodMain_unique_in_P) lemma no_Call_from_ClassMain [dest!]: "PROG P \ ClassMain P sees M:Ts\T = mb in C \ False" by (fastforce dest: sees_method_is_class) lemma no_Call_in_ClassMain [dest!]: "PROG P \ C sees M:Ts\T = mb in ClassMain P \ False" by (fastforce dest: sees_method_idemp) inductive JVMCFG :: "jvm_method \ cfg_node \ (var, val, cname \ mname \ pc, cname \ mname) edge_kind \ cfg_node \ bool" (" _ \ _ -_\ _") and reachable :: "jvm_method \ cfg_node \ bool" (" _ \ \_") where Entry_reachable: "(P, C0, Main) \ \(ClassMain P, MethodMain P, None, Enter)" | reachable_step: "\ P \ \n; P \ n -(e)\ n' \ \ P \ \n'" | Main_to_Call: "(P, C0, Main) \ \(ClassMain P, MethodMain P, \0\, Enter) \ (P, C0, Main) \ (ClassMain P, MethodMain P, \0\, Enter) -\id\ (ClassMain P, MethodMain P, \0\, Normal)" | Main_Call_LFalse: "(P, C0, Main) \ \(ClassMain P, MethodMain P, \0\, Normal) \ (P, C0, Main) \ (ClassMain P, MethodMain P, \0\, Normal) -(\s. False)\<^sub>\\ (ClassMain P, MethodMain P, \0\, Return)" | Main_Call: "\ (P, C0, Main) \ \(ClassMain P, MethodMain P, \0\, Normal); PROG P \ C0 sees Main:[]\T = (mxs, mxl\<^sub>0, is, xt) in D; initParams = [(\s. s Heap),(\s. \Value Null\)]; ek = (\(s, ret). True):(ClassMain P, MethodMain P, 0)\\<^bsub>(D, Main)\<^esub>initParams \ \ (P, C0, Main) \ (ClassMain P, MethodMain P, \0\, Normal) -(ek)\ (D, Main, None, Enter)" | Main_Return_to_Exit: "(P, C0, Main) \ \(ClassMain P, MethodMain P, \0\, Return) \ (P, C0, Main) \ (ClassMain P, MethodMain P, \0\, Return) -(\id)\ (ClassMain P, MethodMain P, None, Return)" | Method_LFalse: "(P, C0, Main) \ \(C, M, None, Enter) \ (P, C0, Main) \ (C, M, None, Enter) -(\s. False)\<^sub>\\ (C, M, None, Return)" | Method_LTrue: "(P, C0, Main) \ \(C, M, None, Enter) \ (P, C0, Main) \ (C, M, None, Enter) -(\s. True)\<^sub>\\ (C, M, \0\, Enter)" | CFG_Load: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Enter); instrs_of (PROG P) C M ! pc = Load n; ek = \(\s. s(Stack (stkLength (P, C, M) pc) := s (Local n))) \ \ (P, C0, Main) \ (C, M, \pc\, Enter) -(ek)\ (C, M, \Suc pc\, Enter)" | CFG_Store: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Enter); instrs_of (PROG P) C M ! pc = Store n; ek = \(\s. s(Local n := s (Stack (stkLength (P, C, M) pc - 1)))) \ \ (P, C0, Main) \ (C, M, \pc\, Enter) -(ek)\ (C, M, \Suc pc\, Enter)" | CFG_Push: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Enter); instrs_of (PROG P) C M ! pc = Push v; ek = \(\s. s(Stack (stkLength (P, C, M) pc) \ Value v)) \ \ (P, C0, Main) \ (C, M, \pc\, Enter) -(ek)\ (C, M, \Suc pc\, Enter)" | CFG_Pop: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Enter); instrs_of (PROG P) C M ! pc = Pop; ek = \id \ \ (P, C0, Main) \ (C, M, \pc\, Enter) -(ek)\ (C, M, \Suc pc\, Enter)" | CFG_IAdd: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Enter); instrs_of (PROG P) C M ! pc = IAdd; ek = \(\s. let i1 = the_Intg (stkAt s (stkLength (P, C, M) pc - 1)); i2 = the_Intg (stkAt s (stkLength (P, C, M) pc - 2)) in s(Stack (stkLength (P, C, M) pc - 2) \ Value (Intg (i1 + i2)))) \ \ (P, C0, Main) \ (C, M, \pc\, Enter) -(ek)\ (C, M, \Suc pc\, Enter)" | CFG_Goto: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Enter); instrs_of (PROG P) C M ! pc = Goto i \ \ (P, C0, Main) \ (C, M, \pc\, Enter) -((\s. True)\<^sub>\)\ (C, M, \nat (int pc + i)\, Enter)" | CFG_CmpEq: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Enter); instrs_of (PROG P) C M ! pc = CmpEq; ek = \(\s. let e1 = stkAt s (stkLength (P, C, M) pc - 1); e2 = stkAt s (stkLength (P, C, M) pc - 2) in s(Stack (stkLength (P, C, M) pc - 2) \ Value (Bool (e1 = e2)))) \ \ (P, C0, Main) \ (C, M, \pc\, Enter) -(ek)\ (C, M, \Suc pc\, Enter)" | CFG_IfFalse_False: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Enter); instrs_of (PROG P) C M ! pc = IfFalse i; i \ 1; ek = (\s. stkAt s (stkLength(P, C, M) pc - 1) = Bool False)\<^sub>\ \ \ (P, C0, Main) \ (C, M, \pc\, Enter) -(ek)\ (C, M, \nat (int pc + i)\, Enter)" | CFG_IfFalse_True: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Enter); instrs_of (PROG P) C M ! pc = IfFalse i; ek = (\s. stkAt s (stkLength(P, C, M) pc - 1) \ Bool False \ i = 1)\<^sub>\ \ \ (P, C0, Main) \ (C, M, \pc\, Enter) -(ek)\ (C, M, \Suc pc\, Enter)" | CFG_New_Check_Normal: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Enter); instrs_of (PROG P) C M ! pc = New Cl; ek = (\s. new_Addr (heap_of s) \ None)\<^sub>\ \ \ (P, C0, Main) \ (C, M, \pc\, Enter) -(ek)\ (C, M, \pc\, Normal)" | CFG_New_Check_Exceptional: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Enter); instrs_of (PROG P) C M ! pc = New Cl; pc' = (case (match_ex_table (PROG P) OutOfMemory pc (ex_table_of (PROG P) C M)) of None \ None | Some (pc'', d) \ \pc''\); ek = (\s. new_Addr (heap_of s) = None)\<^sub>\ \ \ (P, C0, Main) \ (C, M, \pc\, Enter) -(ek)\ (C, M, \pc\, Exceptional pc' Enter)" | CFG_New_Update: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Normal); instrs_of (PROG P) C M ! pc = New Cl; ek = \(\s. let a = the (new_Addr (heap_of s)) - in s(Heap \ Hp ((heap_of s)(a \ blank (PROG P) Cl))) - (Stack (stkLength(P, C, M) pc) \ Value (Addr a))) \ + in s(Heap \ Hp ((heap_of s)(a \ blank (PROG P) Cl)), + Stack (stkLength(P, C, M) pc) \ Value (Addr a))) \ \ (P, C0, Main) \ (C, M, \pc\, Normal) -(ek)\ (C, M, \Suc pc\, Enter)" | CFG_New_Exceptional_prop: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Exceptional None Enter); instrs_of (PROG P) C M ! pc = New Cl; ek = \(\s. s(Exception \ Value (Addr (addr_of_sys_xcpt OutOfMemory)))) \ \ (P, C0, Main) \ (C, M, \pc\, Exceptional None Enter) -(ek)\ (C, M, None, Return)" | CFG_New_Exceptional_handle: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Exceptional \pc'\ Enter); instrs_of (PROG P) C M ! pc = New Cl; - ek = \(\s. s(Exception := None) + ek = \(\s. (s(Exception := None)) (Stack (stkLength (P, C, M) pc' - 1) \ Value (Addr (addr_of_sys_xcpt OutOfMemory)))) \ \ (P, C0, Main) \ (C, M, \pc\, Exceptional \pc'\ Enter) -(ek)\ (C, M, \pc'\, Enter)" | CFG_Getfield_Check_Normal: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Enter); instrs_of (PROG P) C M ! pc = Getfield F Cl; ek = (\s. stkAt s (stkLength (P, C, M) pc - 1) \ Null)\<^sub>\ \ \ (P, C0, Main) \ (C, M, \pc\, Enter) -(ek)\ (C, M, \pc\, Normal)" | CFG_Getfield_Check_Exceptional: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Enter); instrs_of (PROG P) C M ! pc = Getfield F Cl; pc' = (case (match_ex_table (PROG P) NullPointer pc (ex_table_of (PROG P) C M)) of None \ None | Some (pc'', d) \ \pc''\); ek = (\s. stkAt s (stkLength (P, C, M) pc - 1) = Null)\<^sub>\ \ \ (P, C0, Main) \ (C, M, \pc\, Enter) -(ek)\ (C, M, \pc\, Exceptional pc' Enter)" | CFG_Getfield_Update: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Normal); instrs_of (PROG P) C M ! pc = Getfield F Cl; ek = \(\s. let (D, fs) = the (heap_of s (the_Addr (stkAt s (stkLength (P, C, M) pc - 1)))) in s(Stack (stkLength(P, C, M) pc - 1) \ Value (the (fs (F, Cl))))) \ \ (P, C0, Main) \ (C, M, \pc\, Normal) -(ek)\ (C, M, \Suc pc\, Enter)" | CFG_Getfield_Exceptional_prop: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Exceptional None Enter); instrs_of (PROG P) C M ! pc = Getfield F Cl; ek = \(\s. s(Exception \ Value (Addr (addr_of_sys_xcpt NullPointer)))) \ \ (P, C0, Main) \ (C, M, \pc\, Exceptional None Enter) -(ek)\ (C, M, None, Return)" | CFG_Getfield_Exceptional_handle: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Exceptional \pc'\ Enter); instrs_of (PROG P) C M ! pc = Getfield F Cl; - ek = \(\s. s(Exception := None) + ek = \(\s. (s(Exception := None)) (Stack (stkLength (P, C, M) pc' - 1) \ Value (Addr (addr_of_sys_xcpt NullPointer)))) \ \ (P, C0, Main) \ (C, M, \pc\, Exceptional \pc'\ Enter) -(ek)\ (C, M, \pc'\, Enter)" | CFG_Putfield_Check_Normal: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Enter); instrs_of (PROG P) C M ! pc = Putfield F Cl; ek = (\s. stkAt s (stkLength (P, C, M) pc - 2) \ Null)\<^sub>\ \ \ (P, C0, Main) \ (C, M, \pc\, Enter) -(ek)\ (C, M, \pc\, Normal)" | CFG_Putfield_Check_Exceptional: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Enter); instrs_of (PROG P) C M ! pc = Putfield F Cl; pc' = (case (match_ex_table (PROG P) NullPointer pc (ex_table_of (PROG P) C M)) of None \ None | Some (pc'', d) \ \pc''\); ek = (\s. stkAt s (stkLength (P, C, M) pc - 2) = Null)\<^sub>\ \ \ (P, C0, Main) \ (C, M, \pc\, Enter) -(ek)\ (C, M, \pc\, Exceptional pc' Enter)" | CFG_Putfield_Update: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Normal); instrs_of (PROG P) C M ! pc = Putfield F Cl; ek = \(\s. let v = stkAt s (stkLength (P, C, M) pc - 1); r = stkAt s (stkLength (P, C, M) pc - 2); a = the_Addr r; (D, fs) = the (heap_of s a); h' = (heap_of s)(a \ (D, fs((F, Cl) \ v))) in s(Heap \ Hp h')) \ \ (P, C0, Main) \ (C, M, \pc\, Normal) -(ek)\ (C, M, \Suc pc\, Enter)" | CFG_Putfield_Exceptional_prop: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Exceptional None Enter); instrs_of (PROG P) C M ! pc = Putfield F Cl; ek = \(\s. s(Exception \ Value (Addr (addr_of_sys_xcpt NullPointer)))) \ \ (P, C0, Main) \ (C, M, \pc\, Exceptional None Enter) -(ek)\ (C, M, None, Return)" | CFG_Putfield_Exceptional_handle: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Exceptional \pc'\ Enter); instrs_of (PROG P) C M ! pc = Putfield F Cl; - ek = \(\s. s(Exception := None) + ek = \(\s. (s(Exception := None)) (Stack (stkLength (P, C, M) pc' - 1) \ Value (Addr (addr_of_sys_xcpt NullPointer)))) \ \ (P, C0, Main) \ (C, M, \pc\, Exceptional \pc'\ Enter) -(ek)\ (C, M, \pc'\, Enter)" | CFG_Checkcast_Check_Normal: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Enter); instrs_of (PROG P) C M ! pc = Checkcast Cl; ek = (\s. cast_ok (PROG P) Cl (heap_of s) (stkAt s (stkLength (P, C, M) pc - 1)))\<^sub>\ \ \ (P, C0, Main) \ (C, M, \pc\, Enter) -(ek)\ (C, M, \Suc pc\, Enter)" | CFG_Checkcast_Check_Exceptional: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Enter); instrs_of (PROG P) C M ! pc = Checkcast Cl; pc' = (case (match_ex_table (PROG P) ClassCast pc (ex_table_of (PROG P) C M)) of None \ None | Some (pc'', d) \ \pc''\); ek = (\s. \ cast_ok (PROG P) Cl (heap_of s) (stkAt s (stkLength (P, C, M) pc - 1)))\<^sub>\ \ \ (P, C0, Main) \ (C, M, \pc\, Enter) -(ek)\ (C, M, \pc\, Exceptional pc' Enter)" | CFG_Checkcast_Exceptional_prop: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Exceptional None Enter); instrs_of (PROG P) C M ! pc = Checkcast Cl; ek = \(\s. s(Exception \ Value (Addr (addr_of_sys_xcpt ClassCast)))) \ \ (P, C0, Main) \ (C, M, \pc\, Exceptional None Enter) -(ek)\ (C, M, None, Return)" | CFG_Checkcast_Exceptional_handle: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Exceptional \pc'\ Enter); instrs_of (PROG P) C M ! pc = Checkcast Cl; - ek = \(\s. s(Exception := None) + ek = \(\s. (s(Exception := None)) (Stack (stkLength (P, C, M) pc' - 1) \ Value (Addr (addr_of_sys_xcpt ClassCast)))) \ \ (P, C0, Main) \ (C, M, \pc\, Exceptional \pc'\ Enter) -(ek)\ (C, M, \pc'\, Enter)" | CFG_Throw_Check: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Enter); instrs_of (PROG P) C M ! pc = Throw; pc' = None \ match_ex_table (PROG P) Exc pc (ex_table_of (PROG P) C M) = \(the pc', d)\; ek = (\s. let v = stkAt s (stkLength (P, C, M) pc - 1); Cl = if (v = Null) then NullPointer else (cname_of (heap_of s) (the_Addr v)) in case pc' of None \ match_ex_table (PROG P) Cl pc (ex_table_of (PROG P) C M) = None | Some pc'' \ \d. match_ex_table (PROG P) Cl pc (ex_table_of (PROG P) C M) = \(pc'', d)\ )\<^sub>\ \ \ (P, C0, Main) \ (C, M, \pc\, Enter) -(ek)\ (C, M, \pc\, Exceptional pc' Enter)" | CFG_Throw_prop: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Exceptional None Enter); instrs_of (PROG P) C M ! pc = Throw; ek = \(\s. s(Exception \ Value (stkAt s (stkLength (P, C, M) pc - 1)))) \ \ (P, C0, Main) \ (C, M, \pc\, Exceptional None Enter) -(ek)\ (C, M, None, Return)" | CFG_Throw_handle: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Exceptional \pc'\ Enter); pc' \ length (instrs_of (PROG P) C M); instrs_of (PROG P) C M ! pc = Throw; - ek = \(\s. s(Exception := None) + ek = \(\s. (s(Exception := None)) (Stack (stkLength (P, C, M) pc' - 1) \ Value (stkAt s (stkLength (P, C, M) pc - 1)))) \ \ (P, C0, Main) \ (C, M, \pc\, Exceptional \pc'\ Enter) -(ek)\ (C, M, \pc'\, Enter)" | CFG_Invoke_Check_NP_Normal: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Enter); instrs_of (PROG P) C M ! pc = Invoke M' n; ek = (\s. stkAt s (stkLength (P, C, M) pc - Suc n) \ Null)\<^sub>\ \ \ (P, C0, Main) \ (C, M, \pc\, Enter) -(ek)\ (C, M, \pc\, Normal)" | CFG_Invoke_Check_NP_Exceptional: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Enter); instrs_of (PROG P) C M ! pc = Invoke M' n; pc' = (case (match_ex_table (PROG P) NullPointer pc (ex_table_of (PROG P) C M)) of None \ None | Some (pc'', d) \ \pc''\); ek = (\s. stkAt s (stkLength (P, C, M) pc - Suc n) = Null)\<^sub>\ \ \ (P, C0, Main) \ (C, M, \pc\, Enter) -(ek)\ (C, M, \pc\, Exceptional pc' Enter)" | CFG_Invoke_NP_prop: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Exceptional None Enter); instrs_of (PROG P) C M ! pc = Invoke M' n; ek = \(\s. s(Exception \ Value (Addr (addr_of_sys_xcpt NullPointer)))) \ \ (P, C0, Main) \ (C, M, \pc\, Exceptional None Enter) -(ek)\ (C, M, None, Return)" | CFG_Invoke_NP_handle: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Exceptional \pc'\ Enter); instrs_of (PROG P) C M ! pc = Invoke M' n; - ek = \(\s. s(Exception := None) + ek = \(\s. (s(Exception := None)) (Stack (stkLength (P, C, M) pc' - 1) \ Value (Addr (addr_of_sys_xcpt NullPointer)))) \ \ (P, C0, Main) \ (C, M, \pc\, Exceptional \pc'\ Enter) -(ek)\ (C, M, \pc'\, Enter)" | CFG_Invoke_Call: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Normal); instrs_of (PROG P) C M ! pc = Invoke M' n; TYPING P C M ! pc = \(ST, LT)\; ST ! n = Class D'; PROG P \ D' sees M':Ts\T = (mxs, mxl\<^sub>0, is, xt) in D; Q = (\(s, ret). let r = stkAt s (stkLength (P, C, M) pc - Suc n); C' = fst (the (heap_of s (the_Addr r))) in D = fst (method (PROG P) C' M')); paramDefs = (\s. s Heap) # (\s. s (Stack (stkLength (P, C, M) pc - Suc n))) # (rev (map (\i. (\s. s (Stack (stkLength (P, C, M) pc - Suc i)))) [0..\<^bsub>(D,M')\<^esub>paramDefs \ \ (P, C0, Main) \ (C, M, \pc\, Normal) -(ek)\ (D, M', None, Enter)" | CFG_Invoke_False: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Normal); instrs_of (PROG P) C M ! pc = Invoke M' n; ek = (\s. False)\<^sub>\ \ \ (P, C0, Main) \ (C, M, \pc\, Normal) -(ek)\ (C, M, \pc\, Return)" | CFG_Invoke_Return_Check_Normal: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Return); instrs_of (PROG P) C M ! pc = Invoke M' n; (TYPING P) C M ! pc = \(ST, LT)\; ST ! n \ NT; ek = (\s. s Exception = None)\<^sub>\ \ \ (P, C0, Main) \ (C, M, \pc\, Return) -(ek)\ (C, M, \Suc pc\, Enter)" | CFG_Invoke_Return_Check_Exceptional: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Return); instrs_of (PROG P) C M ! pc = Invoke M' n; match_ex_table (PROG P) Exc pc (ex_table_of (PROG P) C M) = \(pc', diff)\; pc' \ length (instrs_of (PROG P) C M); ek = (\s. \v d. s Exception = \v\ \ match_ex_table (PROG P) (cname_of (heap_of s) (the_Addr (the_Value v))) pc (ex_table_of (PROG P) C M) = \(pc', d)\)\<^sub>\ \ \ (P, C0, Main) \ (C, M, \pc\, Return) -(ek)\ (C, M, \pc\, Exceptional \pc'\ Return)" | CFG_Invoke_Return_Exceptional_handle: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Exceptional \pc'\ Return); instrs_of (PROG P) C M ! pc = Invoke M' n; ek = \(\s. s(Exception := None, Stack (stkLength (P, C, M) pc' - 1) := s Exception)) \ \ (P, C0, Main) \ (C, M, \pc\, Exceptional \pc'\ Return) -(ek)\ (C, M, \pc'\, Enter)" | CFG_Invoke_Return_Exceptional_prop: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Return); instrs_of (PROG P) C M ! pc = Invoke M' n; ek = (\s. \v. s Exception = \v\ \ match_ex_table (PROG P) (cname_of (heap_of s) (the_Addr (the_Value v))) pc (ex_table_of (PROG P) C M) = None)\<^sub>\ \ \ (P, C0, Main) \ (C, M, \pc\, Return) -(ek)\ (C, M, None, Return)" | CFG_Return: "\ C \ ClassMain P; (P, C0, Main) \ \(C, M, \pc\, Enter); instrs_of (PROG P) C M ! pc = instr.Return; ek = \(\s. s(Stack 0 := s (Stack (stkLength (P, C, M) pc - 1)))) \ \ (P, C0, Main) \ (C, M, \pc\, Enter) -(ek)\ (C, M, None, Return)" | CFG_Return_from_Method: "\ (P, C0, Main) \ \(C, M, None, Return); (P, C0, Main) \ (C', M', \pc'\, Normal) -(Q':(C', M', pc')\\<^bsub>(C,M)\<^esub>ps)\ (C, M, None, Enter); Q = (\(s, ret). ret = (C', M', pc')); stateUpdate = (\s s'. s'(Heap := s Heap, Exception := s Exception, Stack (stkLength (P, C', M') (Suc pc') - 1) := s (Stack 0)) ); ek = Q\\<^bsub>(C, M)\<^esub>stateUpdate \ \ (P, C0, Main) \ (C, M, None, Return) -(ek)\ (C', M', \pc'\, Return)" (* This takes veeeery long *) lemma JVMCFG_edge_det: "\ P \ n -(et)\ n'; P \ n -(et')\ n' \ \ et = et'" by (erule JVMCFG.cases) (erule JVMCFG.cases, (fastforce dest: sees_method_fun)+)+ lemma sourcenode_reachable: "P \ n -(ek)\ n' \ P \ \n" by (erule JVMCFG.cases, auto) lemma targetnode_reachable: assumes edge: "P \ n -(ek)\ n'" shows "P \ \n'" proof - from edge have "P \ \n" by -(drule sourcenode_reachable) with edge show ?thesis by -(rule JVMCFG_reachable.intros) qed lemmas JVMCFG_reachable_inducts = JVMCFG_reachable.inducts[split_format (complete)] lemma ClassMain_imp_MethodMain: "(P, C0, Main) \ (C', M', pc', nt') -ek\ (ClassMain P, M, pc, nt) \ M = MethodMain P" "(P, C0, Main) \ \(ClassMain P, M, pc, nt) \ M = MethodMain P" proof (induct P=="P" C0\"C0" Main\Main C' M' pc' nt' ek C''=="ClassMain P" M pc nt and P=="P" C0\"C0" Main\Main C'=="ClassMain P" M pc nt rule: JVMCFG_reachable_inducts) case CFG_Return_from_Method thus ?case by (fastforce elim: JVMCFG.cases) qed auto lemma ClassMain_no_Call_target [dest!]: "(P, C0, Main) \ (C, M, pc, nt) -Q:(C', M', pc')\\<^bsub>(D,M'')\<^esub>paramDefs\ (ClassMain P, M''', pc'', nt') \ False" and "(P, C0, Main) \ \(C, M, pc, nt) \ True" by (induct P C0 Main C M pc nt ek=="Q:(C', M', pc')\\<^bsub>(D,M'')\<^esub>paramDefs" C''=="ClassMain P" M''' pc'' nt' and P C0 Main C M pc nt rule: JVMCFG_reachable_inducts) auto lemma method_of_src_and_trg_exists: "\ (P, C0, Main) \ (C', M', pc', nt') -ek\ (C, M, pc, nt); C \ ClassMain P; C' \ ClassMain P \ \ (\Ts T mb. (PROG P) \ C sees M:Ts\T = mb in C) \ (\Ts T mb. (PROG P) \ C' sees M':Ts\T = mb in C')" and method_of_reachable_node_exists: "\ (P, C0, Main) \ \(C, M, pc, nt); C \ ClassMain P \ \ \Ts T mb. (PROG P) \ C sees M:Ts\T = mb in C" proof (induct rule: JVMCFG_reachable_inducts) case CFG_Invoke_Call thus ?case by (blast dest: sees_method_idemp) next case (reachable_step P C0 Main C M pc nt ek C' M' pc' nt') show ?case proof (cases "C = ClassMain P") case True with \(P, C0, Main) \ (C, M, pc, nt) -ek\ (C', M', pc', nt')\ \C' \ ClassMain P\ show ?thesis proof cases case Main_Call thus ?thesis by (blast dest: sees_method_idemp) qed auto next case False with reachable_step show ?thesis by simp qed qed simp_all lemma "\ (P, C0, Main) \ (C', M', pc', nt') -ek\ (C, M, pc, nt); C \ ClassMain P; C' \ ClassMain P \ \ (case pc of None \ True | \pc''\ \ (TYPING P) C M ! pc'' \ None \ pc'' < length (instrs_of (PROG P) C M)) \ (case pc' of None \ True | \pc''\ \ (TYPING P) C' M' ! pc'' \ None \ pc'' < length (instrs_of (PROG P) C' M'))" and instr_of_reachable_node_typable: "\ (P, C0, Main) \ \(C, M, pc, nt); C \ ClassMain P \ \ case pc of None \ True | \pc''\ \ (TYPING P) C M ! pc'' \ None \ pc'' < length (instrs_of (PROG P) C M)" proof (induct rule: JVMCFG_reachable_inducts) case (CFG_Load C P C0 Main M pc n ek) from \(P, C0, Main) \ \(C, M, \pc\, Enter)\ \C \ ClassMain P\ obtain Ts T mxs mxl\<^sub>0 "is" xt where "PROG P \ C sees M:Ts\T = (mxs, mxl\<^sub>0, is, xt) in C" and "instrs_of (PROG P) C M = is" by -(drule method_of_reachable_node_exists, auto) with CFG_Load show ?case by (fastforce dest!: wt_jvm_prog_impl_wt_instr [OF wf_jvmprog_is_wf_typ]) next case (CFG_Store C P C0 Main M pc n ek) from \(P, C0, Main) \ \(C, M, \pc\, Enter)\ \C \ ClassMain P\ obtain Ts T mxs mxl\<^sub>0 "is" xt where "PROG P \ C sees M:Ts\T = (mxs, mxl\<^sub>0, is, xt) in C" and "instrs_of (PROG P) C M = is" by -(drule method_of_reachable_node_exists, auto) with CFG_Store show ?case by (fastforce dest!: wt_jvm_prog_impl_wt_instr [OF wf_jvmprog_is_wf_typ]) next case (CFG_Push C P C0 Main M pc v ek) from \(P, C0, Main) \ \(C, M, \pc\, Enter)\ \C \ ClassMain P\ obtain Ts T mxs mxl\<^sub>0 "is" xt where "PROG P \ C sees M:Ts\T = (mxs, mxl\<^sub>0, is, xt) in C" and "instrs_of (PROG P) C M = is" by -(drule method_of_reachable_node_exists, auto) with CFG_Push show ?case by (fastforce dest!: wt_jvm_prog_impl_wt_instr [OF wf_jvmprog_is_wf_typ]) next case (CFG_Pop C P C0 Main M pc ek) from \(P, C0, Main) \ \(C, M, \pc\, Enter)\ \C \ ClassMain P\ obtain Ts T mxs mxl\<^sub>0 "is" xt where "PROG P \ C sees M:Ts\T = (mxs, mxl\<^sub>0, is, xt) in C" and "instrs_of (PROG P) C M = is" by -(drule method_of_reachable_node_exists, auto) with CFG_Pop show ?case by (fastforce dest!: wt_jvm_prog_impl_wt_instr [OF wf_jvmprog_is_wf_typ]) next case (CFG_IAdd C P C0 Main M pc ek) from \(P, C0, Main) \ \(C, M, \pc\, Enter)\ \C \ ClassMain P\ obtain Ts T mxs mxl\<^sub>0 "is" xt where "PROG P \ C sees M:Ts\T = (mxs, mxl\<^sub>0, is, xt) in C" and "instrs_of (PROG P) C M = is" by -(drule method_of_reachable_node_exists, auto) with CFG_IAdd show ?case by (fastforce dest!: wt_jvm_prog_impl_wt_instr [OF wf_jvmprog_is_wf_typ]) next case (CFG_Goto C P C0 Main M pc i) from \(P, C0, Main) \ \(C, M, \pc\, Enter)\ \C \ ClassMain P\ obtain Ts T mxs mxl\<^sub>0 "is" xt where "PROG P \ C sees M:Ts\T = (mxs, mxl\<^sub>0, is, xt) in C" and "instrs_of (PROG P) C M = is" by -(drule method_of_reachable_node_exists, auto) with CFG_Goto show ?case by (fastforce dest!: wt_jvm_prog_impl_wt_instr [OF wf_jvmprog_is_wf_typ]) next case (CFG_CmpEq C P C0 Main M pc ek) from \(P, C0, Main) \ \(C, M, \pc\, Enter)\ \C \ ClassMain P\ obtain Ts T mxs mxl\<^sub>0 "is" xt where "PROG P \ C sees M:Ts\T = (mxs, mxl\<^sub>0, is, xt) in C" and "instrs_of (PROG P) C M = is" by -(drule method_of_reachable_node_exists, auto) with CFG_CmpEq show ?case by (fastforce dest!: wt_jvm_prog_impl_wt_instr [OF wf_jvmprog_is_wf_typ]) next case (CFG_IfFalse_False C P C0 Main M pc i ek) from \(P, C0, Main) \ \(C, M, \pc\, Enter)\ \C \ ClassMain P\ obtain Ts T mxs mxl\<^sub>0 "is" xt where "PROG P \ C sees M:Ts\T = (mxs, mxl\<^sub>0, is, xt) in C" and "instrs_of (PROG P) C M = is" by -(drule method_of_reachable_node_exists, auto) with CFG_IfFalse_False show ?case by (fastforce dest!: wt_jvm_prog_impl_wt_instr [OF wf_jvmprog_is_wf_typ]) next case (CFG_IfFalse_True C P C0 Main M pc i ek) from \(P, C0, Main) \ \(C, M, \pc\, Enter)\ \C \ ClassMain P\ obtain Ts T mxs mxl\<^sub>0 "is" xt where "PROG P \ C sees M:Ts\T = (mxs, mxl\<^sub>0, is, xt) in C" and "instrs_of (PROG P) C M = is" by -(drule method_of_reachable_node_exists, auto) with CFG_IfFalse_True show ?case using [[simproc del: list_to_set_comprehension]] by (fastforce dest!: wt_jvm_prog_impl_wt_instr [OF wf_jvmprog_is_wf_typ]) next case (CFG_New_Update C P C0 Main M pc Cl ek) from \(P, C0, Main) \ \(C, M, \pc\, Normal)\ \C \ ClassMain P\ obtain Ts T mxs mxl\<^sub>0 "is" xt where "PROG P \ C sees M:Ts\T = (mxs, mxl\<^sub>0, is, xt) in C" and "instrs_of (PROG P) C M = is" by -(drule method_of_reachable_node_exists, auto) with CFG_New_Update show ?case by (fastforce dest!: wt_jvm_prog_impl_wt_instr [OF wf_jvmprog_is_wf_typ]) next case (CFG_New_Exceptional_handle C P C0 Main M pc pc' Cl ek) hence "TYPING P C M ! pc \ None" and "pc < length (instrs_of (PROG P) C M)" by simp_all moreover from \(P, C0, Main) \ \(C, M, \pc\, Exceptional \pc'\ Enter)\ \C \ ClassMain P\ obtain Ts T mxs mxl\<^sub>0 where "PROG P \ C sees M:Ts\T = (mxs, mxl\<^sub>0, instrs_of (PROG P) C M, ex_table_of (PROG P) C M) in C" by (fastforce dest: method_of_reachable_node_exists) with \pc < length (instrs_of (PROG P) C M)\ \instrs_of (PROG P) C M ! pc = New Cl\ have "PROG P,T,mxs,length (instrs_of (PROG P) C M),ex_table_of (PROG P) C M \ New Cl,pc :: TYPING P C M" by (fastforce dest!: wt_jvm_prog_impl_wt_instr [OF wf_jvmprog_is_wf_typ]) moreover from \(P, C0, Main) \ \(C, M, \pc\, Exceptional \pc'\ Enter)\ \C \ ClassMain P\ \instrs_of (PROG P) C M ! pc = New Cl\ obtain d' where "match_ex_table (PROG P) OutOfMemory pc (ex_table_of (PROG P) C M) = \(pc', d')\" by cases (fastforce elim: JVMCFG.cases) hence "\(f, t, D, h, d)\set (ex_table_of (PROG P) C M). matches_ex_entry (PROG P) OutOfMemory pc (f, t, D, h, d) \ h = pc' \ d = d'" by -(drule match_ex_table_SomeD) ultimately show ?case using \instrs_of (PROG P) C M ! pc = New Cl\ by (fastforce simp: relevant_entries_def is_relevant_entry_def matches_ex_entry_def) next case (CFG_Getfield_Update C P C0 Main M pc F Cl ek) from \(P, C0, Main) \ \(C, M, \pc\, Normal)\ \C \ ClassMain P\ obtain Ts T mxs mxl\<^sub>0 "is" xt where "PROG P \ C sees M:Ts\T = (mxs, mxl\<^sub>0, is, xt) in C" and "instrs_of (PROG P) C M = is" by -(drule method_of_reachable_node_exists, auto) with CFG_Getfield_Update show ?case by (fastforce dest!: wt_jvm_prog_impl_wt_instr [OF wf_jvmprog_is_wf_typ]) next case (CFG_Getfield_Exceptional_handle C P C0 Main M pc pc' F Cl ek) hence "TYPING P C M ! pc \ None" and "pc < length (instrs_of (PROG P) C M)" by simp_all moreover from \(P, C0, Main) \ \(C, M, \pc\, Exceptional \pc'\ Enter)\ \C \ ClassMain P\ obtain Ts T mxs mxl\<^sub>0 where "PROG P \ C sees M:Ts\T = (mxs, mxl\<^sub>0, instrs_of (PROG P) C M, ex_table_of (PROG P) C M) in C" by (fastforce dest: method_of_reachable_node_exists) with \pc < length (instrs_of (PROG P) C M)\ \instrs_of (PROG P) C M ! pc = Getfield F Cl\ have "PROG P,T,mxs,length (instrs_of (PROG P) C M),ex_table_of (PROG P) C M \ Getfield F Cl,pc :: TYPING P C M" by (fastforce dest!: wt_jvm_prog_impl_wt_instr [OF wf_jvmprog_is_wf_typ]) moreover from \(P, C0, Main) \ \(C, M, \pc\, Exceptional \pc'\ Enter)\ \C \ ClassMain P\ \instrs_of (PROG P) C M ! pc = Getfield F Cl\ obtain d' where "match_ex_table (PROG P) NullPointer pc (ex_table_of (PROG P) C M) = \(pc', d')\" by cases (fastforce elim: JVMCFG.cases) hence "\(f, t, D, h, d)\set (ex_table_of (PROG P) C M). matches_ex_entry (PROG P) NullPointer pc (f, t, D, h, d) \ h = pc' \ d = d'" by -(drule match_ex_table_SomeD) ultimately show ?case using \instrs_of (PROG P) C M ! pc = Getfield F Cl\ by (fastforce simp: relevant_entries_def is_relevant_entry_def matches_ex_entry_def) next case (CFG_Putfield_Update C P C0 Main M pc F Cl ek) from \(P, C0, Main) \ \(C, M, \pc\, Normal)\ \C \ ClassMain P\ obtain Ts T mxs mxl\<^sub>0 "is" xt where "PROG P \ C sees M:Ts\T = (mxs, mxl\<^sub>0, is, xt) in C" and "instrs_of (PROG P) C M = is" by -(drule method_of_reachable_node_exists, auto) with CFG_Putfield_Update show ?case by (fastforce dest!: wt_jvm_prog_impl_wt_instr [OF wf_jvmprog_is_wf_typ]) next case (CFG_Putfield_Exceptional_handle C P C0 Main M pc pc' F Cl ek) hence "TYPING P C M ! pc \ None" and "pc < length (instrs_of (PROG P) C M)" by simp_all moreover from \(P, C0, Main) \ \(C, M, \pc\, Exceptional \pc'\ Enter)\ \C \ ClassMain P\ obtain Ts T mxs mxl\<^sub>0 where "PROG P \ C sees M:Ts\T = (mxs, mxl\<^sub>0, instrs_of (PROG P) C M, ex_table_of (PROG P) C M) in C" by (fastforce dest: method_of_reachable_node_exists) with \pc < length (instrs_of (PROG P) C M)\ \instrs_of (PROG P) C M ! pc = Putfield F Cl\ have "PROG P,T,mxs,length (instrs_of (PROG P) C M),ex_table_of (PROG P) C M \ Putfield F Cl,pc :: TYPING P C M" by (fastforce dest!: wt_jvm_prog_impl_wt_instr [OF wf_jvmprog_is_wf_typ]) moreover from \(P, C0, Main) \ \(C, M, \pc\, Exceptional \pc'\ Enter)\ \C \ ClassMain P\ \instrs_of (PROG P) C M ! pc = Putfield F Cl\ obtain d' where "match_ex_table (PROG P) NullPointer pc (ex_table_of (PROG P) C M) = \(pc', d')\" by cases (fastforce elim: JVMCFG.cases) hence "\(f, t, D, h, d)\set (ex_table_of (PROG P) C M). matches_ex_entry (PROG P) NullPointer pc (f, t, D, h, d) \ h = pc' \ d = d'" by -(drule match_ex_table_SomeD) ultimately show ?case using \instrs_of (PROG P) C M ! pc = Putfield F Cl\ by (fastforce simp: relevant_entries_def is_relevant_entry_def matches_ex_entry_def) next case (CFG_Checkcast_Check_Normal C P C0 Main M pc Cl ek) from \(P, C0, Main) \ \(C, M, \pc\, Enter)\ \C \ ClassMain P\ obtain Ts T mxs mxl\<^sub>0 "is" xt where "PROG P \ C sees M:Ts\T = (mxs, mxl\<^sub>0, is, xt) in C" and "instrs_of (PROG P) C M = is" by -(drule method_of_reachable_node_exists, auto) with CFG_Checkcast_Check_Normal show ?case by (fastforce dest!: wt_jvm_prog_impl_wt_instr [OF wf_jvmprog_is_wf_typ]) next case (CFG_Checkcast_Exceptional_handle C P C0 Main M pc pc' Cl ek) hence "TYPING P C M ! pc \ None" and "pc < length (instrs_of (PROG P) C M)" by simp_all moreover from \(P, C0, Main) \ \(C, M, \pc\, Exceptional \pc'\ Enter)\ \C \ ClassMain P\ obtain Ts T mxs mxl\<^sub>0 where "PROG P \ C sees M:Ts\T = (mxs, mxl\<^sub>0, instrs_of (PROG P) C M, ex_table_of (PROG P) C M) in C" by (fastforce dest: method_of_reachable_node_exists) with \pc < length (instrs_of (PROG P) C M)\ \instrs_of (PROG P) C M ! pc = Checkcast Cl\ have "PROG P,T,mxs,length (instrs_of (PROG P) C M),ex_table_of (PROG P) C M \ Checkcast Cl,pc :: TYPING P C M" by (fastforce dest!: wt_jvm_prog_impl_wt_instr [OF wf_jvmprog_is_wf_typ]) moreover from \(P, C0, Main) \ \(C, M, \pc\, Exceptional \pc'\ Enter)\ \C \ ClassMain P\ \instrs_of (PROG P) C M ! pc = Checkcast Cl\ obtain d' where "match_ex_table (PROG P) ClassCast pc (ex_table_of (PROG P) C M) = \(pc', d')\" by cases (fastforce elim: JVMCFG.cases) hence "\(f, t, D, h, d)\set (ex_table_of (PROG P) C M). matches_ex_entry (PROG P) ClassCast pc (f, t, D, h, d) \ h = pc' \ d = d'" by -(drule match_ex_table_SomeD) ultimately show ?case using \instrs_of (PROG P) C M ! pc = Checkcast Cl\ by (fastforce simp: relevant_entries_def is_relevant_entry_def matches_ex_entry_def) next case (CFG_Throw_handle C P C0 Main M pc pc' ek) hence "TYPING P C M ! pc \ None" and "pc < length (instrs_of (PROG P) C M)" by simp_all moreover from \(P, C0, Main) \ \(C, M, \pc\, Exceptional \pc'\ Enter)\ \C \ ClassMain P\ obtain Ts T mxs mxl\<^sub>0 where "PROG P \ C sees M:Ts\T = (mxs, mxl\<^sub>0, instrs_of (PROG P) C M, ex_table_of (PROG P) C M) in C" by (fastforce dest: method_of_reachable_node_exists) with \pc < length (instrs_of (PROG P) C M)\ \instrs_of (PROG P) C M ! pc = Throw\ have "PROG P,T,mxs,length (instrs_of (PROG P) C M),ex_table_of (PROG P) C M \ Throw,pc :: TYPING P C M" by (fastforce dest!: wt_jvm_prog_impl_wt_instr [OF wf_jvmprog_is_wf_typ]) moreover from \(P, C0, Main) \ \(C, M, \pc\, Exceptional \pc'\ Enter)\ \C \ ClassMain P\ \instrs_of (PROG P) C M ! pc = Throw\ obtain d' Exc where "match_ex_table (PROG P) Exc pc (ex_table_of (PROG P) C M) = \(pc', d')\" by cases (fastforce elim: JVMCFG.cases) hence "\(f, t, D, h, d)\set (ex_table_of (PROG P) C M). matches_ex_entry (PROG P) Exc pc (f, t, D, h, d) \ h = pc' \ d = d'" by -(drule match_ex_table_SomeD) ultimately show ?case using \instrs_of (PROG P) C M ! pc = Throw\ by (fastforce simp: relevant_entries_def is_relevant_entry_def matches_ex_entry_def) next case (CFG_Invoke_NP_handle C P C0 Main M pc pc' M' n ek) hence "TYPING P C M ! pc \ None" and "pc < length (instrs_of (PROG P) C M)" by simp_all moreover from \(P, C0, Main) \ \(C, M, \pc\, Exceptional \pc'\ Enter)\ \C \ ClassMain P\ obtain Ts T mxs mxl\<^sub>0 where "PROG P \ C sees M:Ts\T = (mxs, mxl\<^sub>0, instrs_of (PROG P) C M, ex_table_of (PROG P) C M) in C" by (fastforce dest: method_of_reachable_node_exists) with \pc < length (instrs_of (PROG P) C M)\ \instrs_of (PROG P) C M ! pc = Invoke M' n\ have "PROG P,T,mxs,length (instrs_of (PROG P) C M),ex_table_of (PROG P) C M \ Invoke M' n,pc :: TYPING P C M" by (fastforce dest!: wt_jvm_prog_impl_wt_instr [OF wf_jvmprog_is_wf_typ]) moreover from \(P, C0, Main) \ \(C, M, \pc\, Exceptional \pc'\ Enter)\ \C \ ClassMain P\ \instrs_of (PROG P) C M ! pc = Invoke M' n\ obtain d' where "match_ex_table (PROG P) NullPointer pc (ex_table_of (PROG P) C M) = \(pc', d')\" by cases (fastforce elim: JVMCFG.cases) hence "\(f, t, D, h, d)\set (ex_table_of (PROG P) C M). matches_ex_entry (PROG P) NullPointer pc (f, t, D, h, d) \ h = pc' \ d = d'" by -(drule match_ex_table_SomeD) ultimately show ?case using \instrs_of (PROG P) C M ! pc = Invoke M' n\ by (fastforce simp: relevant_entries_def is_relevant_entry_def matches_ex_entry_def) next case (CFG_Invoke_Return_Exceptional_handle C P C0 Main M pc pc' M' n ek) hence "TYPING P C M ! pc \ None" and "pc < length (instrs_of (PROG P) C M)" by simp_all moreover from \(P, C0, Main) \ \(C, M, \pc\, Exceptional \pc'\ Return)\ \C \ ClassMain P\ obtain Ts T mxs mxl\<^sub>0 where "PROG P \ C sees M:Ts\T = (mxs, mxl\<^sub>0, instrs_of (PROG P) C M, ex_table_of (PROG P) C M) in C" by (fastforce dest: method_of_reachable_node_exists) with \pc < length (instrs_of (PROG P) C M)\ \instrs_of (PROG P) C M ! pc = Invoke M' n\ have "PROG P,T,mxs,length (instrs_of (PROG P) C M),ex_table_of (PROG P) C M \ Invoke M' n,pc :: TYPING P C M" by (fastforce dest!: wt_jvm_prog_impl_wt_instr [OF wf_jvmprog_is_wf_typ]) moreover from \(P, C0, Main) \ \(C, M, \pc\, Exceptional \pc'\ Return)\ \C \ ClassMain P\ \instrs_of (PROG P) C M ! pc = Invoke M' n\ obtain d' Exc where "match_ex_table (PROG P) Exc pc (ex_table_of (PROG P) C M) = \(pc', d')\" by cases (fastforce elim: JVMCFG.cases) hence "\(f, t, D, h, d)\set (ex_table_of (PROG P) C M). matches_ex_entry (PROG P) Exc pc (f, t, D, h, d) \ h = pc' \ d = d'" by -(drule match_ex_table_SomeD) ultimately show ?case using \instrs_of (PROG P) C M ! pc = Invoke M' n\ by (fastforce simp: relevant_entries_def is_relevant_entry_def matches_ex_entry_def) next case (CFG_Invoke_Return_Check_Normal C P C0 Main M pc M' n ST LT ek) from \(P, C0, Main) \ \(C, M, \pc\, Return)\ \C \ ClassMain P\ obtain Ts T mxs mxl\<^sub>0 "is" xt where "PROG P \ C sees M:Ts\T = (mxs, mxl\<^sub>0, is, xt) in C" and "instrs_of (PROG P) C M = is" by -(drule method_of_reachable_node_exists, auto) with CFG_Invoke_Return_Check_Normal show ?case by (fastforce dest!: wt_jvm_prog_impl_wt_instr [OF wf_jvmprog_is_wf_typ]) next case (Method_LTrue P C0 Main C M) from \(P, C0, Main) \ \(C, M, None, Enter)\ \C \ ClassMain P\ obtain Ts T mxs mxl\<^sub>0 "is" xt where "PROG P \ C sees M:Ts\T = (mxs, mxl\<^sub>0, is, xt) in C" and "instrs_of (PROG P) C M = is" by -(drule method_of_reachable_node_exists, auto) with Method_LTrue show ?case by (fastforce dest!: wt_jvm_prog_impl_wt_start [OF wf_jvmprog_is_wf_typ] simp: wt_start_def) next case (reachable_step P C0 Main C M opc nt ek C' M' opc' nt') thus ?case by (cases "C = ClassMain P") (fastforce elim: JVMCFG.cases, simp) qed simp_all lemma reachable_node_impl_wt_instr: assumes "(P, C0, Main) \ \(C, M, \pc\, nt)" and "C \ ClassMain P" shows "\T mxs mpc xt. PROG P,T,mxs,mpc,xt \ (instrs_of (PROG P) C M ! pc),pc :: TYPING P C M" proof - from \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, nt)\ method_of_reachable_node_exists [of P C0 Main C M "\pc\" nt] instr_of_reachable_node_typable [of P C0 Main C M "\pc\" nt] obtain Ts T mxs mxl\<^sub>0 "is" xt where "PROG P \ C sees M:Ts\T = (mxs, mxl\<^sub>0, is, xt) in C" and "TYPING P C M ! pc \ None" and "pc < length (instrs_of (PROG P) C M)" by fastforce+ with wf_jvmprog_is_wf_typ [of P] have "PROG P,T,mxs,length is,xt \ instrs_of (PROG P) C M ! pc,pc :: TYPING P C M" by (fastforce dest!: wt_jvm_prog_impl_wt_instr) thus ?thesis by blast qed lemma "\ (P, C0, Main) \ (C, M, pc, nt) -ek\ (C', M', pc', nt'); C \ ClassMain P \ C' \ ClassMain P \ \ \T mb D. PROG P \ C0 sees Main:[]\T = mb in D" and reachable_node_impl_Main_ex: "\ (P, C0, Main) \ \(C, M, pc, nt); C \ ClassMain P\ \ \T mb D. PROG P \ C0 sees Main:[]\T = mb in D" by (induct rule: JVMCFG_reachable_inducts) fastforce+ end diff --git a/thys/HRB-Slicing/JinjaVM_Inter/JVMPostdomination.thy b/thys/HRB-Slicing/JinjaVM_Inter/JVMPostdomination.thy --- a/thys/HRB-Slicing/JinjaVM_Inter/JVMPostdomination.thy +++ b/thys/HRB-Slicing/JinjaVM_Inter/JVMPostdomination.thy @@ -1,1244 +1,1244 @@ theory JVMPostdomination imports JVMInterpretation "../StaticInter/Postdomination" begin context CFG begin lemma vp_snocI: "\n -as\\<^sub>\* n'; n' -[a]\* n''; \Q p ret fs. kind a \ Q\\<^bsub>p\<^esub>ret \ \ n -as @ [a]\\<^sub>\* n''" by (cases "kind a") (auto intro: path_Append valid_path_aux_Append simp: vp_def valid_path_def) lemma valid_node_cases' [case_names Source Target, consumes 1]: "\ valid_node n; \e. \ valid_edge e; sourcenode e = n \ \ thesis; \e. \ valid_edge e; targetnode e = n \ \ thesis \ \ thesis" by (auto simp: valid_node_def) end lemma disjE_strong: "\P \ Q; P \ R; \Q; \ P\ \ R\ \ R" by auto lemmas path_intros [intro] = JVMCFG_Interpret.path.Cons_path JVMCFG_Interpret.path.empty_path declare JVMCFG_Interpret.vp_snocI [intro] declare JVMCFG_Interpret.valid_node_def [simp add] valid_edge_def [simp add] JVMCFG_Interpret.intra_path_def [simp add] abbreviation vp_snoc :: "wf_jvmprog \ cname \ mname \ cfg_edge list \ cfg_node \ (var, val, cname \ mname \ pc, cname \ mname) edge_kind \ cfg_node \ bool" where "vp_snoc P C0 Main as n ek n' \ JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) (as @ [(n,ek,n')]) n'" lemma "(P, C0, Main) \ (C, M, pc, nt) -ek\ (C', M', pc', nt') \ (\as. CFG.valid_path' sourcenode targetnode kind (valid_edge (P, C0, Main)) (get_return_edges P) (ClassMain P, MethodMain P, None, Enter) as (C, M, pc, nt)) \ (\as. CFG.valid_path' sourcenode targetnode kind (valid_edge (P, C0, Main)) (get_return_edges P) (ClassMain P, MethodMain P, None, Enter) as (C', M', pc', nt'))" and valid_Entry_path: "(P, C0, Main) \ \(C, M, pc, nt) \ \as. CFG.valid_path' sourcenode targetnode kind (valid_edge (P, C0, Main)) (get_return_edges P) (ClassMain P, MethodMain P, None, Enter) as (C, M, pc, nt)" proof (induct rule: JVMCFG_reachable_inducts) case (Entry_reachable P C0 Main) hence "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) [] (ClassMain P, MethodMain P, None, Enter)" by (fastforce intro: JVMCFG_Interpret.intra_path_vp Method_LTrue JVMCFG_reachable.Entry_reachable) thus ?case by blast next case (reachable_step P C0 Main C M pc nt ek C' M' pc' nt') thus ?case by simp next case (Main_to_Call P C0 Main) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (ClassMain P, MethodMain P, \0\, Enter)" by blast moreover with \(P, C0, Main) \ \(ClassMain P, MethodMain P, \0\, Enter)\ have "vp_snoc P C0 Main as (ClassMain P, MethodMain P, \0\, Enter) \id (ClassMain P, MethodMain P, \0\, Normal)" by (fastforce intro: JVMCFG_reachable.Main_to_Call) ultimately show ?case by blast next case (Main_Call_LFalse P C0 Main) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (ClassMain P, MethodMain P, \0\, Normal)" by blast moreover with \(P, C0, Main) \ \(ClassMain P, MethodMain P, \0\, Normal)\ have "vp_snoc P C0 Main as (ClassMain P, MethodMain P, \0\, Normal) (\s. False)\<^sub>\ (ClassMain P, MethodMain P, \0\, Return)" by (fastforce intro: JVMCFG_reachable.Main_Call_LFalse) ultimately show ?case by blast next case (Main_Call P C0 Main T mxs mxl\<^sub>0 "is" xt D initParams ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (ClassMain P, MethodMain P, \0\, Normal)" by blast moreover with \(P, C0, Main) \ \(ClassMain P, MethodMain P, \0\, Normal)\ \PROG P \ C0 sees Main: []\T = (mxs, mxl\<^sub>0, is, xt) in D\ \initParams = [\s. s Heap, \s. \Value Null\]\ \ek = \(s, ret). True:(ClassMain P, MethodMain P, 0)\\<^bsub>(D, Main)\<^esub>initParams\ have "vp_snoc P C0 Main as (ClassMain P, MethodMain P, \0\, Normal) ((\(s, ret). True):(ClassMain P, MethodMain P, 0)\\<^bsub>(D, Main)\<^esub>[(\s. s Heap),(\s. \Value Null\)]) (D, Main, None, Enter)" by (fastforce intro: JVMCFG_reachable.Main_Call) ultimately show ?case by blast next case (Main_Return_to_Exit P C0 Main) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (ClassMain P, MethodMain P, \0\, nodeType.Return)" by blast moreover with \(P, C0, Main) \ \(ClassMain P, MethodMain P, \0\, nodeType.Return)\ have "vp_snoc P C0 Main as (ClassMain P, MethodMain P, \0\, nodeType.Return) \id (ClassMain P, MethodMain P, None, nodeType.Return)" by (fastforce intro: JVMCFG_reachable.Main_Return_to_Exit) ultimately show ?case by blast next case (Method_LFalse P C0 Main C M) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, None, Enter)" by blast moreover with \(P, C0, Main) \ \(C, M, None, Enter)\ have "vp_snoc P C0 Main as (C, M, None, Enter) (\s. False)\<^sub>\ (C, M, None, Return)" by (fastforce intro: JVMCFG_reachable.Method_LFalse) ultimately show ?case by blast next case (Method_LTrue P C0 Main C M) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, None, Enter)" by blast moreover with \(P, C0, Main) \ \(C, M, None, Enter)\ have "vp_snoc P C0 Main as (C, M, None, Enter) (\s. True)\<^sub>\ (C, M, \0\, Enter)" by (fastforce intro: JVMCFG_reachable.Method_LTrue) ultimately show ?case by blast next case (CFG_Load C P C0 Main M pc n ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, Enter)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, Enter)\ \instrs_of (PROG P) C M ! pc = Load n\ \ek = \\s. s(Stack (stkLength (P, C, M) pc) := s (Local n))\ have "vp_snoc P C0 Main as (C, M, \pc\, Enter) ek (C, M, \Suc pc\, Enter)" by (fastforce intro: JVMCFG_reachable.CFG_Load) ultimately show ?case by blast next case (CFG_Store C P C0 Main M pc n ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, Enter)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, Enter)\ \instrs_of (PROG P) C M ! pc = Store n\ \ek = \\s. s(Local n := s (Stack (stkLength (P, C, M) pc - 1)))\ have "vp_snoc P C0 Main as (C, M, \pc\, Enter) ek (C, M, \Suc pc\, Enter)" by (fastforce intro: JVMCFG_reachable.CFG_Store) ultimately show ?case by blast next case (CFG_Push C P C0 Main M pc v ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, Enter)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, Enter)\ \instrs_of (PROG P) C M ! pc = Push v\ \ek = \\s. s(Stack (stkLength (P, C, M) pc) \ Value v)\ have "vp_snoc P C0 Main as (C, M, \pc\, Enter) ek (C, M, \Suc pc\, Enter)" by (fastforce intro: JVMCFG_reachable.CFG_Push) ultimately show ?case by blast next case (CFG_Pop C P C0 Main M pc ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, Enter)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, Enter)\ \instrs_of (PROG P) C M ! pc = Pop\ \ek = \id\ have "vp_snoc P C0 Main as (C, M, \pc\, Enter) ek (C, M, \Suc pc\, Enter)" by (fastforce intro: JVMCFG_reachable.CFG_Pop) ultimately show ?case by blast next case (CFG_IAdd C P C0 Main M pc ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, Enter)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, Enter)\ \instrs_of (PROG P) C M ! pc = IAdd\ \ek = \\s. let i1 = the_Intg (stkAt s (stkLength (P, C, M) pc - 1)); i2 = the_Intg (stkAt s (stkLength (P, C, M) pc - 2)) in s(Stack (stkLength (P, C, M) pc - 2) \ Value (Intg (i1 + i2)))\ have "vp_snoc P C0 Main as (C, M, \pc\, Enter) ek (C, M, \Suc pc\, Enter)" by (fastforce intro: JVMCFG_reachable.CFG_IAdd) ultimately show ?case by blast next case (CFG_Goto C P C0 Main M pc i) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, Enter)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, Enter)\ \instrs_of (PROG P) C M ! pc = Goto i\ have "vp_snoc P C0 Main as (C, M, \pc\, Enter) (\s. True)\<^sub>\ (C, M, \nat (int pc + i)\, Enter)" by (fastforce intro: JVMCFG_reachable.CFG_Goto) ultimately show ?case by blast next case (CFG_CmpEq C P C0 Main M pc ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, Enter)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, Enter)\ \instrs_of (PROG P) C M ! pc = CmpEq\ \ek = \\s. let e1 = stkAt s (stkLength (P, C, M) pc - 1); e2 = stkAt s (stkLength (P, C, M) pc - 2) in s(Stack (stkLength (P, C, M) pc - 2) \ Value (Bool (e1 = e2)))\ have "vp_snoc P C0 Main as (C, M, \pc\, Enter) ek (C, M, \Suc pc\, Enter)" by (fastforce intro: JVMCFG_reachable.CFG_CmpEq) ultimately show ?case by blast next case (CFG_IfFalse_False C P C0 Main M pc i ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, Enter)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, Enter)\ \instrs_of (PROG P) C M ! pc = IfFalse i\ \i \ 1\ \ek = (\s. stkAt s (stkLength (P, C, M) pc - 1) = Bool False)\<^sub>\\ have "vp_snoc P C0 Main as (C, M, \pc\, Enter) ek (C, M, \nat (int pc + i)\, Enter)" by (fastforce intro: JVMCFG_reachable.CFG_IfFalse_False) ultimately show ?case by blast next case (CFG_IfFalse_True C P C0 Main M pc i ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, Enter)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, Enter)\ \instrs_of (PROG P) C M ! pc = IfFalse i\ \ek = (\s. stkAt s (stkLength (P, C, M) pc - 1) \ Bool False \ i = 1)\<^sub>\\ have "vp_snoc P C0 Main as (C, M, \pc\, Enter) ek (C, M, \Suc pc\, Enter)" by (fastforce intro: JVMCFG_reachable.CFG_IfFalse_True) ultimately show ?case by blast next case (CFG_New_Check_Normal C P C0 Main M pc Cl ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, Enter)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, Enter)\ \instrs_of (PROG P) C M ! pc = New Cl\ \ek = (\s. new_Addr (heap_of s) \ None)\<^sub>\\ have "vp_snoc P C0 Main as (C, M, \pc\, Enter) ek (C, M, \pc\, Normal)" by (fastforce intro: JVMCFG_reachable.CFG_New_Check_Normal) ultimately show ?case by blast next case (CFG_New_Check_Exceptional C P C0 Main M pc Cl pc' ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, Enter)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, Enter)\ \ instrs_of (PROG P) C M ! pc = New Cl\ \pc' = (case match_ex_table (PROG P) OutOfMemory pc (ex_table_of (PROG P) C M) of None \ None | \(pc'', d)\ \ \pc''\)\ \ek = (\s. new_Addr (heap_of s) = None)\<^sub>\\ have "vp_snoc P C0 Main as (C, M, \pc\, Enter) ek (C, M, \pc\, Exceptional pc' Enter)" by (fastforce intro: JVMCFG_reachable.CFG_New_Check_Exceptional) ultimately show ?case by blast next case (CFG_New_Update C P C0 Main M pc Cl ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, Normal)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, Normal)\ \ instrs_of (PROG P) C M ! pc = New Cl\ \ ek = \\s. let a = the (new_Addr (heap_of s)) in - s(Heap \ Hp (heap_of s(a \ blank (PROG P) Cl)), + s(Heap \ Hp ((heap_of s)(a \ blank (PROG P) Cl)), Stack (stkLength (P, C, M) pc) \ Value (Addr a))\ have "vp_snoc P C0 Main as (C, M, \pc\, Normal) ek (C, M, \Suc pc\, Enter)" by (fastforce intro: JVMCFG_reachable.CFG_New_Update) ultimately show ?case by blast next case (CFG_New_Exceptional_prop C P C0 Main M pc Cl ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, Exceptional None Enter)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, Exceptional None Enter)\ \instrs_of (PROG P) C M ! pc = New Cl\ \ek = \\s. s(Exception \ Value (Addr (addr_of_sys_xcpt OutOfMemory)))\ have "vp_snoc P C0 Main as (C, M, \pc\, Exceptional None Enter) ek (C, M, None, Return)" by (fastforce intro: JVMCFG_reachable.CFG_New_Exceptional_prop) ultimately show ?case by blast next case (CFG_New_Exceptional_handle C P C0 Main M pc pc' Cl ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, Exceptional \pc'\ Enter)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, Exceptional \pc'\ Enter)\ \instrs_of (PROG P) C M ! pc = New Cl\ - \ek = \\s. s(Exception := None)(Stack (stkLength (P, C, M) pc' - 1) \ + \ek = \\s. (s(Exception := None))(Stack (stkLength (P, C, M) pc' - 1) \ Value (Addr (addr_of_sys_xcpt OutOfMemory)))\ have "vp_snoc P C0 Main as (C, M, \pc\, Exceptional \pc'\ Enter) ek (C, M, \pc'\, Enter)" by (fastforce intro: JVMCFG_reachable.CFG_New_Exceptional_handle) ultimately show ?case by blast next case (CFG_Getfield_Check_Normal C P C0 Main M pc F Cl ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, Enter)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, Enter)\ \instrs_of (PROG P) C M ! pc = Getfield F Cl\ \ek = (\s. stkAt s (stkLength (P, C, M) pc - 1) \ Null)\<^sub>\\ have "vp_snoc P C0 Main as (C, M, \pc\, Enter) ek (C, M, \pc\, Normal)" by (fastforce intro: JVMCFG_reachable.CFG_Getfield_Check_Normal) ultimately show ?case by blast next case (CFG_Getfield_Check_Exceptional C P C0 Main M pc F Cl pc' ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, Enter)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, Enter)\ \instrs_of (PROG P) C M ! pc = Getfield F Cl\ \pc' = (case match_ex_table (PROG P) NullPointer pc (ex_table_of (PROG P) C M) of None \ None | \(pc'', d)\ \ \pc''\)\ \ek = (\s. stkAt s (stkLength (P, C, M) pc - 1) = Null)\<^sub>\\ have "vp_snoc P C0 Main as (C, M, \pc\, Enter) ek (C, M, \pc\, Exceptional pc' Enter)" by (fastforce intro: JVMCFG_reachable.CFG_Getfield_Check_Exceptional) ultimately show ?case by blast next case (CFG_Getfield_Update C P C0 Main M pc F Cl ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, Normal)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, Normal)\ \instrs_of (PROG P) C M ! pc = Getfield F Cl\ \ek = \\s. let (D, fs) = the (heap_of s (the_Addr (stkAt s (stkLength (P, C, M) pc - 1)))) in s(Stack (stkLength (P, C, M) pc - 1) \ Value (the (fs (F, Cl))))\ have "vp_snoc P C0 Main as (C, M, \pc\, Normal) ek (C, M, \Suc pc\, Enter)" by (fastforce intro: JVMCFG_reachable.CFG_Getfield_Update) ultimately show ?case by blast next case (CFG_Getfield_Exceptional_prop C P C0 Main M pc F Cl ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, Exceptional None Enter)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, Exceptional None Enter)\ \instrs_of (PROG P) C M ! pc = Getfield F Cl\ \ek = \\s. s(Exception \ Value (Addr (addr_of_sys_xcpt NullPointer)))\ have "vp_snoc P C0 Main as (C, M, \pc\, Exceptional None Enter) ek (C, M, None, Return)" by (fastforce intro: JVMCFG_reachable.CFG_Getfield_Exceptional_prop) ultimately show ?case by blast next case (CFG_Getfield_Exceptional_handle C P C0 Main M pc pc' F Cl ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, Exceptional \pc'\ Enter)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, Exceptional \pc'\ Enter)\ \instrs_of (PROG P) C M ! pc = Getfield F Cl\ - \ek = \\s. s(Exception := None)(Stack (stkLength (P, C, M) pc' - 1) \ + \ek = \\s. (s(Exception := None))(Stack (stkLength (P, C, M) pc' - 1) \ Value (Addr (addr_of_sys_xcpt NullPointer)))\ have "vp_snoc P C0 Main as (C, M, \pc\, Exceptional \pc'\ Enter) ek (C, M, \pc'\, Enter)" by (fastforce intro: JVMCFG_reachable.CFG_Getfield_Exceptional_handle) ultimately show ?case by blast next case (CFG_Putfield_Check_Normal C P C0 Main M pc F Cl ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, Enter)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, Enter)\ \instrs_of (PROG P) C M ! pc = Putfield F Cl\ \ek = (\s. stkAt s (stkLength (P, C, M) pc - 2) \ Null)\<^sub>\\ have "vp_snoc P C0 Main as (C, M, \pc\, Enter) ek (C, M, \pc\, Normal)" by (fastforce intro: JVMCFG_reachable.CFG_Putfield_Check_Normal) ultimately show ?case by blast next case (CFG_Putfield_Check_Exceptional C P C0 Main M pc F Cl pc' ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, Enter)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, Enter)\ \instrs_of (PROG P) C M ! pc = Putfield F Cl\ \pc' = (case match_ex_table (PROG P) NullPointer pc (ex_table_of (PROG P) C M) of None \ None | \(pc'', d)\ \ \pc''\)\ \ek = (\s. stkAt s (stkLength (P, C, M) pc - 2) = Null)\<^sub>\\ have "vp_snoc P C0 Main as (C, M, \pc\, Enter) ek (C, M, \pc\, Exceptional pc' Enter)" by (fastforce intro: JVMCFG_reachable.CFG_Putfield_Check_Exceptional) ultimately show ?case by blast next case (CFG_Putfield_Update C P C0 Main M pc F Cl ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, Normal)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, Normal)\ \instrs_of (PROG P) C M ! pc = Putfield F Cl\ \ek = \\s. let v = stkAt s (stkLength (P, C, M) pc - 1); r = stkAt s (stkLength (P, C, M) pc - 2); - a = the_Addr r; (D, fs) = the (heap_of s a); h' = heap_of s(a \ (D, fs((F, Cl) \ v))) + a = the_Addr r; (D, fs) = the (heap_of s a); h' = (heap_of s)(a \ (D, fs((F, Cl) \ v))) in s(Heap \ Hp h')\ have "vp_snoc P C0 Main as (C, M, \pc\, Normal) ek (C, M, \Suc pc\, Enter)" by (fastforce intro: JVMCFG_reachable.CFG_Putfield_Update) ultimately show ?case by blast next case (CFG_Putfield_Exceptional_prop C P C0 Main M pc F Cl ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, Exceptional None Enter)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, Exceptional None Enter)\ \instrs_of (PROG P) C M ! pc = Putfield F Cl\ \ek = \\s. s(Exception \ Value (Addr (addr_of_sys_xcpt NullPointer)))\ have "vp_snoc P C0 Main as (C, M, \pc\, Exceptional None Enter) ek (C, M, None, Return)" by (fastforce intro: JVMCFG_reachable.CFG_Putfield_Exceptional_prop) ultimately show ?case by blast next case (CFG_Putfield_Exceptional_handle C P C0 Main M pc pc' F Cl ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, Exceptional \pc'\ Enter)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, Exceptional \pc'\ Enter)\ \instrs_of (PROG P) C M ! pc = Putfield F Cl\ - \ek = \\s. s(Exception := None)(Stack (stkLength (P, C, M) pc' - 1) \ + \ek = \\s. (s(Exception := None))(Stack (stkLength (P, C, M) pc' - 1) \ Value (Addr (addr_of_sys_xcpt NullPointer)))\ have "vp_snoc P C0 Main as (C, M, \pc\, Exceptional \pc'\ Enter) ek (C, M, \pc'\, Enter)" by (fastforce intro: JVMCFG_reachable.CFG_Putfield_Exceptional_handle) ultimately show ?case by blast next case (CFG_Checkcast_Check_Normal C P C0 Main M pc Cl ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, Enter)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, Enter)\ \instrs_of (PROG P) C M ! pc = Checkcast Cl\ \ek = (\s. cast_ok (PROG P) Cl (heap_of s) (stkAt s (stkLength (P, C, M) pc - 1)))\<^sub>\\ have "vp_snoc P C0 Main as (C, M, \pc\, Enter) ek (C, M, \Suc pc\, Enter)" by (fastforce intro: JVMCFG_reachable.CFG_Checkcast_Check_Normal) ultimately show ?case by blast next case (CFG_Checkcast_Check_Exceptional C P C0 Main M pc Cl pc' ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, Enter)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, Enter)\ \instrs_of (PROG P) C M ! pc = Checkcast Cl\ \pc' = (case match_ex_table (PROG P) ClassCast pc (ex_table_of (PROG P) C M) of None \ None | \(pc'', d)\ \ \pc''\)\ \ek = (\s. \ cast_ok (PROG P) Cl (heap_of s) (stkAt s (stkLength (P, C, M) pc - 1)))\<^sub>\\ have "vp_snoc P C0 Main as (C, M, \pc\, Enter) ek (C, M, \pc\, Exceptional pc' Enter)" by (fastforce intro: JVMCFG_reachable.CFG_Checkcast_Check_Exceptional) ultimately show ?case by blast next case (CFG_Checkcast_Exceptional_prop C P C0 Main M pc Cl ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, Exceptional None Enter)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, Exceptional None Enter)\ \instrs_of (PROG P) C M ! pc = Checkcast Cl\ \ek = \\s. s(Exception \ Value (Addr (addr_of_sys_xcpt ClassCast)))\ have "vp_snoc P C0 Main as (C, M, \pc\, Exceptional None Enter) ek (C, M, None, Return)" by (fastforce intro: JVMCFG_reachable.CFG_Checkcast_Exceptional_prop) ultimately show ?case by blast next case (CFG_Checkcast_Exceptional_handle C P C0 Main M pc pc' Cl ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, Exceptional \pc'\ Enter)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, Exceptional \pc'\ Enter)\ \instrs_of (PROG P) C M ! pc = Checkcast Cl\ - \ek = \\s. s(Exception := None)(Stack (stkLength (P, C, M) pc' - 1) \ + \ek = \\s. (s(Exception := None))(Stack (stkLength (P, C, M) pc' - 1) \ Value (Addr (addr_of_sys_xcpt ClassCast)))\ have "vp_snoc P C0 Main as (C, M, \pc\, Exceptional \pc'\ Enter) ek (C, M, \pc'\, Enter)" by (fastforce intro: JVMCFG_reachable.CFG_Checkcast_Exceptional_handle) ultimately show ?case by blast next case (CFG_Throw_Check C P C0 Main M pc pc' Exc d ek) then obtain as where path_src: "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, Enter)" by blast from \pc' = None \ match_ex_table (PROG P) Exc pc (ex_table_of (PROG P) C M) = \(the pc', d)\\ show ?case proof (elim disjE_strong) assume "pc' = None" with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, Enter)\ \instrs_of (PROG P) C M ! pc = Throw\ \ek = (\s. let v = stkAt s (stkLength (P, C, M) pc - 1); Cl = if v = Null then NullPointer else cname_of (heap_of s) (the_Addr v) in case pc' of None \ match_ex_table (PROG P) Cl pc (ex_table_of (PROG P) C M) = None | \pc''\ \ \d. match_ex_table (PROG P) Cl pc (ex_table_of (PROG P) C M) = \(pc'', d)\)\<^sub>\\ have "(P, C0, Main) \ (C, M, \pc\, Enter) - (\s. (stkAt s (stkLength (P, C, M) pc - Suc 0) = Null \ match_ex_table (PROG P) NullPointer pc (ex_table_of (PROG P) C M) = None) \ (stkAt s (stkLength (P, C, M) pc - Suc 0) \ Null \ match_ex_table (PROG P) (cname_of (heap_of s) (the_Addr (stkAt s (stkLength (P, C, M) pc - Suc 0)))) pc (ex_table_of (PROG P) C M) = None))\<^sub>\\ (C, M, \pc\, Exceptional None Enter)" by -(erule JVMCFG_reachable.CFG_Throw_Check, simp_all) with path_src \pc' = None\ \ek = (\s. let v = stkAt s (stkLength (P, C, M) pc - 1); Cl = if v = Null then NullPointer else cname_of (heap_of s) (the_Addr v) in case pc' of None \ match_ex_table (PROG P) Cl pc (ex_table_of (PROG P) C M) = None | \pc''\ \ \d. match_ex_table (PROG P) Cl pc (ex_table_of (PROG P) C M) = \(pc'', d)\)\<^sub>\\ have "vp_snoc P C0 Main as (C, M, \pc\, Enter) ek (C, M, \pc\, Exceptional None Enter)" by (fastforce intro: JVMCFG_reachable.CFG_Throw_Check) with path_src \pc' = None\ show ?thesis by blast next assume met: "match_ex_table (PROG P) Exc pc (ex_table_of (PROG P) C M) = \(the pc', d)\" and pc': "pc' \ None" with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, Enter)\ \instrs_of (PROG P) C M ! pc = Throw\ \ek = (\s. let v = stkAt s (stkLength (P, C, M) pc - 1); Cl = if v = Null then NullPointer else cname_of (heap_of s) (the_Addr v) in case pc' of None \ match_ex_table (PROG P) Cl pc (ex_table_of (PROG P) C M) = None | \pc''\ \ \d. match_ex_table (PROG P) Cl pc (ex_table_of (PROG P) C M) = \(pc'', d)\)\<^sub>\\ have "(P, C0, Main) \ (C, M, \pc\, Enter) - (\s. (stkAt s (stkLength (P, C, M) pc - Suc 0) = Null \ (\d. match_ex_table (PROG P) NullPointer pc (ex_table_of (PROG P) C M) = \(the pc', d)\)) \ (stkAt s (stkLength (P, C, M) pc - Suc 0) \ Null \ (\d. match_ex_table (PROG P) (cname_of (heap_of s) (the_Addr (stkAt s (stkLength (P, C, M) pc - Suc 0)))) pc (ex_table_of (PROG P) C M) = \(the pc', d)\)))\<^sub>\\ (C, M, \pc\, Exceptional \the pc'\ Enter)" by -(rule JVMCFG_reachable.CFG_Throw_Check, simp_all) with met pc' path_src \ek = (\s. let v = stkAt s (stkLength (P, C, M) pc - 1); Cl = if v = Null then NullPointer else cname_of (heap_of s) (the_Addr v) in case pc' of None \ match_ex_table (PROG P) Cl pc (ex_table_of (PROG P) C M) = None | \pc''\ \ \d. match_ex_table (PROG P) Cl pc (ex_table_of (PROG P) C M) = \(pc'', d)\)\<^sub>\\ have "vp_snoc P C0 Main as (C, M, \pc\, Enter) ek (C, M, \pc\, Exceptional pc' Enter)" by (fastforce intro: JVMCFG_reachable.CFG_Throw_Check) with path_src show ?thesis by blast qed next case (CFG_Throw_prop C P C0 Main M pc ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, Exceptional None Enter)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, Exceptional None Enter)\ \instrs_of (PROG P) C M ! pc = Throw\ \ek = \\s. s(Exception \ Value (stkAt s (stkLength (P, C, M) pc - 1)))\ have "vp_snoc P C0 Main as (C, M, \pc\, Exceptional None Enter) ek (C, M, None, nodeType.Return)" by (fastforce intro: JVMCFG_reachable.CFG_Throw_prop) ultimately show ?case by blast next case (CFG_Throw_handle C P C0 Main M pc pc' ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, Exceptional \pc'\ Enter)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, Exceptional \pc'\ Enter)\ \pc' \ length (instrs_of (PROG P) C M)\ \instrs_of (PROG P) C M ! pc = Throw\ - \ek = \\s. s(Exception := None)(Stack (stkLength (P, C, M) pc' - 1) \ + \ek = \\s. (s(Exception := None))(Stack (stkLength (P, C, M) pc' - 1) \ Value (stkAt s (stkLength (P, C, M) pc - 1)))\ have "vp_snoc P C0 Main as (C, M, \pc\, Exceptional \pc'\ Enter) ek (C, M, \pc'\, Enter)" by (fastforce intro: JVMCFG_reachable.CFG_Throw_handle) ultimately show ?case by blast next case (CFG_Invoke_Check_NP_Normal C P C0 Main M pc M' n ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, Enter)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, Enter)\ \instrs_of (PROG P) C M ! pc = Invoke M' n\ \ek = (\s. stkAt s (stkLength (P, C, M) pc - Suc n) \ Null)\<^sub>\\ have "vp_snoc P C0 Main as (C, M, \pc\, Enter) ek (C, M, \pc\, Normal)" by (fastforce intro: JVMCFG_reachable.CFG_Invoke_Check_NP_Normal) ultimately show ?case by blast next case (CFG_Invoke_Check_NP_Exceptional C P C0 Main M pc M' n pc' ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, Enter)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, Enter)\ \instrs_of (PROG P) C M ! pc = Invoke M' n\ \pc' = (case match_ex_table (PROG P) NullPointer pc (ex_table_of (PROG P) C M) of None \ None | \(pc'', d)\ \ \pc''\)\ \ek = (\s. stkAt s (stkLength (P, C, M) pc - Suc n) = Null)\<^sub>\\ have "vp_snoc P C0 Main as (C, M, \pc\, Enter) ek (C, M, \pc\, Exceptional pc' Enter)" by (fastforce intro: JVMCFG_reachable.CFG_Invoke_Check_NP_Exceptional) ultimately show ?case by blast next case (CFG_Invoke_NP_prop C P C0 Main M pc M' n ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, Exceptional None Enter)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, Exceptional None Enter)\ \instrs_of (PROG P) C M ! pc = Invoke M' n\ \ek = \\s. s(Exception \ Value (Addr (addr_of_sys_xcpt NullPointer)))\ have "vp_snoc P C0 Main as (C, M, \pc\, Exceptional None Enter) ek (C, M, None, Return)" by (fastforce intro: JVMCFG_reachable.CFG_Invoke_NP_prop) ultimately show ?case by blast next case (CFG_Invoke_NP_handle C P C0 Main M pc pc' M' n ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, Exceptional \pc'\ Enter)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, Exceptional \pc'\ Enter)\ \instrs_of (PROG P) C M ! pc = Invoke M' n\ - \ek = \\s. s(Exception := None)(Stack (stkLength (P, C, M) pc' - 1) \ + \ek = \\s. (s(Exception := None))(Stack (stkLength (P, C, M) pc' - 1) \ Value (Addr (addr_of_sys_xcpt NullPointer)))\ have "vp_snoc P C0 Main as (C, M, \pc\, Exceptional \pc'\ Enter) ek (C, M, \pc'\, Enter)" by (fastforce intro: JVMCFG_reachable.CFG_Invoke_NP_handle) ultimately show ?case by blast next case (CFG_Invoke_Call C P C0 Main M pc M' n ST LT D' Ts T mxs mxl\<^sub>0 "is" xt D Q paramDefs ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, Normal)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, Normal)\ \instrs_of (PROG P) C M ! pc = Invoke M' n\ \TYPING P C M ! pc = \(ST, LT)\\ \ST ! n = Class D'\ \PROG P \ D' sees M': Ts\T = (mxs, mxl\<^sub>0, is, xt) in D\ \Q = (\(s, ret). let r = stkAt s (stkLength (P, C, M) pc - Suc n); C' = cname_of (heap_of s) (the_Addr r) in D = fst (method (PROG P) C' M'))\ \paramDefs = (\s. s Heap) # (\s. s (Stack (stkLength (P, C, M) pc - Suc n))) # rev (map (\i s. s (Stack (stkLength (P, C, M) pc - Suc i))) [0.. \ek = Q:(C, M, pc)\\<^bsub>(D, M')\<^esub>paramDefs\ have "vp_snoc P C0 Main as (C, M, \pc\, Normal) ek (D, M', None, Enter)" by (fastforce intro: JVMCFG_reachable.CFG_Invoke_Call) ultimately show ?case by blast next case (CFG_Invoke_False C P C0 Main M pc M' n ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, Normal)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, Normal)\ \instrs_of (PROG P) C M ! pc = Invoke M' n\ \ek = (\s. False)\<^sub>\\ have "vp_snoc P C0 Main as (C, M, \pc\, Normal) ek (C, M, \pc\, Return)" by (fastforce intro: JVMCFG_reachable.CFG_Invoke_False) ultimately show ?case by blast next case (CFG_Invoke_Return_Check_Normal C P C0 Main M pc M' n ST LT ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, nodeType.Return)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, nodeType.Return)\ \instrs_of (PROG P) C M ! pc = Invoke M' n\ \TYPING P C M ! pc = \(ST, LT)\\ \ST ! n \ NT\ \ek = (\s. s Exception = None)\<^sub>\\ have "vp_snoc P C0 Main as (C, M, \pc\, Return) ek (C, M, \Suc pc\, Enter)" by (fastforce intro: JVMCFG_reachable.CFG_Invoke_Return_Check_Normal) ultimately show ?case by blast next case (CFG_Invoke_Return_Check_Exceptional C P C0 Main M pc M' n Exc pc' diff ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, nodeType.Return)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, nodeType.Return)\ \instrs_of (PROG P) C M ! pc = Invoke M' n\ \match_ex_table (PROG P) Exc pc (ex_table_of (PROG P) C M) = \(pc', diff)\\ \pc' \ length (instrs_of (PROG P) C M)\ \ek = (\s. \v d. s Exception = \v\ \ match_ex_table (PROG P) (cname_of (heap_of s) (the_Addr (the_Value v))) pc (ex_table_of (PROG P) C M) = \(pc', d)\)\<^sub>\\ have "vp_snoc P C0 Main as (C, M, \pc\, Return) ek (C, M, \pc\, Exceptional \pc'\ Return)" by (fastforce intro: JVMCFG_reachable.CFG_Invoke_Return_Check_Exceptional) ultimately show ?case by blast next case (CFG_Invoke_Return_Exceptional_handle C P C0 Main M pc pc' M' n ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, Exceptional \pc'\ nodeType.Return)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, Exceptional \pc'\ nodeType.Return)\ \instrs_of (PROG P) C M ! pc = Invoke M' n\ \ek = \\s. s(Exception := None, Stack (stkLength (P, C, M) pc' - 1) := s Exception)\ have "vp_snoc P C0 Main as (C, M, \pc\, Exceptional \pc'\ Return) ek (C, M, \pc'\, Enter)" by (fastforce intro: JVMCFG_reachable.CFG_Invoke_Return_Exceptional_handle) ultimately show ?case by blast next case (CFG_Invoke_Return_Exceptional_prop C P C0 Main M pc M' n ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, nodeType.Return)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, nodeType.Return)\ \instrs_of (PROG P) C M ! pc = Invoke M' n\ \ek = (\s. \v. s Exception = \v\ \ match_ex_table (PROG P) (cname_of (heap_of s) (the_Addr (the_Value v))) pc (ex_table_of (PROG P) C M) = None)\<^sub>\\ have "vp_snoc P C0 Main as (C, M, \pc\, Return) ek (C, M, None, Return)" by (fastforce intro: JVMCFG_reachable.CFG_Invoke_Return_Exceptional_prop) ultimately show ?case by blast next case (CFG_Return C P C0 Main M pc ek) then obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C, M, \pc\, Enter)" by blast moreover with \C \ ClassMain P\ \(P, C0, Main) \ \(C, M, \pc\, Enter)\ \instrs_of (PROG P) C M ! pc = instr.Return\ \ek = \\s. s(Stack 0 := s (Stack (stkLength (P, C, M) pc - 1)))\ have "vp_snoc P C0 Main as (C, M, \pc\, Enter) ek (C, M, None, Return)" by (fastforce intro: JVMCFG_reachable.CFG_Return) ultimately show ?case by blast next case (CFG_Return_from_Method P C0 Main C M C' M' pc' Q' ps Q stateUpdate ek) from \(P, C0, Main) \ (C', M', \pc'\, Normal) -Q':(C', M', pc')\\<^bsub>(C, M)\<^esub>ps\ (C, M, None, Enter)\ show ?case proof cases case Main_Call with CFG_Return_from_Method obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (ClassMain P, MethodMain P, \0\, Normal)" by blast moreover with Main_Call have "vp_snoc P C0 Main as (ClassMain P, MethodMain P, \0\, Normal) (\s. False)\<^sub>\ (ClassMain P, MethodMain P, \0\, Return)" by (fastforce intro: Main_Call_LFalse) ultimately show ?thesis using Main_Call CFG_Return_from_Method by blast next case CFG_Invoke_Call with CFG_Return_from_Method obtain as where "JVMCFG_Interpret.valid_path' P C0 Main (ClassMain P, MethodMain P, None, Enter) as (C', M', \pc'\, Normal)" by blast moreover with CFG_Invoke_Call have "vp_snoc P C0 Main as (C', M', \pc'\, Normal) (\s. False)\<^sub>\ (C', M', \pc'\, Return)" by (fastforce intro: CFG_Invoke_False) ultimately show ?thesis using CFG_Invoke_Call CFG_Return_from_Method by blast qed qed declare JVMCFG_Interpret.vp_snocI [] declare JVMCFG_Interpret.valid_node_def [simp del] valid_edge_def [simp del] JVMCFG_Interpret.intra_path_def [simp del] definition EP :: jvm_prog where "EP = (''C'', Object, [], [(''M'', [], Void, 1::nat, 0::nat, [Push Unit, instr.Return], [])]) # SystemClasses" definition Phi_EP :: ty\<^sub>P where "Phi_EP C M = (if C = ''C'' \ M = ''M'' then [\([],[OK (Class ''C'')])\,\([Void],[OK (Class ''C'')])\] else [])" lemma distinct_classes'': "''C'' \ Object" "''C'' \ NullPointer" "''C'' \ OutOfMemory" "''C'' \ ClassCast" by (simp_all add: Object_def NullPointer_def OutOfMemory_def ClassCast_def) lemmas distinct_classes = distinct_classes distinct_classes'' distinct_classes'' [symmetric] declare distinct_classes [simp add] lemma i_max_2D: "i < Suc (Suc 0) \ i = 0 \ i = 1" by auto lemma EP_wf: "wf_jvm_prog\<^bsub>Phi_EP\<^esub> EP" unfolding wf_jvm_prog_phi_def wf_prog_def proof show "wf_syscls EP" by (simp add: EP_def wf_syscls_def SystemClasses_def sys_xcpts_def ObjectC_def NullPointerC_def OutOfMemoryC_def ClassCastC_def) next have distinct_EP: "distinct_fst EP" by (auto simp: EP_def SystemClasses_def ObjectC_def NullPointerC_def OutOfMemoryC_def ClassCastC_def) moreover have classes_wf: "\c\set EP. wf_cdecl (\P C (M, Ts, T\<^sub>r, mxs, mxl\<^sub>0, is, xt). wt_method P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt (Phi_EP C M)) EP c" proof fix C assume C_in_EP: "C \ set EP" show "wf_cdecl (\P C (M, Ts, T\<^sub>r, mxs, mxl\<^sub>0, is, xt). wt_method P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt (Phi_EP C M)) EP C" proof (cases "C \ set SystemClasses") case True thus ?thesis by (auto simp: wf_cdecl_def SystemClasses_def ObjectC_def NullPointerC_def OutOfMemoryC_def ClassCastC_def EP_def class_def) next case False with C_in_EP have "C = (''C'', the (class EP ''C''))" by (auto simp: EP_def SystemClasses_def class_def) thus ?thesis by (auto dest!: i_max_2D elim: Methods.cases simp: wf_cdecl_def class_def EP_def wf_mdecl_def wt_method_def Phi_EP_def wt_start_def check_types_def states_def JVM_SemiType.sl_def SystemClasses_def stk_esl_def upto_esl_def loc_sl_def SemiType.esl_def ObjectC_def SemiType.sup_def Err.sl_def Err.le_def err_def Listn.sl_def Method_def Err.esl_def Opt.esl_def Product.esl_def relevant_entries_def) qed qed ultimately show "(\c\set EP. wf_cdecl (\P C (M, Ts, T\<^sub>r, mxs, mxl\<^sub>0, is, xt). wt_method P C Ts T\<^sub>r mxs mxl\<^sub>0 is xt (Phi_EP C M)) EP c) \ distinct_fst EP" by simp qed lemma [simp]: "PROG (Abs_wf_jvmprog (EP, Phi_EP)) = EP" proof (cases "(EP, Phi_EP) \ wf_jvmprog") case True thus ?thesis by (simp add: Abs_wf_jvmprog_inverse) next case False with EP_wf show ?thesis by (simp add: wf_jvmprog_def) qed lemma [simp]: "TYPING (Abs_wf_jvmprog (EP, Phi_EP)) = Phi_EP" proof (cases "(EP, Phi_EP) \ wf_jvmprog") case True thus ?thesis by (simp add: Abs_wf_jvmprog_inverse) next case False with EP_wf show ?thesis by (simp add: wf_jvmprog_def) qed lemma method_in_EP_is_M: "EP \ C sees M: Ts\T = (mxs, mxl, is, xt) in D \ C = ''C'' \ M = ''M'' \ Ts = [] \ T = Void \ mxs = 1 \ mxl = 0 \ is = [Push Unit, instr.Return] \ xt = [] \ D = ''C''" by (fastforce elim: Methods.cases simp: class_def SystemClasses_def ObjectC_def NullPointerC_def OutOfMemoryC_def ClassCastC_def if_split_eq1 EP_def Method_def) lemma [simp]: "\T Ts mxs mxl is. (\xt. EP \ ''C'' sees ''M'': Ts\T = (mxs, mxl, is, xt) in ''C'') \ is \ []" using EP_wf by (fastforce dest: mdecl_visible simp: wf_jvm_prog_phi_def EP_def) lemma [simp]: "\T Ts mxs mxl is. (\xt. EP \ ''C'' sees ''M'': Ts\T = (mxs, mxl, is, xt) in ''C'') \ Suc 0 < length is" using EP_wf by (fastforce dest: mdecl_visible simp: wf_jvm_prog_phi_def EP_def) lemma C_sees_M_in_EP [simp]: "EP \ ''C'' sees ''M'': []\Void = (Suc 0, 0, [Push Unit, instr.Return], []) in ''C''" proof - have "EP \ ''C'' sees_methods [''M'' \ (([], Void, 1, 0, [Push Unit, instr.Return], []), ''C'')]" by (fastforce intro: Methods.intros simp: class_def SystemClasses_def ObjectC_def EP_def) thus ?thesis by (fastforce simp: Method_def) qed lemma instrs_of_EP_C_M [simp]: "instrs_of EP ''C'' ''M'' = [Push Unit, instr.Return]" unfolding method_def by (rule theI2 [where P = "\(D, Ts, T, m). EP \ ''C'' sees ''M'': Ts\T = m in D"]) (auto dest: method_in_EP_is_M) lemma ClassMain_not_C [simp]: "ClassMain (Abs_wf_jvmprog (EP, Phi_EP)) \ ''C''" by (fastforce intro: no_Call_in_ClassMain [where P="Abs_wf_jvmprog (EP, Phi_EP)"] C_sees_M_in_EP) lemma method_entry [dest!]: "(Abs_wf_jvmprog (EP, Phi_EP), ''C'', ''M'') \ \(C, M, None, Enter) \ (C = ClassMain (Abs_wf_jvmprog (EP, Phi_EP)) \ M = MethodMain (Abs_wf_jvmprog (EP, Phi_EP))) \ (C = ''C'' \ M = ''M'')" by (fastforce elim: reachable.cases elim!: JVMCFG.cases dest!: method_in_EP_is_M) lemma valid_node_in_EP_D: assumes vn: "JVMCFG_Interpret.valid_node (Abs_wf_jvmprog (EP, Phi_EP)) ''C'' ''M'' n" shows "n \ { (ClassMain (Abs_wf_jvmprog (EP, Phi_EP)), MethodMain (Abs_wf_jvmprog (EP, Phi_EP)), None, Enter), (ClassMain (Abs_wf_jvmprog (EP, Phi_EP)), MethodMain (Abs_wf_jvmprog (EP, Phi_EP)), None, Return), (ClassMain (Abs_wf_jvmprog (EP, Phi_EP)), MethodMain (Abs_wf_jvmprog (EP, Phi_EP)), \0\, Enter), (ClassMain (Abs_wf_jvmprog (EP, Phi_EP)), MethodMain (Abs_wf_jvmprog (EP, Phi_EP)), \0\, Normal), (ClassMain (Abs_wf_jvmprog (EP, Phi_EP)), MethodMain (Abs_wf_jvmprog (EP, Phi_EP)), \0\, Return), (''C'', ''M'', None, Enter), (''C'', ''M'', \0\, Enter), (''C'', ''M'', \1\, Enter), (''C'', ''M'', None, Return) }" using vn proof (cases rule: JVMCFG_Interpret.valid_node_cases') let ?prog = "Abs_wf_jvmprog (EP, Phi_EP)" case (Source e) then obtain C M pc nt ek C' M' pc' nt' where [simp]: "e = ((C, M, pc, nt), ek, (C', M', pc', nt'))" and [simp]: "n = (C, M, pc, nt)" and edge: "(?prog, ''C'', ''M'') \ (C, M, pc, nt) -ek\ (C', M', pc', nt')" by (cases e) (fastforce simp: valid_edge_def) from edge have src_reachable: "(?prog, ''C'', ''M'') \ \(C, M, pc, nt)" by -(drule sourcenode_reachable) show ?thesis proof (cases "C = ClassMain ?prog") case True with src_reachable have "M = MethodMain ?prog" by (fastforce dest: ClassMain_imp_MethodMain) with True edge show ?thesis by clarsimp (erule JVMCFG.cases, simp_all) next case False with src_reachable obtain T Ts mb where "EP \ C sees M:Ts\T = mb in C" by (fastforce dest: method_of_reachable_node_exists) hence [simp]: "C = ''C''" and [simp]: "M = ''M''" and [simp]: "Ts = []" and [simp]: "T = Void" and [simp]: "mb = (1, 0, [Push Unit, instr.Return], [])" by (cases mb, fastforce dest: method_in_EP_is_M)+ from src_reachable False have "pc \ {None, \0\, \1\}" by (fastforce dest: instr_of_reachable_node_typable) show ?thesis proof (cases pc) case None with edge False show ?thesis by clarsimp (erule JVMCFG.cases, simp_all) next case (Some pc') show ?thesis proof (cases pc') case 0 with Some False edge show ?thesis by clarsimp (erule JVMCFG.cases, fastforce+) next case (Suc n) with \pc \ {None, \0\, \1\}\ Some have "pc = \1\" by simp with False edge show ?thesis by clarsimp (erule JVMCFG.cases, fastforce+) qed qed qed next let ?prog = "Abs_wf_jvmprog (EP, Phi_EP)" case (Target e) then obtain C M pc nt ek C' M' pc' nt' where [simp]: "e = ((C, M, pc, nt), ek, (C', M', pc', nt'))" and [simp]: "n = (C', M', pc', nt')" and edge: "(?prog, ''C'', ''M'') \ (C, M, pc, nt) -ek\ (C', M', pc', nt')" by (cases e) (fastforce simp: valid_edge_def) from edge have trg_reachable: "(?prog, ''C'', ''M'') \ \(C', M', pc', nt')" by -(drule targetnode_reachable) show ?thesis proof (cases "C' = ClassMain ?prog") case True with trg_reachable have "M' = MethodMain ?prog" by (fastforce dest: ClassMain_imp_MethodMain) with True edge show ?thesis by -(clarsimp, (erule JVMCFG.cases, simp_all))+ next case False with trg_reachable obtain T Ts mb where "EP \ C' sees M':Ts\T = mb in C'" by (fastforce dest: method_of_reachable_node_exists) hence [simp]: "C' = ''C''" and [simp]: "M' = ''M''" and [simp]: "Ts = []" and [simp]: "T = Void" and [simp]: "mb = (1, 0, [Push Unit, instr.Return], [])" by (cases mb, fastforce dest: method_in_EP_is_M)+ from trg_reachable False have "pc' \ {None, \0\, \1\}" by (fastforce dest: instr_of_reachable_node_typable) show ?thesis proof (cases pc') case None with edge False show ?thesis by clarsimp (erule JVMCFG.cases, simp_all) next case (Some pc'') show ?thesis proof (cases pc'') case 0 with Some False edge show ?thesis by -(clarsimp, (erule JVMCFG.cases, fastforce+))+ next case (Suc n) with \pc' \ {None, \0\, \1\}\ Some have "pc' = \1\" by simp with False edge show ?thesis by -(clarsimp, (erule JVMCFG.cases, fastforce+))+ qed qed qed qed lemma Main_Entry_valid [simp]: "JVMCFG_Interpret.valid_node (Abs_wf_jvmprog (EP, Phi_EP)) ''C'' ''M'' (ClassMain (Abs_wf_jvmprog (EP, Phi_EP)), MethodMain (Abs_wf_jvmprog (EP, Phi_EP)), None, Enter)" proof - have "valid_edge (Abs_wf_jvmprog (EP, Phi_EP), ''C'', ''M'') ((ClassMain (Abs_wf_jvmprog (EP, Phi_EP)), MethodMain (Abs_wf_jvmprog (EP, Phi_EP)), None, Enter), (\s. False)\<^sub>\, (ClassMain (Abs_wf_jvmprog (EP, Phi_EP)), MethodMain (Abs_wf_jvmprog (EP, Phi_EP)), None, Return))" by (auto simp: valid_edge_def intro: JVMCFG_reachable.intros) thus ?thesis by (fastforce simp: JVMCFG_Interpret.valid_node_def) qed lemma main_0_Enter_reachable [simp]: "(P, C0, Main) \ \(ClassMain P, MethodMain P, \0\, Enter)" by (rule reachable_step [where n="(ClassMain P, MethodMain P, None, Enter)"]) (fastforce intro: JVMCFG_reachable.intros)+ lemma main_0_Normal_reachable [simp]: "(P, C0, Main) \ \(ClassMain P, MethodMain P, \0\, Normal)" by (rule reachable_step [where n="(ClassMain P, MethodMain P, \0\, Enter)"], simp) (fastforce intro: JVMCFG_reachable.intros) lemma main_0_Return_reachable [simp]: "(P, C0, Main) \ \(ClassMain P, MethodMain P, \0\, Return)" by (rule reachable_step [where n="(ClassMain P, MethodMain P, \0\, Normal)"], simp) (fastforce intro: JVMCFG_reachable.intros) lemma Exit_reachable [simp]: "(P, C0, Main) \ \(ClassMain P, MethodMain P, None, Return)" by (rule reachable_step [where n="(ClassMain P, MethodMain P, \0\, Return)"], simp) (fastforce intro: JVMCFG_reachable.intros) definition "cfg_wf_prog = {(P, C0, Main). (\n. JVMCFG_Interpret.valid_node P C0 Main n \ (\as. CFG.valid_path' sourcenode targetnode kind (valid_edge (P, C0, Main)) (get_return_edges P) n as (ClassMain P, MethodMain P, None, Return)))}" typedef cfg_wf_prog = cfg_wf_prog unfolding cfg_wf_prog_def proof let ?prog = "(Abs_wf_jvmprog (EP, Phi_EP), ''C'', ''M'')" let ?edge_main0 = "((ClassMain (fst ?prog), MethodMain (fst ?prog), None, Enter), (\s. False)\<^sub>\, (ClassMain (fst ?prog), MethodMain (fst ?prog), None, Return))" let ?edge_main1 = "((ClassMain (fst ?prog), MethodMain (fst ?prog), None, Enter), (\s. True)\<^sub>\, (ClassMain (fst ?prog), MethodMain (fst ?prog), \0\, Enter))" let ?edge_main2 = "((ClassMain (fst ?prog), MethodMain (fst ?prog), \0\, Enter), \id, (ClassMain (fst ?prog), MethodMain (fst ?prog), \0\, Normal))" let ?edge_main3 = "((ClassMain (fst ?prog), MethodMain (fst ?prog), \0\, Normal), (\s. False)\<^sub>\, (ClassMain (fst ?prog), MethodMain (fst ?prog), \0\, Return))" let ?edge_main4 = "((ClassMain (fst ?prog), MethodMain (fst ?prog), \0\, Return), \id, (ClassMain (fst ?prog), MethodMain (fst ?prog), None, Return))" let ?edge_call = "((ClassMain (fst ?prog), MethodMain (fst ?prog), \0\, Normal), ((\(s, ret). True):(ClassMain (fst ?prog), MethodMain (fst ?prog), 0)\\<^bsub>(''C'', ''M'')\<^esub>[(\s. s Heap),(\s. \Value Null\)]), (''C'', ''M'', None, Enter))" let ?edge_C0 = "((''C'', ''M'', None, Enter), (\s. False)\<^sub>\, (''C'', ''M'', None, Return))" let ?edge_C1 = "((''C'', ''M'', None, Enter), (\s. True)\<^sub>\, (''C'', ''M'', \0\, Enter))" let ?edge_C2 = "((''C'', ''M'', \0\, Enter), \(\s. s(Stack 0 \ Value Unit)), (''C'', ''M'', \1\, Enter))" let ?edge_C3 = "((''C'', ''M'', \1\, Enter), \(\s. s(Stack 0 := s (Stack 0))), (''C'', ''M'', None, Return))" let ?edge_return = "((''C'', ''M'', None, Return), (\(s, ret). ret = (ClassMain (fst ?prog), MethodMain (fst ?prog), 0))\\<^bsub>(''C'',''M'')\<^esub>(\s s'. s'(Heap := s Heap, Exception := s Exception, Stack 0 := s (Stack 0))), (ClassMain (fst ?prog), MethodMain (fst ?prog), \0\, Return))" have [simp]: "(Abs_wf_jvmprog (EP, Phi_EP), ''C'', ''M'') \ \(''C'', ''M'', None, Enter)" by (rule reachable_step [where n="(ClassMain (fst ?prog), MethodMain (fst ?prog), \0\, Normal)"] , simp) (fastforce intro: Main_Call C_sees_M_in_EP) hence [simp]: "(Abs_wf_jvmprog (EP, Phi_EP), ''C'', ''M'') \ \(''C'', ''M'', None, nodeType.Return)" by (rule reachable_step [where n="(''C'', ''M'', None, Enter)"]) (fastforce intro: JVMCFG_reachable.intros) have [simp]: "(Abs_wf_jvmprog (EP, Phi_EP), ''C'', ''M'') \ \(''C'', ''M'', \0\, Enter)" by (rule reachable_step [where n="(''C'', ''M'', None, Enter)"], simp) (fastforce intro: JVMCFG_reachable.intros) hence [simp]: "(Abs_wf_jvmprog (EP, Phi_EP), ''C'', ''M'') \ \(''C'', ''M'', \Suc 0\, Enter)" by (fastforce intro: reachable_step [where n="(''C'', ''M'', \0\, Enter)"] CFG_Push simp: ClassMain_not_C [symmetric]) show "?prog \ {(P, C0, Main). \n. CFG.valid_node sourcenode targetnode (valid_edge (P, C0, Main)) n \ (\as. CFG.valid_path' sourcenode targetnode kind (valid_edge (P, C0, Main)) (get_return_edges P) n as (ClassMain P, MethodMain P, None, nodeType.Return))}" proof (auto dest!: valid_node_in_EP_D) have "CFG.valid_path' sourcenode targetnode kind (valid_edge (Abs_wf_jvmprog (EP, Phi_EP), ''C'', ''M'')) (get_return_edges (Abs_wf_jvmprog (EP, Phi_EP))) (ClassMain (Abs_wf_jvmprog (EP, Phi_EP)), MethodMain (Abs_wf_jvmprog (EP, Phi_EP)), None, Enter) [?edge_main0] (ClassMain (Abs_wf_jvmprog (EP, Phi_EP)), MethodMain (Abs_wf_jvmprog (EP, Phi_EP)), None, nodeType.Return)" by (fastforce intro: JVMCFG_Interpret.intra_path_vp JVMCFG_reachable.intros simp: JVMCFG_Interpret.intra_path_def intra_kind_def valid_edge_def) thus " \as. CFG.valid_path' sourcenode targetnode kind (valid_edge (Abs_wf_jvmprog (EP, Phi_EP), ''C'', ''M'')) (get_return_edges (Abs_wf_jvmprog (EP, Phi_EP))) (ClassMain (Abs_wf_jvmprog (EP, Phi_EP)), MethodMain (Abs_wf_jvmprog (EP, Phi_EP)), None, Enter) as (ClassMain (Abs_wf_jvmprog (EP, Phi_EP)), MethodMain (Abs_wf_jvmprog (EP, Phi_EP)), None, nodeType.Return)" by blast next have "CFG.valid_path' sourcenode targetnode kind (valid_edge (Abs_wf_jvmprog (EP, Phi_EP), ''C'', ''M'')) (get_return_edges (Abs_wf_jvmprog (EP, Phi_EP))) (ClassMain (Abs_wf_jvmprog (EP, Phi_EP)), MethodMain (Abs_wf_jvmprog (EP, Phi_EP)), None, nodeType.Return) [] (ClassMain (Abs_wf_jvmprog (EP, Phi_EP)), MethodMain (Abs_wf_jvmprog (EP, Phi_EP)), None, nodeType.Return)" by (fastforce intro: JVMCFG_Interpret.intra_path_vp simp: JVMCFG_Interpret.intra_path_def) thus "\as. CFG.valid_path' sourcenode targetnode kind (valid_edge (Abs_wf_jvmprog (EP, Phi_EP), ''C'', ''M'')) (get_return_edges (Abs_wf_jvmprog (EP, Phi_EP))) (ClassMain (Abs_wf_jvmprog (EP, Phi_EP)), MethodMain (Abs_wf_jvmprog (EP, Phi_EP)), None, nodeType.Return) as (ClassMain (Abs_wf_jvmprog (EP, Phi_EP)), MethodMain (Abs_wf_jvmprog (EP, Phi_EP)), None, nodeType.Return)" by blast next have "CFG.valid_path' sourcenode targetnode kind (valid_edge (Abs_wf_jvmprog (EP, Phi_EP), ''C'', ''M'')) (get_return_edges (Abs_wf_jvmprog (EP, Phi_EP))) (ClassMain (Abs_wf_jvmprog (EP, Phi_EP)), MethodMain (Abs_wf_jvmprog (EP, Phi_EP)), \0\, Enter) [?edge_main2, ?edge_main3, ?edge_main4] (ClassMain (Abs_wf_jvmprog (EP, Phi_EP)), MethodMain (Abs_wf_jvmprog (EP, Phi_EP)), None, nodeType.Return)" by (fastforce intro: JVMCFG_Interpret.intra_path_vp JVMCFG_reachable.intros simp: JVMCFG_Interpret.intra_path_def intra_kind_def valid_edge_def) thus "\as. CFG.valid_path' sourcenode targetnode kind (valid_edge (Abs_wf_jvmprog (EP, Phi_EP), ''C'', ''M'')) (get_return_edges (Abs_wf_jvmprog (EP, Phi_EP))) (ClassMain (Abs_wf_jvmprog (EP, Phi_EP)), MethodMain (Abs_wf_jvmprog (EP, Phi_EP)), \0\, Enter) as (ClassMain (Abs_wf_jvmprog (EP, Phi_EP)), MethodMain (Abs_wf_jvmprog (EP, Phi_EP)), None, nodeType.Return)" by blast next have "CFG.valid_path' sourcenode targetnode kind (valid_edge (Abs_wf_jvmprog (EP, Phi_EP), ''C'', ''M'')) (get_return_edges (Abs_wf_jvmprog (EP, Phi_EP))) (ClassMain (Abs_wf_jvmprog (EP, Phi_EP)), MethodMain (Abs_wf_jvmprog (EP, Phi_EP)), \0\, Normal) [?edge_main3, ?edge_main4] (ClassMain (Abs_wf_jvmprog (EP, Phi_EP)), MethodMain (Abs_wf_jvmprog (EP, Phi_EP)), None, nodeType.Return)" by (fastforce intro: JVMCFG_Interpret.intra_path_vp JVMCFG_reachable.intros simp: JVMCFG_Interpret.intra_path_def intra_kind_def valid_edge_def) thus "\as. CFG.valid_path' sourcenode targetnode kind (valid_edge (Abs_wf_jvmprog (EP, Phi_EP), ''C'', ''M'')) (get_return_edges (Abs_wf_jvmprog (EP, Phi_EP))) (ClassMain (Abs_wf_jvmprog (EP, Phi_EP)), MethodMain (Abs_wf_jvmprog (EP, Phi_EP)), \0\, Normal) as (ClassMain (Abs_wf_jvmprog (EP, Phi_EP)), MethodMain (Abs_wf_jvmprog (EP, Phi_EP)), None, nodeType.Return)" by blast next have "CFG.valid_path' sourcenode targetnode kind (valid_edge (Abs_wf_jvmprog (EP, Phi_EP), ''C'', ''M'')) (get_return_edges (Abs_wf_jvmprog (EP, Phi_EP))) (ClassMain (Abs_wf_jvmprog (EP, Phi_EP)), MethodMain (Abs_wf_jvmprog (EP, Phi_EP)), \0\, nodeType.Return) [?edge_main4] (ClassMain (Abs_wf_jvmprog (EP, Phi_EP)), MethodMain (Abs_wf_jvmprog (EP, Phi_EP)), None, nodeType.Return)" by (fastforce intro: JVMCFG_Interpret.intra_path_vp JVMCFG_reachable.intros simp: JVMCFG_Interpret.intra_path_def intra_kind_def valid_edge_def) thus "\as. CFG.valid_path' sourcenode targetnode kind (valid_edge (Abs_wf_jvmprog (EP, Phi_EP), ''C'', ''M'')) (get_return_edges (Abs_wf_jvmprog (EP, Phi_EP))) (ClassMain (Abs_wf_jvmprog (EP, Phi_EP)), MethodMain (Abs_wf_jvmprog (EP, Phi_EP)), \0\, nodeType.Return) as (ClassMain (Abs_wf_jvmprog (EP, Phi_EP)), MethodMain (Abs_wf_jvmprog (EP, Phi_EP)), None, nodeType.Return)" by blast next have "CFG.valid_path' sourcenode targetnode kind (valid_edge (Abs_wf_jvmprog (EP, Phi_EP), ''C'', ''M'')) (get_return_edges (Abs_wf_jvmprog (EP, Phi_EP))) (''C'', ''M'', None, Enter) [?edge_C0, ?edge_return, ?edge_main4] (ClassMain (Abs_wf_jvmprog (EP, Phi_EP)), MethodMain (Abs_wf_jvmprog (EP, Phi_EP)), None, nodeType.Return)" by (fastforce intro: JVMCFG_reachable.intros C_sees_M_in_EP simp: JVMCFG_Interpret.vp_def valid_edge_def stkLength_def JVMCFG_Interpret.valid_path_def) thus "\as. CFG.valid_path' sourcenode targetnode kind (valid_edge (Abs_wf_jvmprog (EP, Phi_EP), ''C'', ''M'')) (get_return_edges (Abs_wf_jvmprog (EP, Phi_EP))) (''C'', ''M'', None, Enter) as (ClassMain (Abs_wf_jvmprog (EP, Phi_EP)), MethodMain (Abs_wf_jvmprog (EP, Phi_EP)), None, nodeType.Return)" by blast next have "CFG.valid_path' sourcenode targetnode kind (valid_edge (Abs_wf_jvmprog (EP, Phi_EP), ''C'', ''M'')) (get_return_edges (Abs_wf_jvmprog (EP, Phi_EP))) (''C'', ''M'', \0\, Enter) [?edge_C2, ?edge_C3, ?edge_return, ?edge_main4] (ClassMain (Abs_wf_jvmprog (EP, Phi_EP)), MethodMain (Abs_wf_jvmprog (EP, Phi_EP)), None, nodeType.Return)" by (fastforce intro: Main_Return_to_Exit CFG_Return_from_Method Main_Call C_sees_M_in_EP CFG_Return CFG_Push simp: JVMCFG_Interpret.vp_def valid_edge_def stkLength_def Phi_EP_def ClassMain_not_C [symmetric] JVMCFG_Interpret.valid_path_def) thus "\as. CFG.valid_path' sourcenode targetnode kind (valid_edge (Abs_wf_jvmprog (EP, Phi_EP), ''C'', ''M'')) (get_return_edges (Abs_wf_jvmprog (EP, Phi_EP))) (''C'', ''M'', \0\, Enter) as (ClassMain (Abs_wf_jvmprog (EP, Phi_EP)), MethodMain (Abs_wf_jvmprog (EP, Phi_EP)), None, nodeType.Return)" by blast next have "CFG.valid_path' sourcenode targetnode kind (valid_edge (Abs_wf_jvmprog (EP, Phi_EP), ''C'', ''M'')) (get_return_edges (Abs_wf_jvmprog (EP, Phi_EP))) (''C'', ''M'', \Suc 0\, Enter) [?edge_C3, ?edge_return, ?edge_main4] (ClassMain (Abs_wf_jvmprog (EP, Phi_EP)), MethodMain (Abs_wf_jvmprog (EP, Phi_EP)), None, nodeType.Return)" by (fastforce intro: JVMCFG_reachable.intros C_sees_M_in_EP simp: JVMCFG_Interpret.vp_def valid_edge_def stkLength_def Phi_EP_def ClassMain_not_C [symmetric] JVMCFG_Interpret.valid_path_def) thus "\as. CFG.valid_path' sourcenode targetnode kind (valid_edge (Abs_wf_jvmprog (EP, Phi_EP), ''C'', ''M'')) (get_return_edges (Abs_wf_jvmprog (EP, Phi_EP))) (''C'', ''M'', \Suc 0\, Enter) as (ClassMain (Abs_wf_jvmprog (EP, Phi_EP)), MethodMain (Abs_wf_jvmprog (EP, Phi_EP)), None, nodeType.Return)" by blast next have "CFG.valid_path' sourcenode targetnode kind (valid_edge (Abs_wf_jvmprog (EP, Phi_EP), ''C'', ''M'')) (get_return_edges (Abs_wf_jvmprog (EP, Phi_EP))) (''C'', ''M'', None, nodeType.Return) [?edge_return, ?edge_main4] (ClassMain (Abs_wf_jvmprog (EP, Phi_EP)), MethodMain (Abs_wf_jvmprog (EP, Phi_EP)), None, nodeType.Return)" by (fastforce intro: JVMCFG_reachable.intros C_sees_M_in_EP simp: JVMCFG_Interpret.vp_def valid_edge_def JVMCFG_Interpret.valid_path_def stkLength_def) thus "\as. CFG.valid_path' sourcenode targetnode kind (valid_edge (Abs_wf_jvmprog (EP, Phi_EP), ''C'', ''M'')) (get_return_edges (Abs_wf_jvmprog (EP, Phi_EP))) (''C'', ''M'', None, nodeType.Return) as (ClassMain (Abs_wf_jvmprog (EP, Phi_EP)), MethodMain (Abs_wf_jvmprog (EP, Phi_EP)), None, nodeType.Return)" by blast qed qed abbreviation lift_to_cfg_wf_prog :: "(jvm_method \ 'a) \ (cfg_wf_prog \ 'a)" ("_\<^bsub>CFG\<^esub>") where "f\<^bsub>CFG\<^esub> \ (\P. f (Rep_cfg_wf_prog P))" lemma valid_edge_CFG_def: "valid_edge\<^bsub>CFG\<^esub> P = valid_edge (fst\<^bsub>CFG\<^esub> P, fst (snd\<^bsub>CFG\<^esub> P), snd (snd\<^bsub>CFG\<^esub> P))" by (cases P) (clarsimp simp: Abs_cfg_wf_prog_inverse) interpretation JVMCFG_Postdomination: Postdomination "sourcenode" "targetnode" "kind" "valid_edge\<^bsub>CFG\<^esub> P" "(ClassMain (fst\<^bsub>CFG\<^esub> P), MethodMain (fst\<^bsub>CFG\<^esub> P), None, Enter)" "(\(C, M, pc, type). (C, M))" "get_return_edges (fst\<^bsub>CFG\<^esub> P)" "((ClassMain (fst\<^bsub>CFG\<^esub> P), MethodMain (fst\<^bsub>CFG\<^esub> P)),[],[]) # procs (PROG (fst\<^bsub>CFG\<^esub> P))" "(ClassMain (fst\<^bsub>CFG\<^esub> P), MethodMain (fst\<^bsub>CFG\<^esub> P))" "(ClassMain (fst\<^bsub>CFG\<^esub> P), MethodMain (fst\<^bsub>CFG\<^esub> P), None, Return)" for P unfolding valid_edge_CFG_def proof fix n obtain P' C0 Main where [simp]: "fst\<^bsub>CFG\<^esub> P = P'" and [simp]: "fst (snd\<^bsub>CFG\<^esub> P) = C0" and [simp]: "snd (snd\<^bsub>CFG\<^esub> P) = Main" by (cases P) clarsimp assume "CFG.valid_node sourcenode targetnode (valid_edge (fst\<^bsub>CFG\<^esub> P, fst (snd\<^bsub>CFG\<^esub> P), snd (snd\<^bsub>CFG\<^esub> P))) n" thus "\as. CFG.valid_path' sourcenode targetnode kind (valid_edge (fst\<^bsub>CFG\<^esub> P, fst (snd\<^bsub>CFG\<^esub> P), snd (snd\<^bsub>CFG\<^esub> P))) (get_return_edges (fst\<^bsub>CFG\<^esub> P)) (ClassMain (fst\<^bsub>CFG\<^esub> P), MethodMain (fst\<^bsub>CFG\<^esub> P), None, Enter) as n" by (auto dest: sourcenode_reachable targetnode_reachable valid_Entry_path simp: JVMCFG_Interpret.valid_node_def valid_edge_def) next fix n obtain P' C0 Main where [simp]: "fst\<^bsub>CFG\<^esub> P = P'" and [simp]: "fst (snd\<^bsub>CFG\<^esub> P) = C0" and [simp]: "snd (snd\<^bsub>CFG\<^esub> P) = Main" and "(P', C0, Main) \ cfg_wf_prog" by (cases P) (clarsimp simp: Abs_cfg_wf_prog_inverse) assume "CFG.valid_node sourcenode targetnode (valid_edge (fst\<^bsub>CFG\<^esub> P, fst (snd\<^bsub>CFG\<^esub> P), snd (snd\<^bsub>CFG\<^esub> P))) n" with \(P', C0, Main) \ cfg_wf_prog\ show "\as. CFG.valid_path' sourcenode targetnode kind (valid_edge (fst\<^bsub>CFG\<^esub> P, fst (snd\<^bsub>CFG\<^esub> P), snd (snd\<^bsub>CFG\<^esub> P))) (get_return_edges (fst\<^bsub>CFG\<^esub> P)) n as (ClassMain (fst\<^bsub>CFG\<^esub> P), MethodMain (fst\<^bsub>CFG\<^esub> P), None, nodeType.Return)" by (cases n) (fastforce simp: cfg_wf_prog_def) next fix n n' obtain P' C0 Main where [simp]: "fst\<^bsub>CFG\<^esub> P = P'" and [simp]: "fst (snd\<^bsub>CFG\<^esub> P) = C0" and [simp]: "snd (snd\<^bsub>CFG\<^esub> P) = Main" by (cases P) clarsimp assume "CFGExit.method_exit sourcenode kind (valid_edge (fst\<^bsub>CFG\<^esub> P, fst (snd\<^bsub>CFG\<^esub> P), snd (snd\<^bsub>CFG\<^esub> P))) (ClassMain (fst\<^bsub>CFG\<^esub> P), MethodMain (fst\<^bsub>CFG\<^esub> P), None, nodeType.Return) n" and "CFGExit.method_exit sourcenode kind (valid_edge (fst\<^bsub>CFG\<^esub> P, fst (snd\<^bsub>CFG\<^esub> P), snd (snd\<^bsub>CFG\<^esub> P))) (ClassMain (fst\<^bsub>CFG\<^esub> P), MethodMain (fst\<^bsub>CFG\<^esub> P), None, nodeType.Return) n'" and "(\(C, M, pc, type). (C, M)) n = (\(C, M, pc, type). (C, M)) n'" thus "n = n'" by (auto simp: JVMCFG_Exit_Interpret.method_exit_def valid_edge_def) (fastforce elim: JVMCFG.cases)+ qed end diff --git a/thys/MFODL_Monitor_Optimized/Monitor.thy b/thys/MFODL_Monitor_Optimized/Monitor.thy --- a/thys/MFODL_Monitor_Optimized/Monitor.thy +++ b/thys/MFODL_Monitor_Optimized/Monitor.thy @@ -1,5600 +1,5600 @@ (*<*) theory Monitor imports Formula Optimized_Join "MFOTL_Monitor.Abstract_Monitor" "HOL-Library.While_Combinator" "HOL-Library.Mapping" "Deriving.Derive" "Generic_Join.Generic_Join_Correctness" begin (*>*) section \Generic monitoring algorithm\ text \The algorithm defined here abstracts over the implementation of the temporal operators.\ subsection \Monitorable formulas\ definition "mmonitorable \ \ safe_formula \ \ Formula.future_bounded \" definition "mmonitorable_regex b g r \ safe_regex b g r \ Regex.pred_regex Formula.future_bounded r" definition is_simple_eq :: "Formula.trm \ Formula.trm \ bool" where "is_simple_eq t1 t2 = (Formula.is_Const t1 \ (Formula.is_Const t2 \ Formula.is_Var t2) \ Formula.is_Var t1 \ Formula.is_Const t2)" fun mmonitorable_exec :: "Formula.formula \ bool" where "mmonitorable_exec (Formula.Eq t1 t2) = is_simple_eq t1 t2" | "mmonitorable_exec (Formula.Neg (Formula.Eq (Formula.Var x) (Formula.Var y))) = (x = y)" | "mmonitorable_exec (Formula.Pred e ts) = list_all (\t. Formula.is_Var t \ Formula.is_Const t) ts" | "mmonitorable_exec (Formula.Let p \ \) = ({0..} \ Formula.fv \ \ mmonitorable_exec \ \ mmonitorable_exec \)" | "mmonitorable_exec (Formula.Neg \) = (fv \ = {} \ mmonitorable_exec \)" | "mmonitorable_exec (Formula.Or \ \) = (fv \ = fv \ \ mmonitorable_exec \ \ mmonitorable_exec \)" | "mmonitorable_exec (Formula.And \ \) = (mmonitorable_exec \ \ (safe_assignment (fv \) \ \ mmonitorable_exec \ \ fv \ \ fv \ \ (is_constraint \ \ (case \ of Formula.Neg \' \ mmonitorable_exec \' | _ \ False))))" | "mmonitorable_exec (Formula.Ands l) = (let (pos, neg) = partition mmonitorable_exec l in pos \ [] \ list_all mmonitorable_exec (map remove_neg neg) \ \(set (map fv neg)) \ \(set (map fv pos)))" | "mmonitorable_exec (Formula.Exists \) = (mmonitorable_exec \)" | "mmonitorable_exec (Formula.Agg y \ b f \) = (mmonitorable_exec \ \ y + b \ Formula.fv \ \ {0.. Formula.fv \ \ Formula.fv_trm f \ Formula.fv \)" | "mmonitorable_exec (Formula.Prev I \) = (mmonitorable_exec \)" | "mmonitorable_exec (Formula.Next I \) = (mmonitorable_exec \)" | "mmonitorable_exec (Formula.Since \ I \) = (Formula.fv \ \ Formula.fv \ \ (mmonitorable_exec \ \ (case \ of Formula.Neg \' \ mmonitorable_exec \' | _ \ False)) \ mmonitorable_exec \)" | "mmonitorable_exec (Formula.Until \ I \) = (Formula.fv \ \ Formula.fv \ \ right I \ \ \ (mmonitorable_exec \ \ (case \ of Formula.Neg \' \ mmonitorable_exec \' | _ \ False)) \ mmonitorable_exec \)" | "mmonitorable_exec (Formula.MatchP I r) = Regex.safe_regex Formula.fv (\g \. mmonitorable_exec \ \ (g = Lax \ (case \ of Formula.Neg \' \ mmonitorable_exec \' | _ \ False))) Past Strict r" | "mmonitorable_exec (Formula.MatchF I r) = (Regex.safe_regex Formula.fv (\g \. mmonitorable_exec \ \ (g = Lax \ (case \ of Formula.Neg \' \ mmonitorable_exec \' | _ \ False))) Futu Strict r \ right I \ \)" | "mmonitorable_exec _ = False" lemma cases_Neg_iff: "(case \ of formula.Neg \ \ P \ | _ \ False) \ (\\. \ = formula.Neg \ \ P \)" by (cases \) auto lemma safe_formula_mmonitorable_exec: "safe_formula \ \ Formula.future_bounded \ \ mmonitorable_exec \" proof (induct \ rule: safe_formula.induct) case (8 \ \) then show ?case unfolding safe_formula.simps future_bounded.simps mmonitorable_exec.simps by (auto simp: cases_Neg_iff) next case (9 \ \) then show ?case unfolding safe_formula.simps future_bounded.simps mmonitorable_exec.simps by (auto simp: cases_Neg_iff) next case (10 l) from "10.prems"(2) have bounded: "Formula.future_bounded \" if "\ \ set l" for \ using that by (auto simp: list.pred_set) obtain poss negs where posnegs: "(poss, negs) = partition safe_formula l" by simp obtain posm negm where posnegm: "(posm, negm) = partition mmonitorable_exec l" by simp have "set poss \ set posm" proof (rule subsetI) fix x assume "x \ set poss" then have "x \ set l" "safe_formula x" using posnegs by simp_all then have "mmonitorable_exec x" using "10.hyps"(1) bounded by blast then show "x \ set posm" using \x \ set poss\ posnegm posnegs by simp qed then have "set negm \ set negs" using posnegm posnegs by auto obtain "poss \ []" "list_all safe_formula (map remove_neg negs)" "(\x\set negs. fv x) \ (\x\set poss. fv x)" using "10.prems"(1) posnegs by simp then have "posm \ []" using \set poss \ set posm\ by auto moreover have "list_all mmonitorable_exec (map remove_neg negm)" proof - let ?l = "map remove_neg negm" have "\x. x \ set ?l \ mmonitorable_exec x" proof - fix x assume "x \ set ?l" then obtain y where "y \ set negm" "x = remove_neg y" by auto then have "y \ set negs" using \set negm \ set negs\ by blast then have "safe_formula x" unfolding \x = remove_neg y\ using \list_all safe_formula (map remove_neg negs)\ by (simp add: list_all_def) show "mmonitorable_exec x" proof (cases "\z. y = Formula.Neg z") case True then obtain z where "y = Formula.Neg z" by blast then show ?thesis using "10.hyps"(2)[OF posnegs refl] \x = remove_neg y\ \y \ set negs\ posnegs bounded \safe_formula x\ by fastforce next case False then have "remove_neg y = y" by (cases y) simp_all then have "y = x" unfolding \x = remove_neg y\ by simp show ?thesis using "10.hyps"(1) \y \ set negs\ posnegs \safe_formula x\ unfolding \y = x\ by auto qed qed then show ?thesis by (simp add: list_all_iff) qed moreover have "(\x\set negm. fv x) \ (\x\set posm. fv x)" using \\ (fv ` set negs) \ \ (fv ` set poss)\ \set poss \ set posm\ \set negm \ set negs\ by fastforce ultimately show ?case using posnegm by simp next case (15 \ I \) then show ?case unfolding safe_formula.simps future_bounded.simps mmonitorable_exec.simps by (auto simp: cases_Neg_iff) next case (16 \ I \) then show ?case unfolding safe_formula.simps future_bounded.simps mmonitorable_exec.simps by (auto simp: cases_Neg_iff) next case (17 I r) then show ?case by (auto elim!: safe_regex_mono[rotated] simp: cases_Neg_iff regex.pred_set) next case (18 I r) then show ?case by (auto elim!: safe_regex_mono[rotated] simp: cases_Neg_iff regex.pred_set) qed (auto simp add: is_simple_eq_def list.pred_set) lemma safe_assignment_future_bounded: "safe_assignment X \ \ Formula.future_bounded \" unfolding safe_assignment_def by (auto split: formula.splits) lemma is_constraint_future_bounded: "is_constraint \ \ Formula.future_bounded \" by (induction rule: is_constraint.induct) simp_all lemma mmonitorable_exec_mmonitorable: "mmonitorable_exec \ \ mmonitorable \" proof (induct \ rule: mmonitorable_exec.induct) case (7 \ \) then show ?case unfolding mmonitorable_def mmonitorable_exec.simps safe_formula.simps by (auto simp: cases_Neg_iff intro: safe_assignment_future_bounded is_constraint_future_bounded) next case (8 l) obtain poss negs where posnegs: "(poss, negs) = partition safe_formula l" by simp obtain posm negm where posnegm: "(posm, negm) = partition mmonitorable_exec l" by simp have pos_monitorable: "list_all mmonitorable_exec posm" using posnegm by (simp add: list_all_iff) have neg_monitorable: "list_all mmonitorable_exec (map remove_neg negm)" using "8.prems" posnegm by (simp add: list_all_iff) have "set posm \ set poss" using "8.hyps"(1) posnegs posnegm unfolding mmonitorable_def by auto then have "set negs \ set negm" using posnegs posnegm by auto have "poss \ []" using "8.prems" posnegm \set posm \ set poss\ by auto moreover have "list_all safe_formula (map remove_neg negs)" using neg_monitorable "8.hyps"(2)[OF posnegm refl] \set negs \ set negm\ unfolding mmonitorable_def by (auto simp: list_all_iff) moreover have "\ (fv ` set negm) \ \ (fv ` set posm)" using "8.prems" posnegm by simp then have "\ (fv ` set negs) \ \ (fv ` set poss)" using \set posm \ set poss\ \set negs \ set negm\ by fastforce ultimately have "safe_formula (Formula.Ands l)" using posnegs by simp moreover have "Formula.future_bounded (Formula.Ands l)" proof - have "list_all Formula.future_bounded posm" using pos_monitorable "8.hyps"(1) posnegm unfolding mmonitorable_def by (auto elim!: list.pred_mono_strong) moreover have fnegm: "list_all Formula.future_bounded (map remove_neg negm)" using neg_monitorable "8.hyps"(2) posnegm unfolding mmonitorable_def by (auto elim!: list.pred_mono_strong) then have "list_all Formula.future_bounded negm" proof - have "\x. x \ set negm \ Formula.future_bounded x" proof - fix x assume "x \ set negm" have "Formula.future_bounded (remove_neg x)" using fnegm \x \ set negm\ by (simp add: list_all_iff) then show "Formula.future_bounded x" by (cases x) auto qed then show ?thesis by (simp add: list_all_iff) qed ultimately have "list_all Formula.future_bounded l" using posnegm by (auto simp: list_all_iff) then show ?thesis by (auto simp: list_all_iff) qed ultimately show ?case unfolding mmonitorable_def .. next case (13 \ I \) then show ?case unfolding mmonitorable_def mmonitorable_exec.simps safe_formula.simps by (fastforce simp: cases_Neg_iff) next case (14 \ I \) then show ?case unfolding mmonitorable_def mmonitorable_exec.simps safe_formula.simps by (fastforce simp: one_enat_def cases_Neg_iff) next case (15 I r) then show ?case by (auto elim!: safe_regex_mono[rotated] dest: safe_regex_safe[rotated] simp: mmonitorable_regex_def mmonitorable_def cases_Neg_iff regex.pred_set) next case (16 I r) then show ?case by (auto 0 3 elim!: safe_regex_mono[rotated] dest: safe_regex_safe[rotated] simp: mmonitorable_regex_def mmonitorable_def cases_Neg_iff regex.pred_set) qed (auto simp add: mmonitorable_def mmonitorable_regex_def is_simple_eq_def one_enat_def list.pred_set) lemma monitorable_formula_code[code]: "mmonitorable \ = mmonitorable_exec \" using mmonitorable_exec_mmonitorable safe_formula_mmonitorable_exec mmonitorable_def by blast subsection \Handling regular expressions\ datatype mregex = MSkip nat | MTestPos nat | MTestNeg nat | MPlus mregex mregex | MTimes mregex mregex | MStar mregex primrec ok where "ok _ (MSkip n) = True" | "ok m (MTestPos n) = (n < m)" | "ok m (MTestNeg n) = (n < m)" | "ok m (MPlus r s) = (ok m r \ ok m s)" | "ok m (MTimes r s) = (ok m r \ ok m s)" | "ok m (MStar r) = ok m r" primrec from_mregex where "from_mregex (MSkip n) _ = Regex.Skip n" | "from_mregex (MTestPos n) \s = Regex.Test (\s ! n)" | "from_mregex (MTestNeg n) \s = (if safe_formula (Formula.Neg (\s ! n)) then Regex.Test (Formula.Neg (Formula.Neg (Formula.Neg (\s ! n)))) else Regex.Test (Formula.Neg (\s ! n)))" | "from_mregex (MPlus r s) \s = Regex.Plus (from_mregex r \s) (from_mregex s \s)" | "from_mregex (MTimes r s) \s = Regex.Times (from_mregex r \s) (from_mregex s \s)" | "from_mregex (MStar r) \s = Regex.Star (from_mregex r \s)" primrec to_mregex_exec where "to_mregex_exec (Regex.Skip n) xs = (MSkip n, xs)" | "to_mregex_exec (Regex.Test \) xs = (if safe_formula \ then (MTestPos (length xs), xs @ [\]) else case \ of Formula.Neg \' \ (MTestNeg (length xs), xs @ [\']) | _ \ (MSkip 0, xs))" | "to_mregex_exec (Regex.Plus r s) xs = (let (mr, ys) = to_mregex_exec r xs; (ms, zs) = to_mregex_exec s ys in (MPlus mr ms, zs))" | "to_mregex_exec (Regex.Times r s) xs = (let (mr, ys) = to_mregex_exec r xs; (ms, zs) = to_mregex_exec s ys in (MTimes mr ms, zs))" | "to_mregex_exec (Regex.Star r) xs = (let (mr, ys) = to_mregex_exec r xs in (MStar mr, ys))" primrec shift where "shift (MSkip n) k = MSkip n" | "shift (MTestPos i) k = MTestPos (i + k)" | "shift (MTestNeg i) k = MTestNeg (i + k)" | "shift (MPlus r s) k = MPlus (shift r k) (shift s k)" | "shift (MTimes r s) k = MTimes (shift r k) (shift s k)" | "shift (MStar r) k = MStar (shift r k)" primrec to_mregex where "to_mregex (Regex.Skip n) = (MSkip n, [])" | "to_mregex (Regex.Test \) = (if safe_formula \ then (MTestPos 0, [\]) else case \ of Formula.Neg \' \ (MTestNeg 0, [\']) | _ \ (MSkip 0, []))" | "to_mregex (Regex.Plus r s) = (let (mr, ys) = to_mregex r; (ms, zs) = to_mregex s in (MPlus mr (shift ms (length ys)), ys @ zs))" | "to_mregex (Regex.Times r s) = (let (mr, ys) = to_mregex r; (ms, zs) = to_mregex s in (MTimes mr (shift ms (length ys)), ys @ zs))" | "to_mregex (Regex.Star r) = (let (mr, ys) = to_mregex r in (MStar mr, ys))" lemma shift_0: "shift r 0 = r" by (induct r) auto lemma shift_shift: "shift (shift r k) j = shift r (k + j)" by (induct r) auto lemma to_mregex_to_mregex_exec: "case to_mregex r of (mr, \s) \ to_mregex_exec r xs = (shift mr (length xs), xs @ \s)" by (induct r arbitrary: xs) (auto simp: shift_shift ac_simps split: formula.splits prod.splits) lemma to_mregex_to_mregex_exec_Nil[code]: "to_mregex r = to_mregex_exec r []" using to_mregex_to_mregex_exec[where xs="[]" and r = r] by (auto simp: shift_0) lemma ok_mono: "ok m mr \ m \ n \ ok n mr" by (induct mr) auto lemma from_mregex_cong: "ok m mr \ (\i < m. xs ! i = ys ! i) \ from_mregex mr xs = from_mregex mr ys" by (induct mr) auto lemma not_Neg_cases: "(\\. \ \ Formula.Neg \) \ (case \ of formula.Neg \ \ f \ | _ \ x) = x" by (cases \) auto lemma to_mregex_exec_ok: "to_mregex_exec r xs = (mr, ys) \ \zs. ys = xs @ zs \ set zs = atms r \ ok (length ys) mr" proof (induct r arbitrary: xs mr ys) case (Skip x) then show ?case by (auto split: if_splits prod.splits simp: atms_def elim: ok_mono) next case (Test x) show ?case proof (cases "\x'. x = Formula.Neg x'") case True with Test show ?thesis by (auto split: if_splits prod.splits simp: atms_def elim: ok_mono) next case False with Test show ?thesis by (auto split: if_splits prod.splits simp: atms_def not_Neg_cases elim: ok_mono) qed next case (Plus r1 r2) then show ?case by (fastforce split: if_splits prod.splits formula.splits simp: atms_def elim: ok_mono) next case (Times r1 r2) then show ?case by (fastforce split: if_splits prod.splits formula.splits simp: atms_def elim: ok_mono) next case (Star r) then show ?case by (auto split: if_splits prod.splits simp: atms_def elim: ok_mono) qed lemma ok_shift: "ok (i + m) (Monitor.shift r i) \ ok m r" by (induct r) auto lemma to_mregex_ok: "to_mregex r = (mr, ys) \ set ys = atms r \ ok (length ys) mr" proof (induct r arbitrary: mr ys) case (Skip x) then show ?case by (auto simp: atms_def elim: ok_mono split: if_splits prod.splits) next case (Test x) show ?case proof (cases "\x'. x = Formula.Neg x'") case True with Test show ?thesis by (auto split: if_splits prod.splits simp: atms_def elim: ok_mono) next case False with Test show ?thesis by (auto split: if_splits prod.splits simp: atms_def not_Neg_cases elim: ok_mono) qed next case (Plus r1 r2) then show ?case by (fastforce simp: ok_shift atms_def elim: ok_mono split: if_splits prod.splits) next case (Times r1 r2) then show ?case by (fastforce simp: ok_shift atms_def elim: ok_mono split: if_splits prod.splits) next case (Star r) then show ?case by (auto simp: atms_def elim: ok_mono split: if_splits prod.splits) qed lemma from_mregex_shift: "from_mregex (shift r (length xs)) (xs @ ys) = from_mregex r ys" by (induct r) (auto simp: nth_append) lemma from_mregex_to_mregex: "safe_regex m g r \ case_prod from_mregex (to_mregex r) = r" by (induct r rule: safe_regex.induct) (auto dest: to_mregex_ok intro!: from_mregex_cong simp: nth_append from_mregex_shift cases_Neg_iff split: prod.splits modality.splits) lemma from_mregex_eq: "safe_regex m g r \ to_mregex r = (mr, \s) \ from_mregex mr \s = r" using from_mregex_to_mregex[of m g r] by auto lemma from_mregex_to_mregex_exec: "safe_regex m g r \ case_prod from_mregex (to_mregex_exec r xs) = r" proof (induct r arbitrary: xs rule: safe_regex_induct) case (Plus b g r s) from Plus(3) Plus(1)[of xs] Plus(2)[of "snd (to_mregex_exec r xs)"] show ?case by (auto split: prod.splits simp: nth_append dest: to_mregex_exec_ok intro!: from_mregex_cong[where m = "length (snd (to_mregex_exec r xs))"]) next case (TimesF g r s) from TimesF(3) TimesF(2)[of xs] TimesF(1)[of "snd (to_mregex_exec r xs)"] show ?case by (auto split: prod.splits simp: nth_append dest: to_mregex_exec_ok intro!: from_mregex_cong[where m = "length (snd (to_mregex_exec r xs))"]) next case (TimesP g r s) from TimesP(3) TimesP(1)[of xs] TimesP(2)[of "snd (to_mregex_exec r xs)"] show ?case by (auto split: prod.splits simp: nth_append dest: to_mregex_exec_ok intro!: from_mregex_cong[where m = "length (snd (to_mregex_exec r xs))"]) next case (Star b g r) from Star(2) Star(1)[of xs] show ?case by (auto split: prod.splits) qed (auto split: prod.splits simp: cases_Neg_iff) derive linorder mregex subsubsection \LPD\ definition saturate where "saturate f = while (\S. f S \ S) f" lemma saturate_code[code]: "saturate f S = (let S' = f S in if S' = S then S else saturate f S')" unfolding saturate_def Let_def by (subst while_unfold) auto definition "MTimesL r S = MTimes r ` S" definition "MTimesR R s = (\r. MTimes r s) ` R" primrec LPD where "LPD (MSkip n) = (case n of 0 \ {} | Suc m \ {MSkip m})" | "LPD (MTestPos \) = {}" | "LPD (MTestNeg \) = {}" | "LPD (MPlus r s) = (LPD r \ LPD s)" | "LPD (MTimes r s) = MTimesR (LPD r) s \ LPD s" | "LPD (MStar r) = MTimesR (LPD r) (MStar r)" primrec LPDi where "LPDi 0 r = {r}" | "LPDi (Suc i) r = (\s \ LPD r. LPDi i s)" lemma LPDi_Suc_alt: "LPDi (Suc i) r = (\s \ LPDi i r. LPD s)" by (induct i arbitrary: r) fastforce+ definition "LPDs r = (\i. LPDi i r)" lemma LPDs_refl: "r \ LPDs r" by (auto simp: LPDs_def intro: exI[of _ 0]) lemma LPDs_trans: "r \ LPD s \ s \ LPDs t \ r \ LPDs t" by (force simp: LPDs_def LPDi_Suc_alt simp del: LPDi.simps intro: exI[of _ "Suc _"]) lemma LPDi_Test: "LPDi i (MSkip 0) \ {MSkip 0}" "LPDi i (MTestPos \) \ {MTestPos \}" "LPDi i (MTestNeg \) \ {MTestNeg \}" by (induct i) auto lemma LPDs_Test: "LPDs (MSkip 0) \ {MSkip 0}" "LPDs (MTestPos \) \ {MTestPos \}" "LPDs (MTestNeg \) \ {MTestNeg \}" unfolding LPDs_def using LPDi_Test by blast+ lemma LPDi_MSkip: "LPDi i (MSkip n) \ MSkip ` {i. i \ n}" by (induct i arbitrary: n) (auto dest: set_mp[OF LPDi_Test(1)] simp: le_Suc_eq split: nat.splits) lemma LPDs_MSkip: "LPDs (MSkip n) \ MSkip ` {i. i \ n}" unfolding LPDs_def using LPDi_MSkip by auto lemma LPDi_Plus: "LPDi i (MPlus r s) \ {MPlus r s} \ LPDi i r \ LPDi i s" by (induct i arbitrary: r s) auto lemma LPDs_Plus: "LPDs (MPlus r s) \ {MPlus r s} \ LPDs r \ LPDs s" unfolding LPDs_def using LPDi_Plus[of _ r s] by auto lemma LPDi_Times: "LPDi i (MTimes r s) \ {MTimes r s} \ MTimesR (\j \ i. LPDi j r) s \ (\j \ i. LPDi j s)" proof (induct i arbitrary: r s) case (Suc i) show ?case by (fastforce simp: MTimesR_def dest: bspec[of _ _ "Suc _"] dest!: set_mp[OF Suc]) qed simp lemma LPDs_Times: "LPDs (MTimes r s) \ {MTimes r s} \ MTimesR (LPDs r) s \ LPDs s" unfolding LPDs_def using LPDi_Times[of _ r s] by (force simp: MTimesR_def) lemma LPDi_Star: "j \ i \ LPDi j (MStar r) \ {MStar r} \ MTimesR (\j \ i. LPDi j r) (MStar r)" proof (induct i arbitrary: j r) case (Suc i) from Suc(2) show ?case by (auto 0 5 simp: MTimesR_def image_iff le_Suc_eq dest: bspec[of _ _ "Suc 0"] bspec[of _ _ "Suc _"] set_mp[OF Suc(1)] dest!: set_mp[OF LPDi_Times]) qed simp lemma LPDs_Star: "LPDs (MStar r) \ {MStar r} \ MTimesR (LPDs r) (MStar r)" unfolding LPDs_def using LPDi_Star[OF order_refl, of _ r] by (force simp: MTimesR_def) lemma finite_LPDs: "finite (LPDs r)" proof (induct r) case (MSkip n) then show ?case by (intro finite_subset[OF LPDs_MSkip]) simp next case (MTestPos \) then show ?case by (intro finite_subset[OF LPDs_Test(2)]) simp next case (MTestNeg \) then show ?case by (intro finite_subset[OF LPDs_Test(3)]) simp next case (MPlus r s) then show ?case by (intro finite_subset[OF LPDs_Plus]) simp next case (MTimes r s) then show ?case by (intro finite_subset[OF LPDs_Times]) (simp add: MTimesR_def) next case (MStar r) then show ?case by (intro finite_subset[OF LPDs_Star]) (simp add: MTimesR_def) qed context begin private abbreviation (input) "addLPD r \ \S. insert r S \ Set.bind (insert r S) LPD" private lemma mono_addLPD: "mono (addLPD r)" unfolding mono_def Set.bind_def by auto private lemma LPDs_aux1: "lfp (addLPD r) \ LPDs r" by (rule lfp_induct[OF mono_addLPD], auto intro: LPDs_refl LPDs_trans simp: Set.bind_def) private lemma LPDs_aux2: "LPDi i r \ lfp (addLPD r)" proof (induct i) case 0 then show ?case by (subst lfp_unfold[OF mono_addLPD]) auto next case (Suc i) then show ?case by (subst lfp_unfold[OF mono_addLPD]) (auto simp: LPDi_Suc_alt simp del: LPDi.simps) qed lemma LPDs_alt: "LPDs r = lfp (addLPD r)" using LPDs_aux1 LPDs_aux2 by (force simp: LPDs_def) lemma LPDs_code[code]: "LPDs r = saturate (addLPD r) {}" unfolding LPDs_alt saturate_def by (rule lfp_while[OF mono_addLPD _ finite_LPDs, of r]) (auto simp: LPDs_refl LPDs_trans Set.bind_def) end subsubsection \RPD\ primrec RPD where "RPD (MSkip n) = (case n of 0 \ {} | Suc m \ {MSkip m})" | "RPD (MTestPos \) = {}" | "RPD (MTestNeg \) = {}" | "RPD (MPlus r s) = (RPD r \ RPD s)" | "RPD (MTimes r s) = MTimesL r (RPD s) \ RPD r" | "RPD (MStar r) = MTimesL (MStar r) (RPD r)" primrec RPDi where "RPDi 0 r = {r}" | "RPDi (Suc i) r = (\s \ RPD r. RPDi i s)" lemma RPDi_Suc_alt: "RPDi (Suc i) r = (\s \ RPDi i r. RPD s)" by (induct i arbitrary: r) fastforce+ definition "RPDs r = (\i. RPDi i r)" lemma RPDs_refl: "r \ RPDs r" by (auto simp: RPDs_def intro: exI[of _ 0]) lemma RPDs_trans: "r \ RPD s \ s \ RPDs t \ r \ RPDs t" by (force simp: RPDs_def RPDi_Suc_alt simp del: RPDi.simps intro: exI[of _ "Suc _"]) lemma RPDi_Test: "RPDi i (MSkip 0) \ {MSkip 0}" "RPDi i (MTestPos \) \ {MTestPos \}" "RPDi i (MTestNeg \) \ {MTestNeg \}" by (induct i) auto lemma RPDs_Test: "RPDs (MSkip 0) \ {MSkip 0}" "RPDs (MTestPos \) \ {MTestPos \}" "RPDs (MTestNeg \) \ {MTestNeg \}" unfolding RPDs_def using RPDi_Test by blast+ lemma RPDi_MSkip: "RPDi i (MSkip n) \ MSkip ` {i. i \ n}" by (induct i arbitrary: n) (auto dest: set_mp[OF RPDi_Test(1)] simp: le_Suc_eq split: nat.splits) lemma RPDs_MSkip: "RPDs (MSkip n) \ MSkip ` {i. i \ n}" unfolding RPDs_def using RPDi_MSkip by auto lemma RPDi_Plus: "RPDi i (MPlus r s) \ {MPlus r s} \ RPDi i r \ RPDi i s" by (induct i arbitrary: r s) auto lemma RPDi_Suc_RPD_Plus: "RPDi (Suc i) r \ RPDs (MPlus r s)" "RPDi (Suc i) s \ RPDs (MPlus r s)" unfolding RPDs_def by (force intro!: exI[of _ "Suc i"])+ lemma RPDs_Plus: "RPDs (MPlus r s) \ {MPlus r s} \ RPDs r \ RPDs s" unfolding RPDs_def using RPDi_Plus[of _ r s] by auto lemma RPDi_Times: "RPDi i (MTimes r s) \ {MTimes r s} \ MTimesL r (\j \ i. RPDi j s) \ (\j \ i. RPDi j r)" proof (induct i arbitrary: r s) case (Suc i) show ?case by (fastforce simp: MTimesL_def dest: bspec[of _ _ "Suc _"] dest!: set_mp[OF Suc]) qed simp lemma RPDs_Times: "RPDs (MTimes r s) \ {MTimes r s} \ MTimesL r (RPDs s) \ RPDs r" unfolding RPDs_def using RPDi_Times[of _ r s] by (force simp: MTimesL_def) lemma RPDi_Star: "j \ i \ RPDi j (MStar r) \ {MStar r} \ MTimesL (MStar r) (\j \ i. RPDi j r)" proof (induct i arbitrary: j r) case (Suc i) from Suc(2) show ?case by (auto 0 5 simp: MTimesL_def image_iff le_Suc_eq dest: bspec[of _ _ "Suc 0"] bspec[of _ _ "Suc _"] set_mp[OF Suc(1)] dest!: set_mp[OF RPDi_Times]) qed simp lemma RPDs_Star: "RPDs (MStar r) \ {MStar r} \ MTimesL (MStar r) (RPDs r)" unfolding RPDs_def using RPDi_Star[OF order_refl, of _ r] by (force simp: MTimesL_def) lemma finite_RPDs: "finite (RPDs r)" proof (induct r) case MSkip then show ?case by (intro finite_subset[OF RPDs_MSkip]) simp next case (MTestPos \) then show ?case by (intro finite_subset[OF RPDs_Test(2)]) simp next case (MTestNeg \) then show ?case by (intro finite_subset[OF RPDs_Test(3)]) simp next case (MPlus r s) then show ?case by (intro finite_subset[OF RPDs_Plus]) simp next case (MTimes r s) then show ?case by (intro finite_subset[OF RPDs_Times]) (simp add: MTimesL_def) next case (MStar r) then show ?case by (intro finite_subset[OF RPDs_Star]) (simp add: MTimesL_def) qed context begin private abbreviation (input) "addRPD r \ \S. insert r S \ Set.bind (insert r S) RPD" private lemma mono_addRPD: "mono (addRPD r)" unfolding mono_def Set.bind_def by auto private lemma RPDs_aux1: "lfp (addRPD r) \ RPDs r" by (rule lfp_induct[OF mono_addRPD], auto intro: RPDs_refl RPDs_trans simp: Set.bind_def) private lemma RPDs_aux2: "RPDi i r \ lfp (addRPD r)" proof (induct i) case 0 then show ?case by (subst lfp_unfold[OF mono_addRPD]) auto next case (Suc i) then show ?case by (subst lfp_unfold[OF mono_addRPD]) (auto simp: RPDi_Suc_alt simp del: RPDi.simps) qed lemma RPDs_alt: "RPDs r = lfp (addRPD r)" using RPDs_aux1 RPDs_aux2 by (force simp: RPDs_def) lemma RPDs_code[code]: "RPDs r = saturate (addRPD r) {}" unfolding RPDs_alt saturate_def by (rule lfp_while[OF mono_addRPD _ finite_RPDs, of r]) (auto simp: RPDs_refl RPDs_trans Set.bind_def) end subsection \The executable monitor\ type_synonym ts = nat type_synonym 'a mbuf2 = "'a table list \ 'a table list" type_synonym 'a mbufn = "'a table list list" type_synonym 'a msaux = "(ts \ 'a table) list" type_synonym 'a muaux = "(ts \ 'a table \ 'a table) list" type_synonym 'a mr\aux = "(ts \ (mregex, 'a table) mapping) list" type_synonym 'a ml\aux = "(ts \ 'a table list \ 'a table) list" datatype mconstraint = MEq | MLess | MLessEq record args = args_ivl :: \ args_n :: nat args_L :: "nat set" args_R :: "nat set" args_pos :: bool datatype (dead 'msaux, dead 'muaux) mformula = MRel "event_data table" | MPred Formula.name "Formula.trm list" | MLet Formula.name nat "('msaux, 'muaux) mformula" "('msaux, 'muaux) mformula" | MAnd "nat set" "('msaux, 'muaux) mformula" bool "nat set" "('msaux, 'muaux) mformula" "event_data mbuf2" | MAndAssign "('msaux, 'muaux) mformula" "nat \ Formula.trm" | MAndRel "('msaux, 'muaux) mformula" "Formula.trm \ bool \ mconstraint \ Formula.trm" | MAnds "nat set list" "nat set list" "('msaux, 'muaux) mformula list" "event_data mbufn" | MOr "('msaux, 'muaux) mformula" "('msaux, 'muaux) mformula" "event_data mbuf2" | MNeg "('msaux, 'muaux) mformula" | MExists "('msaux, 'muaux) mformula" | MAgg bool nat Formula.agg_op nat "Formula.trm" "('msaux, 'muaux) mformula" | MPrev \ "('msaux, 'muaux) mformula" bool "event_data table list" "ts list" | MNext \ "('msaux, 'muaux) mformula" bool "ts list" | MSince args "('msaux, 'muaux) mformula" "('msaux, 'muaux) mformula" "event_data mbuf2" "ts list" "'msaux" | MUntil args "('msaux, 'muaux) mformula" "('msaux, 'muaux) mformula" "event_data mbuf2" "ts list" "'muaux" | MMatchP \ "mregex" "mregex list" "('msaux, 'muaux) mformula list" "event_data mbufn" "ts list" "event_data mr\aux" | MMatchF \ "mregex" "mregex list" "('msaux, 'muaux) mformula list" "event_data mbufn" "ts list" "event_data ml\aux" record ('msaux, 'muaux) mstate = mstate_i :: nat mstate_m :: "('msaux, 'muaux) mformula" mstate_n :: nat fun eq_rel :: "nat \ Formula.trm \ Formula.trm \ event_data table" where "eq_rel n (Formula.Const x) (Formula.Const y) = (if x = y then unit_table n else empty_table)" | "eq_rel n (Formula.Var x) (Formula.Const y) = singleton_table n x y" | "eq_rel n (Formula.Const x) (Formula.Var y) = singleton_table n y x" | "eq_rel n _ _ = undefined" lemma regex_atms_size: "x \ regex.atms r \ size x < regex.size_regex size r" by (induct r) auto lemma atms_size: assumes "x \ atms r" shows "size x < Regex.size_regex size r" proof - { fix y assume "y \ regex.atms r" "case y of formula.Neg z \ x = z | _ \ False" then have "size x < Regex.size_regex size r" by (cases y rule: formula.exhaust) (auto dest: regex_atms_size) } with assms show ?thesis unfolding atms_def by (auto split: formula.splits dest: regex_atms_size) qed definition init_args :: "\ \ nat \ nat set \ nat set \ bool \ args" where "init_args I n L R pos = \args_ivl = I, args_n = n, args_L = L, args_R = R, args_pos = pos\" locale msaux = fixes valid_msaux :: "args \ ts \ 'msaux \ event_data msaux \ bool" and init_msaux :: "args \ 'msaux" and add_new_ts_msaux :: "args \ ts \ 'msaux \ 'msaux" and join_msaux :: "args \ event_data table \ 'msaux \ 'msaux" and add_new_table_msaux :: "args \ event_data table \ 'msaux \ 'msaux" and result_msaux :: "args \ 'msaux \ event_data table" assumes valid_init_msaux: "L \ R \ valid_msaux (init_args I n L R pos) 0 (init_msaux (init_args I n L R pos)) []" assumes valid_add_new_ts_msaux: "valid_msaux args cur aux auxlist \ nt \ cur \ valid_msaux args nt (add_new_ts_msaux args nt aux) (filter (\(t, rel). enat (nt - t) \ right (args_ivl args)) auxlist)" assumes valid_join_msaux: "valid_msaux args cur aux auxlist \ table (args_n args) (args_L args) rel1 \ valid_msaux args cur (join_msaux args rel1 aux) (map (\(t, rel). (t, join rel (args_pos args) rel1)) auxlist)" assumes valid_add_new_table_msaux: "valid_msaux args cur aux auxlist \ table (args_n args) (args_R args) rel2 \ valid_msaux args cur (add_new_table_msaux args rel2 aux) (case auxlist of [] => [(cur, rel2)] | ((t, y) # ts) \ if t = cur then (t, y \ rel2) # ts else (cur, rel2) # auxlist)" and valid_result_msaux: "valid_msaux args cur aux auxlist \ result_msaux args aux = foldr (\) [rel. (t, rel) \ auxlist, left (args_ivl args) \ cur - t] {}" fun check_before :: "\ \ ts \ (ts \ 'a \ 'b) \ bool" where "check_before I dt (t, a, b) \ enat t + right I < enat dt" fun proj_thd :: "('a \ 'b \ 'c) \ 'c" where "proj_thd (t, a1, a2) = a2" definition update_until :: "args \ event_data table \ event_data table \ ts \ event_data muaux \ event_data muaux" where "update_until args rel1 rel2 nt aux = (map (\x. case x of (t, a1, a2) \ (t, if (args_pos args) then join a1 True rel1 else a1 \ rel1, if mem (nt - t) (args_ivl args) then a2 \ join rel2 (args_pos args) a1 else a2)) aux) @ [(nt, rel1, if left (args_ivl args) = 0 then rel2 else empty_table)]" lemma map_proj_thd_update_until: "map proj_thd (takeWhile (check_before (args_ivl args) nt) auxlist) = map proj_thd (takeWhile (check_before (args_ivl args) nt) (update_until args rel1 rel2 nt auxlist))" proof (induction auxlist) case Nil then show ?case by (simp add: update_until_def) next case (Cons a auxlist) then show ?case by (cases "right (args_ivl args)") (auto simp add: update_until_def split: if_splits prod.splits) qed fun eval_until :: "\ \ ts \ event_data muaux \ event_data table list \ event_data muaux" where "eval_until I nt [] = ([], [])" | "eval_until I nt ((t, a1, a2) # aux) = (if t + right I < nt then (let (xs, aux) = eval_until I nt aux in (a2 # xs, aux)) else ([], (t, a1, a2) # aux))" lemma eval_until_length: "eval_until I nt auxlist = (res, auxlist') \ length auxlist = length res + length auxlist'" by (induction I nt auxlist arbitrary: res auxlist' rule: eval_until.induct) (auto split: if_splits prod.splits) lemma eval_until_res: "eval_until I nt auxlist = (res, auxlist') \ res = map proj_thd (takeWhile (check_before I nt) auxlist)" by (induction I nt auxlist arbitrary: res auxlist' rule: eval_until.induct) (auto split: prod.splits) lemma eval_until_auxlist': "eval_until I nt auxlist = (res, auxlist') \ auxlist' = drop (length res) auxlist" by (induction I nt auxlist arbitrary: res auxlist' rule: eval_until.induct) (auto split: if_splits prod.splits) locale muaux = fixes valid_muaux :: "args \ ts \ 'muaux \ event_data muaux \ bool" and init_muaux :: "args \ 'muaux" and add_new_muaux :: "args \ event_data table \ event_data table \ ts \ 'muaux \ 'muaux" and length_muaux :: "args \ 'muaux \ nat" and eval_muaux :: "args \ ts \ 'muaux \ event_data table list \ 'muaux" assumes valid_init_muaux: "L \ R \ valid_muaux (init_args I n L R pos) 0 (init_muaux (init_args I n L R pos)) []" assumes valid_add_new_muaux: "valid_muaux args cur aux auxlist \ table (args_n args) (args_L args) rel1 \ table (args_n args) (args_R args) rel2 \ nt \ cur \ valid_muaux args nt (add_new_muaux args rel1 rel2 nt aux) (update_until args rel1 rel2 nt auxlist)" assumes valid_length_muaux: "valid_muaux args cur aux auxlist \ length_muaux args aux = length auxlist" assumes valid_eval_muaux: "valid_muaux args cur aux auxlist \ nt \ cur \ eval_muaux args nt aux = (res, aux') \ eval_until (args_ivl args) nt auxlist = (res', auxlist') \ res = res' \ valid_muaux args cur aux' auxlist'" locale maux = msaux valid_msaux init_msaux add_new_ts_msaux join_msaux add_new_table_msaux result_msaux + muaux valid_muaux init_muaux add_new_muaux length_muaux eval_muaux for valid_msaux :: "args \ ts \ 'msaux \ event_data msaux \ bool" and init_msaux :: "args \ 'msaux" and add_new_ts_msaux :: "args \ ts \ 'msaux \ 'msaux" and join_msaux :: "args \ event_data table \ 'msaux \ 'msaux" and add_new_table_msaux :: "args \ event_data table \ 'msaux \ 'msaux" and result_msaux :: "args \ 'msaux \ event_data table" and valid_muaux :: "args \ ts \ 'muaux \ event_data muaux \ bool" and init_muaux :: "args \ 'muaux" and add_new_muaux :: "args \ event_data table \ event_data table \ ts \ 'muaux \ 'muaux" and length_muaux :: "args \ 'muaux \ nat" and eval_muaux :: "args \ nat \ 'muaux \ event_data table list \ 'muaux" fun split_assignment :: "nat set \ Formula.formula \ nat \ Formula.trm" where "split_assignment X (Formula.Eq t1 t2) = (case (t1, t2) of (Formula.Var x, Formula.Var y) \ if x \ X then (y, t1) else (x, t2) | (Formula.Var x, _) \ (x, t2) | (_, Formula.Var y) \ (y, t1))" | "split_assignment _ _ = undefined" fun split_constraint :: "Formula.formula \ Formula.trm \ bool \ mconstraint \ Formula.trm" where "split_constraint (Formula.Eq t1 t2) = (t1, True, MEq, t2)" | "split_constraint (Formula.Less t1 t2) = (t1, True, MLess, t2)" | "split_constraint (Formula.LessEq t1 t2) = (t1, True, MLessEq, t2)" | "split_constraint (Formula.Neg (Formula.Eq t1 t2)) = (t1, False, MEq, t2)" | "split_constraint (Formula.Neg (Formula.Less t1 t2)) = (t1, False, MLess, t2)" | "split_constraint (Formula.Neg (Formula.LessEq t1 t2)) = (t1, False, MLessEq, t2)" | "split_constraint _ = undefined" function (in maux) (sequential) minit0 :: "nat \ Formula.formula \ ('msaux, 'muaux) mformula" where "minit0 n (Formula.Neg \) = (if fv \ = {} then MNeg (minit0 n \) else MRel empty_table)" | "minit0 n (Formula.Eq t1 t2) = MRel (eq_rel n t1 t2)" | "minit0 n (Formula.Pred e ts) = MPred e ts" | "minit0 n (Formula.Let p \ \) = MLet p (Formula.nfv \) (minit0 (Formula.nfv \) \) (minit0 n \)" | "minit0 n (Formula.Or \ \) = MOr (minit0 n \) (minit0 n \) ([], [])" | "minit0 n (Formula.And \ \) = (if safe_assignment (fv \) \ then MAndAssign (minit0 n \) (split_assignment (fv \) \) else if safe_formula \ then MAnd (fv \) (minit0 n \) True (fv \) (minit0 n \) ([], []) else if is_constraint \ then MAndRel (minit0 n \) (split_constraint \) else (case \ of Formula.Neg \ \ MAnd (fv \) (minit0 n \) False (fv \) (minit0 n \) ([], [])))" | "minit0 n (Formula.Ands l) = (let (pos, neg) = partition safe_formula l in let mpos = map (minit0 n) pos in let mneg = map (minit0 n) (map remove_neg neg) in let vpos = map fv pos in let vneg = map fv neg in MAnds vpos vneg (mpos @ mneg) (replicate (length l) []))" | "minit0 n (Formula.Exists \) = MExists (minit0 (Suc n) \)" | "minit0 n (Formula.Agg y \ b f \) = MAgg (fv \ \ {0.. b f (minit0 (b + n) \)" | "minit0 n (Formula.Prev I \) = MPrev I (minit0 n \) True [] []" | "minit0 n (Formula.Next I \) = MNext I (minit0 n \) True []" | "minit0 n (Formula.Since \ I \) = (if safe_formula \ then MSince (init_args I n (Formula.fv \) (Formula.fv \) True) (minit0 n \) (minit0 n \) ([], []) [] (init_msaux (init_args I n (Formula.fv \) (Formula.fv \) True)) else (case \ of Formula.Neg \ \ MSince (init_args I n (Formula.fv \) (Formula.fv \) False) (minit0 n \) (minit0 n \) ([], []) [] (init_msaux (init_args I n (Formula.fv \) (Formula.fv \) False)) | _ \ undefined))" | "minit0 n (Formula.Until \ I \) = (if safe_formula \ then MUntil (init_args I n (Formula.fv \) (Formula.fv \) True) (minit0 n \) (minit0 n \) ([], []) [] (init_muaux (init_args I n (Formula.fv \) (Formula.fv \) True)) else (case \ of Formula.Neg \ \ MUntil (init_args I n (Formula.fv \) (Formula.fv \) False) (minit0 n \) (minit0 n \) ([], []) [] (init_muaux (init_args I n (Formula.fv \) (Formula.fv \) False)) | _ \ undefined))" | "minit0 n (Formula.MatchP I r) = (let (mr, \s) = to_mregex r in MMatchP I mr (sorted_list_of_set (RPDs mr)) (map (minit0 n) \s) (replicate (length \s) []) [] [])" | "minit0 n (Formula.MatchF I r) = (let (mr, \s) = to_mregex r in MMatchF I mr (sorted_list_of_set (LPDs mr)) (map (minit0 n) \s) (replicate (length \s) []) [] [])" | "minit0 n _ = undefined" by pat_completeness auto termination (in maux) by (relation "measure (\(_, \). size \)") (auto simp: less_Suc_eq_le size_list_estimation' size_remove_neg dest!: to_mregex_ok[OF sym] atms_size) definition (in maux) minit :: "Formula.formula \ ('msaux, 'muaux) mstate" where "minit \ = (let n = Formula.nfv \ in \mstate_i = 0, mstate_m = minit0 n \, mstate_n = n\)" definition (in maux) minit_safe where "minit_safe \ = (if mmonitorable_exec \ then minit \ else undefined)" fun mprev_next :: "\ \ event_data table list \ ts list \ event_data table list \ event_data table list \ ts list" where "mprev_next I [] ts = ([], [], ts)" | "mprev_next I xs [] = ([], xs, [])" | "mprev_next I xs [t] = ([], xs, [t])" | "mprev_next I (x # xs) (t # t' # ts) = (let (ys, zs) = mprev_next I xs (t' # ts) in ((if mem (t' - t) I then x else empty_table) # ys, zs))" fun mbuf2_add :: "event_data table list \ event_data table list \ event_data mbuf2 \ event_data mbuf2" where "mbuf2_add xs' ys' (xs, ys) = (xs @ xs', ys @ ys')" fun mbuf2_take :: "(event_data table \ event_data table \ 'b) \ event_data mbuf2 \ 'b list \ event_data mbuf2" where "mbuf2_take f (x # xs, y # ys) = (let (zs, buf) = mbuf2_take f (xs, ys) in (f x y # zs, buf))" | "mbuf2_take f (xs, ys) = ([], (xs, ys))" fun mbuf2t_take :: "(event_data table \ event_data table \ ts \ 'b \ 'b) \ 'b \ event_data mbuf2 \ ts list \ 'b \ event_data mbuf2 \ ts list" where "mbuf2t_take f z (x # xs, y # ys) (t # ts) = mbuf2t_take f (f x y t z) (xs, ys) ts" | "mbuf2t_take f z (xs, ys) ts = (z, (xs, ys), ts)" lemma size_list_length_diff1: "xs \ [] \ [] \ set xs \ size_list (\xs. length xs - Suc 0) xs < size_list length xs" proof (induct xs) case (Cons x xs) then show ?case by (cases xs) auto qed simp fun mbufn_add :: "event_data table list list \ event_data mbufn \ event_data mbufn" where "mbufn_add xs' xs = List.map2 (@) xs xs'" function mbufn_take :: "(event_data table list \ 'b \ 'b) \ 'b \ event_data mbufn \ 'b \ event_data mbufn" where "mbufn_take f z buf = (if buf = [] \ [] \ set buf then (z, buf) else mbufn_take f (f (map hd buf) z) (map tl buf))" by pat_completeness auto termination by (relation "measure (\(_, _, buf). size_list length buf)") (auto simp: comp_def Suc_le_eq size_list_length_diff1) fun mbufnt_take :: "(event_data table list \ ts \ 'b \ 'b) \ 'b \ event_data mbufn \ ts list \ 'b \ event_data mbufn \ ts list" where "mbufnt_take f z buf ts = (if [] \ set buf \ ts = [] then (z, buf, ts) else mbufnt_take f (f (map hd buf) (hd ts) z) (map tl buf) (tl ts))" fun match :: "Formula.trm list \ event_data list \ (nat \ event_data) option" where "match [] [] = Some Map.empty" | "match (Formula.Const x # ts) (y # ys) = (if x = y then match ts ys else None)" | "match (Formula.Var x # ts) (y # ys) = (case match ts ys of None \ None | Some f \ (case f x of None \ Some (f(x \ y)) | Some z \ if y = z then Some f else None))" | "match _ _ = None" fun meval_trm :: "Formula.trm \ event_data tuple \ event_data" where "meval_trm (Formula.Var x) v = the (v ! x)" | "meval_trm (Formula.Const x) v = x" | "meval_trm (Formula.Plus x y) v = meval_trm x v + meval_trm y v" | "meval_trm (Formula.Minus x y) v = meval_trm x v - meval_trm y v" | "meval_trm (Formula.UMinus x) v = - meval_trm x v" | "meval_trm (Formula.Mult x y) v = meval_trm x v * meval_trm y v" | "meval_trm (Formula.Div x y) v = meval_trm x v div meval_trm y v" | "meval_trm (Formula.Mod x y) v = meval_trm x v mod meval_trm y v" | "meval_trm (Formula.F2i x) v = EInt (integer_of_event_data (meval_trm x v))" | "meval_trm (Formula.I2f x) v = EFloat (double_of_event_data (meval_trm x v))" definition eval_agg :: "nat \ bool \ nat \ Formula.agg_op \ nat \ Formula.trm \ event_data table \ event_data table" where "eval_agg n g0 y \ b f rel = (if g0 \ rel = empty_table then singleton_table n y (eval_agg_op \ {}) else (\k. let group = Set.filter (\x. drop b x = k) rel; M = (\y. (y, ecard (Set.filter (\x. meval_trm f x = y) group))) ` meval_trm f ` group in k[y:=Some (eval_agg_op \ M)]) ` (drop b) ` rel)" definition (in maux) update_since :: "args \ event_data table \ event_data table \ ts \ 'msaux \ event_data table \ 'msaux" where "update_since args rel1 rel2 nt aux = (let aux0 = join_msaux args rel1 (add_new_ts_msaux args nt aux); aux' = add_new_table_msaux args rel2 aux0 in (result_msaux args aux', aux'))" definition "lookup = Mapping.lookup_default empty_table" fun \_lax where "\_lax guard \s (MSkip n) = (if n = 0 then guard else empty_table)" | "\_lax guard \s (MTestPos i) = join guard True (\s ! i)" | "\_lax guard \s (MTestNeg i) = join guard False (\s ! i)" | "\_lax guard \s (MPlus r s) = \_lax guard \s r \ \_lax guard \s s" | "\_lax guard \s (MTimes r s) = join (\_lax guard \s r) True (\_lax guard \s s)" | "\_lax guard \s (MStar r) = guard" fun r\_strict where "r\_strict n \s (MSkip m) = (if m = 0 then unit_table n else empty_table)" | "r\_strict n \s (MTestPos i) = \s ! i" | "r\_strict n \s (MTestNeg i) = (if \s ! i = empty_table then unit_table n else empty_table)" | "r\_strict n \s (MPlus r s) = r\_strict n \s r \ r\_strict n \s s" | "r\_strict n \s (MTimes r s) = \_lax (r\_strict n \s r) \s s" | "r\_strict n \s (MStar r) = unit_table n" fun l\_strict where "l\_strict n \s (MSkip m) = (if m = 0 then unit_table n else empty_table)" | "l\_strict n \s (MTestPos i) = \s ! i" | "l\_strict n \s (MTestNeg i) = (if \s ! i = empty_table then unit_table n else empty_table)" | "l\_strict n \s (MPlus r s) = l\_strict n \s r \ l\_strict n \s s" | "l\_strict n \s (MTimes r s) = \_lax (l\_strict n \s s) \s r" | "l\_strict n \s (MStar r) = unit_table n" fun r\ :: "(mregex \ mregex) \ (mregex, 'a table) mapping \ 'a table list \ mregex \ 'a table" where "r\ \ X \s (MSkip n) = (case n of 0 \ empty_table | Suc m \ lookup X (\ (MSkip m)))" | "r\ \ X \s (MTestPos i) = empty_table" | "r\ \ X \s (MTestNeg i) = empty_table" | "r\ \ X \s (MPlus r s) = r\ \ X \s r \ r\ \ X \s s" | "r\ \ X \s (MTimes r s) = r\ (\t. \ (MTimes r t)) X \s s \ \_lax (r\ \ X \s r) \s s" | "r\ \ X \s (MStar r) = r\ (\t. \ (MTimes (MStar r) t)) X \s r" fun l\ :: "(mregex \ mregex) \ (mregex, 'a table) mapping \ 'a table list \ mregex \ 'a table" where "l\ \ X \s (MSkip n) = (case n of 0 \ empty_table | Suc m \ lookup X (\ (MSkip m)))" | "l\ \ X \s (MTestPos i) = empty_table" | "l\ \ X \s (MTestNeg i) = empty_table" | "l\ \ X \s (MPlus r s) = l\ \ X \s r \ l\ \ X \s s" | "l\ \ X \s (MTimes r s) = l\ (\t. \ (MTimes t s)) X \s r \ \_lax (l\ \ X \s s) \s r" | "l\ \ X \s (MStar r) = l\ (\t. \ (MTimes t (MStar r))) X \s r" lift_definition mrtabulate :: "mregex list \ (mregex \ 'b table) \ (mregex, 'b table) mapping" is "\ks f. (map_of (List.map_filter (\k. let fk = f k in if fk = empty_table then None else Some (k, fk)) ks))" . lemma lookup_tabulate: "distinct xs \ lookup (mrtabulate xs f) x = (if x \ set xs then f x else empty_table)" unfolding lookup_default_def lookup_def by transfer (auto simp: Let_def map_filter_def map_of_eq_None_iff o_def image_image dest!: map_of_SomeD split: if_splits option.splits) definition update_matchP :: "nat \ \ \ mregex \ mregex list \ event_data table list \ ts \ event_data mr\aux \ event_data table \ event_data mr\aux" where "update_matchP n I mr mrs rels nt aux = (let aux = (case [(t, mrtabulate mrs (\mr. r\ id rel rels mr \ (if t = nt then r\_strict n rels mr else {}))). (t, rel) \ aux, enat (nt - t) \ right I] of [] \ [(nt, mrtabulate mrs (r\_strict n rels))] | x # aux' \ (if fst x = nt then x # aux' else (nt, mrtabulate mrs (r\_strict n rels)) # x # aux')) in (foldr (\) [lookup rel mr. (t, rel) \ aux, left I \ nt - t] {}, aux))" definition update_matchF_base where "update_matchF_base n I mr mrs rels nt = (let X = mrtabulate mrs (l\_strict n rels) in ([(nt, rels, if left I = 0 then lookup X mr else empty_table)], X))" definition update_matchF_step where "update_matchF_step I mr mrs nt = (\(t, rels', rel) (aux', X). (let Y = mrtabulate mrs (l\ id X rels') in ((t, rels', if mem (nt - t) I then rel \ lookup Y mr else rel) # aux', Y)))" definition update_matchF :: "nat \ \ \ mregex \ mregex list \ event_data table list \ ts \ event_data ml\aux \ event_data ml\aux" where "update_matchF n I mr mrs rels nt aux = fst (foldr (update_matchF_step I mr mrs nt) aux (update_matchF_base n I mr mrs rels nt))" fun eval_matchF :: "\ \ mregex \ ts \ event_data ml\aux \ event_data table list \ event_data ml\aux" where "eval_matchF I mr nt [] = ([], [])" | "eval_matchF I mr nt ((t, rels, rel) # aux) = (if t + right I < nt then (let (xs, aux) = eval_matchF I mr nt aux in (rel # xs, aux)) else ([], (t, rels, rel) # aux))" primrec map_split where "map_split f [] = ([], [])" | "map_split f (x # xs) = (let (y, z) = f x; (ys, zs) = map_split f xs in (y # ys, z # zs))" fun eval_assignment :: "nat \ Formula.trm \ event_data tuple \ event_data tuple" where "eval_assignment (x, t) y = (y[x:=Some (meval_trm t y)])" fun eval_constraint0 :: "mconstraint \ event_data \ event_data \ bool" where "eval_constraint0 MEq x y = (x = y)" | "eval_constraint0 MLess x y = (x < y)" | "eval_constraint0 MLessEq x y = (x \ y)" fun eval_constraint :: "Formula.trm \ bool \ mconstraint \ Formula.trm \ event_data tuple \ bool" where "eval_constraint (t1, p, c, t2) x = (eval_constraint0 c (meval_trm t1 x) (meval_trm t2 x) = p)" primrec (in maux) meval :: "nat \ ts \ Formula.database \ ('msaux, 'muaux) mformula \ event_data table list \ ('msaux, 'muaux) mformula" where "meval n t db (MRel rel) = ([rel], MRel rel)" | "meval n t db (MPred e ts) = (map (\X. (\f. Table.tabulate f 0 n) ` Option.these (match ts ` X)) (case Mapping.lookup db e of None \ [{}] | Some xs \ xs), MPred e ts)" | "meval n t db (MLet p m \ \) = (let (xs, \) = meval m t db \; (ys, \) = meval n t (Mapping.update p (map (image (map the)) xs) db) \ in (ys, MLet p m \ \))" | "meval n t db (MAnd A_\ \ pos A_\ \ buf) = (let (xs, \) = meval n t db \; (ys, \) = meval n t db \; (zs, buf) = mbuf2_take (\r1 r2. bin_join n A_\ r1 pos A_\ r2) (mbuf2_add xs ys buf) in (zs, MAnd A_\ \ pos A_\ \ buf))" | "meval n t db (MAndAssign \ conf) = (let (xs, \) = meval n t db \ in (map (\r. eval_assignment conf ` r) xs, MAndAssign \ conf))" | "meval n t db (MAndRel \ conf) = (let (xs, \) = meval n t db \ in (map (Set.filter (eval_constraint conf)) xs, MAndRel \ conf))" | "meval n t db (MAnds A_pos A_neg L buf) = (let R = map (meval n t db) L in let buf = mbufn_add (map fst R) buf in let (zs, buf) = mbufn_take (\xs zs. zs @ [mmulti_join n A_pos A_neg xs]) [] buf in (zs, MAnds A_pos A_neg (map snd R) buf))" | "meval n t db (MOr \ \ buf) = (let (xs, \) = meval n t db \; (ys, \) = meval n t db \; (zs, buf) = mbuf2_take (\r1 r2. r1 \ r2) (mbuf2_add xs ys buf) in (zs, MOr \ \ buf))" | "meval n t db (MNeg \) = (let (xs, \) = meval n t db \ in (map (\r. (if r = empty_table then unit_table n else empty_table)) xs, MNeg \))" | "meval n t db (MExists \) = (let (xs, \) = meval (Suc n) t db \ in (map (\r. tl ` r) xs, MExists \))" | "meval n t db (MAgg g0 y \ b f \) = (let (xs, \) = meval (b + n) t db \ in (map (eval_agg n g0 y \ b f) xs, MAgg g0 y \ b f \))" | "meval n t db (MPrev I \ first buf nts) = (let (xs, \) = meval n t db \; (zs, buf, nts) = mprev_next I (buf @ xs) (nts @ [t]) in (if first then empty_table # zs else zs, MPrev I \ False buf nts))" | "meval n t db (MNext I \ first nts) = (let (xs, \) = meval n t db \; (xs, first) = (case (xs, first) of (_ # xs, True) \ (xs, False) | a \ a); (zs, _, nts) = mprev_next I xs (nts @ [t]) in (zs, MNext I \ first nts))" | "meval n t db (MSince args \ \ buf nts aux) = (let (xs, \) = meval n t db \; (ys, \) = meval n t db \; ((zs, aux), buf, nts) = mbuf2t_take (\r1 r2 t (zs, aux). let (z, aux) = update_since args r1 r2 t aux in (zs @ [z], aux)) ([], aux) (mbuf2_add xs ys buf) (nts @ [t]) in (zs, MSince args \ \ buf nts aux))" | "meval n t db (MUntil args \ \ buf nts aux) = (let (xs, \) = meval n t db \; (ys, \) = meval n t db \; (aux, buf, nts) = mbuf2t_take (add_new_muaux args) aux (mbuf2_add xs ys buf) (nts @ [t]); (zs, aux) = eval_muaux args (case nts of [] \ t | nt # _ \ nt) aux in (zs, MUntil args \ \ buf nts aux))" | "meval n t db (MMatchP I mr mrs \s buf nts aux) = (let (xss, \s) = map_split id (map (meval n t db) \s); ((zs, aux), buf, nts) = mbufnt_take (\rels t (zs, aux). let (z, aux) = update_matchP n I mr mrs rels t aux in (zs @ [z], aux)) ([], aux) (mbufn_add xss buf) (nts @ [t]) in (zs, MMatchP I mr mrs \s buf nts aux))" | "meval n t db (MMatchF I mr mrs \s buf nts aux) = (let (xss, \s) = map_split id (map (meval n t db) \s); (aux, buf, nts) = mbufnt_take (update_matchF n I mr mrs) aux (mbufn_add xss buf) (nts @ [t]); (zs, aux) = eval_matchF I mr (case nts of [] \ t | nt # _ \ nt) aux in (zs, MMatchF I mr mrs \s buf nts aux))" definition (in maux) mstep :: "Formula.database \ ts \ ('msaux, 'muaux) mstate \ (nat \ event_data table) list \ ('msaux, 'muaux) mstate" where "mstep tdb st = (let (xs, m) = meval (mstate_n st) (snd tdb) (fst tdb) (mstate_m st) in (List.enumerate (mstate_i st) xs, \mstate_i = mstate_i st + length xs, mstate_m = m, mstate_n = mstate_n st\))" subsection \Verdict delay\ context fixes \ :: Formula.trace begin fun progress :: "(Formula.name \ nat) \ Formula.formula \ nat \ nat" where "progress P (Formula.Pred e ts) j = (case P e of None \ j | Some k \ k)" | "progress P (Formula.Let p \ \) j = progress (P(p \ progress P \ j)) \ j" | "progress P (Formula.Eq t1 t2) j = j" | "progress P (Formula.Less t1 t2) j = j" | "progress P (Formula.LessEq t1 t2) j = j" | "progress P (Formula.Neg \) j = progress P \ j" | "progress P (Formula.Or \ \) j = min (progress P \ j) (progress P \ j)" | "progress P (Formula.And \ \) j = min (progress P \ j) (progress P \ j)" | "progress P (Formula.Ands l) j = (if l = [] then j else Min (set (map (\\. progress P \ j) l)))" | "progress P (Formula.Exists \) j = progress P \ j" | "progress P (Formula.Agg y \ b f \) j = progress P \ j" | "progress P (Formula.Prev I \) j = (if j = 0 then 0 else min (Suc (progress P \ j)) j)" | "progress P (Formula.Next I \) j = progress P \ j - 1" | "progress P (Formula.Since \ I \) j = min (progress P \ j) (progress P \ j)" | "progress P (Formula.Until \ I \) j = Inf {i. \k. k < j \ k \ min (progress P \ j) (progress P \ j) \ \ \ i + right I \ \ \ k}" | "progress P (Formula.MatchP I r) j = min_regex_default (progress P) r j" | "progress P (Formula.MatchF I r) j = Inf {i. \k. k < j \ k \ min_regex_default (progress P) r j \ \ \ i + right I \ \ \ k}" definition "progress_regex P = min_regex_default (progress P)" declare progress.simps[simp del] lemmas progress_simps[simp] = progress.simps[folded progress_regex_def[THEN fun_cong, THEN fun_cong]] end definition "pred_mapping Q = pred_fun (\_. True) (pred_option Q)" definition "rel_mapping Q = rel_fun (=) (rel_option Q)" lemma pred_mapping_alt: "pred_mapping Q P = (\p \ dom P. Q (the (P p)))" unfolding pred_mapping_def pred_fun_def option.pred_set dom_def by (force split: option.splits) lemma rel_mapping_alt: "rel_mapping Q P P' = (dom P = dom P' \ (\p \ dom P. Q (the (P p)) (the (P' p))))" unfolding rel_mapping_def rel_fun_def rel_option_iff dom_def by (force split: option.splits) lemma rel_mapping_map_upd[simp]: "Q x y \ rel_mapping Q P P' \ rel_mapping Q (P(p \ x)) (P'(p \ y))" by (auto simp: rel_mapping_alt) lemma pred_mapping_map_upd[simp]: "Q x \ pred_mapping Q P \ pred_mapping Q (P(p \ x))" by (auto simp: pred_mapping_alt) lemma pred_mapping_empty[simp]: "pred_mapping Q Map.empty" by (auto simp: pred_mapping_alt) lemma pred_mapping_mono: "pred_mapping Q P \ Q \ R \ pred_mapping R P" by (auto simp: pred_mapping_alt) lemma pred_mapping_mono_strong: "pred_mapping Q P \ (\p. p \ dom P \ Q (the (P p)) \ R (the (P p))) \ pred_mapping R P" by (auto simp: pred_mapping_alt) lemma progress_mono_gen: "j \ j' \ rel_mapping (\) P P' \ progress \ P \ j \ progress \ P' \ j'" proof (induction \ arbitrary: P P') case (Pred e ts) then show ?case by (force simp: rel_mapping_alt dom_def split: option.splits) next case (Ands l) then show ?case by (auto simp: image_iff intro!: Min.coboundedI[THEN order_trans]) next case (Until \ I \) from Until(1,2)[of P P'] Until.prems show ?case by (cases "right I") (auto dest: trans_le_add1[OF \_mono] intro!: cInf_superset_mono) next case (MatchF I r) from MatchF(1)[of _ P P'] MatchF.prems show ?case by (cases "right I"; cases "regex.atms r = {}") (auto 0 3 simp: Min_le_iff progress_regex_def dest: trans_le_add1[OF \_mono] intro!: cInf_superset_mono elim!: less_le_trans order_trans) qed (force simp: Min_le_iff progress_regex_def split: option.splits)+ lemma rel_mapping_reflp: "reflp Q \ rel_mapping Q P P" by (auto simp: rel_mapping_alt reflp_def) lemmas progress_mono = progress_mono_gen[OF _ rel_mapping_reflp[unfolded reflp_def], simplified] lemma progress_le_gen: "pred_mapping (\x. x \ j) P \ progress \ P \ j \ j" proof (induction \ arbitrary: P) case (Pred e ts) then show ?case by (auto simp: pred_mapping_alt dom_def split: option.splits) next case (Ands l) then show ?case by (auto simp: image_iff intro!: Min.coboundedI[where a="progress \ P (hd l) j", THEN order_trans]) next case (Until \ I \) then show ?case by (cases "right I") (auto intro: trans_le_add1[OF \_mono] intro!: cInf_lower) next case (MatchF I r) then show ?case by (cases "right I") (auto intro: trans_le_add1[OF \_mono] intro!: cInf_lower) qed (force simp: Min_le_iff progress_regex_def split: option.splits)+ lemma progress_le: "progress \ Map.empty \ j \ j" using progress_le_gen[of _ Map.empty] by auto lemma progress_0_gen[simp]: "pred_mapping (\x. x = 0) P \ progress \ P \ 0 = 0" using progress_le_gen[of 0 P] by auto lemma progress_0[simp]: "progress \ Map.empty \ 0 = 0" by (auto simp: pred_mapping_alt) definition max_mapping :: "('b \ 'a option) \ ('b \ 'a option) \ 'b \ ('a :: linorder) option" where "max_mapping P P' x = (case (P x, P' x) of (None, None) \ None | (Some x, None) \ None | (None, Some x) \ None | (Some x, Some y) \ Some (max x y))" definition Max_mapping :: "('b \ 'a option) set \ 'b \ ('a :: linorder) option" where "Max_mapping Ps x = (if (\P \ Ps. P x \ None) then Some (Max ((\P. the (P x)) ` Ps)) else None)" lemma dom_max_mapping[simp]: "dom (max_mapping P1 P2) = dom P1 \ dom P2" unfolding max_mapping_def by (auto split: option.splits) lemma dom_Max_mapping[simp]: "dom (Max_mapping X) = (\P \ X. dom P)" unfolding Max_mapping_def by (auto split: if_splits) lemma Max_mapping_coboundedI: assumes "finite X" "\Q \ X. dom Q = dom P" "P \ X" shows "rel_mapping (\) P (Max_mapping X)" unfolding rel_mapping_alt proof (intro conjI ballI) from assms(3) have "X \ {}" by auto then show "dom P = dom (Max_mapping X)" using assms(2) by auto next fix p assume "p \ dom P" with assms show "the (P p) \ the (Max_mapping X p)" by (force simp add: Max_mapping_def intro!: Max.coboundedI imageI) qed lemma rel_mapping_trans: "P OO Q \ R \ rel_mapping P P1 P2 \ rel_mapping Q P2 P3 \ rel_mapping R P1 P3" by (force simp: rel_mapping_alt dom_def set_eq_iff) abbreviation range_mapping :: "nat \ nat \ ('b \ nat option) \ bool" where "range_mapping i j P \ pred_mapping (\x. i \ x \ x \ j) P" lemma range_mapping_relax: "range_mapping i j P \ i' \ i \ j' \ j \ range_mapping i' j' P" by (auto simp: pred_mapping_alt dom_def set_eq_iff max_mapping_def split: option.splits) lemma range_mapping_max_mapping[simp]: "range_mapping i j1 P1 \ range_mapping i j2 P2 \ range_mapping i (max j1 j2) (max_mapping P1 P2)" by (auto simp: pred_mapping_alt dom_def set_eq_iff max_mapping_def split: option.splits) lemma range_mapping_Max_mapping[simp]: "finite X \ X \ {} \ \x\X. range_mapping i (j x) (P x) \ range_mapping i (Max (j ` X)) (Max_mapping (P ` X))" by (force simp: pred_mapping_alt Max_mapping_def dom_def image_iff intro!: Max_ge_iff[THEN iffD2] split: if_splits) lemma pred_mapping_le: "pred_mapping ((\) i) P1 \ rel_mapping (\) P1 P2 \ pred_mapping ((\) (i :: nat)) P2" by (force simp: rel_mapping_alt pred_mapping_alt dom_def set_eq_iff) lemma pred_mapping_le': "pred_mapping ((\) j) P1 \ i \ j \ rel_mapping (\) P1 P2 \ pred_mapping ((\) (i :: nat)) P2" by (force simp: rel_mapping_alt pred_mapping_alt dom_def set_eq_iff) lemma max_mapping_cobounded1: "dom P1 \ dom P2 \ rel_mapping (\) P1 (max_mapping P1 P2)" unfolding max_mapping_def rel_mapping_alt by (auto simp: dom_def split: option.splits) lemma max_mapping_cobounded2: "dom P2 \ dom P1 \ rel_mapping (\) P2 (max_mapping P1 P2)" unfolding max_mapping_def rel_mapping_alt by (auto simp: dom_def split: option.splits) lemma max_mapping_fun_upd2[simp]: - "max_mapping P1 (P2(p := y))(p \ x) = (max_mapping P1 P2)(p \ x)" + "(max_mapping P1 (P2(p := y)))(p \ x) = (max_mapping P1 P2)(p \ x)" by (auto simp: max_mapping_def) lemma rel_mapping_max_mapping_fun_upd: "dom P2 \ dom P1 \ p \ dom P2 \ the (P2 p) \ y \ - rel_mapping (\) P2 (max_mapping P1 P2(p \ y))" + rel_mapping (\) P2 ((max_mapping P1 P2)(p \ y))" by (auto simp: rel_mapping_alt max_mapping_def split: option.splits) lemma progress_ge_gen: "Formula.future_bounded \ \ \P j. dom P = S \ range_mapping i j P \ i \ progress \ P \ j" proof (induction \ arbitrary: i S) case (Pred e ts) then show ?case by (intro exI[of _ "\e. if e \ S then Some i else None"]) (auto split: option.splits if_splits simp: rel_mapping_alt pred_mapping_alt dom_def) next case (Let p \ \) from Let.prems obtain P2 j2 where P2: "dom P2 = insert p S" "range_mapping i j2 P2" "i \ progress \ P2 \ j2" by (atomize_elim, intro Let(2)) (force simp: pred_mapping_alt rel_mapping_alt dom_def)+ from Let.prems obtain P1 j1 where P1: "dom P1 = S" "range_mapping (the (P2 p)) j1 P1" "the (P2 p) \ progress \ P1 \ j1" by (atomize_elim, intro Let(1)) auto let ?P12 = "max_mapping P1 P2" from P1 P2 have le1: "progress \ P1 \ j1 \ progress \ (?P12(p := P1 p)) \ (max j1 j2)" by (intro progress_mono_gen) (auto simp: rel_mapping_alt max_mapping_def) from P1 P2 have le2: "progress \ P2 \ j2 \ progress \ (?P12(p \ progress \ P1 \ j1)) \ (max j1 j2)" by (intro progress_mono_gen) (auto simp: rel_mapping_alt max_mapping_def) show ?case unfolding progress.simps proof (intro exI[of _ "?P12(p := P1 p)"] exI[of _ "max j1 j2"] conjI) show "dom (?P12(p := P1 p)) = S" using P1 P2 by (auto simp: dom_def max_mapping_def) next show "range_mapping i (max j1 j2) (?P12(p := P1 p))" using P1 P2 by (force simp add: pred_mapping_alt dom_def max_mapping_def split: option.splits) next have "i \ progress \ P2 \ j2" by fact also have "... \ progress \ (?P12(p \ progress \ P1 \ j1)) \ (max j1 j2)" using le2 by blast - also have "... \ progress \ (?P12(p := P1 p)(p\progress \ (?P12(p := P1 p)) \ (max j1 j2))) \ (max j1 j2)" + also have "... \ progress \ ((?P12(p := P1 p))(p\progress \ (?P12(p := P1 p)) \ (max j1 j2))) \ (max j1 j2)" by (auto intro!: progress_mono_gen simp: le1 rel_mapping_alt) finally show "i \ ..." . qed next case (Eq _ _) then show ?case by (intro exI[of _ "\e. if e \ S then Some i else None"]) (auto split: if_splits simp: pred_mapping_alt) next case (Less _ _) then show ?case by (intro exI[of _ "\e. if e \ S then Some i else None"]) (auto split: if_splits simp: pred_mapping_alt) next case (LessEq _ _) then show ?case by (intro exI[of _ "\e. if e \ S then Some i else None"]) (auto split: if_splits simp: pred_mapping_alt) next case (Or \1 \2) from Or(3) obtain P1 j1 where P1: "dom P1 = S" "range_mapping i j1 P1" "i \ progress \ P1 \1 j1" using Or(1)[of S i] by auto moreover from Or(3) obtain P2 j2 where P2: "dom P2 = S" "range_mapping i j2 P2" "i \ progress \ P2 \2 j2" using Or(2)[of S i] by auto ultimately have "i \ progress \ (max_mapping P1 P2) (Formula.Or \1 \2) (max j1 j2)" by (auto 0 3 elim!: order.trans[OF _ progress_mono_gen] intro: max_mapping_cobounded1 max_mapping_cobounded2) with P1 P2 show ?case by (intro exI[of _ "max_mapping P1 P2"] exI[of _ "max j1 j2"]) auto next case (And \1 \2) from And(3) obtain P1 j1 where P1: "dom P1 = S" "range_mapping i j1 P1" "i \ progress \ P1 \1 j1" using And(1)[of S i] by auto moreover from And(3) obtain P2 j2 where P2: "dom P2 = S" "range_mapping i j2 P2" "i \ progress \ P2 \2 j2" using And(2)[of S i] by auto ultimately have "i \ progress \ (max_mapping P1 P2) (Formula.And \1 \2) (max j1 j2)" by (auto 0 3 elim!: order.trans[OF _ progress_mono_gen] intro: max_mapping_cobounded1 max_mapping_cobounded2) with P1 P2 show ?case by (intro exI[of _ "max_mapping P1 P2"] exI[of _ "max j1 j2"]) auto next case (Ands l) show ?case proof (cases "l = []") case True then show ?thesis by (intro exI[of _ "\e. if e \ S then Some i else None"]) (auto split: if_splits simp: pred_mapping_alt) next case False then obtain \ where "\ \ set l" by (cases l) auto from Ands.prems have "\\\set l. Formula.future_bounded \" by (simp add: list.pred_set) { fix \ assume "\ \ set l" with Ands.prems obtain P j where "dom P = S" "range_mapping i j P" "i \ progress \ P \ j" by (atomize_elim, intro Ands(1)[of \ S i]) (auto simp: list.pred_set) then have "\Pj. dom (fst Pj) = S \ range_mapping i (snd Pj) (fst Pj) \ i \ progress \ (fst Pj) \ (snd Pj)" (is "\Pj. ?P Pj") by (intro exI[of _ "(P, j)"]) auto } then have "\\\set l. \Pj. dom (fst Pj) = S \ range_mapping i (snd Pj) (fst Pj) \ i \ progress \ (fst Pj) \ (snd Pj)" (is "\\\set l. \Pj. ?P Pj \") by blast with Ands(1) Ands.prems False have "\Pjf. \\\set l. ?P (Pjf \) \" by (auto simp: Ball_def intro: choice) then obtain Pjf where Pjf: "\\\set l. ?P (Pjf \) \" .. define Pf where "Pf = fst o Pjf" define jf where "jf = snd o Pjf" have *: "dom (Pf \) = S" "range_mapping i (jf \) (Pf \)" "i \ progress \ (Pf \) \ (jf \)" if "\ \ set l" for \ using Pjf[THEN bspec, OF that] unfolding Pf_def jf_def by auto with False show ?thesis unfolding progress.simps eq_False[THEN iffD2, OF False] if_False by ((subst Min_ge_iff; simp add: False), intro exI[where x="MAX \\set l. jf \"] exI[where x="Max_mapping (Pf ` set l)"] conjI ballI order.trans[OF *(3) progress_mono_gen] Max_mapping_coboundedI) (auto simp: False *[OF \\ \ set l\] \\ \ set l\) qed next case (Exists \) then show ?case by simp next case (Prev I \) then obtain P j where "dom P = S" "range_mapping i j P" "i \ progress \ P \ j" by (atomize_elim, intro Prev(1)) (auto simp: pred_mapping_alt dom_def) with Prev(2) have "dom P = S \ range_mapping i (max i j) P \ i \ progress \ P (formula.Prev I \) (max i j)" by (auto simp: le_Suc_eq max_def pred_mapping_alt split: if_splits elim: order.trans[OF _ progress_mono]) then show ?case by blast next case (Next I \) then obtain P j where "dom P = S" "range_mapping (Suc i) j P" "Suc i \ progress \ P \ j" by (atomize_elim, intro Next(1)) (auto simp: pred_mapping_alt dom_def) then show ?case by (intro exI[of _ P] exI[of _ j]) (auto 0 3 simp: pred_mapping_alt dom_def) next case (Since \1 I \2) from Since(3) obtain P1 j1 where P1: "dom P1 = S" "range_mapping i j1 P1" "i \ progress \ P1 \1 j1" using Since(1)[of S i] by auto moreover from Since(3) obtain P2 j2 where P2: "dom P2 = S" "range_mapping i j2 P2" "i \ progress \ P2 \2 j2" using Since(2)[of S i] by auto ultimately have "i \ progress \ (max_mapping P1 P2) (Formula.Since \1 I \2) (max j1 j2)" by (auto elim!: order.trans[OF _ progress_mono_gen] simp: max_mapping_cobounded1 max_mapping_cobounded2) with P1 P2 show ?case by (intro exI[of _ "max_mapping P1 P2"] exI[of _ "max j1 j2"]) (auto elim!: pred_mapping_le intro: max_mapping_cobounded1) next case (Until \1 I \2) from Until.prems obtain b where [simp]: "right I = enat b" by (cases "right I") (auto) obtain i' where "i < i'" and "\ \ i + b + 1 \ \ \ i'" using ex_le_\[where x="\ \ i + b + 1"] by (auto simp add: less_eq_Suc_le) then have 1: "\ \ i + b < \ \ i'" by simp from Until.prems obtain P1 j1 where P1: "dom P1 = S" "range_mapping (Suc i') j1 P1" "Suc i' \ progress \ P1 \1 j1" by (atomize_elim, intro Until(1)) (auto simp: pred_mapping_alt dom_def) from Until.prems obtain P2 j2 where P2: "dom P2 = S" "range_mapping (Suc i') j2 P2" "Suc i' \ progress \ P2 \2 j2" by (atomize_elim, intro Until(2)) (auto simp: pred_mapping_alt dom_def) let ?P12 = "max_mapping P1 P2" have "i \ progress \ ?P12 (Formula.Until \1 I \2) (max j1 j2)" unfolding progress.simps proof (intro cInf_greatest, goal_cases nonempty greatest) case nonempty then show ?case by (auto simp: trans_le_add1[OF \_mono] intro!: exI[of _ "max j1 j2"]) next case (greatest x) from P1(2,3) have "i' < j1" by (auto simp: less_eq_Suc_le intro!: progress_le_gen elim!: order.trans pred_mapping_mono_strong) then have "i' < max j1 j2" by simp have "progress \ P1 \1 j1 \ progress \ ?P12 \1 (max j1 j2)" using P1(1) P2(1) by (auto intro!: progress_mono_gen max_mapping_cobounded1) moreover have "progress \ P2 \2 j2 \ progress \ ?P12 \2 (max j1 j2)" using P1(1) P2(1) by (auto intro!: progress_mono_gen max_mapping_cobounded2) ultimately have "i' \ min (progress \ ?P12 \1 (max j1 j2)) (progress \ ?P12 \2 (max j1 j2))" using P1(3) P2(3) by simp with greatest \i' < max j1 j2\ have "\ \ i' \ \ \ x + b" by simp with 1 have "\ \ i < \ \ x" by simp then show ?case by (auto dest!: less_\D) qed with P1 P2 \i < i'\ show ?case by (intro exI[of _ "max_mapping P1 P2"] exI[of _ "max j1 j2"]) (auto simp: range_mapping_relax) next case (MatchP I r) then show ?case proof (cases "regex.atms r = {}") case True with MatchP.prems show ?thesis unfolding progress.simps by (intro exI[of _ "\e. if e \ S then Some i else None"] exI[of _ i]) (auto split: if_splits simp: pred_mapping_alt regex.pred_set) next case False define pick where "pick = (\\. SOME Pj. dom (fst Pj) = S \ range_mapping i (snd Pj) (fst Pj) \ i \ progress \ (fst Pj) \ (snd Pj))" let ?pickP = "fst o pick" let ?pickj = "snd o pick" from MatchP have pick: "\ \ regex.atms r \ dom (?pickP \) = S \ range_mapping i (?pickj \) (?pickP \) \ i \ progress \ (?pickP \) \ (?pickj \)" for \ unfolding pick_def o_def future_bounded.simps regex.pred_set by (intro someI_ex[where P = "\Pj. dom (fst Pj) = S \ range_mapping i (snd Pj) (fst Pj) \ i \ progress \ (fst Pj) \ (snd Pj)"]) auto with False show ?thesis unfolding progress.simps by (intro exI[of _ "Max_mapping (?pickP ` regex.atms r)"] exI[of _ "Max (?pickj ` regex.atms r)"]) (auto simp: Max_mapping_coboundedI order_trans[OF pick[THEN conjunct2, THEN conjunct2] progress_mono_gen]) qed next case (MatchF I r) from MatchF.prems obtain b where [simp]: "right I = enat b" by auto obtain i' where i': "i < i'" "\ \ i + b + 1 \ \ \ i'" using ex_le_\[where x="\ \ i + b + 1"] by (auto simp add: less_eq_Suc_le) then have 1: "\ \ i + b < \ \ i'" by simp have ix: "i \ x" if "\ \ i' \ b + \ \ x" "b + \ \ i < \ \ i'" for x using less_\D[of \ i] that less_le_trans by fastforce show ?case proof (cases "regex.atms r = {}") case True with MatchF.prems i' ix show ?thesis unfolding progress.simps by (intro exI[of _ "\e. if e \ S then Some (Suc i') else None"] exI[of _ "Suc i'"]) (auto split: if_splits simp: pred_mapping_alt regex.pred_set add.commute less_Suc_eq intro!: cInf_greatest dest!: spec[of _ i'] less_imp_le[THEN \_mono, of _ i' \]) next case False then obtain \ where \: "\ \ regex.atms r" by auto define pick where "pick = (\\. SOME Pj. dom (fst Pj) = S \ range_mapping (Suc i') (snd Pj) (fst Pj) \ Suc i' \ progress \ (fst Pj) \ (snd Pj))" define pickP where "pickP = fst o pick" define pickj where "pickj = snd o pick" from MatchF have pick: "\ \ regex.atms r \ dom (pickP \) = S \ range_mapping (Suc i') (pickj \) (pickP \) \ Suc i' \ progress \ (pickP \) \ (pickj \)" for \ unfolding pick_def o_def future_bounded.simps regex.pred_set pickj_def pickP_def by (intro someI_ex[where P = "\Pj. dom (fst Pj) = S \ range_mapping (Suc i') (snd Pj) (fst Pj) \ Suc i' \ progress \ (fst Pj) \ (snd Pj)"]) auto let ?P = "Max_mapping (pickP ` regex.atms r)" let ?j = "Max (pickj ` regex.atms r)" from pick[OF \] False \ have "Suc i' \ ?j" by (intro order_trans[OF pick[THEN conjunct2, THEN conjunct2], OF \] order_trans[OF progress_le_gen]) (auto simp: Max_ge_iff dest: range_mapping_relax[of _ _ _ 0, OF _ _ order_refl, simplified]) moreover note i' 1 ix moreover from MatchF.prems have "Regex.pred_regex Formula.future_bounded r" by auto ultimately show ?thesis using \_mono[of _ ?j \] less_\D[of \ i] pick False by (intro exI[of _ "?j"] exI[of _ "?P"]) (auto 0 3 intro!: cInf_greatest order_trans[OF le_SucI[OF order_refl] order_trans[OF pick[THEN conjunct2, THEN conjunct2] progress_mono_gen]] range_mapping_Max_mapping[OF _ _ ballI[OF range_mapping_relax[of "Suc i'" _ _ i, OF _ _ order_refl]]] simp: ac_simps Suc_le_eq trans_le_add2 Max_mapping_coboundedI progress_regex_def dest: spec[of _ "i'"] spec[of _ ?j]) qed qed (auto split: option.splits) lemma progress_ge: "Formula.future_bounded \ \ \j. i \ progress \ Map.empty \ j" using progress_ge_gen[of \ "{}" i \] by auto lemma cInf_restrict_nat: fixes x :: nat assumes "x \ A" shows "Inf A = Inf {y \ A. y \ x}" using assms by (auto intro!: antisym intro: cInf_greatest cInf_lower Inf_nat_def1) lemma progress_time_conv: assumes "\i \ i = \ \' i" shows "progress \ P \ j = progress \' P \ j" using assms proof (induction \ arbitrary: P) case (Until \1 I \2) have *: "i \ j - 1 \ i < j" if "j \ 0" for i using that by auto with Until show ?case proof (cases "right I") case (enat b) then show ?thesis proof (cases "j") case (Suc n) with enat * Until show ?thesis using \_mono[THEN trans_le_add1] by (auto 8 0 intro!: box_equals[OF arg_cong[where f=Inf] cInf_restrict_nat[symmetric, where x=n] cInf_restrict_nat[symmetric, where x=n]]) qed simp qed simp next case (MatchF I r) have *: "i \ j - 1 \ i < j" if "j \ 0" for i using that by auto with MatchF show ?case using \_mono[THEN trans_le_add1] by (cases "right I"; cases j) ((auto 6 0 simp: progress_le_gen progress_regex_def intro!: box_equals[OF arg_cong[where f=Inf] cInf_restrict_nat[symmetric, where x="j-1"] cInf_restrict_nat[symmetric, where x="j-1"]]) [])+ qed (auto simp: progress_regex_def) lemma Inf_UNIV_nat: "(Inf UNIV :: nat) = 0" by (simp add: cInf_eq_minimum) lemma progress_prefix_conv: assumes "prefix_of \ \" and "prefix_of \ \'" shows "progress \ P \ (plen \) = progress \' P \ (plen \)" using assms by (auto intro: progress_time_conv \_prefix_conv) lemma bounded_rtranclp_mono: fixes n :: "'x :: linorder" assumes "\i j. R i j \ j < n \ S i j" "\i j. R i j \ i \ j" shows "rtranclp R i j \ j < n \ rtranclp S i j" proof (induct rule: rtranclp_induct) case (step y z) then show ?case using assms(1,2)[of y z] by (auto elim!: rtrancl_into_rtrancl[to_pred, rotated]) qed auto lemma sat_prefix_conv_gen: assumes "prefix_of \ \" and "prefix_of \ \'" shows "i < progress \ P \ (plen \) \ dom V = dom V' \ dom P = dom V \ pred_mapping (\x. x \ plen \) P \ (\p i \. p \ dom V \ i < the (P p) \ the (V p) i = the (V' p) i) \ Formula.sat \ V v i \ \ Formula.sat \' V' v i \" proof (induction \ arbitrary: P V V' v i) case (Pred e ts) from Pred.prems(1,4) have "i < plen \" by (blast intro!: order.strict_trans2 progress_le_gen) show ?case proof (cases "V e") case None then have "V' e = None" using \dom V = dom V'\ by auto with None \_prefix_conv[OF assms(1,2) \i < plen \\] show ?thesis by simp next case (Some a) obtain a' where "V' e = Some a'" using Some \dom V = dom V'\ by auto then have "i < the (P e)" using Pred.prems(1-3) by (auto split: option.splits) then have "the (V e) i = the (V' e) i" using Some by (intro Pred.prems(5)) (simp_all add: domI) with Some \V' e = Some a'\ show ?thesis by simp qed next case (Let p \ \) let ?V = "\V \. (V(p \ \i. {v. length v = Formula.nfv \ \ Formula.sat \ V v i \}))" show ?case unfolding sat.simps proof (rule Let.IH(2)) from Let.prems show "i < progress \ (P(p \ progress \ P \ (plen \))) \ (plen \)" by simp from Let.prems show "dom (?V V \) = dom (?V V' \')" by simp from Let.prems show "dom (P(p \ progress \ P \ (plen \))) = dom (?V V \)" by simp from Let.prems show "pred_mapping (\x. x \ plen \) (P(p \ progress \ P \ (plen \)))" by (auto intro!: pred_mapping_map_upd elim!: progress_le_gen) next fix p' i \' assume 1: "p' \ dom (?V V \)" and 2: "i < the ((P(p \ progress \ P \ (plen \))) p')" show "the (?V V \ p') i = the (?V V' \' p') i" proof (cases "p' = p") case True with Let 2 show ?thesis by auto next case False with 1 2 show ?thesis by (auto intro!: Let.prems(5)) qed qed next case (Eq t1 t2) show ?case by simp next case (Neg \) then show ?case by simp next case (Or \1 \2) then show ?case by auto next case (Ands l) from Ands.prems have "\\\set l. i < progress \ P \ (plen \)" by (cases l) simp_all with Ands show ?case unfolding sat_Ands by blast next case (Exists \) then show ?case by simp next case (Prev I \) with \_prefix_conv[OF assms(1,2)] show ?case by (cases i) (auto split: if_splits) next case (Next I \) then have "Suc i < plen \" by (auto intro: order.strict_trans2[OF _ progress_le_gen[of _ P \ \]]) with Next.prems \_prefix_conv[OF assms(1,2)] show ?case unfolding sat.simps by (intro conj_cong Next) auto next case (Since \1 I \2) then have "i < plen \" by (auto elim!: order.strict_trans2[OF _ progress_le_gen]) with Since.prems \_prefix_conv[OF assms(1,2)] show ?case unfolding sat.simps by (intro conj_cong ex_cong ball_cong Since) auto next case (Until \1 I \2) from Until.prems obtain b where right[simp]: "right I = enat b" by (cases "right I") (auto simp add: Inf_UNIV_nat) from Until.prems obtain j where "\ \ i + b + 1 \ \ \ j" "j \ progress \ P \1 (plen \)" "j \ progress \ P \2 (plen \)" by atomize_elim (auto 0 4 simp add: less_eq_Suc_le not_le intro: Suc_leI dest: spec[of _ "i"] dest!: le_cInf_iff[THEN iffD1, rotated -1]) then have 1: "k < progress \ P \1 (plen \)" and 2: "k < progress \ P \2 (plen \)" if "\ \ k \ \ \ i + b" for k using that by (fastforce elim!: order.strict_trans2[rotated] intro: less_\D[of \])+ have 3: "k < plen \" if "\ \ k \ \ \ i + b" for k using 1[OF that] Until(6) by (auto simp only: less_eq_Suc_le order.trans[OF _ progress_le_gen]) from Until.prems have "i < progress \' P (Formula.Until \1 I \2) (plen \)" unfolding progress_prefix_conv[OF assms(1,2)] by blast then obtain j where "\ \' i + b + 1 \ \ \' j" "j \ progress \' P \1 (plen \)" "j \ progress \' P \2 (plen \)" by atomize_elim (auto 0 4 simp add: less_eq_Suc_le not_le intro: Suc_leI dest: spec[of _ "i"] dest!: le_cInf_iff[THEN iffD1, rotated -1]) then have 11: "k < progress \ P \1 (plen \)" and 21: "k < progress \ P \2 (plen \)" if "\ \' k \ \ \' i + b" for k unfolding progress_prefix_conv[OF assms(1,2)] using that by (fastforce elim!: order.strict_trans2[rotated] intro: less_\D[of \'])+ have 31: "k < plen \" if "\ \' k \ \ \' i + b" for k using 11[OF that] Until(6) by (auto simp only: less_eq_Suc_le order.trans[OF _ progress_le_gen]) show ?case unfolding sat.simps proof ((intro ex_cong iffI; elim conjE), goal_cases LR RL) case (LR j) with Until(1)[OF 1] Until(2)[OF 2] \_prefix_conv[OF assms(1,2) 3] Until.prems show ?case by (auto 0 4 simp: le_diff_conv add.commute dest: less_imp_le order.trans[OF \_mono, rotated]) next case (RL j) with Until(1)[OF 11] Until(2)[OF 21] \_prefix_conv[OF assms(1,2) 31] Until.prems show ?case by (auto 0 4 simp: le_diff_conv add.commute dest: less_imp_le order.trans[OF \_mono, rotated]) qed next case (MatchP I r) then have "i < plen \" by (force simp: pred_mapping_alt elim!: order.strict_trans2[OF _ progress_le_gen]) with MatchP.prems \_prefix_conv[OF assms(1,2)] show ?case unfolding sat.simps by (intro ex_cong conj_cong match_cong_strong MatchP) (auto simp: progress_regex_def split: if_splits) next case (MatchF I r) from MatchF.prems obtain b where right[simp]: "right I = enat b" by (cases "right I") (auto simp add: Inf_UNIV_nat) show ?case proof (cases "regex.atms r = {}") case True from MatchF.prems(1) obtain j where "\ \ i + b + 1 \ \ \ j" "j \ plen \" by atomize_elim (auto 0 4 simp add: less_eq_Suc_le not_le dest!: le_cInf_iff[THEN iffD1, rotated -1]) then have 1: "k < plen \" if "\ \ k \ \ \ i + b" for k by (meson \_mono discrete not_le order.strict_trans2 that) from MatchF.prems have "i < progress \' P (Formula.MatchF I r) (plen \)" unfolding progress_prefix_conv[OF assms(1,2)] by blast then obtain j where "\ \' i + b + 1 \ \ \' j" "j \ plen \" by atomize_elim (auto 0 4 simp add: less_eq_Suc_le not_le dest!: le_cInf_iff[THEN iffD1, rotated -1]) then have 2: "k < plen \" if "\ \' k \ \ \' i + b" for k by (meson \_mono discrete not_le order.strict_trans2 that) from MatchF.prems(1,4) True show ?thesis unfolding sat.simps conj_commute[of "left I \ _" "_ \ _"] proof (intro ex_cong conj_cong match_cong_strong, goal_cases left right sat) case (left j) then show ?case by (intro iffI) ((subst (1 2) \_prefix_conv[OF assms(1,2) 1, symmetric]; auto elim: order.trans[OF \_mono, rotated]), (subst (1 2) \_prefix_conv[OF assms(1,2) 2]; auto elim: order.trans[OF \_mono, rotated])) next case (right j) then show ?case by (intro iffI) ((subst (1 2) \_prefix_conv[OF assms(1,2) 2, symmetric]; auto elim: order.trans[OF \_mono, rotated]), (subst (1 2) \_prefix_conv[OF assms(1,2) 2]; auto elim: order.trans[OF \_mono, rotated])) qed auto next case False from MatchF.prems(1) False obtain j where "\ \ i + b + 1 \ \ \ j" "(\x\regex.atms r. j \ progress \ P x (plen \))" by atomize_elim (auto 0 6 simp add: less_eq_Suc_le not_le progress_regex_def dest!: le_cInf_iff[THEN iffD1, rotated -1]) then have 1: "\ \ regex.atms r \ k < progress \ P \ (plen \)" if "\ \ k \ \ \ i + b" for k \ using that by (fastforce elim!: order.strict_trans2[rotated] intro: less_\D[of \]) then have 2: "k < plen \" if "\ \ k \ \ \ i + b" "regex.atms r \ {}" for k using that by (fastforce intro: order.strict_trans2[OF _ progress_le_gen[OF MatchF(5), of \], of k]) from MatchF.prems have "i < progress \' P (Formula.MatchF I r) (plen \)" unfolding progress_prefix_conv[OF assms(1,2)] by blast with False obtain j where "\ \' i + b + 1 \ \ \' j" "(\x\regex.atms r. j \ progress \' P x (plen \))" by atomize_elim (auto 0 6 simp add: less_eq_Suc_le not_le progress_regex_def dest!: le_cInf_iff[THEN iffD1, rotated -1]) then have 11: "\ \ regex.atms r \ k < progress \ P \ (plen \)" if "\ \' k \ \ \' i + b" for k \ using that using progress_prefix_conv[OF assms(1,2)] by (auto 0 3 elim!: order.strict_trans2[rotated] intro: less_\D[of \']) have 21: "k < plen \" if "\ \' k \ \ \' i + b" for k using 11[OF that(1)] False by (fastforce intro: order.strict_trans2[OF _ progress_le_gen[OF MatchF(5), of \], of k]) show ?thesis unfolding sat.simps conj_commute[of "left I \ _" "_ \ _"] proof ((intro ex_cong conj_cong match_cong_strong MatchF(1)[OF _ _ MatchF(3-6)]; assumption?), goal_cases right left progress) case (right j) with False show ?case by (intro iffI) ((subst (1 2) \_prefix_conv[OF assms(1,2) 2, symmetric]; auto elim: order.trans[OF \_mono, rotated]), (subst (1 2) \_prefix_conv[OF assms(1,2) 21]; auto elim: order.trans[OF \_mono, rotated])) next case (left j) with False show ?case unfolding right enat_ord_code le_diff_conv add.commute[of b] by (intro iffI) ((subst (1 2) \_prefix_conv[OF assms(1,2) 21, symmetric]; auto elim: order.trans[OF \_mono, rotated]), (subst (1 2) \_prefix_conv[OF assms(1,2) 21]; auto elim: order.trans[OF \_mono, rotated])) next case (progress j k z) with False show ?case unfolding right enat_ord_code le_diff_conv add.commute[of b] by (elim 1[rotated]) (subst (1 2) \_prefix_conv[OF assms(1,2) 21]; auto elim!: order.trans[OF \_mono, rotated]) qed qed qed auto lemma sat_prefix_conv: assumes "prefix_of \ \" and "prefix_of \ \'" shows "i < progress \ Map.empty \ (plen \) \ Formula.sat \ Map.empty v i \ \ Formula.sat \' Map.empty v i \" by (erule sat_prefix_conv_gen[OF assms]) auto lemma progress_remove_neg[simp]: "progress \ P (remove_neg \) j = progress \ P \ j" by (cases \) simp_all lemma safe_progress_get_and: "safe_formula \ \ Min ((\\. progress \ P \ j) ` set (get_and_list \)) = progress \ P \ j" by (induction \ rule: get_and_list.induct) auto lemma progress_convert_multiway: "safe_formula \ \ progress \ P (convert_multiway \) j = progress \ P \ j" proof (induction \ arbitrary: P rule: safe_formula_induct) case (And_safe \ \) let ?c = "convert_multiway (Formula.And \ \)" let ?c\ = "convert_multiway \" let ?c\ = "convert_multiway \" have c_eq: "?c = Formula.Ands (get_and_list ?c\ @ get_and_list ?c\)" using And_safe by simp from \safe_formula \\ have "safe_formula ?c\" by (rule safe_convert_multiway) moreover from \safe_formula \\ have "safe_formula ?c\" by (rule safe_convert_multiway) ultimately show ?case unfolding c_eq using And_safe.IH by (auto simp: get_and_nonempty Min.union safe_progress_get_and) next case (And_Not \ \) let ?c = "convert_multiway (Formula.And \ (Formula.Neg \))" let ?c\ = "convert_multiway \" let ?c\ = "convert_multiway \" have c_eq: "?c = Formula.Ands (Formula.Neg ?c\ # get_and_list ?c\)" using And_Not by simp from \safe_formula \\ have "safe_formula ?c\" by (rule safe_convert_multiway) moreover from \safe_formula \\ have "safe_formula ?c\" by (rule safe_convert_multiway) ultimately show ?case unfolding c_eq using And_Not.IH by (auto simp: get_and_nonempty Min.union safe_progress_get_and) next case (MatchP I r) from MatchP show ?case unfolding progress.simps regex.map convert_multiway.simps regex.set_map image_image by (intro if_cong arg_cong[of _ _ Min] image_cong) (auto 0 4 simp: atms_def elim!: disjE_Not2 dest: safe_regex_safe_formula) next case (MatchF I r) from MatchF show ?case unfolding progress.simps regex.map convert_multiway.simps regex.set_map image_image by (intro if_cong arg_cong[of _ _ Min] arg_cong[of _ _ Inf] arg_cong[of _ _ "(\) _"] image_cong Collect_cong all_cong1 imp_cong conj_cong image_cong) (auto 0 4 simp: atms_def elim!: disjE_Not2 dest: safe_regex_safe_formula) qed auto subsection \Specification\ definition pprogress :: "Formula.formula \ Formula.prefix \ nat" where "pprogress \ \ = (THE n. \\. prefix_of \ \ \ progress \ Map.empty \ (plen \) = n)" lemma pprogress_eq: "prefix_of \ \ \ pprogress \ \ = progress \ Map.empty \ (plen \)" unfolding pprogress_def using progress_prefix_conv by blast locale future_bounded_mfodl = fixes \ :: Formula.formula assumes future_bounded: "Formula.future_bounded \" sublocale future_bounded_mfodl \ sliceable_timed_progress "Formula.nfv \" "Formula.fv \" "relevant_events \" "\\ v i. Formula.sat \ Map.empty v i \" "pprogress \" proof (unfold_locales, goal_cases) case (1 x) then show ?case by (simp add: fvi_less_nfv) next case (2 v v' \ i) then show ?case by (simp cong: sat_fv_cong[rule_format]) next case (3 v S \ i) then show ?case using sat_slice_iff[symmetric] by simp next case (4 \ \') moreover obtain \ where "prefix_of \' \" using ex_prefix_of .. moreover have "prefix_of \ \" using prefix_of_antimono[OF \\ \ \'\ \prefix_of \' \\] . ultimately show ?case by (simp add: pprogress_eq plen_mono progress_mono) next case (5 \ x) obtain j where "x \ progress \ Map.empty \ j" using future_bounded progress_ge by blast then have "x \ pprogress \ (take_prefix j \)" by (simp add: pprogress_eq[of _ \]) then show ?case by force next case (6 \ \ \' i v) then have "i < progress \ Map.empty \ (plen \)" by (simp add: pprogress_eq) with 6 show ?case using sat_prefix_conv by blast next case (7 \ \') then have "plen \ = plen \'" by transfer (simp add: list_eq_iff_nth_eq) moreover obtain \ \' where "prefix_of \ \" "prefix_of \' \'" using ex_prefix_of by blast+ moreover have "\i < plen \. \ \ i = \ \' i" using 7 calculation by transfer (simp add: list_eq_iff_nth_eq) ultimately show ?case by (simp add: pprogress_eq progress_time_conv) qed locale verimon_spec = fixes \ :: Formula.formula assumes monitorable: "mmonitorable \" sublocale verimon_spec \ future_bounded_mfodl using monitorable by unfold_locales (simp add: mmonitorable_def) subsection \Correctness\ subsubsection \Invariants\ definition wf_mbuf2 :: "nat \ nat \ nat \ (nat \ event_data table \ bool) \ (nat \ event_data table \ bool) \ event_data mbuf2 \ bool" where "wf_mbuf2 i ja jb P Q buf \ i \ ja \ i \ jb \ (case buf of (xs, ys) \ list_all2 P [i.. list_all2 Q [i.. 'b \ 'c \ bool) \ 'a list \ 'b list \ 'c list \ bool" for P :: "('a \ 'b \ 'c \ bool)" where "list_all3 P [] [] []" | "P a1 a2 a3 \ list_all3 P q1 q2 q3 \ list_all3 P (a1 # q1) (a2 # q2) (a3 # q3)" lemma list_all3_list_all2D: "list_all3 P xs ys zs \ (length xs = length ys \ list_all2 (case_prod P) (zip xs ys) zs)" by (induct xs ys zs rule: list_all3.induct) auto lemma list_all2_list_all3I: "length xs = length ys \ list_all2 (case_prod P) (zip xs ys) zs \ list_all3 P xs ys zs" by (induct xs ys arbitrary: zs rule: list_induct2) (auto simp: list_all2_Cons1 intro: list_all3.intros) lemma list_all3_list_all2_eq: "list_all3 P xs ys zs \ (length xs = length ys \ list_all2 (case_prod P) (zip xs ys) zs)" using list_all2_list_all3I list_all3_list_all2D by blast lemma list_all3_mapD: "list_all3 P (map f xs) (map g ys) (map h zs) \ list_all3 (\x y z. P (f x) (g y) (h z)) xs ys zs" by (induct "map f xs" "map g ys" "map h zs" arbitrary: xs ys zs rule: list_all3.induct) (auto intro: list_all3.intros) lemma list_all3_mapI: "list_all3 (\x y z. P (f x) (g y) (h z)) xs ys zs \ list_all3 P (map f xs) (map g ys) (map h zs)" by (induct xs ys zs rule: list_all3.induct) (auto intro: list_all3.intros) lemma list_all3_map_iff: "list_all3 P (map f xs) (map g ys) (map h zs) \ list_all3 (\x y z. P (f x) (g y) (h z)) xs ys zs" using list_all3_mapD list_all3_mapI by blast lemmas list_all3_map = list_all3_map_iff[where g=id and h=id, unfolded list.map_id id_apply] list_all3_map_iff[where f=id and h=id, unfolded list.map_id id_apply] list_all3_map_iff[where f=id and g=id, unfolded list.map_id id_apply] lemma list_all3_conv_all_nth: "list_all3 P xs ys zs = (length xs = length ys \ length ys = length zs \ (\i < length xs. P (xs!i) (ys!i) (zs!i)))" by (auto simp add: list_all3_list_all2_eq list_all2_conv_all_nth) lemma list_all3_refl [intro?]: "(\x. x \ set xs \ P x x x) \ list_all3 P xs xs xs" by (simp add: list_all3_conv_all_nth) definition wf_mbufn :: "nat \ nat list \ (nat \ event_data table \ bool) list \ event_data mbufn \ bool" where "wf_mbufn i js Ps buf \ list_all3 (\P j xs. i \ j \ list_all2 P [i.. _ \ _ \ nat \ nat \ event_data list set \ Formula.formula \ Formula.formula \ event_data mbuf2 \ bool" where "wf_mbuf2' \ P V j n R \ \ buf \ wf_mbuf2 (min (progress \ P \ j) (progress \ P \ j)) (progress \ P \ j) (progress \ P \ j) (\i. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \)) (\i. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \)) buf" definition wf_mbufn' :: "Formula.trace \ _ \ _ \ nat \ nat \ event_data list set \ Formula.formula Regex.regex \ event_data mbufn \ bool" where "wf_mbufn' \ P V j n R r buf \ (case to_mregex r of (mr, \s) \ wf_mbufn (progress_regex \ P r j) (map (\\. progress \ P \ j) \s) (map (\\ i. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \)) \s) buf)" lemma wf_mbuf2'_UNIV_alt: "wf_mbuf2' \ P V j n UNIV \ \ buf \ (case buf of (xs, ys) \ list_all2 (\i. wf_table n (Formula.fv \) (\v. Formula.sat \ V (map the v) i \)) [min (progress \ P \ j) (progress \ P \ j) ..< (progress \ P \ j)] xs \ list_all2 (\i. wf_table n (Formula.fv \) (\v. Formula.sat \ V (map the v) i \)) [min (progress \ P \ j) (progress \ P \ j) ..< (progress \ P \ j)] ys)" unfolding wf_mbuf2'_def wf_mbuf2_def by (simp add: mem_restr_UNIV[THEN eqTrueI, abs_def] split: prod.split) definition wf_ts :: "Formula.trace \ _ \ nat \ Formula.formula \ Formula.formula \ ts list \ bool" where "wf_ts \ P j \ \ ts \ list_all2 (\i t. t = \ \ i) [min (progress \ P \ j) (progress \ P \ j).. _ \ nat \ Formula.formula Regex.regex \ ts list \ bool" where "wf_ts_regex \ P j r ts \ list_all2 (\i t. t = \ \ i) [progress_regex \ P r j.. I \ \ Formula.Since (if pos then \ else Formula.Neg \) I \" definition (in msaux) wf_since_aux :: "Formula.trace \ _ \ event_data list set \ args \ Formula.formula \ Formula.formula \ 'msaux \ nat \ bool" where "wf_since_aux \ V R args \ \ aux ne \ Formula.fv \ \ Formula.fv \ \ (\cur auxlist. valid_msaux args cur aux auxlist \ cur = (if ne = 0 then 0 else \ \ (ne - 1)) \ sorted_wrt (\x y. fst x > fst y) auxlist \ (\t X. (t, X) \ set auxlist \ ne \ 0 \ t \ \ \ (ne - 1) \ \ \ (ne - 1) - t \ right (args_ivl args) \ (\i. \ \ i = t) \ qtable (args_n args) (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) (ne-1) (Sincep (args_pos args) \ (point (\ \ (ne - 1) - t)) \)) X) \ (\t. ne \ 0 \ t \ \ \ (ne - 1) \ \ \ (ne - 1) - t \ right (args_ivl args) \ (\i. \ \ i = t) \ (\X. (t, X) \ set auxlist)))" definition wf_matchP_aux :: "Formula.trace \ _ \ nat \ event_data list set \ \ \ Formula.formula Regex.regex \ event_data mr\aux \ nat \ bool" where "wf_matchP_aux \ V n R I r aux ne \ sorted_wrt (\x y. fst x > fst y) aux \ (\t X. (t, X) \ set aux \ ne \ 0 \ t \ \ \ (ne-1) \ \ \ (ne-1) - t \ right I \ (\i. \ \ i = t) \ (case to_mregex r of (mr, \s) \ (\ms \ RPDs mr. qtable n (Formula.fv_regex r) (mem_restr R) (\v. Formula.sat \ V (map the v) (ne-1) (Formula.MatchP (point (\ \ (ne-1) - t)) (from_mregex ms \s))) (lookup X ms)))) \ (\t. ne \ 0 \ t \ \ \ (ne-1) \ \ \ (ne-1) - t \ right I \ (\i. \ \ i = t) \ (\X. (t, X) \ set aux))" lemma qtable_mem_restr_UNIV: "qtable n A(mem_restr UNIV) Q X = wf_table n A Q X" unfolding qtable_def by auto lemma (in msaux) wf_since_aux_UNIV_alt: "wf_since_aux \ V UNIV args \ \ aux ne \ Formula.fv \ \ Formula.fv \ \ (\cur auxlist. valid_msaux args cur aux auxlist \ cur = (if ne = 0 then 0 else \ \ (ne - 1)) \ sorted_wrt (\x y. fst x > fst y) auxlist \ (\t X. (t, X) \ set auxlist \ ne \ 0 \ t \ \ \ (ne - 1) \ \ \ (ne - 1) - t \ right (args_ivl args) \ (\i. \ \ i = t) \ wf_table (args_n args) (Formula.fv \) (\v. Formula.sat \ V (map the v) (ne - 1) (Sincep (args_pos args) \ (point (\ \ (ne - 1) - t)) \)) X) \ (\t. ne \ 0 \ t \ \ \ (ne - 1) \ \ \ (ne - 1) - t \ right (args_ivl args) \ (\i. \ \ i = t) \ (\X. (t, X) \ set auxlist)))" unfolding wf_since_aux_def qtable_mem_restr_UNIV .. definition wf_until_auxlist :: "Formula.trace \ _ \ nat \ event_data list set \ bool \ Formula.formula \ \ \ Formula.formula \ event_data muaux \ nat \ bool" where "wf_until_auxlist \ V n R pos \ I \ auxlist ne \ list_all2 (\x i. case x of (t, r1, r2) \ t = \ \ i \ qtable n (Formula.fv \) (mem_restr R) (\v. if pos then (\k\{i.. V (map the v) k \) else (\k\{i.. V (map the v) k \)) r1 \ qtable n (Formula.fv \) (mem_restr R) (\v. (\j. i \ j \ j < ne + length auxlist \ mem (\ \ j - \ \ i) I \ Formula.sat \ V (map the v) j \ \ (\k\{i.. V (map the v) k \ else \ Formula.sat \ V (map the v) k \))) r2) auxlist [ne.. _ \ event_data list set \ args \ Formula.formula \ Formula.formula \ 'muaux \ nat \ bool" where "wf_until_aux \ V R args \ \ aux ne \ Formula.fv \ \ Formula.fv \ \ (\cur auxlist. valid_muaux args cur aux auxlist \ cur = (if ne + length auxlist = 0 then 0 else \ \ (ne + length auxlist - 1)) \ wf_until_auxlist \ V (args_n args) R (args_pos args) \ (args_ivl args) \ auxlist ne)" lemma (in muaux) wf_until_aux_UNIV_alt: "wf_until_aux \ V UNIV args \ \ aux ne \ Formula.fv \ \ Formula.fv \ \ (\cur auxlist. valid_muaux args cur aux auxlist \ cur = (if ne + length auxlist = 0 then 0 else \ \ (ne + length auxlist - 1)) \ list_all2 (\x i. case x of (t, r1, r2) \ t = \ \ i \ wf_table (args_n args) (Formula.fv \) (\v. if (args_pos args) then (\k\{i.. V (map the v) k \) else (\k\{i.. V (map the v) k \)) r1 \ wf_table (args_n args) (Formula.fv \) (\v. \j. i \ j \ j < ne + length auxlist \ mem (\ \ j - \ \ i) (args_ivl args) \ Formula.sat \ V (map the v) j \ \ (\k\{i.. V (map the v) k \ else \ Formula.sat \ V (map the v) k \)) r2) auxlist [ne.. _ \ nat \ event_data list set \ \ \ Formula.formula Regex.regex \ event_data ml\aux \ nat \ nat \ bool" where "wf_matchF_aux \ V n R I r aux ne k \ (case to_mregex r of (mr, \s) \ list_all2 (\x i. case x of (t, rels, rel) \ t = \ \ i \ list_all2 (\\. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \)) \s rels \ qtable n (Formula.fv_regex r) (mem_restr R) (\v. (\j. i \ j \ j < ne + length aux + k \ mem (\ \ j - \ \ i) I \ Regex.match (Formula.sat \ V (map the v)) r i j)) rel) aux [ne.. V n R I r st i = (case st of (aux, Y) \ aux \ [] \ wf_matchF_aux \ V n R I r aux i 0 \ (case to_mregex r of (mr, \s) \ \ms \ LPDs mr. qtable n (Formula.fv_regex r) (mem_restr R) (\v. Regex.match (Formula.sat \ V (map the v)) (from_mregex ms \s) i (i + length aux - 1)) (lookup Y ms)))" definition lift_envs' :: "nat \ event_data list set \ event_data list set" where "lift_envs' b R = (\(xs,ys). xs @ ys) ` ({xs. length xs = b} \ R)" fun formula_of_constraint :: "Formula.trm \ bool \ mconstraint \ Formula.trm \ Formula.formula" where "formula_of_constraint (t1, True, MEq, t2) = Formula.Eq t1 t2" | "formula_of_constraint (t1, True, MLess, t2) = Formula.Less t1 t2" | "formula_of_constraint (t1, True, MLessEq, t2) = Formula.LessEq t1 t2" | "formula_of_constraint (t1, False, MEq, t2) = Formula.Neg (Formula.Eq t1 t2)" | "formula_of_constraint (t1, False, MLess, t2) = Formula.Neg (Formula.Less t1 t2)" | "formula_of_constraint (t1, False, MLessEq, t2) = Formula.Neg (Formula.LessEq t1 t2)" inductive (in maux) wf_mformula :: "Formula.trace \ nat \ _ \ _ \ nat \ event_data list set \ ('msaux, 'muaux) mformula \ Formula.formula \ bool" for \ j where Eq: "is_simple_eq t1 t2 \ \x\Formula.fv_trm t1. x < n \ \x\Formula.fv_trm t2. x < n \ wf_mformula \ j P V n R (MRel (eq_rel n t1 t2)) (Formula.Eq t1 t2)" | neq_Var: "x < n \ wf_mformula \ j P V n R (MRel empty_table) (Formula.Neg (Formula.Eq (Formula.Var x) (Formula.Var x)))" | Pred: "\x\Formula.fv (Formula.Pred e ts). x < n \ \t\set ts. Formula.is_Var t \ Formula.is_Const t \ wf_mformula \ j P V n R (MPred e ts) (Formula.Pred e ts)" | Let: "wf_mformula \ j P V m UNIV \ \' \ wf_mformula \ j (P(p \ progress \ P \' j)) (V(p \ \i. {v. length v = m \ Formula.sat \ V v i \'})) n R \ \' \ {0.. Formula.fv \' \ b \ m \ m = Formula.nfv \' \ wf_mformula \ j P V n R (MLet p m \ \) (Formula.Let p \' \')" | And: "wf_mformula \ j P V n R \ \' \ wf_mformula \ j P V n R \ \' \ if pos then \ = Formula.And \' \' else \ = Formula.And \' (Formula.Neg \') \ Formula.fv \' \ Formula.fv \' \ wf_mbuf2' \ P V j n R \' \' buf \ wf_mformula \ j P V n R (MAnd (fv \') \ pos (fv \') \ buf) \" | AndAssign: "wf_mformula \ j P V n R \ \' \ x < n \ x \ Formula.fv \' \ Formula.fv_trm t \ Formula.fv \' \ (x, t) = conf \ \' = Formula.Eq (Formula.Var x) t \ \' = Formula.Eq t (Formula.Var x) \ wf_mformula \ j P V n R (MAndAssign \ conf) (Formula.And \' \')" | AndRel: "wf_mformula \ j P V n R \ \' \ \' = formula_of_constraint conf \ (let (t1, _, _, t2) = conf in Formula.fv_trm t1 \ Formula.fv_trm t2 \ Formula.fv \') \ wf_mformula \ j P V n R (MAndRel \ conf) (Formula.And \' \')" | Ands: "list_all2 (\\ \'. wf_mformula \ j P V n R \ \') l (l_pos @ map remove_neg l_neg) \ wf_mbufn (progress \ P (Formula.Ands l') j) (map (\\. progress \ P \ j) (l_pos @ map remove_neg l_neg)) (map (\\ i. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \)) (l_pos @ map remove_neg l_neg)) buf \ (l_pos, l_neg) = partition safe_formula l' \ l_pos \ [] \ list_all safe_formula (map remove_neg l_neg) \ A_pos = map fv l_pos \ A_neg = map fv l_neg \ \(set A_neg) \ \(set A_pos) \ wf_mformula \ j P V n R (MAnds A_pos A_neg l buf) (Formula.Ands l')" | Or: "wf_mformula \ j P V n R \ \' \ wf_mformula \ j P V n R \ \' \ Formula.fv \' = Formula.fv \' \ wf_mbuf2' \ P V j n R \' \' buf \ wf_mformula \ j P V n R (MOr \ \ buf) (Formula.Or \' \')" | Neg: "wf_mformula \ j P V n R \ \' \ Formula.fv \' = {} \ wf_mformula \ j P V n R (MNeg \) (Formula.Neg \')" | Exists: "wf_mformula \ j P V (Suc n) (lift_envs R) \ \' \ wf_mformula \ j P V n R (MExists \) (Formula.Exists \')" | Agg: "wf_mformula \ j P V (b + n) (lift_envs' b R) \ \' \ y < n \ y + b \ Formula.fv \' \ {0.. Formula.fv \' \ Formula.fv_trm f \ Formula.fv \' \ g0 = (Formula.fv \' \ {0.. wf_mformula \ j P V n R (MAgg g0 y \ b f \) (Formula.Agg y \ b f \')" | Prev: "wf_mformula \ j P V n R \ \' \ first \ j = 0 \ list_all2 (\i. qtable n (Formula.fv \') (mem_restr R) (\v. Formula.sat \ V (map the v) i \')) [min (progress \ P \' j) (j-1).. P \' j] buf \ list_all2 (\i t. t = \ \ i) [min (progress \ P \' j) (j-1).. wf_mformula \ j P V n R (MPrev I \ first buf nts) (Formula.Prev I \')" | Next: "wf_mformula \ j P V n R \ \' \ first \ progress \ P \' j = 0 \ list_all2 (\i t. t = \ \ i) [progress \ P \' j - 1.. wf_mformula \ j P V n R (MNext I \ first nts) (Formula.Next I \')" | Since: "wf_mformula \ j P V n R \ \' \ wf_mformula \ j P V n R \ \' \ if args_pos args then \'' = \' else \'' = Formula.Neg \' \ safe_formula \'' = args_pos args \ args_ivl args = I \ args_n args = n \ args_L args = Formula.fv \' \ args_R args = Formula.fv \' \ Formula.fv \' \ Formula.fv \' \ wf_mbuf2' \ P V j n R \' \' buf \ wf_ts \ P j \' \' nts \ wf_since_aux \ V R args \' \' aux (progress \ P (Formula.Since \'' I \') j) \ wf_mformula \ j P V n R (MSince args \ \ buf nts aux) (Formula.Since \'' I \')" | Until: "wf_mformula \ j P V n R \ \' \ wf_mformula \ j P V n R \ \' \ if args_pos args then \'' = \' else \'' = Formula.Neg \' \ safe_formula \'' = args_pos args \ args_ivl args = I \ args_n args = n \ args_L args = Formula.fv \' \ args_R args = Formula.fv \' \ Formula.fv \' \ Formula.fv \' \ wf_mbuf2' \ P V j n R \' \' buf \ wf_ts \ P j \' \' nts \ wf_until_aux \ V R args \' \' aux (progress \ P (Formula.Until \'' I \') j) \ progress \ P (Formula.Until \'' I \') j + length_muaux args aux = min (progress \ P \' j) (progress \ P \' j) \ wf_mformula \ j P V n R (MUntil args \ \ buf nts aux) (Formula.Until \'' I \')" | MatchP: "(case to_mregex r of (mr', \s') \ list_all2 (wf_mformula \ j P V n R) \s \s' \ mr = mr') \ mrs = sorted_list_of_set (RPDs mr) \ safe_regex Past Strict r \ wf_mbufn' \ P V j n R r buf \ wf_ts_regex \ P j r nts \ wf_matchP_aux \ V n R I r aux (progress \ P (Formula.MatchP I r) j) \ wf_mformula \ j P V n R (MMatchP I mr mrs \s buf nts aux) (Formula.MatchP I r)" | MatchF: "(case to_mregex r of (mr', \s') \ list_all2 (wf_mformula \ j P V n R) \s \s' \ mr = mr') \ mrs = sorted_list_of_set (LPDs mr) \ safe_regex Futu Strict r \ wf_mbufn' \ P V j n R r buf \ wf_ts_regex \ P j r nts \ wf_matchF_aux \ V n R I r aux (progress \ P (Formula.MatchF I r) j) 0 \ progress \ P (Formula.MatchF I r) j + length aux = progress_regex \ P r j \ wf_mformula \ j P V n R (MMatchF I mr mrs \s buf nts aux) (Formula.MatchF I r)" definition (in maux) wf_mstate :: "Formula.formula \ Formula.prefix \ event_data list set \ ('msaux, 'muaux) mstate \ bool" where "wf_mstate \ \ R st \ mstate_n st = Formula.nfv \ \ (\\. prefix_of \ \ \ mstate_i st = progress \ Map.empty \ (plen \) \ wf_mformula \ (plen \) Map.empty Map.empty (mstate_n st) R (mstate_m st) \)" subsubsection \Initialisation\ lemma wf_mbuf2'_0: "pred_mapping (\x. x = 0) P \ wf_mbuf2' \ P V 0 n R \ \ ([], [])" unfolding wf_mbuf2'_def wf_mbuf2_def by simp lemma wf_mbufn'_0: "to_mregex r = (mr, \s) \ pred_mapping (\x. x = 0) P \ wf_mbufn' \ P V 0 n R r (replicate (length \s) [])" unfolding wf_mbufn'_def wf_mbufn_def map_replicate_const[symmetric] by (auto simp: list_all3_map intro: list_all3_refl simp: Min_eq_iff progress_regex_def) lemma wf_ts_0: "wf_ts \ P 0 \ \ []" unfolding wf_ts_def by simp lemma wf_ts_regex_0: "wf_ts_regex \ P 0 r []" unfolding wf_ts_regex_def by simp lemma (in msaux) wf_since_aux_Nil: "Formula.fv \' \ Formula.fv \' \ wf_since_aux \ V R (init_args I n (Formula.fv \') (Formula.fv \') b) \' \' (init_msaux (init_args I n (Formula.fv \') (Formula.fv \') b)) 0" unfolding wf_since_aux_def by (auto intro!: valid_init_msaux) lemma (in muaux) wf_until_aux_Nil: "Formula.fv \' \ Formula.fv \' \ wf_until_aux \ V R (init_args I n (Formula.fv \') (Formula.fv \') b) \' \' (init_muaux (init_args I n (Formula.fv \') (Formula.fv \') b)) 0" unfolding wf_until_aux_def wf_until_auxlist_def by (auto intro: valid_init_muaux) lemma wf_matchP_aux_Nil: "wf_matchP_aux \ V n R I r [] 0" unfolding wf_matchP_aux_def by simp lemma wf_matchF_aux_Nil: "wf_matchF_aux \ V n R I r [] 0 k" unfolding wf_matchF_aux_def by simp lemma fv_regex_alt: "safe_regex m g r \ Formula.fv_regex r = (\\ \ atms r. Formula.fv \)" unfolding fv_regex_alt atms_def by (auto 0 3 dest: safe_regex_safe_formula) lemmas to_mregex_atms = to_mregex_ok[THEN conjunct1, THEN equalityD1, THEN set_mp, rotated] lemma (in maux) wf_minit0: "safe_formula \ \ \x\Formula.fv \. x < n \ pred_mapping (\x. x = 0) P \ wf_mformula \ 0 P V n R (minit0 n \) \" proof (induction arbitrary: n R P V rule: safe_formula_induct) case (Eq_Const c d) then show ?case by (auto simp add: is_simple_eq_def simp del: eq_rel.simps intro!: wf_mformula.Eq) next case (Eq_Var1 c x) then show ?case by (auto simp add: is_simple_eq_def simp del: eq_rel.simps intro!: wf_mformula.Eq) next case (Eq_Var2 c x) then show ?case by (auto simp add: is_simple_eq_def simp del: eq_rel.simps intro!: wf_mformula.Eq) next case (neq_Var x y) then show ?case by (auto intro!: wf_mformula.neq_Var) next case (Pred e ts) then show ?case by (auto intro!: wf_mformula.Pred) next case (Let p \ \) with fvi_less_nfv show ?case by (auto simp: pred_mapping_alt dom_def intro!: wf_mformula.Let Let(4,5)) next case (And_assign \ \) then have 1: "\x\fv \. x < n" by simp from 1 \safe_assignment (fv \) \\ obtain x t where "x < n" "x \ fv \" "fv_trm t \ fv \" "\ = Formula.Eq (Formula.Var x) t \ \ = Formula.Eq t (Formula.Var x)" unfolding safe_assignment_def by (force split: formula.splits trm.splits) with And_assign show ?case by (auto intro!: wf_mformula.AndAssign split: trm.splits) next case (And_safe \ \) then show ?case by (auto intro!: wf_mformula.And wf_mbuf2'_0) next case (And_constraint \ \) from \fv \ \ fv \\ \is_constraint \\ obtain t1 p c t2 where "(t1, p, c, t2) = split_constraint \" "formula_of_constraint (split_constraint \) = \" "fv_trm t1 \ fv_trm t2 \ fv \" by (induction rule: is_constraint.induct) auto with And_constraint show ?case by (auto 0 3 intro!: wf_mformula.AndRel) next case (And_Not \ \) then show ?case by (auto intro!: wf_mformula.And wf_mbuf2'_0) next case (Ands l pos neg) note posneg = "Ands.hyps"(1) let ?wf_minit = "\x. wf_mformula \ 0 P V n R (minit0 n x)" let ?pos = "filter safe_formula l" let ?neg = "filter (Not \ safe_formula) l" have "list_all2 ?wf_minit ?pos pos" using Ands.IH(1) Ands.prems posneg by (auto simp: list_all_iff intro!: list.rel_refl_strong) moreover have "list_all2 ?wf_minit (map remove_neg ?neg) (map remove_neg neg)" using Ands.IH(2) Ands.prems posneg by (auto simp: list.rel_map list_all_iff intro!: list.rel_refl_strong) moreover have "list_all3 (\_ _ _. True) (?pos @ map remove_neg ?neg) (?pos @ map remove_neg ?neg) l" by (auto simp: list_all3_conv_all_nth comp_def sum_length_filter_compl) moreover have "l \ [] \ (MIN \\set l. (0 :: nat)) = 0" by (cases l) (auto simp: Min_eq_iff) ultimately show ?case using Ands.hyps Ands.prems(2) by (auto simp: wf_mbufn_def list_all3_map list.rel_map map_replicate_const[symmetric] subset_eq map_map[symmetric] map_append[symmetric] simp del: map_map map_append intro!: wf_mformula.Ands list_all2_appendI) next case (Neg \) then show ?case by (auto intro!: wf_mformula.Neg) next case (Or \ \) then show ?case by (auto intro!: wf_mformula.Or wf_mbuf2'_0) next case (Exists \) then show ?case by (auto simp: fvi_Suc_bound intro!: wf_mformula.Exists) next case (Agg y \ b f \) then show ?case by (auto intro!: wf_mformula.Agg Agg.IH fvi_plus_bound) next case (Prev I \) thm wf_mformula.Prev[where P=P] then show ?case by (auto intro!: wf_mformula.Prev) next case (Next I \) then show ?case by (auto intro!: wf_mformula.Next) next case (Since \ I \) then show ?case using wf_since_aux_Nil by (auto simp add: init_args_def intro!: wf_mformula.Since wf_mbuf2'_0 wf_ts_0) next case (Not_Since \ I \) then show ?case using wf_since_aux_Nil by (auto simp add: init_args_def intro!: wf_mformula.Since wf_mbuf2'_0 wf_ts_0) next case (Until \ I \) then show ?case using valid_length_muaux[OF valid_init_muaux[OF Until(1)]] wf_until_aux_Nil by (auto simp add: init_args_def simp del: progress_simps intro!: wf_mformula.Until wf_mbuf2'_0 wf_ts_0) next case (Not_Until \ I \) then show ?case using valid_length_muaux[OF valid_init_muaux[OF Not_Until(1)]] wf_until_aux_Nil by (auto simp add: init_args_def simp del: progress_simps intro!: wf_mformula.Until wf_mbuf2'_0 wf_ts_0) next case (MatchP I r) then show ?case by (auto simp: list.rel_map fv_regex_alt simp del: progress_simps split: prod.split intro!: wf_mformula.MatchP list.rel_refl_strong wf_mbufn'_0 wf_ts_regex_0 wf_matchP_aux_Nil dest!: to_mregex_atms) next case (MatchF I r) then show ?case by (auto simp: list.rel_map fv_regex_alt progress_le Min_eq_iff progress_regex_def simp del: progress_simps split: prod.split intro!: wf_mformula.MatchF list.rel_refl_strong wf_mbufn'_0 wf_ts_regex_0 wf_matchF_aux_Nil dest!: to_mregex_atms) qed lemma (in maux) wf_mstate_minit: "safe_formula \ \ wf_mstate \ pnil R (minit \)" unfolding wf_mstate_def minit_def Let_def by (auto intro!: wf_minit0 fvi_less_nfv) subsubsection \Evaluation\ lemma match_wf_tuple: "Some f = match ts xs \ wf_tuple n (\t\set ts. Formula.fv_trm t) (Table.tabulate f 0 n)" by (induction ts xs arbitrary: f rule: match.induct) (fastforce simp: wf_tuple_def split: if_splits option.splits)+ lemma match_fvi_trm_None: "Some f = match ts xs \ \t\set ts. x \ Formula.fv_trm t \ f x = None" by (induction ts xs arbitrary: f rule: match.induct) (auto split: if_splits option.splits) lemma match_fvi_trm_Some: "Some f = match ts xs \ t \ set ts \ x \ Formula.fv_trm t \ f x \ None" by (induction ts xs arbitrary: f rule: match.induct) (auto split: if_splits option.splits) lemma match_eval_trm: "\t\set ts. \i\Formula.fv_trm t. i < n \ Some f = match ts xs \ map (Formula.eval_trm (Table.tabulate (\i. the (f i)) 0 n)) ts = xs" proof (induction ts xs arbitrary: f rule: match.induct) case (3 x ts y ys) from 3(1)[symmetric] 3(2,3) show ?case by (auto 0 3 dest: match_fvi_trm_Some sym split: option.splits if_splits intro!: eval_trm_fv_cong) qed (auto split: if_splits) lemma wf_tuple_tabulate_Some: "wf_tuple n A (Table.tabulate f 0 n) \ x \ A \ x < n \ \y. f x = Some y" unfolding wf_tuple_def by auto lemma ex_match: "wf_tuple n (\t\set ts. Formula.fv_trm t) v \ \t\set ts. (\x\Formula.fv_trm t. x < n) \ (Formula.is_Var t \ Formula.is_Const t) \ \f. match ts (map (Formula.eval_trm (map the v)) ts) = Some f \ v = Table.tabulate f 0 n" proof (induction ts "map (Formula.eval_trm (map the v)) ts" arbitrary: v rule: match.induct) case (3 x ts y ys) then show ?case proof (cases "x \ (\t\set ts. Formula.fv_trm t)") case True with 3 show ?thesis by (auto simp: insert_absorb dest!: wf_tuple_tabulate_Some meta_spec[of _ v]) next case False with 3(3,4) have *: "map (Formula.eval_trm (map the v)) ts = map (Formula.eval_trm (map the (v[x := None]))) ts" by (auto simp: wf_tuple_def nth_list_update intro!: eval_trm_fv_cong) from False 3(2-4) obtain f where "match ts (map (Formula.eval_trm (map the v)) ts) = Some f" "v[x := None] = Table.tabulate f 0 n" unfolding * by (atomize_elim, intro 3(1)[of "v[x := None]"]) (auto simp: wf_tuple_def nth_list_update intro!: eval_trm_fv_cong) moreover from False this have "f x = None" "length v = n" by (auto dest: match_fvi_trm_None[OF sym] arg_cong[of _ _ length]) ultimately show ?thesis using 3(3) by (auto simp: list_eq_iff_nth_eq wf_tuple_def) qed qed (auto simp: wf_tuple_def intro: nth_equalityI) lemma eq_rel_eval_trm: "v \ eq_rel n t1 t2 \ is_simple_eq t1 t2 \ \x\Formula.fv_trm t1 \ Formula.fv_trm t2. x < n \ Formula.eval_trm (map the v) t1 = Formula.eval_trm (map the v) t2" by (cases t1; cases t2) (simp_all add: is_simple_eq_def singleton_table_def split: if_splits) lemma in_eq_rel: "wf_tuple n (Formula.fv_trm t1 \ Formula.fv_trm t2) v \ is_simple_eq t1 t2 \ Formula.eval_trm (map the v) t1 = Formula.eval_trm (map the v) t2 \ v \ eq_rel n t1 t2" by (cases t1; cases t2) (auto simp: is_simple_eq_def singleton_table_def wf_tuple_def unit_table_def intro!: nth_equalityI split: if_splits) lemma table_eq_rel: "is_simple_eq t1 t2 \ table n (Formula.fv_trm t1 \ Formula.fv_trm t2) (eq_rel n t1 t2)" by (cases t1; cases t2; simp add: is_simple_eq_def) lemma wf_tuple_Suc_fviD: "wf_tuple (Suc n) (Formula.fvi b \) v \ wf_tuple n (Formula.fvi (Suc b) \) (tl v)" unfolding wf_tuple_def by (simp add: fvi_Suc nth_tl) lemma table_fvi_tl: "table (Suc n) (Formula.fvi b \) X \ table n (Formula.fvi (Suc b) \) (tl ` X)" unfolding table_def by (auto intro: wf_tuple_Suc_fviD) lemma wf_tuple_Suc_fvi_SomeI: "0 \ Formula.fvi b \ \ wf_tuple n (Formula.fvi (Suc b) \) v \ wf_tuple (Suc n) (Formula.fvi b \) (Some x # v)" unfolding wf_tuple_def by (auto simp: fvi_Suc less_Suc_eq_0_disj) lemma wf_tuple_Suc_fvi_NoneI: "0 \ Formula.fvi b \ \ wf_tuple n (Formula.fvi (Suc b) \) v \ wf_tuple (Suc n) (Formula.fvi b \) (None # v)" unfolding wf_tuple_def by (auto simp: fvi_Suc less_Suc_eq_0_disj) lemma qtable_project_fv: "qtable (Suc n) (fv \) (mem_restr (lift_envs R)) P X \ qtable n (Formula.fvi (Suc 0) \) (mem_restr R) (\v. \x. P ((if 0 \ fv \ then Some x else None) # v)) (tl ` X)" using neq0_conv by (fastforce simp: image_iff Bex_def fvi_Suc elim!: qtable_cong dest!: qtable_project) lemma mem_restr_lift_envs'_append[simp]: "length xs = b \ mem_restr (lift_envs' b R) (xs @ ys) = mem_restr R ys" unfolding mem_restr_def lift_envs'_def by (auto simp: list_all2_append list.rel_map intro!: exI[where x="map the xs"] list.rel_refl) lemma nth_list_update_alt: "xs[i := x] ! j = (if i < length xs \ i = j then x else xs ! j)" by auto lemma wf_tuple_upd_None: "wf_tuple n A xs \ A - {i} = B \ wf_tuple n B (xs[i:=None])" unfolding wf_tuple_def by (auto simp: nth_list_update_alt) lemma mem_restr_upd_None: "mem_restr R xs \ mem_restr R (xs[i:=None])" unfolding mem_restr_def by (auto simp: list_all2_conv_all_nth nth_list_update_alt) lemma mem_restr_dropI: "mem_restr (lift_envs' b R) xs \ mem_restr R (drop b xs)" unfolding mem_restr_def lift_envs'_def by (auto simp: append_eq_conv_conj list_all2_append2) lemma mem_restr_dropD: assumes "b \ length xs" and "mem_restr R (drop b xs)" shows "mem_restr (lift_envs' b R) xs" proof - let ?R = "\a b. a \ None \ a = Some b" from assms(2) obtain v where "v \ R" and "list_all2 ?R (drop b xs) v" unfolding mem_restr_def .. show ?thesis unfolding mem_restr_def proof have "list_all2 ?R (take b xs) (map the (take b xs))" by (auto simp: list.rel_map intro!: list.rel_refl) moreover note \list_all2 ?R (drop b xs) v\ ultimately have "list_all2 ?R (take b xs @ drop b xs) (map the (take b xs) @ v)" by (rule list_all2_appendI) then show "list_all2 ?R xs (map the (take b xs) @ v)" by simp show "map the (take b xs) @ v \ lift_envs' b R" unfolding lift_envs'_def using assms(1) \v \ R\ by auto qed qed lemma wf_tuple_append: "wf_tuple a {x \ A. x < a} xs \ wf_tuple b {x - a | x. x \ A \ x \ a} ys \ wf_tuple (a + b) A (xs @ ys)" unfolding wf_tuple_def by (auto simp: nth_append eq_diff_iff) lemma wf_tuple_map_Some: "length xs = n \ {0.. A \ wf_tuple n A (map Some xs)" unfolding wf_tuple_def by auto lemma wf_tuple_drop: "wf_tuple (b + n) A xs \ {x - b | x. x \ A \ x \ b} = B \ wf_tuple n B (drop b xs)" unfolding wf_tuple_def by force lemma ecard_image: "inj_on f A \ ecard (f ` A) = ecard A" unfolding ecard_def by (auto simp: card_image dest: finite_imageD) lemma meval_trm_eval_trm: "wf_tuple n A x \ fv_trm t \ A \ \i\A. i < n \ meval_trm t x = Formula.eval_trm (map the x) t" unfolding wf_tuple_def by (induction t) simp_all lemma list_update_id: "xs ! i = z \ xs[i:=z] = xs" by (induction xs arbitrary: i) (auto split: nat.split) lemma qtable_wf_tupleD: "qtable n A P Q X \ \x\X. wf_tuple n A x" unfolding qtable_def table_def by blast lemma qtable_eval_agg: assumes inner: "qtable (b + n) (Formula.fv \) (mem_restr (lift_envs' b R)) (\v. Formula.sat \ V (map the v) i \) rel" and n: "\x\Formula.fv (Formula.Agg y \ b f \). x < n" and fresh: "y + b \ Formula.fv \" and b_fv: "{0.. Formula.fv \" and f_fv: "Formula.fv_trm f \ Formula.fv \" and g0: "g0 = (Formula.fv \ \ {0.. b f \)) (mem_restr R) (\v. Formula.sat \ V (map the v) i (Formula.Agg y \ b f \)) (eval_agg n g0 y \ b f rel)" (is "qtable _ ?fv _ ?Q ?rel'") proof - define M where "M = (\v. {(x, ecard Zs) | x Zs. Zs = {zs. length zs = b \ Formula.sat \ V (zs @ v) i \ \ Formula.eval_trm (zs @ v) f = x} \ Zs \ {}})" have f_fvi: "Formula.fvi_trm b f \ Formula.fvi b \" using f_fv by (auto simp: fvi_trm_iff_fv_trm[where b=b] fvi_iff_fv[where b=b]) show ?thesis proof (cases "g0 \ rel = empty_table") case True then have [simp]: "Formula.fvi b \ = {}" by (auto simp: g0 fvi_iff_fv(1)[where b=b]) then have [simp]: "Formula.fvi_trm b f = {}" using f_fvi by auto show ?thesis proof (rule qtableI) show "table n ?fv ?rel'" by (simp add: eval_agg_def True) next fix v assume "wf_tuple n ?fv v" "mem_restr R v" have "\ Formula.sat \ V (zs @ map the v) i \" if [simp]: "length zs = b" for zs proof - let ?zs = "map2 (\z i. if i \ Formula.fv \ then Some z else None) zs [0.. fv \. x < b} ?zs" by (simp add: wf_tuple_def) then have "wf_tuple (b + n) (Formula.fv \) (?zs @ v[y:=None])" using \wf_tuple n ?fv v\ True by (auto simp: g0 intro!: wf_tuple_append wf_tuple_upd_None) then have "\ Formula.sat \ V (map the (?zs @ v[y:=None])) i \" using True \mem_restr R v\ by (auto simp del: map_append dest!: in_qtableI[OF inner, rotated -1] intro!: mem_restr_upd_None) also have "Formula.sat \ V (map the (?zs @ v[y:=None])) i \ \ Formula.sat \ V (zs @ map the v) i \" using True by (auto simp: g0 nth_append intro!: sat_fv_cong) finally show ?thesis . qed then have M_empty: "M (map the v) = {}" unfolding M_def by blast show "Formula.sat \ V (map the v) i (Formula.Agg y \ b f \)" if "v \ eval_agg n g0 y \ b f rel" using M_empty True that n by (simp add: M_def eval_agg_def g0 singleton_table_def) have "v \ singleton_table n y (the (v ! y))" "length v = n" using \wf_tuple n ?fv v\ unfolding wf_tuple_def singleton_table_def by (auto simp add: tabulate_alt map_nth intro!: trans[OF map_cong[where g="(!) v", simplified nth_map, OF refl], symmetric]) then show "v \ eval_agg n g0 y \ b f rel" if "Formula.sat \ V (map the v) i (Formula.Agg y \ b f \)" using M_empty True that n by (simp add: M_def eval_agg_def g0) qed next case non_default_case: False have union_fv: "{0.. (\x. x + b) ` Formula.fvi b \ = fv \" using b_fv by (auto simp: fvi_iff_fv(1)[where b=b] intro!: image_eqI[where b=x and x="x - b" for x]) have b_n: "\x\fv \. x < b + n" proof fix x assume "x \ fv \" show "x < b + n" proof (cases "x \ b") case True with \x \ fv \\ have "x - b \ ?fv" by (simp add: fvi_iff_fv(1)[where b=b]) then show ?thesis using n f_fvi by (auto simp: Un_absorb2) qed simp qed define M' where "M' = (\k. let group = Set.filter (\x. drop b x = k) rel; images = meval_trm f ` group in (\y. (y, ecard (Set.filter (\x. meval_trm f x = y) group))) ` images)" have M'_M: "M' (drop b x) = M (map the (drop b x))" if "x \ rel" "mem_restr (lift_envs' b R) x" for x proof - from that have wf_x: "wf_tuple (b + n) (fv \) x" by (auto elim!: in_qtableE[OF inner]) then have wf_zs_x: "wf_tuple (b + n) (fv \) (map Some zs @ drop b x)" if "length zs = b" for zs using that b_fv by (auto intro!: wf_tuple_append wf_tuple_map_Some wf_tuple_drop) have 1: "(length zs = b \ Formula.sat \ V (zs @ map the (drop b x)) i \ \ Formula.eval_trm (zs @ map the (drop b x)) f = y) \ (\a. a \ rel \ take b a = map Some zs \ drop b a = drop b x \ meval_trm f a = y)" (is "?A \ (\a. ?B a)") for y zs proof (intro iffI conjI) assume ?A then have "?B (map Some zs @ drop (length zs) x)" using in_qtableI[OF inner wf_zs_x] \mem_restr (lift_envs' b R) x\ meval_trm_eval_trm[OF wf_zs_x f_fv b_n] by (auto intro!: mem_restr_dropI) then show "\a. ?B a" .. next assume "\a. ?B a" then obtain a where "?B a" .. then have "a \ rel" and a_eq: "a = map Some zs @ drop b x" using append_take_drop_id[of b a] by auto then have "length a = b + n" using inner unfolding qtable_def table_def by (blast intro!: wf_tuple_length) then show "length zs = b" using wf_tuple_length[OF wf_x] unfolding a_eq by simp then have "mem_restr (lift_envs' b R) a" using \mem_restr _ x\ unfolding a_eq by (auto intro!: mem_restr_dropI) then show "Formula.sat \ V (zs @ map the (drop b x)) i \" using in_qtableE[OF inner \a \ rel\] by (auto simp: a_eq sat_fv_cong[THEN iffD1, rotated -1]) from \?B a\ show "Formula.eval_trm (zs @ map the (drop b x)) f = y" using meval_trm_eval_trm[OF wf_zs_x f_fv b_n, OF \length zs = b\] unfolding a_eq by simp qed have 2: "map Some (map the (take b a)) = take b a" if "a \ rel" for a using that b_fv inner[THEN qtable_wf_tupleD] unfolding table_def wf_tuple_def by (auto simp: list_eq_iff_nth_eq) have 3: "ecard {zs. \a. a \ rel \ take b a = map Some zs \ drop b a = drop b x \ P a} = ecard {a. a \ rel \ drop b a = drop b x \ P a}" (is "ecard ?A = ecard ?B") for P proof - have "ecard ?A = ecard ((\zs. map Some zs @ drop b x) ` ?A)" by (auto intro!: ecard_image[symmetric] inj_onI) also have "(\zs. map Some zs @ drop b x) ` ?A = ?B" by (subst (1 2) eq_commute) (auto simp: image_iff, metis "2" append_take_drop_id) finally show ?thesis . qed show ?thesis unfolding M_def M'_def by (auto simp: non_default_case Let_def image_def Set.filter_def 1 3, metis "2") qed have drop_lift: "mem_restr (lift_envs' b R) x" if "x \ rel" "mem_restr R ((drop b x)[y:=z])" for x z proof - have "(drop b x)[y:=None] = (drop b x)[y:=drop b x ! y]" proof - from \x \ rel\ have "drop b x ! y = None" using fresh n inner[THEN qtable_wf_tupleD] by (simp add: add.commute wf_tuple_def) then show ?thesis by simp qed then have "(drop b x)[y:=None] = drop b x" by simp moreover from \x \ rel\ have "length x = b + n" using inner[THEN qtable_wf_tupleD] by (simp add: wf_tuple_def) moreover from that(2) have "mem_restr R ((drop b x)[y:=z, y:=None])" by (rule mem_restr_upd_None) ultimately show ?thesis by (auto intro!: mem_restr_dropD) qed { fix v assume "mem_restr R v" have "v \ (\k. k[y:=Some (eval_agg_op \ (M' k))]) ` drop b ` rel \ v \ (\k. k[y:=Some (eval_agg_op \ (M (map the k)))]) ` drop b ` rel" (is "v \ ?A \ v \ ?B") proof assume "v \ ?A" then obtain v' where *: "v' \ rel" "v = (drop b v')[y:=Some (eval_agg_op \ (M' (drop b v')))]" by auto then have "M' (drop b v') = M (map the (drop b v'))" using \mem_restr R v\ by (auto intro!: M'_M drop_lift) with * show "v \ ?B" by simp next assume "v \ ?B" then obtain v' where *: "v' \ rel" "v = (drop b v')[y:=Some (eval_agg_op \ (M (map the (drop b v'))))]" by auto then have "M (map the (drop b v')) = M' (drop b v')" using \mem_restr R v\ by (auto intro!: M'_M[symmetric] drop_lift) with * show "v \ ?A" by simp qed then have "v \ eval_agg n g0 y \ b f rel \ v \ (\k. k[y:=Some (eval_agg_op \ (M (map the k)))]) ` drop b ` rel" by (simp add: non_default_case eval_agg_def M'_def Let_def) } note alt = this show ?thesis proof (rule qtableI) show "table n ?fv ?rel'" using inner[THEN qtable_wf_tupleD] n f_fvi by (auto simp: eval_agg_def non_default_case table_def wf_tuple_def Let_def nth_list_update fvi_iff_fv[where b=b] add.commute) next fix v assume "wf_tuple n ?fv v" "mem_restr R v" then have length_v: "length v = n" by (simp add: wf_tuple_def) show "Formula.sat \ V (map the v) i (Formula.Agg y \ b f \)" if "v \ eval_agg n g0 y \ b f rel" proof - from that obtain v' where "v' \ rel" "v = (drop b v')[y:=Some (eval_agg_op \ (M (map the (drop b v'))))]" using alt[OF \mem_restr R v\] by blast then have length_v': "length v' = b + n" using inner[THEN qtable_wf_tupleD] by (simp add: wf_tuple_def) have "Formula.sat \ V (map the v') i \" using \v' \ rel\ \mem_restr R v\ by (auto simp: \v = _\ elim!: in_qtableE[OF inner] intro!: drop_lift \v' \ rel\) then have "Formula.sat \ V (map the (take b v') @ map the v) i \" proof (rule sat_fv_cong[THEN iffD1, rotated], intro ballI) fix x assume "x \ fv \" then have "x \ y + b" using fresh by blast moreover have "x < length v'" using \x \ fv \\ b_n by (simp add: length_v') ultimately show "map the v' ! x = (map the (take b v') @ map the v) ! x" by (auto simp: \v = _\ nth_append) qed then have 1: "M (map the v) \ {}" by (force simp: M_def length_v') have "y < length (drop b v')" using n by (simp add: length_v') moreover have "Formula.sat \ V (zs @ map the v) i \ \ Formula.sat \ V (zs @ map the (drop b v')) i \" if "length zs = b" for zs proof (intro sat_fv_cong ballI) fix x assume "x \ fv \" then have "x \ y + b" using fresh by blast moreover have "x < length v'" using \x \ fv \\ b_n by (simp add: length_v') ultimately show "(zs @ map the v) ! x = (zs @ map the (drop b v')) ! x" by (auto simp: \v = _\ that nth_append) qed moreover have "Formula.eval_trm (zs @ map the v) f = Formula.eval_trm (zs @ map the (drop b v')) f" if "length zs = b" for zs proof (intro eval_trm_fv_cong ballI) fix x assume "x \ fv_trm f" then have "x \ y + b" using f_fv fresh by blast moreover have "x < length v'" using \x \ fv_trm f\ f_fv b_n by (auto simp: length_v') ultimately show "(zs @ map the v) ! x = (zs @ map the (drop b v')) ! x" by (auto simp: \v = _\ that nth_append) qed ultimately have "map the v ! y = eval_agg_op \ (M (map the v))" by (simp add: M_def \v = _\ conj_commute cong: conj_cong) with 1 show ?thesis by (auto simp: M_def) qed show "v \ eval_agg n g0 y \ b f rel" if sat_Agg: "Formula.sat \ V (map the v) i (Formula.Agg y \ b f \)" proof - obtain zs where "length zs = b" and "map Some zs @ v[y:=None] \ rel" proof (cases "fv \ \ {0.. empty_table" by (simp add: g0) then obtain x where "x \ rel" by auto have "(\i < n. (v[y:=None]) ! i = None)" using True \wf_tuple n ?fv v\ f_fv by (fastforce simp: wf_tuple_def fvi_iff_fv[where b=b] fvi_trm_iff_fv_trm[where b=b]) moreover have x: "(\i < n. drop b x ! i = None) \ length x = b + n" using True \x \ rel\ inner[THEN qtable_wf_tupleD] f_fv by (auto simp: wf_tuple_def) ultimately have "v[y:=None] = drop b x" unfolding list_eq_iff_nth_eq by (auto simp: length_v) with \x \ rel\ have "take b x @ v[y:=None] \ rel" by simp moreover have "map (Some \ the) (take b x) = take b x" using True \x \ rel\ inner[THEN qtable_wf_tupleD] b_fv by (subst map_cong[where g=id, OF refl]) (auto simp: wf_tuple_def in_set_conv_nth) ultimately have "map Some (map the (take b x)) @ v[y:=None] \ rel" by simp then show thesis using x[THEN conjunct2] by (fastforce intro!: that[rotated]) next case False with sat_Agg obtain zs where "length zs = b" and "Formula.sat \ V (zs @ map the v) i \" by auto then have "Formula.sat \ V (zs @ map the (v[y:=None])) i \" using fresh by (auto simp: map_update not_less nth_append elim!: sat_fv_cong[THEN iffD1, rotated] intro!: nth_list_update_neq[symmetric]) then have "map Some zs @ v[y:=None] \ rel" using b_fv f_fv fresh by (auto intro!: in_qtableI[OF inner] wf_tuple_append wf_tuple_map_Some wf_tuple_upd_None \wf_tuple n ?fv v\ mem_restr_upd_None \mem_restr R v\ simp: \length zs = b\ set_eq_iff fvi_iff_fv[where b=b] fvi_trm_iff_fv_trm[where b=b]) force+ with that \length zs = b\ show thesis by blast qed then have 1: "v[y:=None] \ drop b ` rel" by (intro image_eqI) auto have y_length: "y < length v" using n by (simp add: length_v) moreover have "Formula.sat \ V (zs @ map the (v[y:=None])) i \ \ Formula.sat \ V (zs @ map the v) i \" if "length zs = b" for zs proof (intro sat_fv_cong ballI) fix x assume "x \ fv \" then have "x \ y + b" using fresh by blast moreover have "x < b + length v" using \x \ fv \\ b_n by (simp add: length_v) ultimately show "(zs @ map the (v[y:=None])) ! x = (zs @ map the v) ! x" by (auto simp: that nth_append) qed moreover have "Formula.eval_trm (zs @ map the (v[y:=None])) f = Formula.eval_trm (zs @ map the v) f" if "length zs = b" for zs proof (intro eval_trm_fv_cong ballI) fix x assume "x \ fv_trm f" then have "x \ y + b" using f_fv fresh by blast moreover have "x < b + length v" using \x \ fv_trm f\ f_fv b_n by (auto simp: length_v) ultimately show "(zs @ map the (v[y:=None])) ! x = (zs @ map the v) ! x" by (auto simp: that nth_append) qed ultimately have "map the v ! y = eval_agg_op \ (M (map the (v[y:=None])))" using sat_Agg by (simp add: M_def cong: conj_cong) (simp cong: rev_conj_cong) then have 2: "v ! y = Some (eval_agg_op \ (M (map the (v[y:=None]))))" using \wf_tuple n ?fv v\ y_length by (auto simp add: wf_tuple_def) show ?thesis unfolding alt[OF \mem_restr R v\] by (rule image_eqI[where x="v[y:=None]"]) (use 1 2 in \auto simp: y_length list_update_id\) qed qed qed qed lemma mprev: "mprev_next I xs ts = (ys, xs', ts') \ list_all2 P [i.. list_all2 (\i t. t = \ \ i) [i.. i \ j' \ i < j \ list_all2 (\i X. if mem (\ \ (Suc i) - \ \ i) I then P i X else X = empty_table) [i.. list_all2 P [min j' (j-1).. list_all2 (\i t. t = \ \ i) [min j' (j-1).. list_all2 P [Suc i.. list_all2 (\i t. t = \ \ i) [i.. Suc i \ j' \ i < j \ list_all2 (\i X. if mem (\ \ (Suc i) - \ \ i) I then P (Suc i) X else X = empty_table) [i.. list_all2 P [Suc (min (j'-1) (j-1)).. list_all2 (\i t. t = \ \ i) [min (j'-1) (j-1).. A \ A \ set xs \ x \ foldr (\) xs {}" by (induction xs) auto lemma in_foldr_UnE: "x \ foldr (\) xs {} \ (\A. A \ set xs \ x \ A \ P) \ P" by (induction xs) auto lemma sat_the_restrict: "fv \ \ A \ Formula.sat \ V (map the (restrict A v)) i \ = Formula.sat \ V (map the v) i \" by (rule sat_fv_cong) (auto intro!: map_the_restrict) lemma eps_the_restrict: "fv_regex r \ A \ Regex.eps (Formula.sat \ V (map the (restrict A v))) i r = Regex.eps (Formula.sat \ V (map the v)) i r" by (rule eps_fv_cong) (auto intro!: map_the_restrict) lemma sorted_wrt_filter[simp]: "sorted_wrt R xs \ sorted_wrt R (filter P xs)" by (induct xs) auto lemma concat_map_filter[simp]: "concat (map f (filter P xs)) = concat (map (\x. if P x then f x else []) xs)" by (induct xs) auto lemma map_filter_alt: "map f (filter P xs) = concat (map (\x. if P x then [f x] else []) xs)" by (induct xs) auto lemma (in maux) update_since: assumes pre: "wf_since_aux \ V R args \ \ aux ne" and qtable1: "qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) ne \) rel1" and qtable2: "qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) ne \) rel2" and result_eq: "(rel, aux') = update_since args rel1 rel2 (\ \ ne) aux" and fvi_subset: "Formula.fv \ \ Formula.fv \" and args_ivl: "args_ivl args = I" and args_n: "args_n args = n" and args_L: "args_L args = Formula.fv \" and args_R: "args_R args = Formula.fv \" and args_pos: "args_pos args = pos" shows "wf_since_aux \ V R args \ \ aux' (Suc ne)" and "qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) ne (Sincep pos \ I \)) rel" proof - let ?wf_tuple = "\v. wf_tuple n (Formula.fv \) v" note sat.simps[simp del] from pre[unfolded wf_since_aux_def] obtain cur auxlist where aux: "valid_msaux args cur aux auxlist" "sorted_wrt (\x y. fst y < fst x) auxlist" "\t X. (t, X) \ set auxlist \ ne \ 0 \ t \ \ \ (ne - 1) \ \ \ (ne - 1) - t \ right I \ (\i. \ \ i = t) \ qtable n (fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) (ne - 1) (Sincep pos \ (point (\ \ (ne - 1) - t)) \)) X" "\t. ne \ 0 \ t \ \ \ (ne - 1) \ \ \ (ne - 1) - t \ right I \ (\i. \ \ i = t) \ (\X. (t, X) \ set auxlist)" and cur_def: "cur = (if ne = 0 then 0 else \ \ (ne - 1))" unfolding args_ivl args_n args_pos by blast from pre[unfolded wf_since_aux_def] have fv_sub: "Formula.fv \ \ Formula.fv \" by simp define aux0 where "aux0 = join_msaux args rel1 (add_new_ts_msaux args (\ \ ne) aux)" define auxlist0 where "auxlist0 = [(t, join rel pos rel1). (t, rel) \ auxlist, \ \ ne - t \ right I]" have tabL: "table (args_n args) (args_L args) rel1" using qtable1[unfolded qtable_def] unfolding args_n[symmetric] args_L[symmetric] by simp have cur_le: "cur \ \ \ ne" unfolding cur_def by auto have valid0: "valid_msaux args (\ \ ne) aux0 auxlist0" unfolding aux0_def auxlist0_def using valid_join_msaux[OF valid_add_new_ts_msaux[OF aux(1)], OF cur_le tabL] by (auto simp: args_ivl args_pos cur_def map_filter_alt split_beta cong: map_cong) from aux(2) have sorted_auxlist0: "sorted_wrt (\x y. fst x > fst y) auxlist0" unfolding auxlist0_def by (induction auxlist) (auto simp add: sorted_wrt_append) have in_auxlist0_1: "(t, X) \ set auxlist0 \ ne \ 0 \ t \ \ \ (ne-1) \ \ \ ne - t \ right I \ (\i. \ \ i = t) \ qtable n (Formula.fv \) (mem_restr R) (\v. (Formula.sat \ V (map the v) (ne-1) (Sincep pos \ (point (\ \ (ne-1) - t)) \) \ (if pos then Formula.sat \ V (map the v) ne \ else \ Formula.sat \ V (map the v) ne \))) X" for t X unfolding auxlist0_def using fvi_subset by (auto 0 1 elim!: qtable_join[OF _ qtable1] simp: sat_the_restrict dest!: aux(3)) then have in_auxlist0_le_\: "(t, X) \ set auxlist0 \ t \ \ \ ne" for t X by (meson \_mono diff_le_self le_trans) have in_auxlist0_2: "ne \ 0 \ t \ \ \ (ne-1) \ \ \ ne - t \ right I \ \i. \ \ i = t \ \X. (t, X) \ set auxlist0" for t proof - fix t assume "ne \ 0" "t \ \ \ (ne-1)" "\ \ ne - t \ right I" "\i. \ \ i = t" then obtain X where "(t, X) \ set auxlist" by (atomize_elim, intro aux(4)) (auto simp: gr0_conv_Suc elim!: order_trans[rotated] intro!: diff_le_mono \_mono) with \\ \ ne - t \ right I\ have "(t, join X pos rel1) \ set auxlist0" unfolding auxlist0_def by (auto elim!: bexI[rotated] intro!: exI[of _ X]) then show "\X. (t, X) \ set auxlist0" by blast qed have auxlist0_Nil: "auxlist0 = [] \ ne = 0 \ ne \ 0 \ (\t. t \ \ \ (ne-1) \ \ \ ne - t \ right I \ (\i. \ \ i = t))" using in_auxlist0_2 by (auto) have aux'_eq: "aux' = add_new_table_msaux args rel2 aux0" using result_eq unfolding aux0_def update_since_def Let_def by simp define auxlist' where auxlist'_eq: "auxlist' = (case auxlist0 of [] \ [(\ \ ne, rel2)] | x # auxlist' \ (if fst x = \ \ ne then (fst x, snd x \ rel2) # auxlist' else (\ \ ne, rel2) # x # auxlist'))" have tabR: "table (args_n args) (args_R args) rel2" using qtable2[unfolded qtable_def] unfolding args_n[symmetric] args_R[symmetric] by simp have valid': "valid_msaux args (\ \ ne) aux' auxlist'" unfolding aux'_eq auxlist'_eq using valid_add_new_table_msaux[OF valid0 tabR] by (auto simp: not_le split: list.splits option.splits if_splits) have sorted_auxlist': "sorted_wrt (\x y. fst x > fst y) auxlist'" unfolding auxlist'_eq using sorted_auxlist0 in_auxlist0_le_\ by (cases auxlist0) fastforce+ have in_auxlist'_1: "t \ \ \ ne \ \ \ ne - t \ right I \ (\i. \ \ i = t) \ qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) ne (Sincep pos \ (point (\ \ ne - t)) \)) X" if auxlist': "(t, X) \ set auxlist'" for t X proof (cases auxlist0) case Nil with auxlist' show ?thesis unfolding auxlist'_eq using qtable2 auxlist0_Nil by (auto simp: zero_enat_def[symmetric] sat_Since_rec[where i=ne] dest: spec[of _ "\ \ (ne-1)"] elim!: qtable_cong[OF _ refl]) next case (Cons a as) show ?thesis proof (cases "t = \ \ ne") case t: True show ?thesis proof (cases "fst a = \ \ ne") case True with auxlist' Cons t have "X = snd a \ rel2" unfolding auxlist'_eq using sorted_auxlist0 by (auto split: if_splits) moreover from in_auxlist0_1[of "fst a" "snd a"] Cons have "ne \ 0" "fst a \ \ \ (ne - 1)" "\ \ ne - fst a \ right I" "\i. \ \ i = fst a" "qtable n (fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) (ne - 1) (Sincep pos \ (point (\ \ (ne - 1) - fst a)) \) \ (if pos then Formula.sat \ V (map the v) ne \ else \ Formula.sat \ V (map the v) ne \)) (snd a)" by (auto simp: True[symmetric] zero_enat_def[symmetric]) ultimately show ?thesis using qtable2 t True by (auto simp: sat_Since_rec[where i=ne] sat.simps(6) elim!: qtable_union) next case False with auxlist' Cons t have "X = rel2" unfolding auxlist'_eq using sorted_auxlist0 in_auxlist0_le_\[of "fst a" "snd a"] by (auto split: if_splits) with auxlist' Cons t False show ?thesis unfolding auxlist'_eq using qtable2 in_auxlist0_2[of "\ \ (ne-1)"] in_auxlist0_le_\[of "fst a" "snd a"] sorted_auxlist0 by (auto simp: sat_Since_rec[where i=ne] sat.simps(3) zero_enat_def[symmetric] enat_0_iff not_le elim!: qtable_cong[OF _ refl] dest!: le_\_less meta_mp) qed next case False with auxlist' Cons have "(t, X) \ set auxlist0" unfolding auxlist'_eq by (auto split: if_splits) then have "ne \ 0" "t \ \ \ (ne - 1)" "\ \ ne - t \ right I" "\i. \ \ i = t" "qtable n (fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) (ne - 1) (Sincep pos \ (point (\ \ (ne - 1) - t)) \) \ (if pos then Formula.sat \ V (map the v) ne \ else \ Formula.sat \ V (map the v) ne \)) X" using in_auxlist0_1 by blast+ with False auxlist' Cons show ?thesis unfolding auxlist'_eq using qtable2 by (fastforce simp: sat_Since_rec[where i=ne] sat.simps(6) diff_diff_right[where i="\ \ ne" and j="\ \ _ + \ \ ne" and k="\ \ (ne - 1)", OF trans_le_add2, simplified] elim!: qtable_cong[OF _ refl] order_trans dest: le_\_less) qed qed have in_auxlist'_2: "\X. (t, X) \ set auxlist'" if "t \ \ \ ne" "\ \ ne - t \ right I" "\i. \ \ i = t" for t proof (cases "t = \ \ ne") case True then show ?thesis proof (cases auxlist0) case Nil with True show ?thesis unfolding auxlist'_eq by (simp add: zero_enat_def[symmetric]) next case (Cons a as) with True show ?thesis unfolding auxlist'_eq by (cases "fst a = \ \ ne") (auto simp: zero_enat_def[symmetric]) qed next case False with that have "ne \ 0" using le_\_less neq0_conv by blast moreover from False that have "t \ \ \ (ne-1)" by (metis One_nat_def Suc_leI Suc_pred \_mono diff_is_0_eq' order.antisym neq0_conv not_le) ultimately obtain X where "(t, X) \ set auxlist0" using \\ \ ne - t \ right I\ \\i. \ \ i = t\ using \_mono[of "ne - 1" "ne" \] by (atomize_elim, cases "right I") (auto intro!: in_auxlist0_2 simp del: \_mono) then show ?thesis unfolding auxlist'_eq using False \\ \ ne - t \ right I\ by (auto intro: exI[of _ X] split: list.split) qed show "wf_since_aux \ V R args \ \ aux' (Suc ne)" unfolding wf_since_aux_def args_ivl args_n args_pos by (auto simp add: fv_sub dest: in_auxlist'_1 intro: sorted_auxlist' in_auxlist'_2 intro!: exI[of _ auxlist'] valid') have "rel = result_msaux args aux'" using result_eq by (auto simp add: update_since_def Let_def) with valid' have rel_eq: "rel = foldr (\) [rel. (t, rel) \ auxlist', left I \ \ \ ne - t] {}" by (auto simp add: args_ivl valid_result_msaux intro!: arg_cong[where f = "\x. foldr (\) (concat x) {}"] split: option.splits) have rel_alt: "rel = (\(t, rel) \ set auxlist'. if left I \ \ \ ne - t then rel else empty_table)" unfolding rel_eq by (auto elim!: in_foldr_UnE bexI[rotated] intro!: in_foldr_UnI) show "qtable n (fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) ne (Sincep pos \ I \)) rel" unfolding rel_alt proof (rule qtable_Union[where Qi="\(t, X) v. left I \ \ \ ne - t \ Formula.sat \ V (map the v) ne (Sincep pos \ (point (\ \ ne - t)) \)"], goal_cases finite qtable equiv) case (equiv v) show ?case proof (rule iffI, erule sat_Since_point, goal_cases left right) case (left j) then show ?case using in_auxlist'_2[of "\ \ j", OF _ _ exI, OF _ _ refl] by auto next case right then show ?case by (auto elim!: sat_Since_pointD dest: in_auxlist'_1) qed qed (auto dest!: in_auxlist'_1 intro!: qtable_empty) qed lemma fv_regex_from_mregex: "ok (length \s) mr \ fv_regex (from_mregex mr \s) \ (\\ \ set \s. fv \)" by (induct mr) (auto simp: Bex_def in_set_conv_nth)+ lemma qtable_\_lax: assumes "ok (length \s) mr" and "list_all2 (\\ rel. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \) rel) \s rels" and "fv_regex (from_mregex mr \s) \ A" and "qtable n A (mem_restr R) Q guard" shows "qtable n A (mem_restr R) (\v. Regex.eps (Formula.sat \ V (map the v)) i (from_mregex mr \s) \ Q v) (\_lax guard rels mr)" using assms proof (induct mr) case (MPlus mr1 mr2) from MPlus(3-6) show ?case by (auto intro!: qtable_union[OF MPlus(1,2)]) next case (MTimes mr1 mr2) then have "fv_regex (from_mregex mr1 \s) \ A" "fv_regex (from_mregex mr2 \s) \ A" using fv_regex_from_mregex[of \s mr1] fv_regex_from_mregex[of \s mr2] by (auto simp: subset_eq) with MTimes(3-6) show ?case by (auto simp: eps_the_restrict restrict_idle intro!: qtable_join[OF MTimes(1,2)]) qed (auto split: prod.splits if_splits simp: qtable_empty_iff list_all2_conv_all_nth in_set_conv_nth restrict_idle sat_the_restrict intro: in_qtableI qtableI elim!: qtable_join[where A=A and C=A]) lemma nullary_qtable_cases: "qtable n {} P Q X \ (X = empty_table \ X = unit_table n)" by (simp add: qtable_def table_empty) lemma qtable_empty_unit_table: "qtable n {} R P empty_table \ qtable n {} R (\v. \ P v) (unit_table n)" by (auto intro: qtable_unit_table simp add: qtable_empty_iff) lemma qtable_unit_empty_table: "qtable n {} R P (unit_table n) \ qtable n {} R (\v. \ P v) empty_table" by (auto intro!: qtable_empty elim: in_qtableE simp add: wf_tuple_empty unit_table_def) lemma qtable_nonempty_empty_table: "qtable n {} R P X \ x \ X \ qtable n {} R (\v. \ P v) empty_table" by (frule nullary_qtable_cases) (auto dest: qtable_unit_empty_table) lemma qtable_r\_strict: assumes "safe_regex Past Strict (from_mregex mr \s)" "ok (length \s) mr" "A = fv_regex (from_mregex mr \s)" and "list_all2 (\\ rel. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \) rel) \s rels" shows "qtable n A (mem_restr R) (\v. Regex.eps (Formula.sat \ V (map the v)) i (from_mregex mr \s)) (r\_strict n rels mr)" using assms proof (hypsubst, induct Past Strict "from_mregex mr \s" arbitrary: mr rule: safe_regex_induct) case (Skip n) then show ?case by (cases mr) (auto simp: qtable_empty_iff qtable_unit_table split: if_splits) next case (Test \) then show ?case by (cases mr) (auto simp: list_all2_conv_all_nth qtable_empty_unit_table dest!: qtable_nonempty_empty_table split: if_splits) next case (Plus r s) then show ?case by (cases mr) (fastforce intro: qtable_union split: if_splits)+ next case (TimesP r s) then show ?case by (cases mr) (auto intro: qtable_cong[OF qtable_\_lax] split: if_splits)+ next case (Star r) then show ?case by (cases mr) (auto simp: qtable_unit_table split: if_splits) qed lemma qtable_l\_strict: assumes "safe_regex Futu Strict (from_mregex mr \s)" "ok (length \s) mr" "A = fv_regex (from_mregex mr \s)" and "list_all2 (\\ rel. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \) rel) \s rels" shows "qtable n A (mem_restr R) (\v. Regex.eps (Formula.sat \ V (map the v)) i (from_mregex mr \s)) (l\_strict n rels mr)" using assms proof (hypsubst, induct Futu Strict "from_mregex mr \s" arbitrary: mr rule: safe_regex_induct) case (Skip n) then show ?case by (cases mr) (auto simp: qtable_empty_iff qtable_unit_table split: if_splits) next case (Test \) then show ?case by (cases mr) (auto simp: list_all2_conv_all_nth qtable_empty_unit_table dest!: qtable_nonempty_empty_table split: if_splits) next case (Plus r s) then show ?case by (cases mr) (fastforce intro: qtable_union split: if_splits)+ next case (TimesF r s) then show ?case by (cases mr) (auto intro: qtable_cong[OF qtable_\_lax] split: if_splits)+ next case (Star r) then show ?case by (cases mr) (auto simp: qtable_unit_table split: if_splits) qed lemma rtranclp_False: "(\i j. False)\<^sup>*\<^sup>* = (=)" proof - have "(\i j. False)\<^sup>*\<^sup>* i j \ i = j" for i j :: 'a by (induct i j rule: rtranclp.induct) auto then show ?thesis by (auto intro: exI[of _ 0]) qed inductive ok_rctxt for \s where "ok_rctxt \s id id" | "ok_rctxt \s \ \' \ ok_rctxt \s (\t. \ (MTimes mr t)) (\t. \' (Regex.Times (from_mregex mr \s) t))" lemma ok_rctxt_swap: "ok_rctxt \s \ \' \ from_mregex (\ mr) \s = \' (from_mregex mr \s)" by (induct \ \' arbitrary: mr rule: ok_rctxt.induct) auto lemma ok_rctxt_cong: "ok_rctxt \s \ \' \ Regex.match (Formula.sat \ V v) r = Regex.match (Formula.sat \ V v) s \ Regex.match (Formula.sat \ V v) (\' r) i j = Regex.match (Formula.sat \ V v) (\' s) i j" by (induct \ \' arbitrary: r s rule: ok_rctxt.induct) simp_all lemma qtable_r\\: assumes "ok (length \s) mr" "fv_regex (from_mregex mr \s) \ A" and "list_all2 (\\ rel. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) j \) rel) \s rels" and "ok_rctxt \s \ \'" and "\ms \ \ ` RPD mr. qtable n A (mem_restr R) (\v. Q (map the v) (from_mregex ms \s)) (lookup rel ms)" shows "qtable n A (mem_restr R) (\v. \s \ Regex.rpd\ \' (Formula.sat \ V (map the v)) j (from_mregex mr \s). Q (map the v) s) (r\ \ rel rels mr)" using assms proof (induct mr arbitrary: \ \') case MSkip then show ?case by (auto simp: rtranclp_False ok_rctxt_swap qtable_empty_iff elim!: qtable_cong[OF _ _ ok_rctxt_cong[of _ \ \']] split: nat.splits) next case (MPlus mr1 mr2) from MPlus(3-7) show ?case by (auto intro!: qtable_union[OF MPlus(1,2)]) next case (MTimes mr1 mr2) from MTimes(3-7) show ?case by (auto intro!: qtable_union[OF MTimes(2) qtable_\_lax[OF _ _ _ MTimes(1)]] elim!: ok_rctxt.intros(2) simp: MTimesL_def Ball_def) next case (MStar mr) from MStar(2-6) show ?case by (auto intro!: qtable_cong[OF MStar(1)] intro: ok_rctxt.intros simp: MTimesL_def Ball_def) qed (auto simp: qtable_empty_iff) lemmas qtable_r\ = qtable_r\\[OF _ _ _ ok_rctxt.intros(1), unfolded rpd\_rpd image_id id_apply] inductive ok_lctxt for \s where "ok_lctxt \s id id" | "ok_lctxt \s \ \' \ ok_lctxt \s (\t. \ (MTimes t mr)) (\t. \' (Regex.Times t (from_mregex mr \s)))" lemma ok_lctxt_swap: "ok_lctxt \s \ \' \ from_mregex (\ mr) \s = \' (from_mregex mr \s)" by (induct \ \' arbitrary: mr rule: ok_lctxt.induct) auto lemma ok_lctxt_cong: "ok_lctxt \s \ \' \ Regex.match (Formula.sat \ V v) r = Regex.match (Formula.sat \ V v) s \ Regex.match (Formula.sat \ V v) (\' r) i j = Regex.match (Formula.sat \ V v) (\' s) i j" by (induct \ \' arbitrary: r s rule: ok_lctxt.induct) simp_all lemma qtable_l\\: assumes "ok (length \s) mr" "fv_regex (from_mregex mr \s) \ A" and "list_all2 (\\ rel. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) j \) rel) \s rels" and "ok_lctxt \s \ \'" and "\ms \ \ ` LPD mr. qtable n A (mem_restr R) (\v. Q (map the v) (from_mregex ms \s)) (lookup rel ms)" shows "qtable n A (mem_restr R) (\v. \s \ Regex.lpd\ \' (Formula.sat \ V (map the v)) j (from_mregex mr \s). Q (map the v) s) (l\ \ rel rels mr)" using assms proof (induct mr arbitrary: \ \') case MSkip then show ?case by (auto simp: rtranclp_False ok_lctxt_swap qtable_empty_iff elim!: qtable_cong[OF _ _ ok_rctxt_cong[of _ \ \']] split: nat.splits) next case (MPlus mr1 mr2) from MPlus(3-7) show ?case by (auto intro!: qtable_union[OF MPlus(1,2)]) next case (MTimes mr1 mr2) from MTimes(3-7) show ?case by (auto intro!: qtable_union[OF MTimes(1) qtable_\_lax[OF _ _ _ MTimes(2)]] elim!: ok_lctxt.intros(2) simp: MTimesR_def Ball_def) next case (MStar mr) from MStar(2-6) show ?case by (auto intro!: qtable_cong[OF MStar(1)] intro: ok_lctxt.intros simp: MTimesR_def Ball_def) qed (auto simp: qtable_empty_iff) lemmas qtable_l\ = qtable_l\\[OF _ _ _ ok_lctxt.intros(1), unfolded lpd\_lpd image_id id_apply] lemma RPD_fv_regex_le: "ms \ RPD mr \ fv_regex (from_mregex ms \s) \ fv_regex (from_mregex mr \s)" by (induct mr arbitrary: ms) (auto simp: MTimesL_def split: nat.splits)+ lemma RPD_safe: "safe_regex Past g (from_mregex mr \s) \ ms \ RPD mr \ safe_regex Past g (from_mregex ms \s)" proof (induct Past g "from_mregex mr \s" arbitrary: mr ms rule: safe_regex_induct) case Skip then show ?case by (cases mr) (auto split: nat.splits) next case (Test g \) then show ?case by (cases mr) auto next case (Plus g r s mrs) then show ?case proof (cases mrs) case (MPlus mr ms) with Plus(3-5) show ?thesis by (auto dest!: Plus(1,2)) qed auto next case (TimesP g r s mrs) then show ?case proof (cases mrs) case (MTimes mr ms) with TimesP(3-5) show ?thesis by (cases g) (auto 0 4 simp: MTimesL_def dest: RPD_fv_regex_le TimesP(1,2)) qed auto next case (Star g r) then show ?case proof (cases mr) case (MStar x6) with Star(2-4) show ?thesis by (cases g) (auto 0 4 simp: MTimesL_def dest: RPD_fv_regex_le elim!: safe_cosafe[rotated] dest!: Star(1)) qed auto qed lemma RPDi_safe: "safe_regex Past g (from_mregex mr \s) \ ms \ RPDi n mr ==> safe_regex Past g (from_mregex ms \s)" by (induct n arbitrary: ms mr) (auto dest: RPD_safe) lemma RPDs_safe: "safe_regex Past g (from_mregex mr \s) \ ms \ RPDs mr ==> safe_regex Past g (from_mregex ms \s)" unfolding RPDs_def by (auto dest: RPDi_safe) lemma RPD_safe_fv_regex: "safe_regex Past Strict (from_mregex mr \s) \ ms \ RPD mr \ fv_regex (from_mregex ms \s) = fv_regex (from_mregex mr \s)" proof (induct Past Strict "from_mregex mr \s" arbitrary: mr rule: safe_regex_induct) case (Skip n) then show ?case by (cases mr) (auto split: nat.splits) next case (Test \) then show ?case by (cases mr) auto next case (Plus r s) then show ?case by (cases mr) auto next case (TimesP r s) then show ?case by (cases mr) (auto 0 3 simp: MTimesL_def dest: RPD_fv_regex_le split: modality.splits) next case (Star r) then show ?case by (cases mr) (auto 0 3 simp: MTimesL_def dest: RPD_fv_regex_le) qed lemma RPDi_safe_fv_regex: "safe_regex Past Strict (from_mregex mr \s) \ ms \ RPDi n mr ==> fv_regex (from_mregex ms \s) = fv_regex (from_mregex mr \s)" by (induct n arbitrary: ms mr) (auto 5 0 dest: RPD_safe_fv_regex RPD_safe) lemma RPDs_safe_fv_regex: "safe_regex Past Strict (from_mregex mr \s) \ ms \ RPDs mr ==> fv_regex (from_mregex ms \s) = fv_regex (from_mregex mr \s)" unfolding RPDs_def by (auto dest: RPDi_safe_fv_regex) lemma RPD_ok: "ok m mr \ ms \ RPD mr \ ok m ms" proof (induct mr arbitrary: ms) case (MPlus mr1 mr2) from MPlus(3,4) show ?case by (auto elim: MPlus(1,2)) next case (MTimes mr1 mr2) from MTimes(3,4) show ?case by (auto elim: MTimes(1,2) simp: MTimesL_def) next case (MStar mr) from MStar(2,3) show ?case by (auto elim: MStar(1) simp: MTimesL_def) qed (auto split: nat.splits) lemma RPDi_ok: "ok m mr \ ms \ RPDi n mr \ ok m ms" by (induct n arbitrary: ms mr) (auto intro: RPD_ok) lemma RPDs_ok: "ok m mr \ ms \ RPDs mr \ ok m ms" unfolding RPDs_def by (auto intro: RPDi_ok) lemma LPD_fv_regex_le: "ms \ LPD mr \ fv_regex (from_mregex ms \s) \ fv_regex (from_mregex mr \s)" by (induct mr arbitrary: ms) (auto simp: MTimesR_def split: nat.splits)+ lemma LPD_safe: "safe_regex Futu g (from_mregex mr \s) \ ms \ LPD mr ==> safe_regex Futu g (from_mregex ms \s)" proof (induct Futu g "from_mregex mr \s" arbitrary: mr ms rule: safe_regex_induct) case Skip then show ?case by (cases mr) (auto split: nat.splits) next case (Test g \) then show ?case by (cases mr) auto next case (Plus g r s mrs) then show ?case proof (cases mrs) case (MPlus mr ms) with Plus(3-5) show ?thesis by (auto dest!: Plus(1,2)) qed auto next case (TimesF g r s mrs) then show ?case proof (cases mrs) case (MTimes mr ms) with TimesF(3-5) show ?thesis by (cases g) (auto 0 4 simp: MTimesR_def dest: LPD_fv_regex_le split: modality.splits dest: TimesF(1,2)) qed auto next case (Star g r) then show ?case proof (cases mr) case (MStar x6) with Star(2-4) show ?thesis by (cases g) (auto 0 4 simp: MTimesR_def dest: LPD_fv_regex_le elim!: safe_cosafe[rotated] dest!: Star(1)) qed auto qed lemma LPDi_safe: "safe_regex Futu g (from_mregex mr \s) \ ms \ LPDi n mr ==> safe_regex Futu g (from_mregex ms \s)" by (induct n arbitrary: ms mr) (auto dest: LPD_safe) lemma LPDs_safe: "safe_regex Futu g (from_mregex mr \s) \ ms \ LPDs mr ==> safe_regex Futu g (from_mregex ms \s)" unfolding LPDs_def by (auto dest: LPDi_safe) lemma LPD_safe_fv_regex: "safe_regex Futu Strict (from_mregex mr \s) \ ms \ LPD mr ==> fv_regex (from_mregex ms \s) = fv_regex (from_mregex mr \s)" proof (induct Futu Strict "from_mregex mr \s" arbitrary: mr rule: safe_regex_induct) case Skip then show ?case by (cases mr) (auto split: nat.splits) next case (Test \) then show ?case by (cases mr) auto next case (Plus r s) then show ?case by (cases mr) auto next case (TimesF r s) then show ?case by (cases mr) (auto 0 3 simp: MTimesR_def dest: LPD_fv_regex_le split: modality.splits) next case (Star r) then show ?case by (cases mr) (auto 0 3 simp: MTimesR_def dest: LPD_fv_regex_le) qed lemma LPDi_safe_fv_regex: "safe_regex Futu Strict (from_mregex mr \s) \ ms \ LPDi n mr ==> fv_regex (from_mregex ms \s) = fv_regex (from_mregex mr \s)" by (induct n arbitrary: ms mr) (auto 5 0 dest: LPD_safe_fv_regex LPD_safe) lemma LPDs_safe_fv_regex: "safe_regex Futu Strict (from_mregex mr \s) \ ms \ LPDs mr ==> fv_regex (from_mregex ms \s) = fv_regex (from_mregex mr \s)" unfolding LPDs_def by (auto dest: LPDi_safe_fv_regex) lemma LPD_ok: "ok m mr \ ms \ LPD mr \ ok m ms" proof (induct mr arbitrary: ms) case (MPlus mr1 mr2) from MPlus(3,4) show ?case by (auto elim: MPlus(1,2)) next case (MTimes mr1 mr2) from MTimes(3,4) show ?case by (auto elim: MTimes(1,2) simp: MTimesR_def) next case (MStar mr) from MStar(2,3) show ?case by (auto elim: MStar(1) simp: MTimesR_def) qed (auto split: nat.splits) lemma LPDi_ok: "ok m mr \ ms \ LPDi n mr \ ok m ms" by (induct n arbitrary: ms mr) (auto intro: LPD_ok) lemma LPDs_ok: "ok m mr \ ms \ LPDs mr \ ok m ms" unfolding LPDs_def by (auto intro: LPDi_ok) lemma update_matchP: assumes pre: "wf_matchP_aux \ V n R I r aux ne" and safe: "safe_regex Past Strict r" and mr: "to_mregex r = (mr, \s)" and mrs: "mrs = sorted_list_of_set (RPDs mr)" and qtables: "list_all2 (\\ rel. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) ne \) rel) \s rels" and result_eq: "(rel, aux') = update_matchP n I mr mrs rels (\ \ ne) aux" shows "wf_matchP_aux \ V n R I r aux' (Suc ne)" and "qtable n (Formula.fv_regex r) (mem_restr R) (\v. Formula.sat \ V (map the v) ne (Formula.MatchP I r)) rel" proof - let ?wf_tuple = "\v. wf_tuple n (Formula.fv_regex r) v" let ?update = "\rel t. mrtabulate mrs (\mr. r\ id rel rels mr \ (if t = \ \ ne then r\_strict n rels mr else {}))" note sat.simps[simp del] define aux0 where "aux0 = [(t, ?update rel t). (t, rel) \ aux, enat (\ \ ne - t) \ right I]" have sorted_aux0: "sorted_wrt (\x y. fst x > fst y) aux0" using pre[unfolded wf_matchP_aux_def, THEN conjunct1] unfolding aux0_def by (induction aux) (auto simp add: sorted_wrt_append) { fix ms assume "ms \ RPDs mr" then have "fv_regex (from_mregex ms \s) = fv_regex r" "safe_regex Past Strict (from_mregex ms \s)" "ok (length \s) ms" "RPD ms \ RPDs mr" using safe RPDs_safe RPDs_safe_fv_regex mr from_mregex_to_mregex RPDs_ok to_mregex_ok RPDs_trans by fastforce+ } note * = this have **: "\ \ ne - (\ \ i + \ \ ne - \ \ (ne - Suc 0)) = \ \ (ne - Suc 0) - \ \ i" for i by (metis (no_types, lifting) Nat.diff_diff_right \_mono add.commute add_diff_cancel_left diff_le_self le_add2 order_trans) have ***: "\ \ i = \ \ ne" if "\ \ ne \ \ \ i" "\ \ i \ \ \ (ne - Suc 0)" "ne > 0" for i by (metis (no_types, lifting) Suc_pred \_mono diff_le_self le_\_less le_antisym not_less_eq that) then have in_aux0_1: "(t, X) \ set aux0 \ ne \ 0 \ t \ \ \ ne \ \ \ ne - t \ right I \ (\i. \ \ i = t) \ (\ms\RPDs mr. qtable n (fv_regex r) (mem_restr R) (\v. Formula.sat \ V (map the v) ne (Formula.MatchP (point (\ \ ne - t)) (from_mregex ms \s))) (lookup X ms))" for t X unfolding aux0_def using safe mr mrs by (auto simp: lookup_tabulate map_of_map_restrict restrict_map_def finite_RPDs * ** RPDs_trans diff_le_mono2 intro!: sat_MatchP_rec[of \ _ _ ne, THEN iffD2] qtable_union[OF qtable_r\[OF _ _ qtables] qtable_r\_strict[OF _ _ _ qtables], of ms "fv_regex r" "\v r. Formula.sat \ V v (ne - Suc 0) (Formula.MatchP (point 0) r)" _ ms for ms] qtable_cong[OF qtable_r\[OF _ _ qtables], of ms "fv_regex r" "\v r. Formula.sat \ V v (ne - Suc 0) (Formula.MatchP (point (\ \ (ne - Suc 0) - \ \ i)) r)" _ _ "(\v. Formula.sat \ V (map the v) ne (Formula.MatchP (point (\ \ ne - \ \ i)) (from_mregex ms \s)))" for ms i] dest!: assms(1)[unfolded wf_matchP_aux_def, THEN conjunct2, THEN conjunct1, rule_format] sat_MatchP_rec["of" \ _ _ ne, THEN iffD1] elim!: bspec order.trans[OF _ \_mono] bexI[rotated] split: option.splits if_splits) (* slow 7 sec *) then have in_aux0_le_\: "(t, X) \ set aux0 \ t \ \ \ ne" for t X by (meson \_mono diff_le_self le_trans) have in_aux0_2: "ne \ 0 \ t \ \ \ (ne-1) \ \ \ ne - t \ right I \ \i. \ \ i = t \ \X. (t, X) \ set aux0" for t proof - fix t assume "ne \ 0" "t \ \ \ (ne-1)" "\ \ ne - t \ right I" "\i. \ \ i = t" then obtain X where "(t, X) \ set aux" by (atomize_elim, intro assms(1)[unfolded wf_matchP_aux_def, THEN conjunct2, THEN conjunct2, rule_format]) (auto simp: gr0_conv_Suc elim!: order_trans[rotated] intro!: diff_le_mono \_mono) with \\ \ ne - t \ right I\ have "(t, ?update X t) \ set aux0" unfolding aux0_def by (auto simp: id_def elim!: bexI[rotated] intro!: exI[of _ X]) then show "\X. (t, X) \ set aux0" by blast qed have aux0_Nil: "aux0 = [] \ ne = 0 \ ne \ 0 \ (\t. t \ \ \ (ne-1) \ \ \ ne - t \ right I \ (\i. \ \ i = t))" using in_aux0_2 by (cases "ne = 0") (auto) have aux'_eq: "aux' = (case aux0 of [] \ [(\ \ ne, mrtabulate mrs (r\_strict n rels))] | x # aux' \ (if fst x = \ \ ne then x # aux' else (\ \ ne, mrtabulate mrs (r\_strict n rels)) # x # aux'))" using result_eq unfolding aux0_def update_matchP_def Let_def by simp have sorted_aux': "sorted_wrt (\x y. fst x > fst y) aux'" unfolding aux'_eq using sorted_aux0 in_aux0_le_\ by (cases aux0) (fastforce)+ have in_aux'_1: "t \ \ \ ne \ \ \ ne - t \ right I \ (\i. \ \ i = t) \ (\ms\RPDs mr. qtable n (Formula.fv_regex r) (mem_restr R) (\v. Formula.sat \ V (map the v) ne (Formula.MatchP (point (\ \ ne - t)) (from_mregex ms \s))) (lookup X ms))" if aux': "(t, X) \ set aux'" for t X proof (cases aux0) case Nil with aux' show ?thesis unfolding aux'_eq using safe mrs qtables aux0_Nil * by (auto simp: zero_enat_def[symmetric] sat_MatchP_rec[where i=ne] lookup_tabulate finite_RPDs split: option.splits intro!: qtable_cong[OF qtable_r\_strict] dest: spec[of _ "\ \ (ne-1)"]) next case (Cons a as) show ?thesis proof (cases "t = \ \ ne") case t: True show ?thesis proof (cases "fst a = \ \ ne") case True with aux' Cons t have "X = snd a" unfolding aux'_eq using sorted_aux0 by auto moreover from in_aux0_1[of "fst a" "snd a"] Cons have "ne \ 0" "fst a \ \ \ ne" "\ \ ne - fst a \ right I" "\i. \ \ i = fst a" "\ms \ RPDs mr. qtable n (fv_regex r) (mem_restr R) (\v. Formula.sat \ V (map the v) ne (Formula.MatchP (point (\ \ ne - fst a)) (from_mregex ms \s))) (lookup (snd a) ms)" by auto ultimately show ?thesis using t True by auto next case False with aux' Cons t have "X = mrtabulate mrs (r\_strict n rels)" unfolding aux'_eq using sorted_aux0 in_aux0_le_\[of "fst a" "snd a"] by auto with aux' Cons t False show ?thesis unfolding aux'_eq using safe mrs qtables * in_aux0_2[of "\ \ (ne-1)"] in_aux0_le_\[of "fst a" "snd a"] sorted_aux0 by (auto simp: sat_MatchP_rec[where i=ne] zero_enat_def[symmetric] enat_0_iff not_le lookup_tabulate finite_RPDs split: option.splits intro!: qtable_cong[OF qtable_r\_strict] dest!: le_\_less meta_mp) qed next case False with aux' Cons have "(t, X) \ set aux0" unfolding aux'_eq by (auto split: if_splits) then have "ne \ 0" "t \ \ \ ne" "\ \ ne - t \ right I" "\i. \ \ i = t" "\ms \ RPDs mr. qtable n (fv_regex r) (mem_restr R) (\v. Formula.sat \ V (map the v) ne (Formula.MatchP (point (\ \ ne - t)) (from_mregex ms \s))) (lookup X ms)" using in_aux0_1 by blast+ with False aux' Cons show ?thesis unfolding aux'_eq by auto qed qed have in_aux'_2: "\X. (t, X) \ set aux'" if "t \ \ \ ne" "\ \ ne - t \ right I" "\i. \ \ i = t" for t proof (cases "t = \ \ ne") case True then show ?thesis proof (cases aux0) case Nil with True show ?thesis unfolding aux'_eq by simp next case (Cons a as) with True show ?thesis unfolding aux'_eq using eq_fst_iff[of t a] by (cases "fst a = \ \ ne") auto qed next case False with that have "ne \ 0" using le_\_less neq0_conv by blast moreover from False that have "t \ \ \ (ne-1)" by (metis One_nat_def Suc_leI Suc_pred \_mono diff_is_0_eq' order.antisym neq0_conv not_le) ultimately obtain X where "(t, X) \ set aux0" using \\ \ ne - t \ right I\ \\i. \ \ i = t\ by atomize_elim (auto intro!: in_aux0_2) then show ?thesis unfolding aux'_eq using False by (auto intro: exI[of _ X] split: list.split) qed show "wf_matchP_aux \ V n R I r aux' (Suc ne)" unfolding wf_matchP_aux_def using mr by (auto dest: in_aux'_1 intro: sorted_aux' in_aux'_2) have rel_eq: "rel = foldr (\) [lookup rel mr. (t, rel) \ aux', left I \ \ \ ne - t] {}" unfolding aux'_eq aux0_def using result_eq by (simp add: update_matchP_def Let_def) have rel_alt: "rel = (\(t, rel) \ set aux'. if left I \ \ \ ne - t then lookup rel mr else empty_table)" unfolding rel_eq by (auto elim!: in_foldr_UnE bexI[rotated] intro!: in_foldr_UnI) show "qtable n (fv_regex r) (mem_restr R) (\v. Formula.sat \ V (map the v) ne (Formula.MatchP I r)) rel" unfolding rel_alt proof (rule qtable_Union[where Qi="\(t, X) v. left I \ \ \ ne - t \ Formula.sat \ V (map the v) ne (Formula.MatchP (point (\ \ ne - t)) r)"], goal_cases finite qtable equiv) case (equiv v) show ?case proof (rule iffI, erule sat_MatchP_point, goal_cases left right) case (left j) then show ?case using in_aux'_2[of "\ \ j", OF _ _ exI, OF _ _ refl] by auto next case right then show ?case by (auto elim!: sat_MatchP_pointD dest: in_aux'_1) qed qed (auto dest!: in_aux'_1 intro!: qtable_empty dest!: bspec[OF _ RPDs_refl] simp: from_mregex_eq[OF safe mr]) qed lemma length_update_until: "length (update_until args rel1 rel2 nt aux) = Suc (length aux)" unfolding update_until_def by simp lemma wf_update_until_auxlist: assumes pre: "wf_until_auxlist \ V n R pos \ I \ auxlist ne" and qtable1: "qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) (ne + length auxlist) \) rel1" and qtable2: "qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) (ne + length auxlist) \) rel2" and fvi_subset: "Formula.fv \ \ Formula.fv \" and args_ivl: "args_ivl args = I" and args_n: "args_n args = n" and args_pos: "args_pos args = pos" shows "wf_until_auxlist \ V n R pos \ I \ (update_until args rel1 rel2 (\ \ (ne + length auxlist)) auxlist) ne" unfolding wf_until_auxlist_def length_update_until unfolding update_until_def list.rel_map add_Suc_right upt.simps eqTrueI[OF le_add1] if_True proof (rule list_all2_appendI, unfold list.rel_map, goal_cases old new) case old show ?case proof (rule list.rel_mono_strong[OF assms(1)[unfolded wf_until_auxlist_def]]; safe, goal_cases mono1 mono2) case (mono1 i X Y v) then show ?case by (fastforce simp: args_ivl args_n args_pos sat_the_restrict less_Suc_eq elim!: qtable_join[OF _ qtable1] qtable_union[OF _ qtable1]) next case (mono2 i X Y v) then show ?case using fvi_subset by (auto 0 3 simp: args_ivl args_n args_pos sat_the_restrict less_Suc_eq split: if_splits elim!: qtable_union[OF _ qtable_join_fixed[OF qtable2]] elim: qtable_cong[OF _ refl] intro: exI[of _ "ne + length auxlist"]) (* slow 8 sec*) qed next case new then show ?case by (auto intro!: qtable_empty qtable1 qtable2[THEN qtable_cong] exI[of _ "ne + length auxlist"] simp: args_ivl args_n args_pos less_Suc_eq zero_enat_def[symmetric]) qed lemma (in muaux) wf_update_until: assumes pre: "wf_until_aux \ V R args \ \ aux ne" and qtable1: "qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) (ne + length_muaux args aux) \) rel1" and qtable2: "qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) (ne + length_muaux args aux) \) rel2" and fvi_subset: "Formula.fv \ \ Formula.fv \" and args_ivl: "args_ivl args = I" and args_n: "args_n args = n" and args_L: "args_L args = Formula.fv \" and args_R: "args_R args = Formula.fv \" and args_pos: "args_pos args = pos" shows "wf_until_aux \ V R args \ \ (add_new_muaux args rel1 rel2 (\ \ (ne + length_muaux args aux)) aux) ne \ length_muaux args (add_new_muaux args rel1 rel2 (\ \ (ne + length_muaux args aux)) aux) = Suc (length_muaux args aux)" proof - from pre obtain cur auxlist where valid_aux: "valid_muaux args cur aux auxlist" and cur: "cur = (if ne + length auxlist = 0 then 0 else \ \ (ne + length auxlist - 1))" and pre_list: "wf_until_auxlist \ V n R pos \ I \ auxlist ne" unfolding wf_until_aux_def args_ivl args_n args_pos by auto have length_aux: "length_muaux args aux = length auxlist" using valid_length_muaux[OF valid_aux] . define nt where "nt \ \ \ (ne + length_muaux args aux)" have nt_mono: "cur \ nt" unfolding cur nt_def length_aux by simp define auxlist' where "auxlist' \ update_until args rel1 rel2 (\ \ (ne + length auxlist)) auxlist" have length_auxlist': "length auxlist' = Suc (length auxlist)" unfolding auxlist'_def by (auto simp add: length_update_until) have tab1: "table (args_n args) (args_L args) rel1" using qtable1 unfolding args_n[symmetric] args_L[symmetric] by (auto simp add: qtable_def) have tab2: "table (args_n args) (args_R args) rel2" using qtable2 unfolding args_n[symmetric] args_R[symmetric] by (auto simp add: qtable_def) have fv_sub: "fv \ \ fv \" using pre unfolding wf_until_aux_def by auto moreover have valid_add_new_auxlist: "valid_muaux args nt (add_new_muaux args rel1 rel2 nt aux) auxlist'" using valid_add_new_muaux[OF valid_aux tab1 tab2 nt_mono] unfolding auxlist'_def nt_def length_aux . moreover have "length_muaux args (add_new_muaux args rel1 rel2 nt aux) = Suc (length_muaux args aux)" using valid_length_muaux[OF valid_add_new_auxlist] unfolding length_auxlist' length_aux[symmetric] . moreover have "wf_until_auxlist \ V n R pos \ I \ auxlist' ne" using wf_update_until_auxlist[OF pre_list qtable1[unfolded length_aux] qtable2[unfolded length_aux] fv_sub args_ivl args_n args_pos] unfolding auxlist'_def . moreover have "\ \ (ne + length auxlist) = (if ne + length auxlist' = 0 then 0 else \ \ (ne + length auxlist' - 1))" unfolding cur length_auxlist' by auto ultimately show ?thesis unfolding wf_until_aux_def nt_def length_aux args_ivl args_n args_pos by fast qed lemma length_update_matchF_base: "length (fst (update_matchF_base I mr mrs nt entry st)) = Suc 0" by (auto simp: Let_def update_matchF_base_def) lemma length_update_matchF_step: "length (fst (update_matchF_step I mr mrs nt entry st)) = Suc (length (fst st))" by (auto simp: Let_def update_matchF_step_def split: prod.splits) lemma length_foldr_update_matchF_step: "length (fst (foldr (update_matchF_step I mr mrs nt) aux base)) = length aux + length (fst base)" by (induct aux arbitrary: base) (auto simp: Let_def length_update_matchF_step) lemma length_update_matchF: "length (update_matchF n I mr mrs rels nt aux) = Suc (length aux)" unfolding update_matchF_def update_matchF_base_def length_foldr_update_matchF_step by (auto simp: Let_def) lemma wf_update_matchF_base_invar: assumes safe: "safe_regex Futu Strict r" and mr: "to_mregex r = (mr, \s)" and mrs: "mrs = sorted_list_of_set (LPDs mr)" and qtables: "list_all2 (\\ rel. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) j \) rel) \s rels" shows "wf_matchF_invar \ V n R I r (update_matchF_base n I mr mrs rels (\ \ j)) j" proof - have from_mregex: "from_mregex mr \s = r" using safe mr using from_mregex_eq by blast { fix ms assume "ms \ LPDs mr" then have "fv_regex (from_mregex ms \s) = fv_regex r" "safe_regex Futu Strict (from_mregex ms \s)" "ok (length \s) ms" "LPD ms \ LPDs mr" using safe LPDs_safe LPDs_safe_fv_regex mr from_mregex_to_mregex LPDs_ok to_mregex_ok LPDs_trans by fastforce+ } note * = this show ?thesis unfolding wf_matchF_invar_def wf_matchF_aux_def update_matchF_base_def mr prod.case Let_def mrs using safe by (auto simp: * from_mregex qtables qtable_empty_iff zero_enat_def[symmetric] lookup_tabulate finite_LPDs eps_match less_Suc_eq LPDs_refl intro!: qtable_cong[OF qtable_l\_strict[where \s=\s]] intro: qtables exI[of _ j] split: option.splits) qed lemma Un_empty_table[simp]: "rel \ empty_table = rel" "empty_table \ rel = rel" unfolding empty_table_def by auto lemma wf_matchF_invar_step: assumes wf: "wf_matchF_invar \ V n R I r st (Suc i)" and safe: "safe_regex Futu Strict r" and mr: "to_mregex r = (mr, \s)" and mrs: "mrs = sorted_list_of_set (LPDs mr)" and qtables: "list_all2 (\\ rel. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \) rel) \s rels" and rel: "qtable n (Formula.fv_regex r) (mem_restr R) (\v. (\j. i \ j \ j < i + length (fst st) \ mem (\ \ j - \ \ i) I \ Regex.match (Formula.sat \ V (map the v)) r i j)) rel" and entry: "entry = (\ \ i, rels, rel)" and nt: "nt = \ \ (i + length (fst st))" shows "wf_matchF_invar \ V n R I r (update_matchF_step I mr mrs nt entry st) i" proof - have from_mregex: "from_mregex mr \s = r" using safe mr using from_mregex_eq by blast { fix ms assume "ms \ LPDs mr" then have "fv_regex (from_mregex ms \s) = fv_regex r" "safe_regex Futu Strict (from_mregex ms \s)" "ok (length \s) ms" "LPD ms \ LPDs mr" using safe LPDs_safe LPDs_safe_fv_regex mr from_mregex_to_mregex LPDs_ok to_mregex_ok LPDs_trans by fastforce+ } note * = this { fix aux X ms assume "st = (aux, X)" "ms \ LPDs mr" with wf mr have "qtable n (fv_regex r) (mem_restr R) (\v. Regex.match (Formula.sat \ V (map the v)) (from_mregex ms \s) i (i + length aux)) (l\ (\x. x) X rels ms)" by (intro qtable_cong[OF qtable_l\[where \s=\s and A="fv_regex r" and Q="\v r. Regex.match (Formula.sat \ V v) r (Suc i) (i + length aux)", OF _ _ qtables]]) (auto simp: wf_matchF_invar_def * LPDs_trans lpd_match[of i] elim!: bspec) } note l\ = this have "lookup (mrtabulate mrs f) ms = f ms" if "ms \ LPDs mr" for ms and f :: "mregex \ 'a table" using that mrs by (fastforce simp: lookup_tabulate finite_LPDs split: option.splits)+ then show ?thesis using wf mr mrs entry nt LPDs_trans by (auto 0 3 simp: Let_def wf_matchF_invar_def update_matchF_step_def wf_matchF_aux_def mr * LPDs_refl list_all2_Cons1 append_eq_Cons_conv upt_eq_Cons_conv Suc_le_eq qtables lookup_tabulate finite_LPDs id_def l\ from_mregex less_Suc_eq intro!: qtable_union[OF rel l\] qtable_cong[OF rel] intro: exI[of _ "i + length _"] split: if_splits prod.splits) qed lemma wf_update_matchF_invar: assumes pre: "wf_matchF_aux \ V n R I r aux ne (length (fst st) - 1)" and wf: "wf_matchF_invar \ V n R I r st (ne + length aux)" and safe: "safe_regex Futu Strict r" and mr: "to_mregex r = (mr, \s)" and mrs: "mrs = sorted_list_of_set (LPDs mr)" and j: "j = ne + length aux + length (fst st) - 1" shows "wf_matchF_invar \ V n R I r (foldr (update_matchF_step I mr mrs (\ \ j)) aux st) ne" using pre wf unfolding j proof (induct aux arbitrary: ne) case (Cons entry aux) from Cons(1)[of "Suc ne"] Cons(2,3) show ?case unfolding foldr.simps o_apply by (intro wf_matchF_invar_step[where rels = "fst (snd entry)" and rel = "snd (snd entry)"]) (auto simp: safe mr mrs wf_matchF_aux_def wf_matchF_invar_def list_all2_Cons1 append_eq_Cons_conv Suc_le_eq upt_eq_Cons_conv length_foldr_update_matchF_step add.assoc split: if_splits) qed simp lemma wf_update_matchF: assumes pre: "wf_matchF_aux \ V n R I r aux ne 0" and safe: "safe_regex Futu Strict r" and mr: "to_mregex r = (mr, \s)" and mrs: "mrs = sorted_list_of_set (LPDs mr)" and nt: "nt = \ \ (ne + length aux)" and qtables: "list_all2 (\\ rel. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) (ne + length aux) \) rel) \s rels" shows "wf_matchF_aux \ V n R I r (update_matchF n I mr mrs rels nt aux) ne 0" unfolding update_matchF_def using wf_update_matchF_base_invar[OF safe mr mrs qtables, of I] unfolding nt by (intro wf_update_matchF_invar[OF _ _ safe mr mrs, unfolded wf_matchF_invar_def split_beta, THEN conjunct2, THEN conjunct1]) (auto simp: length_update_matchF_base wf_matchF_invar_def update_matchF_base_def Let_def pre) lemma wf_until_aux_Cons: "wf_until_auxlist \ V n R pos \ I \ (a # aux) ne \ wf_until_auxlist \ V n R pos \ I \ aux (Suc ne)" unfolding wf_until_auxlist_def by (simp add: upt_conv_Cons del: upt_Suc cong: if_cong) lemma wf_matchF_aux_Cons: "wf_matchF_aux \ V n R I r (entry # aux) ne i \ wf_matchF_aux \ V n R I r aux (Suc ne) i" unfolding wf_matchF_aux_def by (simp add: upt_conv_Cons del: upt_Suc cong: if_cong split: prod.splits) lemma wf_until_aux_Cons1: "wf_until_auxlist \ V n R pos \ I \ ((t, a1, a2) # aux) ne \ t = \ \ ne" unfolding wf_until_auxlist_def by (simp add: upt_conv_Cons del: upt_Suc) lemma wf_matchF_aux_Cons1: "wf_matchF_aux \ V n R I r ((t, rels, rel) # aux) ne i \ t = \ \ ne" unfolding wf_matchF_aux_def by (simp add: upt_conv_Cons del: upt_Suc split: prod.splits) lemma wf_until_aux_Cons3: "wf_until_auxlist \ V n R pos \ I \ ((t, a1, a2) # aux) ne \ qtable n (Formula.fv \) (mem_restr R) (\v. (\j. ne \ j \ j < Suc (ne + length aux) \ mem (\ \ j - \ \ ne) I \ Formula.sat \ V (map the v) j \ \ (\k\{ne.. V (map the v) k \ else \ Formula.sat \ V (map the v) k \))) a2" unfolding wf_until_auxlist_def by (simp add: upt_conv_Cons del: upt_Suc) lemma wf_matchF_aux_Cons3: "wf_matchF_aux \ V n R I r ((t, rels, rel) # aux) ne i \ qtable n (Formula.fv_regex r) (mem_restr R) (\v. (\j. ne \ j \ j < Suc (ne + length aux + i) \ mem (\ \ j - \ \ ne) I \ Regex.match (Formula.sat \ V (map the v)) r ne j)) rel" unfolding wf_matchF_aux_def by (simp add: upt_conv_Cons del: upt_Suc split: prod.splits) lemma upt_append: "a \ b \ b \ c \ [a.. ja'" "jb \ jb'" shows "wf_mbuf2 i ja' jb' P Q (mbuf2_add xs ys buf)" using assms unfolding wf_mbuf2_def by (auto 0 3 simp: list_all2_append2 upt_append dest: list_all2_lengthD intro: exI[where x="[i..j j'. [j..) js js'" shows "wf_mbufn i js' Ps (mbufn_add xss buf)" unfolding wf_mbufn_def list_all3_conv_all_nth proof safe show "length Ps = length js'" "length js' = length (mbufn_add xss buf)" using assms unfolding wf_mbufn_def list_all3_conv_all_nth list_all2_conv_all_nth by auto next fix k assume k: "k < length Ps" then show "i \ js' ! k" using assms unfolding wf_mbufn_def list_all3_conv_all_nth list_all2_conv_all_nth by (auto 0 4 dest: spec[of _ i]) from k have " [i..i z. \x y. P i x \ Q i y \ z = f x y) [i.. xs = [] \ ys = []" "list_all3 P xs [] zs \ xs = [] \ zs = []" "list_all3 P [] ys zs \ ys = [] \ zs = []" unfolding list_all3_conv_all_nth by auto lemma list_all3_Cons: "list_all3 P xs ys (z # zs) \ (\x xs' y ys'. xs = x # xs' \ ys = y # ys' \ P x y z \ list_all3 P xs' ys' zs)" "list_all3 P xs (y # ys) zs \ (\x xs' z zs'. xs = x # xs' \ zs = z # zs' \ P x y z \ list_all3 P xs' ys zs')" "list_all3 P (x # xs) ys zs \ (\y ys' z zs'. ys = y # ys' \ zs = z # zs' \ P x y z \ list_all3 P xs ys' zs')" unfolding list_all3_conv_all_nth by (auto simp: length_Suc_conv Suc_length_conv nth_Cons split: nat.splits) lemma list_all3_mono_strong: "list_all3 P xs ys zs \ (\x y z. x \ set xs \ y \ set ys \ z \ set zs \ P x y z \ Q x y z) \ list_all3 Q xs ys zs" by (induct xs ys zs rule: list_all3.induct) (auto intro: list_all3.intros) definition Mini where "Mini i js = (if js = [] then i else Min (set js))" lemma wf_mbufn_in_set_Mini: assumes "wf_mbufn i js Ps buf" shows "[] \ set buf \ Mini i js = i" unfolding in_set_conv_nth proof (elim exE conjE) fix k have "i \ j" if "j \ set js" for j using that assms unfolding wf_mbufn_def list_all3_conv_all_nth in_set_conv_nth by auto moreover assume "k < length buf" "buf ! k = []" ultimately show ?thesis using assms unfolding Mini_def wf_mbufn_def list_all3_conv_all_nth by (auto 0 4 dest!: spec[of _ k] intro: Min_eqI simp: in_set_conv_nth) qed lemma wf_mbufn_notin_set: assumes "wf_mbufn i js Ps buf" shows "[] \ set buf \ j \ set js \ i < j" using assms unfolding wf_mbufn_def list_all3_conv_all_nth by (cases "i \ set js") (auto intro: le_neq_implies_less simp: in_set_conv_nth) lemma wf_mbufn_map_tl: "wf_mbufn i js Ps buf \ [] \ set buf \ wf_mbufn (Suc i) js Ps (map tl buf)" by (auto simp: wf_mbufn_def list_all3_map Suc_le_eq dest: rel_funD[OF tl_transfer] elim!: list_all3_mono_strong le_neq_implies_less) lemma list_all3_list_all2I: "list_all3 (\x y z. Q x z) xs ys zs \ list_all2 Q xs zs" by (induct xs ys zs rule: list_all3.induct) auto lemma mbuf2t_take_eqD: assumes "mbuf2t_take f z buf nts = (z', buf', nts')" and "wf_mbuf2 i ja jb P Q buf" and "list_all2 R [i.. j" "jb \ j" shows "wf_mbuf2 (min ja jb) ja jb P Q buf'" and "list_all2 R [min ja jb.. set buf") case True from rec.prems(2) have "\j\set js. i \ j" by (auto simp: in_set_conv_nth list_all3_conv_all_nth) moreover from True rec.prems(2) have "i \ set js" by (fastforce simp: in_set_conv_nth list_all3_conv_all_nth list_all2_iff) ultimately have "Mini i js = i" unfolding Mini_def by (auto intro!: antisym[OF Min.coboundedI Min.boundedI]) with rec.prems nonempty True show ?thesis by simp next case False from nonempty rec.prems(2) have "Mini i js = Mini (Suc i) js" unfolding Mini_def by auto show ?thesis unfolding \Mini i js = Mini (Suc i) js\ proof (rule rec.IH) show "\ (buf = [] \ [] \ set buf)" using nonempty False by simp show "list_all3 (\P j xs. Suc i \ j \ list_all2 P [Suc i..k. k \ set js \ k \ j" and "k = Mini (i + length nts) js" shows "wf_mbufn k js Ps buf'" and "list_all2 R [k.. []" using that by (fastforce simp flip: length_greater_0_conv) with 2 show ?case using wf_mbufn_in_set_Mini[OF 2(2)] wf_mbufn_notin_set[OF 2(2)] by (subst (asm) mbufnt_take.simps) (force simp: Mini_def wf_mbufn_def dest!: IH(2)[rotated, OF _ wf_mbufn_map_tl[OF 2(2)] *] split: if_splits prod.splits) qed lemma mbuf2t_take_induct[consumes 5, case_names base step]: assumes "mbuf2t_take f z buf nts = (z', buf', nts')" and "wf_mbuf2 i ja jb P Q buf" and "list_all2 R [i.. j" "jb \ j" and "U i z" and "\k x y t z. i \ k \ Suc k \ ja \ Suc k \ jb \ P k x \ Q k y \ R k t \ U k z \ U (Suc k) (f x y t z)" shows "U (min ja jb) z'" using assms unfolding wf_mbuf2_def by (induction f z buf nts arbitrary: i z' buf' nts' rule: mbuf2t_take.induct) (auto simp add: list_all2_Cons2 upt_eq_Cons_conv less_eq_Suc_le min_absorb1 min_absorb2 elim!: arg_cong2[of _ _ _ _ U, OF _ refl, THEN iffD1, rotated] split: prod.split) lemma list_all2_hdD: assumes "list_all2 P [i.. []" shows "P i (hd xs)" "i < j" using assms unfolding list_all2_conv_all_nth by (auto simp: hd_conv_nth intro: zero_less_diff[THEN iffD1] dest!: spec[of _ 0]) lemma mbufn_take_induct[consumes 3, case_names base step]: assumes "mbufn_take f z buf = (z', buf')" and "wf_mbufn i js Ps buf" and "U i z" and "\k xs z. i \ k \ Suc k \ Mini i js \ list_all2 (\P x. P k x) Ps xs \ U k z \ U (Suc k) (f xs z)" shows "U (Mini i js) z'" using assms unfolding wf_mbufn_def proof (induction f z buf arbitrary: i z' buf' rule: mbufn_take.induct) case rec: (1 f z buf) show ?case proof (cases "buf = []") case True with rec.prems show ?thesis unfolding Mini_def by simp next case nonempty: False show ?thesis proof (cases "[] \ set buf") case True from rec.prems(2) have "\j\set js. i \ j" by (auto simp: in_set_conv_nth list_all3_conv_all_nth) moreover from True rec.prems(2) have "i \ set js" by (fastforce simp: in_set_conv_nth list_all3_conv_all_nth list_all2_iff) ultimately have "Mini i js = i" unfolding Mini_def by (auto intro!: antisym[OF Min.coboundedI Min.boundedI]) with rec.prems nonempty True show ?thesis by simp next case False with nonempty rec.prems(2) have more: "Suc i \ Mini i js" using diff_is_0_eq not_le unfolding Mini_def by (fastforce simp: in_set_conv_nth list_all3_conv_all_nth list_all2_iff) then have "Mini i js = Mini (Suc i) js" unfolding Mini_def by auto show ?thesis unfolding \Mini i js = Mini (Suc i) js\ proof (rule rec.IH) show "\ (buf = [] \ [] \ set buf)" using nonempty False by simp show "mbufn_take f (f (map hd buf) z) (map tl buf) = (z', buf')" using nonempty False rec.prems by simp show "list_all3 (\P j xs. Suc i \ j \ list_all2 P [Suc i..k xs z. Suc i \ k \ Suc k \ Mini (Suc i) js \ list_all2 (\P. P k) Ps xs \ U k z \ U (Suc k) (f xs z)" by (rule rec.prems(4)) (auto simp: Mini_def) qed qed qed qed lemma mbufnt_take_induct[consumes 5, case_names base step]: assumes "mbufnt_take f z buf nts = (z', buf', nts')" and "wf_mbufn i js Ps buf" and "list_all2 R [i..k. k \ set js \ k \ j" and "U i z" and "\k xs t z. i \ k \ Suc k \ Mini j js \ list_all2 (\P x. P k x) Ps xs \ R k t \ U k z \ U (Suc k) (f xs t z)" shows "U (Mini (i + length nts) js) z'" using assms proof (induction f z buf nts arbitrary: i z' buf' nts' rule: mbufnt_take.induct) case (1 f z buf nts) then have *: "list_all2 R [Suc i.. []" "nts = []" using that unfolding wf_mbufn_def using wf_mbufn_in_set_Mini[OF 1(3)] by (fastforce simp: Mini_def list_all3_Cons neq_Nil_conv) with 1(2-7) list_all2_hdD[OF 1(4)] show ?case unfolding wf_mbufn_def using wf_mbufn_in_set_Mini[OF 1(3)] wf_mbufn_notin_set[OF 1(3)] by (subst (asm) mbufnt_take.simps) (auto simp add: Mini_def list.rel_map Suc_le_eq elim!: arg_cong2[of _ _ _ _ U, OF _ refl, THEN iffD1, rotated] list_all3_mono_strong[THEN list_all3_list_all2I[of _ _ js]] list_all2_hdD dest!: 1(1)[rotated, OF _ wf_mbufn_map_tl[OF 1(3)] * _ 1(7)] split: prod.split if_splits) qed lemma mbuf2_take_add': assumes eq: "mbuf2_take f (mbuf2_add xs ys buf) = (zs, buf')" and pre: "wf_mbuf2' \ P V j n R \ \ buf" and rm: "rel_mapping (\) P P'" and xs: "list_all2 (\i. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \)) [progress \ P \ j.. P' \ j'] xs" and ys: "list_all2 (\i. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \)) [progress \ P \ j.. P' \ j'] ys" and "j \ j'" shows "wf_mbuf2' \ P' V j' n R \ \ buf'" and "list_all2 (\i Z. \X Y. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \) X \ qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \) Y \ Z = f X Y) [min (progress \ P \ j) (progress \ P \ j).. P' \ j') (progress \ P' \ j')] zs" using pre rm unfolding wf_mbuf2'_def by (force intro!: mbuf2_take_eqD[OF eq] wf_mbuf2_add[OF _ xs ys] progress_mono_gen[OF \j \ j'\ rm])+ lemma mbuf2t_take_add': assumes eq: "mbuf2t_take f z (mbuf2_add xs ys buf) nts = (z', buf', nts')" and bounded: "pred_mapping (\x. x \ j) P" "pred_mapping (\x. x \ j') P'" and rm: "rel_mapping (\) P P'" and pre_buf: "wf_mbuf2' \ P V j n R \ \ buf" and pre_nts: "list_all2 (\i t. t = \ \ i) [min (progress \ P \ j) (progress \ P \ j)..i. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \)) [progress \ P \ j.. P' \ j'] xs" and ys: "list_all2 (\i. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \)) [progress \ P \ j.. P' \ j'] ys" and "j \ j'" shows "wf_mbuf2' \ P' V j' n R \ \ buf'" and "wf_ts \ P' j' \ \ nts'" using pre_buf pre_nts bounded rm unfolding wf_mbuf2'_def wf_ts_def by (auto intro!: mbuf2t_take_eqD[OF eq] wf_mbuf2_add[OF _ xs ys] progress_mono_gen[OF \j \ j'\ rm] progress_le_gen) lemma ok_0_atms: "ok 0 mr \ regex.atms (from_mregex mr []) = {}" by (induct mr) auto lemma ok_0_progress: "ok 0 mr \ progress_regex \ P (from_mregex mr []) j = j" by (drule ok_0_atms) (auto simp: progress_regex_def) lemma atms_empty_atms: "safe_regex m g r \ atms r = {} \ regex.atms r = {}" by (induct r rule: safe_regex_induct) (auto split: if_splits simp: cases_Neg_iff) lemma atms_empty_progress: "safe_regex m g r \ atms r = {} \ progress_regex \ P r j = j" by (drule atms_empty_atms) (auto simp: progress_regex_def) lemma to_mregex_empty_progress: "safe_regex m g r \ to_mregex r = (mr, []) \ progress_regex \ P r j = j" using from_mregex_eq ok_0_progress to_mregex_ok atms_empty_atms by fastforce lemma progress_regex_le: "pred_mapping (\x. x \ j) P \ progress_regex \ P r j \ j" by (auto intro!: progress_le_gen simp: Min_le_iff progress_regex_def) lemma Neg_acyclic: "formula.Neg x = x \ P" by (induct x) auto lemma case_Neg_in_iff: "x \ (case y of formula.Neg \' \ {\'} | _ \ {}) \ y = formula.Neg x" by (cases y) auto lemma atms_nonempty_progress: "safe_regex m g r \ atms r \ {} \ (\\. progress \ P \ j) ` atms r = (\\. progress \ P \ j) ` regex.atms r" by (frule atms_empty_atms; simp) (auto 0 3 simp: atms_def image_iff case_Neg_in_iff elim!: disjE_Not2 dest: safe_regex_safe_formula) lemma to_mregex_nonempty_progress: "safe_regex m g r \ to_mregex r = (mr, \s) \ \s \ [] \ progress_regex \ P r j = (MIN \\set \s. progress \ P \ j)" using atms_nonempty_progress to_mregex_ok unfolding progress_regex_def by fastforce lemma to_mregex_progress: "safe_regex m g r \ to_mregex r = (mr, \s) \ progress_regex \ P r j = (if \s = [] then j else (MIN \\set \s. progress \ P \ j))" using to_mregex_nonempty_progress to_mregex_empty_progress unfolding progress_regex_def by auto lemma mbufnt_take_add': assumes eq: "mbufnt_take f z (mbufn_add xss buf) nts = (z', buf', nts')" and bounded: "pred_mapping (\x. x \ j) P" "pred_mapping (\x. x \ j') P'" and rm: "rel_mapping (\) P P'" and safe: "safe_regex m g r" and mr: "to_mregex r = (mr, \s)" and pre_buf: "wf_mbufn' \ P V j n R r buf" and pre_nts: "list_all2 (\i t. t = \ \ i) [progress_regex \ P r j..\ i. qtable n (fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \)) \s) (map2 upt (map (\\. progress \ P \ j) \s) (map (\\. progress \ P' \ j') \s)) xss" and "j \ j'" shows "wf_mbufn' \ P' V j' n R r buf'" and "wf_ts_regex \ P' j' r nts'" using pre_buf pre_nts bounded rm mr safe xss \j \ j'\ unfolding wf_mbufn'_def wf_ts_regex_def using atms_empty_progress[of m g r] to_mregex_ok[OF mr] by (auto 0 3 simp: list.rel_map to_mregex_empty_progress to_mregex_nonempty_progress Mini_def intro: progress_mono_gen[OF \j \ j'\ rm] list.rel_refl_strong progress_le_gen dest: list_all2_lengthD elim!: mbufnt_take_eqD[OF eq wf_mbufn_add]) lemma mbuf2t_take_add_induct'[consumes 6, case_names base step]: assumes eq: "mbuf2t_take f z (mbuf2_add xs ys buf) nts = (z', buf', nts')" and bounded: "pred_mapping (\x. x \ j) P" "pred_mapping (\x. x \ j') P'" and rm: "rel_mapping (\) P P'" and pre_buf: "wf_mbuf2' \ P V j n R \ \ buf" and pre_nts: "list_all2 (\i t. t = \ \ i) [min (progress \ P \ j) (progress \ P \ j)..i. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \)) [progress \ P \ j.. P' \ j'] xs" and ys: "list_all2 (\i. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \)) [progress \ P \ j.. P' \ j'] ys" and "j \ j'" and base: "U (min (progress \ P \ j) (progress \ P \ j)) z" and step: "\k X Y z. min (progress \ P \ j) (progress \ P \ j) \ k \ Suc k \ progress \ P' \ j' \ Suc k \ progress \ P' \ j' \ qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) k \) X \ qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) k \) Y \ U k z \ U (Suc k) (f X Y (\ \ k) z)" shows "U (min (progress \ P' \ j') (progress \ P' \ j')) z'" using pre_buf pre_nts bounded rm unfolding wf_mbuf2'_def by (blast intro!: mbuf2t_take_induct[OF eq] wf_mbuf2_add[OF _ xs ys] progress_mono_gen[OF \j \ j'\ rm] progress_le_gen base step) lemma mbufnt_take_add_induct'[consumes 6, case_names base step]: assumes eq: "mbufnt_take f z (mbufn_add xss buf) nts = (z', buf', nts')" and bounded: "pred_mapping (\x. x \ j) P" "pred_mapping (\x. x \ j') P'" and rm: "rel_mapping (\) P P'" and safe: "safe_regex m g r" and mr: "to_mregex r = (mr, \s)" and pre_buf: "wf_mbufn' \ P V j n R r buf" and pre_nts: "list_all2 (\i t. t = \ \ i) [progress_regex \ P r j..\ i. qtable n (fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \)) \s) (map2 upt (map (\\. progress \ P \ j) \s) (map (\\. progress \ P' \ j') \s)) xss" and "j \ j'" and base: "U (progress_regex \ P r j) z" and step: "\k Xs z. progress_regex \ P r j \ k \ Suc k \ progress_regex \ P' r j' \ list_all2 (\\. qtable n (Formula.fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) k \)) \s Xs \ U k z \ U (Suc k) (f Xs (\ \ k) z)" shows "U (progress_regex \ P' r j') z'" using pre_buf pre_nts bounded rm \j \ j'\ to_mregex_progress[OF safe mr, of \ P' j'] to_mregex_empty_progress[OF safe, of mr \ P j] base unfolding wf_mbufn'_def mr prod.case by (fastforce dest!: mbufnt_take_induct[OF eq wf_mbufn_add[OF _ xss] pre_nts, of U] simp: list.rel_map le_imp_diff_is_add ac_simps Mini_def intro: progress_mono_gen[OF \j \ j'\ rm] list.rel_refl_strong progress_le_gen intro!: base step dest: list_all2_lengthD split: if_splits) lemma progress_Until_le: "progress \ P (Formula.Until \ I \) j \ min (progress \ P \ j) (progress \ P \ j)" by (cases "right I") (auto simp: trans_le_add1 intro!: cInf_lower) lemma progress_MatchF_le: "progress \ P (Formula.MatchF I r) j \ progress_regex \ P r j" by (cases "right I") (auto simp: trans_le_add1 progress_regex_def intro!: cInf_lower) lemma list_all2_upt_Cons: "P a x \ list_all2 P [Suc a.. Suc a \ b \ list_all2 P [a.. list_all2 P [b.. a \ b \ b \ c \ list_all2 P [a..x. R x x) xs ys" by (auto simp: list_all3_conv_all_nth list_all2_conv_all_nth) lemma map_split_map: "map_split f (map g xs) = map_split (f o g) xs" by (induct xs) auto lemma map_split_alt: "map_split f xs = (map (fst o f) xs, map (snd o f) xs)" by (induct xs) (auto split: prod.splits) lemma fv_formula_of_constraint: "fv (formula_of_constraint (t1, p, c, t2)) = fv_trm t1 \ fv_trm t2" by (induction "(t1, p, c, t2)" rule: formula_of_constraint.induct) simp_all lemma (in maux) wf_mformula_wf_set: "wf_mformula \ j P V n R \ \' \ wf_set n (Formula.fv \')" unfolding wf_set_def proof (induction rule: wf_mformula.induct) case (AndRel P V n R \ \' \' conf) then show ?case by (auto simp: fv_formula_of_constraint dest!: subsetD) next case (Ands P V n R l l_pos l_neg l' buf A_pos A_neg) from Ands.IH have "\\'\set (l_pos @ map remove_neg l_neg). \x\fv \'. x < n" by (simp add: list_all2_conv_all_nth all_set_conv_all_nth[of "_ @ _"] del: set_append) then have "\\'\set (l_pos @ l_neg). \x\fv \'. x < n" by (auto dest: bspec[where x="remove_neg _"]) then show ?case using Ands.hyps(2) by auto next case (Agg P V b n R \ \' y f g0 \) then have "Formula.fvi_trm b f \ Formula.fvi b \'" by (auto simp: fvi_trm_iff_fv_trm[where b=b] fvi_iff_fv[where b=b]) with Agg show ?case by (auto 0 3 simp: Un_absorb2 fvi_iff_fv[where b=b]) next case (MatchP r P V n R \s mr mrs buf nts I aux) then obtain \s' where conv: "to_mregex r = (mr, \s')" by blast with MatchP have "\\'\set \s'. \x\fv \'. x < n" by (simp add: list_all2_conv_all_nth all_set_conv_all_nth[of \s']) with conv show ?case by (simp add: to_mregex_ok[THEN conjunct1] fv_regex_alt[OF \safe_regex _ _ r\]) next case (MatchF r P V n R \s mr mrs buf nts I aux) then obtain \s' where conv: "to_mregex r = (mr, \s')" by blast with MatchF have "\\'\set \s'. \x\fv \'. x < n" by (simp add: list_all2_conv_all_nth all_set_conv_all_nth[of \s']) with conv show ?case by (simp add: to_mregex_ok[THEN conjunct1] fv_regex_alt[OF \safe_regex _ _ r\]) qed (auto simp: fvi_Suc split: if_splits) lemma qtable_mmulti_join: assumes pos: "list_all3 (\A Qi X. qtable n A P Qi X \ wf_set n A) A_pos Q_pos L_pos" and neg: "list_all3 (\A Qi X. qtable n A P Qi X \ wf_set n A) A_neg Q_neg L_neg" and C_eq: "C = \(set A_pos)" and L_eq: "L = L_pos @ L_neg" and "A_pos \ []" and fv_subset: "\(set A_neg) \ \(set A_pos)" and restrict_pos: "\x. wf_tuple n C x \ P x \ list_all (\A. P (restrict A x)) A_pos" and restrict_neg: "\x. wf_tuple n C x \ P x \ list_all (\A. P (restrict A x)) A_neg" and Qs: "\x. wf_tuple n C x \ P x \ Q x \ list_all2 (\A Qi. Qi (restrict A x)) A_pos Q_pos \ list_all2 (\A Qi. \ Qi (restrict A x)) A_neg Q_neg" shows "qtable n C P Q (mmulti_join n A_pos A_neg L)" proof (rule qtableI) from pos have 1: "list_all2 (\A X. table n A X \ wf_set n A) A_pos L_pos" by (simp add: list_all3_conv_all_nth list_all2_conv_all_nth qtable_def) moreover from neg have "list_all2 (\A X. table n A X \ wf_set n A) A_neg L_neg" by (simp add: list_all3_conv_all_nth list_all2_conv_all_nth qtable_def) ultimately have L: "list_all2 (\A X. table n A X \ wf_set n A) (A_pos @ A_neg) (L_pos @ L_neg)" by (rule list_all2_appendI) note in_join_iff = mmulti_join_correct[OF \A_pos \ []\ L] from 1 have take_eq: "take (length A_pos) (L_pos @ L_neg) = L_pos" by (auto dest!: list_all2_lengthD) from 1 have drop_eq: "drop (length A_pos) (L_pos @ L_neg) = L_neg" by (auto dest!: list_all2_lengthD) note mmulti_join.simps[simp del] show "table n C (mmulti_join n A_pos A_neg L)" unfolding C_eq L_eq table_def by (simp add: in_join_iff) show "Q x" if "x \ mmulti_join n A_pos A_neg L" "wf_tuple n C x" "P x" for x using that(2,3) proof (rule Qs[THEN iffD2, OF _ _ conjI]) have pos': "list_all2 (\A. (\) (restrict A x)) A_pos L_pos" and neg': "list_all2 (\A. (\) (restrict A x)) A_neg L_neg" using that(1) unfolding L_eq in_join_iff take_eq drop_eq by simp_all show "list_all2 (\A Qi. Qi (restrict A x)) A_pos Q_pos" using pos pos' restrict_pos that(2,3) by (simp add: list_all3_conv_all_nth list_all2_conv_all_nth list_all_length qtable_def) have fv_subset': "\i. i < length A_neg \ A_neg ! i \ C" using fv_subset unfolding C_eq by (auto simp: Sup_le_iff) show "list_all2 (\A Qi. \ Qi (restrict A x)) A_neg Q_neg" using neg neg' restrict_neg that(2,3) by (auto simp: list_all3_conv_all_nth list_all2_conv_all_nth list_all_length qtable_def wf_tuple_restrict_simple[OF _ fv_subset']) qed show "x \ mmulti_join n A_pos A_neg L" if "wf_tuple n C x" "P x" "Q x" for x unfolding L_eq in_join_iff take_eq drop_eq proof (intro conjI) from that have pos': "list_all2 (\A Qi. Qi (restrict A x)) A_pos Q_pos" and neg': "list_all2 (\A Qi. \ Qi (restrict A x)) A_neg Q_neg" using Qs[THEN iffD1] by auto show "wf_tuple n (\A\set A_pos. A) x" using \wf_tuple n C x\ unfolding C_eq by simp show "list_all2 (\A. (\) (restrict A x)) A_pos L_pos" using pos pos' restrict_pos that(1,2) by (simp add: list_all3_conv_all_nth list_all2_conv_all_nth list_all_length qtable_def C_eq wf_tuple_restrict_simple[OF _ Sup_upper]) show "list_all2 (\A. (\) (restrict A x)) A_neg L_neg" using neg neg' restrict_neg that(1,2) by (auto simp: list_all3_conv_all_nth list_all2_conv_all_nth list_all_length qtable_def) qed qed lemma nth_filter: "i < length (filter P xs) \ (\i'. i' < length xs \ P (xs ! i') \ Q (xs ! i')) \ Q (filter P xs ! i)" by (metis (lifting) in_set_conv_nth set_filter mem_Collect_eq) lemma nth_partition: "i < length xs \ (\i'. i' < length (filter P xs) \ Q (filter P xs ! i')) \ (\i'. i' < length (filter (Not \ P) xs) \ Q (filter (Not \ P) xs ! i')) \ Q (xs ! i)" by (metis (no_types, lifting) in_set_conv_nth set_filter mem_Collect_eq comp_apply) lemma qtable_bin_join: assumes "qtable n A P Q1 X" "qtable n B P Q2 Y" "\ b \ B \ A" "C = A \ B" "\x. wf_tuple n C x \ P x \ P (restrict A x) \ P (restrict B x)" "\x. b \ wf_tuple n C x \ P x \ Q x \ Q1 (restrict A x) \ Q2 (restrict B x)" "\x. \ b \ wf_tuple n C x \ P x \ Q x \ Q1 (restrict A x) \ \ Q2 (restrict B x)" shows "qtable n C P Q (bin_join n A X b B Y)" using qtable_join[OF assms] bin_join_table[of n A X B Y b] assms(1,2) by (auto simp add: qtable_def) lemma restrict_update: "y \ A \ y < length x \ restrict A (x[y:=z]) = restrict A x" unfolding restrict_def by (auto simp add: nth_list_update) lemma qtable_assign: assumes "qtable n A P Q X" "y < n" "insert y A = A'" "y \ A" "\x'. wf_tuple n A' x' \ P x' \ P (restrict A x')" "\x. wf_tuple n A x \ P x \ Q x \ Q' (x[y:=Some (f x)])" "\x'. wf_tuple n A' x' \ P x' \ Q' x' \ Q (restrict A x') \ x' ! y = Some (f (restrict A x'))" shows "qtable n A' P Q' ((\x. x[y:=Some (f x)]) ` X)" (is "qtable _ _ _ _ ?Y") proof (rule qtableI) from assms(1) have "table n A X" unfolding qtable_def by simp then show "table n A' ?Y" unfolding table_def wf_tuple_def using assms(2,3) by (auto simp: nth_list_update) next fix x' assume "x' \ ?Y" "wf_tuple n A' x'" "P x'" then obtain x where "x \ X" and x'_eq: "x' = x[y:=Some (f x)]" by blast then have "wf_tuple n A x" using assms(1) unfolding qtable_def table_def by blast then have "y < length x" using assms(2) by (simp add: wf_tuple_def) with \wf_tuple n A x\ have "restrict A x' = x" unfolding x'_eq by (simp add: restrict_update[OF assms(4)] restrict_idle) with \wf_tuple n A' x'\ \P x'\ have "P x" using assms(5) by blast with \wf_tuple n A x\ \x \ X\ have "Q x" using assms(1) by (elim in_qtableE) with \wf_tuple n A x\ \P x\ show "Q' x'" unfolding x'_eq by (rule assms(6)) next fix x' assume "wf_tuple n A' x'" "P x'" "Q' x'" then have "wf_tuple n A (restrict A x')" using assms(3) by (auto intro!: wf_tuple_restrict_simple) moreover have "P (restrict A x')" using \wf_tuple n A' x'\ \P x'\ by (rule assms(5)) moreover have "Q (restrict A x')" and y: "x' ! y = Some (f (restrict A x'))" using \wf_tuple n A' x'\ \P x'\ \Q' x'\ by (auto dest!: assms(7)) ultimately have "restrict A x' \ X" by (intro in_qtableI[OF assms(1)]) moreover have "x' = (restrict A x')[y:=Some (f (restrict A x'))]" using y assms(2,3) \wf_tuple n A (restrict A x')\ \wf_tuple n A' x'\ by (auto simp: list_eq_iff_nth_eq wf_tuple_def nth_list_update nth_restrict) ultimately show "x' \ ?Y" by simp qed lemma sat_the_update: "y \ fv \ \ Formula.sat \ V (map the (x[y:=z])) i \ = Formula.sat \ V (map the x) i \" by (rule sat_fv_cong) (metis map_update nth_list_update_neq) lemma progress_constraint: "progress \ P (formula_of_constraint c) j = j" by (induction rule: formula_of_constraint.induct) simp_all lemma qtable_filter: assumes "qtable n A P Q X" "\x. wf_tuple n A x \ P x \ Q x \ R x \ Q' x" shows "qtable n A P Q' (Set.filter R X)" (is "qtable _ _ _ _ ?Y") proof (rule qtableI) from assms(1) have "table n A X" unfolding qtable_def by simp then show "table n A ?Y" unfolding table_def wf_tuple_def by simp next fix x assume "x \ ?Y" "wf_tuple n A x" "P x" with assms show "Q' x" by (auto elim!: in_qtableE) next fix x assume "wf_tuple n A x" "P x" "Q' x" with assms show "x \ Set.filter R X" by (auto intro!: in_qtableI) qed lemma eval_constraint_sat_eq: "wf_tuple n A x \ fv_trm t1 \ A \ fv_trm t2 \ A \ \i\A. i < n \ eval_constraint (t1, p, c, t2) x = Formula.sat \ V (map the x) i (formula_of_constraint (t1, p, c, t2))" by (induction "(t1, p, c, t2)" rule: formula_of_constraint.induct) (simp_all add: meval_trm_eval_trm) declare progress_le_gen[simp] definition "wf_envs \ j P P' V db = (dom V = dom P \ Mapping.keys db = dom P \ {p. p \ fst ` \ \ j} \ rel_mapping (\) P P' \ pred_mapping (\i. i \ j) P \ pred_mapping (\i. i \ Suc j) P' \ (\p \ Mapping.keys db - dom P. the (Mapping.lookup db p) = [{ts. (p, ts) \ \ \ j}]) \ (\p \ dom P. list_all2 (\i X. X = the (V p) i) [the (P p).. event_data list) set \ Formula.database" is "\X p. (if p \ fst ` X then Some [{ts. (p, ts) \ X}] else None)" . lemma wf_envs_mk_db: "wf_envs \ j Map.empty Map.empty Map.empty (mk_db (\ \ j))" unfolding wf_envs_def mk_db_def by transfer (force split: if_splits simp: image_iff rel_mapping_alt) lemma wf_envs_update: assumes wf_envs: "wf_envs \ j P P' V db" and m_eq: "m = Formula.nfv \" and in_fv: "{0 ..< m} \ fv \" and xs: "list_all2 (\i. qtable m (Formula.fv \) (mem_restr UNIV) (\v. Formula.sat \ V (map the v) i \)) [progress \ P \ j.. P' \ (Suc j)] xs" shows "wf_envs \ j (P(p \ progress \ P \ j)) (P'(p \ progress \ P' \ (Suc j))) (V(p \ \i. {v. length v = m \ Formula.sat \ V v i \})) (Mapping.update p (map (image (map the)) xs) db)" unfolding wf_envs_def proof (intro conjI ballI, goal_cases) case 3 from assms show ?case by (auto simp: wf_envs_def pred_mapping_alt progress_le progress_mono_gen intro!: rel_mapping_map_upd) next case (6 p') with assms show ?case by (cases "p' \ dom P") (auto simp: wf_envs_def lookup_update') next case (7 p') from xs in_fv have "list_all2 (\x y. map the ` y = {v. length v = m \ Formula.sat \ V v x \}) [progress \ P \ j.. P' \ (Suc j)] xs" by (elim list.rel_mono_strong) (auto 0 3 simp: wf_tuple_def nth_append elim!: in_qtableE in_qtableI intro!: image_eqI[where x="map Some _"]) moreover have "list_all2 (\i X. X = the (V p') i) [the (P p').. p'" proof - from that 7 have "p' \ dom P" by simp with wf_envs show ?thesis by (simp add: wf_envs_def) qed ultimately show ?case by (simp add: list.rel_map image_iff lookup_update') qed (use assms in \auto simp: wf_envs_def\) lemma wf_envs_P_simps[simp]: "wf_envs \ j P P' V db \ pred_mapping (\i. i \ j) P" "wf_envs \ j P P' V db \ pred_mapping (\i. i \ Suc j) P'" "wf_envs \ j P P' V db \ rel_mapping (\) P P'" unfolding wf_envs_def by auto lemma wf_envs_progress_le[simp]: "wf_envs \ j P P' V db \ progress \ P \ j \ j" "wf_envs \ j P P' V db \ progress \ P' \ (Suc j) \ Suc j" unfolding wf_envs_def by auto lemma wf_envs_progress_regex_le[simp]: "wf_envs \ j P P' V db \ progress_regex \ P r j \ j" "wf_envs \ j P P' V db \ progress_regex \ P' r (Suc j) \ Suc j" unfolding wf_envs_def by (auto simp: progress_regex_le) lemma wf_envs_progress_mono[simp]: "wf_envs \ j P P' V db \ a \ b \ progress \ P \ a \ progress \ P' \ b" unfolding wf_envs_def by (auto simp: progress_mono_gen) lemma qtable_wf_tuple_cong: "qtable n A P Q X \ A = B \ (\v. wf_tuple n A v \ P v \ Q v = Q' v) \ qtable n B P Q' X" unfolding qtable_def table_def by blast lemma (in maux) meval: assumes "wf_mformula \ j P V n R \ \'" "wf_envs \ j P P' V db" shows "case meval n (\ \ j) db \ of (xs, \\<^sub>n) \ wf_mformula \ (Suc j) P' V n R \\<^sub>n \' \ list_all2 (\i. qtable n (Formula.fv \') (mem_restr R) (\v. Formula.sat \ V (map the v) i \')) [progress \ P \' j.. P' \' (Suc j)] xs" using assms proof (induction \ arbitrary: db P P' V n R \') case (MRel rel) then show ?case by (cases rule: wf_mformula.cases) (auto simp add: ball_Un intro: wf_mformula.intros table_eq_rel eq_rel_eval_trm in_eq_rel qtable_empty qtable_unit_table intro!: qtableI) next case (MPred e ts) then show ?case proof (cases "e \ dom P") case True with MPred(2) have "e \ Mapping.keys db" "e \ dom P'" "e \ dom V" "list_all2 (\i X. X = the (V e) i) [the (P e).. dom P'" "e \ dom V" unfolding wf_envs_def rel_mapping_alt by auto moreover from False MPred(2) have *: "e \ fst ` \ \ j \ e \ Mapping.keys db" unfolding wf_envs_def by auto from False MPred(2) have "e \ Mapping.keys db \ Mapping.lookup db e = Some [{ts. (e, ts) \ \ \ j}]" unfolding wf_envs_def keys_dom_lookup by (metis Diff_iff domD option.sel) with * have "(case Mapping.lookup db e of None \ [{}] | Some xs \ xs) = [{ts. (e, ts) \ \ \ j}]" by (cases "e \ fst ` \ \ j") (auto simp: image_iff keys_dom_lookup split: option.splits) ultimately show ?thesis by (cases rule: wf_mformula.cases) (fastforce intro!: wf_mformula.Pred qtableI bexI[where P="\x. _ = tabulate x 0 n", OF refl] elim!: list.rel_mono_strong bexI[rotated] dest: ex_match simp: list.rel_map table_def match_wf_tuple in_these_eq match_eval_trm image_iff list.map_comp) qed next case (MLet p m \1 \2) from MLet.prems(1) obtain \1' \2' where Let: "\' = Formula.Let p \1' \2'" and 1: "wf_mformula \ j P V m UNIV \1 \1'" and fv: "m = Formula.nfv \1'" "{0.. fv \1'" and 2: "wf_mformula \ j (P(p \ progress \ P \1' j)) (V(p \ \i. {v. length v = m \ Formula.sat \ V v i \1'})) n R \2 \2'" by (cases rule: wf_mformula.cases) auto obtain xs \1_new where e1: "meval m (\ \ j) db \1 = (xs, \1_new)" and wf1: "wf_mformula \ (Suc j) P' V m UNIV \1_new \1'" and res1: "list_all2 (\i. qtable m (fv \1') (mem_restr UNIV) (\v. Formula.sat \ V (map the v) i \1')) [progress \ P \1' j.. P' \1' (Suc j)] xs" using MLet(1)[OF 1(1) MLet.prems(2)] by (auto simp: eqTrueI[OF mem_restr_UNIV, abs_def]) from MLet(2)[OF 2 wf_envs_update[OF MLet.prems(2) fv res1]] wf1 e1 fv show ?case unfolding Let by (auto simp: fun_upd_def intro!: wf_mformula.Let) next case (MAnd A_\ \ pos A_\ \ buf) from MAnd.prems show ?case by (cases rule: wf_mformula.cases) (auto simp: sat_the_restrict simp del: bin_join.simps dest!: MAnd.IH split: if_splits prod.splits intro!: wf_mformula.And qtable_bin_join elim: mbuf2_take_add'(1) list.rel_mono_strong[OF mbuf2_take_add'(2)]) next case (MAndAssign \ conf) from MAndAssign.prems obtain \'' x t \'' where wf_envs: "wf_envs \ j P P' V db" and \'_eq: "\' = formula.And \'' \''" and wf_\: "wf_mformula \ j P V n R \ \''" and "x < n" and "x \ fv \''" and fv_t_subset: "fv_trm t \ fv \''" and conf: "(x, t) = conf" and \''_eqs: "\'' = formula.Eq (trm.Var x) t \ \'' = formula.Eq t (trm.Var x)" by (cases rule: wf_mformula.cases) from wf_\ wf_envs obtain xs \\<^sub>n where meval_eq: "meval n (\ \ j) db \ = (xs, \\<^sub>n)" and wf_\\<^sub>n: "wf_mformula \ (Suc j) P' V n R \\<^sub>n \''" and xs: "list_all2 (\i. qtable n (fv \'') (mem_restr R) (\v. Formula.sat \ V (map the v) i \'')) [progress \ P \'' j.. P' \'' (Suc j)] xs" by (auto dest!: MAndAssign.IH) have progress_eqs: "progress \ P \' j = progress \ P \'' j" "progress \ P' \' (Suc j) = progress \ P' \'' (Suc j)" using \''_eqs wf_envs_progress_le[OF wf_envs] by (auto simp: \'_eq) show ?case proof (simp add: meval_eq, intro conjI) show "wf_mformula \ (Suc j) P' V n R (MAndAssign \\<^sub>n conf) \'" unfolding \'_eq by (rule wf_mformula.AndAssign) fact+ next show "list_all2 (\i. qtable n (fv \') (mem_restr R) (\v. Formula.sat \ V (map the v) i \')) [progress \ P \' j.. P' \' (Suc j)] (map ((`) (eval_assignment conf)) xs)" unfolding list.rel_map progress_eqs conf[symmetric] eval_assignment.simps using xs proof (rule list.rel_mono_strong) fix i X assume qtable: "qtable n (fv \'') (mem_restr R) (\v. Formula.sat \ V (map the v) i \'') X" then show "qtable n (fv \') (mem_restr R) (\v. Formula.sat \ V (map the v) i \') ((\y. y[x := Some (meval_trm t y)]) ` X)" proof (rule qtable_assign) show "x < n" by fact show "insert x (fv \'') = fv \'" using \''_eqs fv_t_subset by (auto simp: \'_eq) show "x \ fv \''" by fact next fix v assume wf_v: "wf_tuple n (fv \') v" and "mem_restr R v" then show "mem_restr R (restrict (fv \'') v)" by simp assume sat: "Formula.sat \ V (map the v) i \'" then have A: "Formula.sat \ V (map the (restrict (fv \'') v)) i \''" (is ?A) by (simp add: \'_eq sat_the_restrict) have "map the v ! x = Formula.eval_trm (map the v) t" using sat \''_eqs by (auto simp: \'_eq) also have "... = Formula.eval_trm (map the (restrict (fv \'') v)) t" using fv_t_subset by (auto simp: map_the_restrict intro!: eval_trm_fv_cong) finally have "map the v ! x = meval_trm t (restrict (fv \'') v)" using meval_trm_eval_trm[of n "fv \''" "restrict (fv \'') v" t] fv_t_subset wf_v wf_mformula_wf_set[unfolded wf_set_def, OF wf_\] by (fastforce simp: \'_eq intro!: wf_tuple_restrict) then have B: "v ! x = Some (meval_trm t (restrict (fv \'') v))" (is ?B) using \''_eqs wf_v \x < n\ by (auto simp: wf_tuple_def \'_eq) from A B show "?A \ ?B" .. next fix v assume wf_v: "wf_tuple n (fv \'') v" and "mem_restr R v" and sat: "Formula.sat \ V (map the v) i \''" let ?v = "v[x := Some (meval_trm t v)]" from sat have A: "Formula.sat \ V (map the ?v) i \''" using \x \ fv \''\ by (simp add: sat_the_update) have "y \ fv_trm t \ x \ y" for y using fv_t_subset \x \ fv \''\ by auto then have B: "Formula.sat \ V (map the ?v) i \''" using \''_eqs meval_trm_eval_trm[of n "fv \''" v t] \x < n\ fv_t_subset wf_v wf_mformula_wf_set[unfolded wf_set_def, OF wf_\] by (auto simp: wf_tuple_def map_update intro!: eval_trm_fv_cong) from A B show "Formula.sat \ V (map the ?v) i \'" by (simp add: \'_eq) qed qed qed next case (MAndRel \ conf) from MAndRel.prems show ?case by (cases rule: wf_mformula.cases) (auto simp: progress_constraint progress_le list.rel_map fv_formula_of_constraint Un_absorb2 wf_mformula_wf_set[unfolded wf_set_def] split: prod.splits dest!: MAndRel.IH[where db=db and P=P and P'=P'] eval_constraint_sat_eq[THEN iffD2] intro!: wf_mformula.AndRel elim!: list.rel_mono_strong qtable_filter eval_constraint_sat_eq[THEN iffD1]) next case (MAnds A_pos A_neg l buf) note mbufn_take.simps[simp del] mbufn_add.simps[simp del] mmulti_join.simps[simp del] from MAnds.prems obtain pos neg l' where wf_l: "list_all2 (wf_mformula \ j P V n R) l (pos @ map remove_neg neg)" and wf_buf: "wf_mbufn (progress \ P (formula.Ands l') j) (map (\\. progress \ P \ j) (pos @ map remove_neg neg)) (map (\\ i. qtable n (fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \)) (pos @ map remove_neg neg)) buf" and posneg: "(pos, neg) = partition safe_formula l'" and "pos \ []" and safe_neg: "list_all safe_formula (map remove_neg neg)" and A_eq: "A_pos = map fv pos" "A_neg = map fv neg" and fv_subset: "\ (set A_neg) \ \ (set A_pos)" and "\' = Formula.Ands l'" by (cases rule: wf_mformula.cases) simp have progress_eq: "progress \ P' (formula.Ands l') (Suc j) = Mini (progress \ P (formula.Ands l') j) (map (\\. progress \ P' \ (Suc j)) (pos @ map remove_neg neg))" using \pos \ []\ posneg by (auto simp: Mini_def image_Un[symmetric] Collect_disj_eq[symmetric] intro!: arg_cong[where f=Min]) have join_ok: "qtable n (\ (fv ` set l')) (mem_restr R) (\v. list_all (Formula.sat \ V (map the v) k) l') (mmulti_join n A_pos A_neg L)" if args_ok: "list_all2 (\x. qtable n (fv x) (mem_restr R) (\v. Formula.sat \ V (map the v) k x)) (pos @ map remove_neg neg) L" for k L proof (rule qtable_mmulti_join) let ?ok = "\A Qi X. qtable n A (mem_restr R) Qi X \ wf_set n A" let ?L_pos = "take (length A_pos) L" let ?L_neg = "drop (length A_pos) L" have 1: "length pos \ length L" using args_ok by (auto dest!: list_all2_lengthD) show "list_all3 ?ok A_pos (map (\\ v. Formula.sat \ V (map the v) k \) pos) ?L_pos" using args_ok wf_l unfolding A_eq by (auto simp add: list_all3_conv_all_nth list_all2_conv_all_nth nth_append split: if_splits intro!: wf_mformula_wf_set[of \ j P V n R] dest: order.strict_trans2[OF _ 1]) from args_ok have prems_neg: "list_all2 (\\. qtable n (fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) k (remove_neg \))) neg ?L_neg" by (auto simp: A_eq list_all2_append1 list.rel_map) show "list_all3 ?ok A_neg (map (\\ v. Formula.sat \ V (map the v) k (remove_neg \)) neg) ?L_neg" using prems_neg wf_l unfolding A_eq by (auto simp add: list_all3_conv_all_nth list_all2_conv_all_nth list_all_length nth_append less_diff_conv split: if_splits intro!: wf_mformula_wf_set[of \ j P V n R _ "remove_neg _", simplified]) show "\(fv ` set l') = \(set A_pos)" using fv_subset posneg unfolding A_eq by auto show "L = take (length A_pos) L @ drop (length A_pos) L" by simp show "A_pos \ []" using \pos \ []\ A_eq by simp fix x :: "event_data tuple" assume "wf_tuple n (\ (fv ` set l')) x" and "mem_restr R x" then show "list_all (\A. mem_restr R (restrict A x)) A_pos" and "list_all (\A. mem_restr R (restrict A x)) A_neg" by (simp_all add: list.pred_set) have "list_all Formula.is_Neg neg" using posneg safe_neg by (auto 0 3 simp add: list.pred_map elim!: list.pred_mono_strong intro: formula.exhaust[of \ "Formula.is_Neg \" for \]) then have "list_all (\\. Formula.sat \ V (map the v) i (remove_neg \) \ \ Formula.sat \ V (map the v) i \) neg" for v i by (fastforce simp: Formula.is_Neg_def elim!: list.pred_mono_strong) then show "list_all (Formula.sat \ V (map the x) k) l' = (list_all2 (\A Qi. Qi (restrict A x)) A_pos (map (\\ v. Formula.sat \ V (map the v) k \) pos) \ list_all2 (\A Qi. \ Qi (restrict A x)) A_neg (map (\\ v. Formula.sat \ V (map the v) k (remove_neg \)) neg))" using posneg by (auto simp add: A_eq list_all2_conv_all_nth list_all_length sat_the_restrict elim: nth_filter nth_partition[where P=safe_formula and Q="Formula.sat _ _ _ _"]) qed fact from MAnds.prems(2) show ?case unfolding \\' = Formula.Ands l'\ by (auto 0 3 simp add: list.rel_map progress_eq map2_map_map list_all3_map list_all3_list_all2_conv list.pred_map simp del: set_append map_append progress_simps split: prod.splits intro!: wf_mformula.Ands[OF _ _ posneg \pos \ []\ safe_neg A_eq fv_subset] list.rel_mono_strong[OF wf_l] wf_mbufn_add[OF wf_buf] list.rel_flip[THEN iffD1, OF list.rel_mono_strong, OF wf_l] list.rel_refl join_ok[unfolded list.pred_set] dest!: MAnds.IH[OF _ _ MAnds.prems(2), rotated] elim!: wf_mbufn_take list_all2_appendI elim: mbufn_take_induct[OF _ wf_mbufn_add[OF wf_buf]]) next case (MOr \ \ buf) from MOr.prems show ?case by (cases rule: wf_mformula.cases) (auto dest!: MOr.IH split: if_splits prod.splits intro!: wf_mformula.Or qtable_union elim: mbuf2_take_add'(1) list.rel_mono_strong[OF mbuf2_take_add'(2)]) next case (MNeg \) have *: "qtable n {} (mem_restr R) (\v. P v) X \ \ qtable n {} (mem_restr R) (\v. \ P v) empty_table \ x \ X \ False" for P x X using nullary_qtable_cases qtable_unit_empty_table by fastforce from MNeg.prems show ?case by (cases rule: wf_mformula.cases) (auto 0 4 intro!: wf_mformula.Neg dest!: MNeg.IH simp add: list.rel_map dest: nullary_qtable_cases qtable_unit_empty_table intro!: qtable_empty_unit_table elim!: list.rel_mono_strong elim: *) next case (MExists \) from MExists.prems show ?case by (cases rule: wf_mformula.cases) (force simp: list.rel_map fvi_Suc sat_fv_cong nth_Cons' intro!: wf_mformula.Exists dest!: MExists.IH qtable_project_fv elim!: list.rel_mono_strong table_fvi_tl qtable_cong sat_fv_cong[THEN iffD1, rotated -1] split: if_splits)+ next case (MAgg g0 y \ b f \) from MAgg.prems show ?case using wf_mformula_wf_set[OF MAgg.prems(1), unfolded wf_set_def] by (cases rule: wf_mformula.cases) (auto 0 3 simp add: list.rel_map simp del: sat.simps fvi.simps split: prod.split intro!: wf_mformula.Agg qtable_eval_agg dest!: MAgg.IH elim!: list.rel_mono_strong) next case (MPrev I \ first buf nts) from MPrev.prems show ?case proof (cases rule: wf_mformula.cases) case (Prev \) let ?xs = "fst (meval n (\ \ j) db \)" let ?\ = "snd (meval n (\ \ j) db \)" let ?ls = "fst (mprev_next I (buf @ ?xs) (nts @ [\ \ j]))" let ?rs = "fst (snd (mprev_next I (buf @ ?xs) (nts @ [\ \ j])))" let ?ts = "snd (snd (mprev_next I (buf @ ?xs) (nts @ [\ \ j])))" let ?P = "\i X. qtable n (fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \) X" let ?min = "min (progress \ P' \ (Suc j)) (Suc j - 1)" from Prev MPrev.IH[OF _ MPrev.prems(2), of n R \] have IH: "wf_mformula \ (Suc j) P' V n R ?\ \" and "list_all2 ?P [progress \ P \ j.. P' \ (Suc j)] ?xs" by auto with Prev(4,5) MPrev.prems(2) have "list_all2 (\i X. if mem (\ \ (Suc i) - \ \ i) I then ?P i X else X = empty_table) [min (progress \ P \ j) (j - 1).. list_all2 ?P [?min.. P' \ (Suc j)] ?rs \ list_all2 (\i t. t = \ \ i) [?min.. P \ j)) j = Suc (min (progress \ P \ j) (j-1))" if "j > 0" using that by auto ultimately show ?thesis using Prev(1,3) MPrev.prems(2) IH by (auto simp: map_Suc_upt[symmetric] upt_Suc[of 0] list.rel_map qtable_empty_iff simp del: upt_Suc elim!: wf_mformula.Prev list.rel_mono_strong split: prod.split if_split_asm) qed next case (MNext I \ first nts) from MNext.prems show ?case proof (cases rule: wf_mformula.cases) case (Next \) have min[simp]: "min (progress \ P \ j - Suc 0) (j - Suc 0) = progress \ P \ j - Suc 0" "min (progress \ P' \ (Suc j) - Suc 0) j = progress \ P' \ (Suc j) - Suc 0" using wf_envs_progress_le[OF MNext.prems(2), of \] by auto let ?xs = "fst (meval n (\ \ j) db \)" let ?ys = "case (?xs, first) of (_ # xs, True) \ xs | _ \ ?xs" let ?\ = "snd (meval n (\ \ j) db \)" let ?ls = "fst (mprev_next I ?ys (nts @ [\ \ j]))" let ?rs = "fst (snd (mprev_next I ?ys (nts @ [\ \ j])))" let ?ts = "snd (snd (mprev_next I ?ys (nts @ [\ \ j])))" let ?P = "\i X. qtable n (fv \) (mem_restr R) (\v. Formula.sat \ V (map the v) i \) X" let ?min = "min (progress \ P' \ (Suc j) - 1) (Suc j - 1)" from Next MNext.IH[OF _ MNext.prems(2), of n R \] have IH: "wf_mformula \ (Suc j) P' V n R ?\ \" "list_all2 ?P [progress \ P \ j.. P' \ (Suc j)] ?xs" by auto with Next have "list_all2 (\i X. if mem (\ \ (Suc i) - \ \ i) I then ?P (Suc i) X else X = empty_table) [progress \ P \ j - 1.. list_all2 ?P [Suc ?min.. P' \ (Suc j)] ?rs \ list_all2 (\i t. t = \ \ i) [?min.. P \ j < progress \ P' \ (Suc j)" using that wf_envs_progress_le[OF MNext.prems(2), of \] by (intro mnext) (auto simp: list_all2_Cons2 upt_eq_Cons_conv intro!: list_all2_upt_append list_all2_appendI split: list.splits) then show ?thesis using wf_envs_progress_le[OF MNext.prems(2), of \] wf_envs_progress_mono[OF MNext.prems(2), of j "Suc j" \, simplified] Next IH by (cases "progress \ P' \ (Suc j) > progress \ P \ j") (auto 0 3 simp: qtable_empty_iff le_Suc_eq le_diff_conv elim!: wf_mformula.Next list.rel_mono_strong list_all2_appendI split: prod.split list.splits if_split_asm) (* slow 5 sec*) qed next case (MSince args \ \ buf nts aux) note sat.simps[simp del] from MSince.prems obtain \'' \''' \'' I where Since_eq: "\' = Formula.Since \''' I \''" and pos: "if args_pos args then \''' = \'' else \''' = Formula.Neg \''" and pos_eq: "safe_formula \''' = args_pos args" and \: "wf_mformula \ j P V n R \ \''" and \: "wf_mformula \ j P V n R \ \''" and fvi_subset: "Formula.fv \'' \ Formula.fv \''" and buf: "wf_mbuf2' \ P V j n R \'' \'' buf" and nts: "wf_ts \ P j \'' \'' nts" and aux: "wf_since_aux \ V R args \'' \'' aux (progress \ P (Formula.Since \''' I \'') j)" and args_ivl: "args_ivl args = I" and args_n: "args_n args = n" and args_L: "args_L args = Formula.fv \''" and args_R: "args_R args = Formula.fv \''" by (cases rule: wf_mformula.cases) (auto) have \''': "Formula.fv \''' = Formula.fv \''" "progress \ P \''' j = progress \ P \'' j" "progress \ P' \''' j = progress \ P' \'' j" for j using pos by (simp_all split: if_splits) from MSince.prems(2) have nts_snoc: "list_all2 (\i t. t = \ \ i) [min (progress \ P \'' j) (progress \ P \'' j).. \ j])" using nts unfolding wf_ts_def by (auto simp add: wf_envs_progress_le[THEN min.coboundedI1] intro: list_all2_appendI) have update: "wf_since_aux \ V R args \'' \'' (snd (zs, aux')) (progress \ P' (Formula.Since \''' I \'') (Suc j)) \ list_all2 (\i. qtable n (Formula.fv \''' \ Formula.fv \'') (mem_restr R) (\v. Formula.sat \ V (map the v) i (Formula.Since \''' I \''))) [progress \ P (Formula.Since \''' I \'') j.. P' (Formula.Since \''' I \'') (Suc j)] (fst (zs, aux'))" if eval_\: "fst (meval n (\ \ j) db \) = xs" and eval_\: "fst (meval n (\ \ j) db \) = ys" and eq: "mbuf2t_take (\r1 r2 t (zs, aux). case update_since args r1 r2 t aux of (z, x) \ (zs @ [z], x)) ([], aux) (mbuf2_add xs ys buf) (nts @ [\ \ j]) = ((zs, aux'), buf', nts')" for xs ys zs aux' buf' nts' unfolding progress_simps \''' proof (rule mbuf2t_take_add_induct'[where j=j and j'="Suc j" and z'="(zs, aux')", OF eq wf_envs_P_simps[OF MSince.prems(2)] buf nts_snoc], goal_cases xs ys _ base step) case xs then show ?case using MSince.IH(1)[OF \ MSince.prems(2)] eval_\ by auto next case ys then show ?case using MSince.IH(2)[OF \ MSince.prems(2)] eval_\ by auto next case base then show ?case using aux by (simp add: \''') next case (step k X Y z) then show ?case using fvi_subset pos by (auto 0 3 simp: args_ivl args_n args_L args_R Un_absorb1 elim!: update_since(1) list_all2_appendI dest!: update_since(2) split: prod.split if_splits) qed simp with MSince.IH(1)[OF \ MSince.prems(2)] MSince.IH(2)[OF \ MSince.prems(2)] show ?case by (auto 0 3 simp: Since_eq split: prod.split intro: wf_mformula.Since[OF _ _ pos pos_eq args_ivl args_n args_L args_R fvi_subset] elim: mbuf2t_take_add'(1)[OF _ wf_envs_P_simps[OF MSince.prems(2)] buf nts_snoc] mbuf2t_take_add'(2)[OF _ wf_envs_P_simps[OF MSince.prems(2)] buf nts_snoc]) next case (MUntil args \ \ buf nts aux) note sat.simps[simp del] progress_simps[simp del] from MUntil.prems obtain \'' \''' \'' I where Until_eq: "\' = Formula.Until \''' I \''" and pos: "if args_pos args then \''' = \'' else \''' = Formula.Neg \''" and pos_eq: "safe_formula \''' = args_pos args" and \: "wf_mformula \ j P V n R \ \''" and \: "wf_mformula \ j P V n R \ \''" and fvi_subset: "Formula.fv \'' \ Formula.fv \''" and buf: "wf_mbuf2' \ P V j n R \'' \'' buf" and nts: "wf_ts \ P j \'' \'' nts" and aux: "wf_until_aux \ V R args \'' \'' aux (progress \ P (Formula.Until \''' I \'') j)" and args_ivl: "args_ivl args = I" and args_n: "args_n args = n" and args_L: "args_L args = Formula.fv \''" and args_R: "args_R args = Formula.fv \''" and length_aux: "progress \ P (Formula.Until \''' I \'') j + length_muaux args aux = min (progress \ P \'' j) (progress \ P \'' j)" by (cases rule: wf_mformula.cases) (auto) define pos where args_pos: "pos = args_pos args" have \''': "progress \ P \''' j = progress \ P \'' j" "progress \ P' \''' j = progress \ P' \'' j" for j using pos by (simp_all add: progress.simps split: if_splits) from MUntil.prems(2) have nts_snoc: "list_all2 (\i t. t = \ \ i) [min (progress \ P \'' j) (progress \ P \'' j).. \ j])" using nts unfolding wf_ts_def by (auto simp add: wf_envs_progress_le[THEN min.coboundedI1] intro: list_all2_appendI) { fix xs ys zs aux' aux'' buf' nts' assume eval_\: "fst (meval n (\ \ j) db \) = xs" and eval_\: "fst (meval n (\ \ j) db \) = ys" and eq1: "mbuf2t_take (add_new_muaux args) aux (mbuf2_add xs ys buf) (nts @ [\ \ j]) = (aux', buf', nts')" and eq2: "eval_muaux args (case nts' of [] \ \ \ j | nt # _ \ nt) aux' = (zs, aux'')" define ne where "ne \ progress \ P (Formula.Until \''' I \'') j" have update1: "wf_until_aux \ V R args \'' \'' aux' (progress \ P (Formula.Until \''' I \'') j) \ ne + length_muaux args aux' = min (progress \ P' \''' (Suc j)) (progress \ P' \'' (Suc j))" using MUntil.IH(1)[OF \ MUntil.prems(2)] eval_\ MUntil.IH(2)[OF \ MUntil.prems(2)] eval_\ nts_snoc nts_snoc length_aux aux fvi_subset unfolding \''' by (elim mbuf2t_take_add_induct'[where j'="Suc j", OF eq1 wf_envs_P_simps[OF MUntil.prems(2)] buf]) (auto simp: args_n args_L args_R ne_def wf_update_until) then obtain cur auxlist' where valid_aux': "valid_muaux args cur aux' auxlist'" and cur: "cur = (if ne + length auxlist' = 0 then 0 else \ \ (ne + length auxlist' - 1))" and wf_auxlist': "wf_until_auxlist \ V n R pos \'' I \'' auxlist' (progress \ P (Formula.Until \''' I \'') j)" unfolding wf_until_aux_def ne_def args_ivl args_n args_pos by auto have length_aux': "length_muaux args aux' = length auxlist'" using valid_length_muaux[OF valid_aux'] . have nts': "wf_ts \ P' (Suc j) \'' \'' nts'" using MUntil.IH(1)[OF \ MUntil.prems(2)] eval_\ MUntil.IH(2)[OF \ MUntil.prems(2)] MUntil.prems(2) eval_\ nts_snoc unfolding wf_ts_def by (intro mbuf2t_take_eqD(2)[OF eq1]) (auto intro: wf_mbuf2_add buf[unfolded wf_mbuf2'_def]) define zs'' where "zs'' = fst (eval_until I (case nts' of [] \ \ \ j | nt # x \ nt) auxlist')" define auxlist'' where "auxlist'' = snd (eval_until I (case nts' of [] \ \ \ j | nt # x \ nt) auxlist')" have current_w_le: "cur \ (case nts' of [] \ \ \ j | nt # x \ nt)" proof (cases nts') case Nil have p_le: "min (progress \ P' \''' (Suc j)) (progress \ P' \'' (Suc j)) - 1 \ j" using wf_envs_progress_le[OF MUntil.prems(2)] by (auto simp: min_def le_diff_conv) then show ?thesis unfolding cur conjunct2[OF update1, unfolded length_aux'] using Nil by auto next case (Cons nt x) have progress_\''': "progress \ P' \'' (Suc j) = progress \ P' \''' (Suc j)" using pos by (auto simp add: progress.simps split: if_splits) have "nt = \ \ (min (progress \ P' \'' (Suc j)) (progress \ P' \'' (Suc j)))" using nts'[unfolded wf_ts_def Cons] unfolding list_all2_Cons2 upt_eq_Cons_conv by auto then show ?thesis unfolding cur conjunct2[OF update1, unfolded length_aux'] Cons progress_\''' by (auto split: if_splits list.splits intro!: \_mono) qed have valid_aux'': "valid_muaux args cur aux'' auxlist''" using valid_eval_muaux[OF valid_aux' current_w_le eq2, of zs'' auxlist''] by (auto simp add: args_ivl zs''_def auxlist''_def) have length_aux'': "length_muaux args aux'' = length auxlist''" using valid_length_muaux[OF valid_aux''] . have eq2': "eval_until I (case nts' of [] \ \ \ j | nt # _ \ nt) auxlist' = (zs, auxlist'')" using valid_eval_muaux[OF valid_aux' current_w_le eq2, of zs'' auxlist''] by (auto simp add: args_ivl zs''_def auxlist''_def) have length_aux'_aux'': "length_muaux args aux' = length zs + length_muaux args aux''" using eval_until_length[OF eq2'] unfolding length_aux' length_aux'' . have "i \ progress \ P' (Formula.Until \''' I \'') (Suc j) \ wf_until_auxlist \ V n R pos \'' I \'' auxlist' i \ i + length auxlist' = min (progress \ P' \''' (Suc j)) (progress \ P' \'' (Suc j)) \ wf_until_auxlist \ V n R pos \'' I \'' auxlist'' (progress \ P' (Formula.Until \''' I \'') (Suc j)) \ i + length zs = progress \ P' (Formula.Until \''' I \'') (Suc j) \ i + length zs + length auxlist'' = min (progress \ P' \''' (Suc j)) (progress \ P' \'' (Suc j)) \ list_all2 (\i. qtable n (Formula.fv \'') (mem_restr R) (\v. Formula.sat \ V (map the v) i (Formula.Until (if pos then \'' else Formula.Neg \'') I \''))) [i.. V n R pos \'' I \'' aux' (Suc i)" by (rule wf_until_aux_Cons) from Cons.prems(2) have 1: "t = \ \ i" unfolding \a = (t, a1, a2)\ by (rule wf_until_aux_Cons1) from Cons.prems(2) have 3: "qtable n (Formula.fv \'') (mem_restr R) (\v. (\j\i. j < Suc (i + length aux') \ mem (\ \ j - \ \ i) I \ Formula.sat \ V (map the v) j \'' \ (\k\{i.. V (map the v) k \'' else \ Formula.sat \ V (map the v) k \''))) a2" unfolding \a = (t, a1, a2)\ by (rule wf_until_aux_Cons3) from Cons.prems(3) have Suc_i_aux': "Suc i + length aux' = min (progress \ P' \''' (Suc j)) (progress \ P' \'' (Suc j))" by simp have "i \ progress \ P' (Formula.Until \''' I \'') (Suc j)" if "enat (case nts' of [] \ \ \ j | nt # x \ nt) \ enat t + right I" using that nts' unfolding wf_ts_def progress.simps by (auto simp add: 1 list_all2_Cons2 upt_eq_Cons_conv \''' intro!: cInf_lower \_mono elim!: order.trans[rotated] simp del: upt_Suc split: if_splits list.splits) moreover have "Suc i \ progress \ P' (Formula.Until \''' I \'') (Suc j)" if "enat t + right I < enat (case nts' of [] \ \ \ j | nt # x \ nt)" proof - from that obtain m where m: "right I = enat m" by (cases "right I") auto have \_min: "\ \ (min j k) = min (\ \ j) (\ \ k)" for k by (simp add: min_of_mono monoI) have le_progress_iff[simp]: "(Suc j) \ progress \ P' \ (Suc j) \ progress \ P' \ (Suc j) = (Suc j)" for \ using wf_envs_progress_le[OF MUntil.prems(2), of \] by auto have min_Suc[simp]: "min j (Suc j) = j" by auto let ?X = "{i. \k. k < Suc j \ k \min (progress \ P' \''' (Suc j)) (progress \ P' \'' (Suc j)) \ enat (\ \ k) \ enat (\ \ i) + right I}" let ?min = "min j (min (progress \ P' \'' (Suc j)) (progress \ P' \'' (Suc j)))" have "\ \ ?min \ \ \ j" by (rule \_mono) auto from m have "?X \ {}" by (auto dest!: \_mono[of _ "progress \ P' \'' (Suc j)" \] simp: not_le not_less \''' intro!: exI[of _ "progress \ P' \'' (Suc j)"]) from m show ?thesis using that nts' unfolding wf_ts_def progress.simps by (intro cInf_greatest[OF \?X \ {}\]) (auto simp: 1 \''' not_le not_less list_all2_Cons2 upt_eq_Cons_conv less_Suc_eq simp del: upt_Suc split: list.splits if_splits dest!: spec[of _ "?min"] less_le_trans[of "\ \ i + m" "\ \ _" "\ \ _ + m"] less_\D) qed moreover have *: "k < progress \ P' \ (Suc j)" if "enat (\ \ i) + right I < enat (case nts' of [] \ \ \ j | nt # x \ nt)" "enat (\ \ k - \ \ i) \ right I" "\ = \'' \ \ = \''" for k \ proof - from that(1,2) obtain m where "right I = enat m" "\ \ i + m < (case nts' of [] \ \ \ j | nt # x \ nt)" "\ \ k - \ \ i \ m" by (cases "right I") auto with that(3) nts' progress_le[of \ \'' "Suc j"] progress_le[of \ \'' "Suc j"] show ?thesis unfolding wf_ts_def le_diff_conv by (auto simp: not_le list_all2_Cons2 upt_eq_Cons_conv less_Suc_eq add.commute simp del: upt_Suc split: list.splits if_splits dest!: le_less_trans[of "\ \ k"] less_\D) qed ultimately show ?case using Cons.prems Suc_i_aux'[simplified] unfolding \a = (t, a1, a2)\ by (auto simp: \''' 1 sat.simps upt_conv_Cons dest!: Cons.IH[OF _ aux' Suc_i_aux'] simp del: upt_Suc split: if_splits prod.splits intro!: iff_exI qtable_cong[OF 3 refl]) qed thm this note wf_aux'' = this[OF wf_envs_progress_mono[OF MUntil.prems(2) le_SucI[OF order_refl]] wf_auxlist' conjunct2[OF update1, unfolded ne_def length_aux']] have "progress \ P (formula.Until \''' I \'') j + length auxlist' = progress \ P' (formula.Until \''' I \'') (Suc j) + length auxlist''" using wf_aux'' valid_aux'' length_aux'_aux'' by (auto simp add: ne_def length_aux' length_aux'') then have "cur = (if progress \ P' (formula.Until \''' I \'') (Suc j) + length auxlist'' = 0 then 0 else \ \ (progress \ P' (formula.Until \''' I \'') (Suc j) + length auxlist'' - 1))" unfolding cur ne_def by auto then have "wf_until_aux \ V R args \'' \'' aux'' (progress \ P' (formula.Until \''' I \'') (Suc j)) \ progress \ P (formula.Until \''' I \'') j + length zs = progress \ P' (formula.Until \''' I \'') (Suc j) \ progress \ P (formula.Until \''' I \'') j + length zs + length_muaux args aux'' = min (progress \ P' \''' (Suc j)) (progress \ P' \'' (Suc j)) \ list_all2 (\i. qtable n (fv \'') (mem_restr R) (\v. Formula.sat \ V (map the v) i (formula.Until (if pos then \'' else formula.Neg \'') I \''))) [progress \ P (formula.Until \''' I \'') j.. P (formula.Until \''' I \'') j + length zs] zs" using wf_aux'' valid_aux'' fvi_subset unfolding wf_until_aux_def length_aux'' args_ivl args_n args_pos by (auto simp only: length_aux'') } note update = this from MUntil.IH(1)[OF \ MUntil.prems(2)] MUntil.IH(2)[OF \ MUntil.prems(2)] pos pos_eq fvi_subset show ?case by (auto 0 4 simp: args_ivl args_n args_pos Until_eq \''' progress.simps(6) split: prod.split if_splits dest!: update[OF refl refl, rotated] intro!: wf_mformula.Until[OF _ _ _ _ args_ivl args_n args_L args_R fvi_subset] elim!: list.rel_mono_strong qtable_cong elim: mbuf2t_take_add'(1)[OF _ wf_envs_P_simps[OF MUntil.prems(2)] buf nts_snoc] mbuf2t_take_add'(2)[OF _ wf_envs_P_simps[OF MUntil.prems(2)] buf nts_snoc]) next case (MMatchP I mr mrs \s buf nts aux) note sat.simps[simp del] mbufnt_take.simps[simp del] mbufn_add.simps[simp del] from MMatchP.prems obtain r \s where eq: "\' = Formula.MatchP I r" and safe: "safe_regex Past Strict r" and mr: "to_mregex r = (mr, \s)" and mrs: "mrs = sorted_list_of_set (RPDs mr)" and \s: "list_all2 (wf_mformula \ j P V n R) \s \s" and buf: "wf_mbufn' \ P V j n R r buf" and nts: "wf_ts_regex \ P j r nts" and aux: "wf_matchP_aux \ V n R I r aux (progress \ P (Formula.MatchP I r) j)" by (cases rule: wf_mformula.cases) (auto) have nts_snoc: "list_all2 (\i t. t = \ \ i) [progress_regex \ P r j.. \ j])" using nts unfolding wf_ts_regex_def by (auto simp add: wf_envs_progress_regex_le[OF MMatchP.prems(2)] intro: list_all2_appendI) have update: "wf_matchP_aux \ V n R I r (snd (zs, aux')) (progress \ P' (Formula.MatchP I r) (Suc j)) \ list_all2 (\i. qtable n (Formula.fv_regex r) (mem_restr R) (\v. Formula.sat \ V (map the v) i (Formula.MatchP I r))) [progress \ P (Formula.MatchP I r) j.. P' (Formula.MatchP I r) (Suc j)] (fst (zs, aux'))" if eval: "map (fst o meval n (\ \ j) db) \s = xss" and eq: "mbufnt_take (\rels t (zs, aux). case update_matchP n I mr mrs rels t aux of (z, x) \ (zs @ [z], x)) ([], aux) (mbufn_add xss buf) (nts @ [\ \ j]) = ((zs, aux'), buf', nts')" for xss zs aux' buf' nts' unfolding progress_simps proof (rule mbufnt_take_add_induct'[where j'="Suc j" and z'="(zs, aux')", OF eq wf_envs_P_simps[OF MMatchP.prems(2)] safe mr buf nts_snoc], goal_cases xss _ base step) case xss then show ?case using eval \s by (auto simp: list_all3_map map2_map_map list_all3_list_all2_conv list.rel_map list.rel_flip[symmetric, of _ \s \s] dest!: MMatchP.IH(1)[OF _ _ MMatchP.prems(2)] elim!: list.rel_mono_strong split: prod.splits) next case base then show ?case using aux by auto next case (step k Xs z) then show ?case by (auto simp: Un_absorb1 mrs safe mr elim!: update_matchP(1) list_all2_appendI dest!: update_matchP(2) split: prod.split) qed simp then show ?case using \s by (auto simp: eq mr mrs safe map_split_alt list.rel_flip[symmetric, of _ \s \s] list_all3_map map2_map_map list_all3_list_all2_conv list.rel_map intro!: wf_mformula.intros elim!: list.rel_mono_strong mbufnt_take_add'(1)[OF _ wf_envs_P_simps[OF MMatchP.prems(2)] safe mr buf nts_snoc] mbufnt_take_add'(2)[OF _ wf_envs_P_simps[OF MMatchP.prems(2)] safe mr buf nts_snoc] dest!: MMatchP.IH[OF _ _ MMatchP.prems(2)] split: prod.splits) next case (MMatchF I mr mrs \s buf nts aux) note sat.simps[simp del] mbufnt_take.simps[simp del] mbufn_add.simps[simp del] progress_simps[simp del] from MMatchF.prems obtain r \s where eq: "\' = Formula.MatchF I r" and safe: "safe_regex Futu Strict r" and mr: "to_mregex r = (mr, \s)" and mrs: "mrs = sorted_list_of_set (LPDs mr)" and \s: "list_all2 (wf_mformula \ j P V n R) \s \s" and buf: "wf_mbufn' \ P V j n R r buf" and nts: "wf_ts_regex \ P j r nts" and aux: "wf_matchF_aux \ V n R I r aux (progress \ P (Formula.MatchF I r) j) 0" and length_aux: "progress \ P (Formula.MatchF I r) j + length aux = progress_regex \ P r j" by (cases rule: wf_mformula.cases) (auto) have nts_snoc: "list_all2 (\i t. t = \ \ i) [progress_regex \ P r j.. \ j])" using nts unfolding wf_ts_regex_def by (auto simp add: wf_envs_progress_regex_le[OF MMatchF.prems(2)] intro: list_all2_appendI) { fix xss zs aux' aux'' buf' nts' assume eval: "map (fst o meval n (\ \ j) db) \s = xss" and eq1: "mbufnt_take (update_matchF n I mr mrs) aux (mbufn_add xss buf) (nts @ [\ \ j]) = (aux', buf', nts')" and eq2: "eval_matchF I mr (case nts' of [] \ \ \ j | nt # _ \ nt) aux' = (zs, aux'')" have update1: "wf_matchF_aux \ V n R I r aux' (progress \ P (Formula.MatchF I r) j) 0 \ progress \ P (Formula.MatchF I r) j + length aux' = progress_regex \ P' r (Suc j)" using eval nts_snoc nts_snoc length_aux aux \s by (elim mbufnt_take_add_induct'[where j'="Suc j", OF eq1 wf_envs_P_simps[OF MMatchF.prems(2)] safe mr buf]) (auto simp: length_update_matchF list_all3_map map2_map_map list_all3_list_all2_conv list.rel_map list.rel_flip[symmetric, of _ \s \s] dest!: MMatchF.IH[OF _ _ MMatchF.prems(2)] elim: wf_update_matchF[OF _ safe mr mrs] elim!: list.rel_mono_strong) from MMatchF.prems(2) have nts': "wf_ts_regex \ P' (Suc j) r nts'" using eval eval nts_snoc \s unfolding wf_ts_regex_def by (intro mbufnt_take_eqD(2)[OF eq1 wf_mbufn_add[where js'="map (\\. progress \ P' \ (Suc j)) \s", OF buf[unfolded wf_mbufn'_def mr prod.case]]]) (auto simp: to_mregex_progress[OF safe mr] Mini_def list_all3_map map2_map_map list_all3_list_all2_conv list.rel_map list.rel_flip[symmetric, of _ \s \s] list_all2_Cons1 elim!: list.rel_mono_strong intro!: list.rel_refl_strong dest!: MMatchF.IH[OF _ _ MMatchF.prems(2)]) have "i \ progress \ P' (Formula.MatchF I r) (Suc j) \ wf_matchF_aux \ V n R I r aux' i 0 \ i + length aux' = progress_regex \ P' r (Suc j) \ wf_matchF_aux \ V n R I r aux'' (progress \ P' (Formula.MatchF I r) (Suc j)) 0 \ i + length zs = progress \ P' (Formula.MatchF I r) (Suc j) \ i + length zs + length aux'' = progress_regex \ P' r (Suc j) \ list_all2 (\i. qtable n (Formula.fv_regex r) (mem_restr R) (\v. Formula.sat \ V (map the v) i (Formula.MatchF I r))) [i.. V n R I r aux' (Suc i) 0" by (rule wf_matchF_aux_Cons) from Cons.prems(2) have 1: "t = \ \ i" unfolding \a = (t, rels, rel)\ by (rule wf_matchF_aux_Cons1) from Cons.prems(2) have 3: "qtable n (Formula.fv_regex r) (mem_restr R) (\v. (\j\i. j < Suc (i + length aux') \ mem (\ \ j - \ \ i) I \ Regex.match (Formula.sat \ V (map the v)) r i j)) rel" unfolding \a = (t, rels, rel)\ using wf_matchF_aux_Cons3 by force from Cons.prems(3) have Suc_i_aux': "Suc i + length aux' = progress_regex \ P' r (Suc j)" by simp have "i \ progress \ P' (Formula.MatchF I r) (Suc j)" if "enat (case nts' of [] \ \ \ j | nt # x \ nt) \ enat t + right I" using that nts' unfolding wf_ts_regex_def progress_simps by (auto simp add: 1 list_all2_Cons2 upt_eq_Cons_conv intro!: cInf_lower \_mono elim!: order.trans[rotated] simp del: upt_Suc split: if_splits list.splits) moreover have "Suc i \ progress \ P' (Formula.MatchF I r) (Suc j)" if "enat t + right I < enat (case nts' of [] \ \ \ j | nt # x \ nt)" proof - from that obtain m where m: "right I = enat m" by (cases "right I") auto have \_min: "\ \ (min j k) = min (\ \ j) (\ \ k)" for k by (simp add: min_of_mono monoI) have le_progress_iff[simp]: "Suc j \ progress \ P' \ (Suc j) \ progress \ P' \ (Suc j) = (Suc j)" for \ using wf_envs_progress_le[OF MMatchF.prems(2), of \] by auto have min_Suc[simp]: "min j (Suc j) = j" by auto let ?X = "{i. \k. k < Suc j \ k \ progress_regex \ P' r (Suc j) \ enat (\ \ k) \ enat (\ \ i) + right I}" let ?min = "min j (progress_regex \ P' r (Suc j))" have "\ \ ?min \ \ \ j" by (rule \_mono) auto from m have "?X \ {}" by (auto dest!: less_\D add_lessD1 simp: not_le not_less) from m show ?thesis using that nts' wf_envs_progress_regex_le[OF MMatchF.prems(2), of r] unfolding wf_ts_regex_def progress_simps by (intro cInf_greatest[OF \?X \ {}\]) (auto simp: 1 not_le not_less list_all2_Cons2 upt_eq_Cons_conv less_Suc_eq simp del: upt_Suc split: list.splits if_splits dest!: spec[of _ "?min"] less_le_trans[of "\ \ i + m" "\ \ _" "\ \ _ + m"] less_\D) qed moreover have *: "k < progress_regex \ P' r (Suc j)" if "enat (\ \ i) + right I < enat (case nts' of [] \ \ \ j | nt # x \ nt)" "enat (\ \ k - \ \ i) \ right I" for k proof - from that(1,2) obtain m where "right I = enat m" "\ \ i + m < (case nts' of [] \ \ \ j | nt # x \ nt)" "\ \ k - \ \ i \ m" by (cases "right I") auto with nts' wf_envs_progress_regex_le[OF MMatchF.prems(2), of r] show ?thesis unfolding wf_ts_regex_def le_diff_conv by (auto simp: not_le list_all2_Cons2 upt_eq_Cons_conv less_Suc_eq add.commute simp del: upt_Suc split: list.splits if_splits dest!: le_less_trans[of "\ \ k"] less_\D) qed ultimately show ?case using Cons.prems Suc_i_aux'[simplified] unfolding \a = (t, rels, rel)\ by (auto simp: 1 sat.simps upt_conv_Cons dest!: Cons.IH[OF _ aux' Suc_i_aux'] simp del: upt_Suc split: if_splits prod.splits intro!: iff_exI qtable_cong[OF 3 refl]) qed note this[OF progress_mono_gen[OF le_SucI, OF order.refl] conjunct1[OF update1] conjunct2[OF update1]] } note update = this[OF refl, rotated] with MMatchF.prems(2) show ?case using \s by (auto simp: eq mr mrs safe map_split_alt list.rel_flip[symmetric, of _ \s \s] list_all3_map map2_map_map list_all3_list_all2_conv list.rel_map intro!: wf_mformula.intros elim!: list.rel_mono_strong mbufnt_take_add'(1)[OF _ wf_envs_P_simps[OF MMatchF.prems(2)] safe mr buf nts_snoc] mbufnt_take_add'(2)[OF _ wf_envs_P_simps[OF MMatchF.prems(2)] safe mr buf nts_snoc] dest!: MMatchF.IH[OF _ _ MMatchF.prems(2)] update split: prod.splits) qed subsubsection \Monitor step\ lemma (in maux) wf_mstate_mstep: "wf_mstate \ \ R st \ last_ts \ \ snd tdb \ wf_mstate \ (psnoc \ tdb) R (snd (mstep (map_prod mk_db id tdb) st))" unfolding wf_mstate_def mstep_def Let_def by (fastforce simp add: progress_mono le_imp_diff_is_add split: prod.splits elim!: prefix_of_psnocE dest: meval[OF _ wf_envs_mk_db] list_all2_lengthD) definition "flatten_verdicts Vs = (\ (set (map (\(i, X). (\v. (i, v)) ` X) Vs)))" lemma flatten_verdicts_append[simp]: "flatten_verdicts (Vs @ Us) = flatten_verdicts Vs \ flatten_verdicts Us" by (induct Vs) (auto simp: flatten_verdicts_def) lemma (in maux) mstep_output_iff: assumes "wf_mstate \ \ R st" "last_ts \ \ snd tdb" "prefix_of (psnoc \ tdb) \" "mem_restr R v" shows "(i, v) \ flatten_verdicts (fst (mstep (map_prod mk_db id tdb) st)) \ progress \ Map.empty \ (plen \) \ i \ i < progress \ Map.empty \ (Suc (plen \)) \ wf_tuple (Formula.nfv \) (Formula.fv \) v \ Formula.sat \ Map.empty (map the v) i \" proof - from prefix_of_psnocE[OF assms(3,2)] have "prefix_of \ \" "\ \ (plen \) = fst tdb" "\ \ (plen \) = snd tdb" by auto moreover from assms(1) \prefix_of \ \\ have "mstate_n st = Formula.nfv \" "mstate_i st = progress \ Map.empty \ (plen \)" "wf_mformula \ (plen \) Map.empty Map.empty (mstate_n st) R (mstate_m st) \" unfolding wf_mstate_def by blast+ moreover from meval[OF \wf_mformula \ (plen \) Map.empty Map.empty (mstate_n st) R (mstate_m st) \\ wf_envs_mk_db] obtain Vs st' where "meval (mstate_n st) (\ \ (plen \)) (mk_db (\ \ (plen \))) (mstate_m st) = (Vs, st')" "wf_mformula \ (Suc (plen \)) Map.empty Map.empty (mstate_n st) R st' \" "list_all2 (\i. qtable (mstate_n st) (fv \) (mem_restr R) (\v. Formula.sat \ Map.empty (map the v) i \)) [progress \ Map.empty \ (plen \).. Map.empty \ (Suc (plen \))] Vs" by blast moreover from this assms(4) have "qtable (mstate_n st) (fv \) (mem_restr R) (\v. Formula.sat \ Map.empty (map the v) i \) (Vs ! (i - progress \ Map.empty \ (plen \)))" if "progress \ Map.empty \ (plen \) \ i" "i < progress \ Map.empty \ (Suc (plen \))" using that by (auto simp: list_all2_conv_all_nth dest!: spec[of _ "(i - progress \ Map.empty \ (plen \))"]) ultimately show ?thesis using assms(4) unfolding mstep_def Let_def flatten_verdicts_def by (auto simp: in_set_enumerate_eq list_all2_conv_all_nth progress_mono le_imp_diff_is_add elim!: in_qtableE in_qtableI intro!: bexI[of _ "(i, Vs ! (i - progress \ Map.empty \ (plen \)))"]) qed subsubsection \Monitor function\ locale verimon = verimon_spec + maux lemma (in verimon) mstep_mverdicts: assumes wf: "wf_mstate \ \ R st" and le[simp]: "last_ts \ \ snd tdb" and restrict: "mem_restr R v" shows "(i, v) \ flatten_verdicts (fst (mstep (map_prod mk_db id tdb) st)) \ (i, v) \ M (psnoc \ tdb) - M \" proof - obtain \ where p2: "prefix_of (psnoc \ tdb) \" using ex_prefix_of by blast with le have p1: "prefix_of \ \" by (blast elim!: prefix_of_psnocE) show ?thesis unfolding M_def by (auto 0 3 simp: p2 progress_prefix_conv[OF _ p1] sat_prefix_conv[OF _ p1] not_less pprogress_eq[OF p1] pprogress_eq[OF p2] dest: mstep_output_iff[OF wf le p2 restrict, THEN iffD1] spec[of _ \] mstep_output_iff[OF wf le _ restrict, THEN iffD1] progress_sat_cong[OF p1] intro: mstep_output_iff[OF wf le p2 restrict, THEN iffD2] p1) qed context maux begin primrec msteps0 where "msteps0 [] st = ([], st)" | "msteps0 (tdb # \) st = (let (V', st') = mstep (map_prod mk_db id tdb) st; (V'', st'') = msteps0 \ st' in (V' @ V'', st''))" primrec msteps0_stateless where "msteps0_stateless [] st = []" | "msteps0_stateless (tdb # \) st = (let (V', st') = mstep (map_prod mk_db id tdb) st in V' @ msteps0_stateless \ st')" lemma msteps0_msteps0_stateless: "fst (msteps0 w st) = msteps0_stateless w st" by (induct w arbitrary: st) (auto simp: split_beta) lift_definition msteps :: "Formula.prefix \ ('msaux, 'muaux) mstate \ (nat \ event_data table) list \ ('msaux, 'muaux) mstate" is msteps0 . lift_definition msteps_stateless :: "Formula.prefix \ ('msaux, 'muaux) mstate \ (nat \ event_data table) list" is msteps0_stateless . lemma msteps_msteps_stateless: "fst (msteps w st) = msteps_stateless w st" by transfer (rule msteps0_msteps0_stateless) lemma msteps0_snoc: "msteps0 (\ @ [tdb]) st = (let (V', st') = msteps0 \ st; (V'', st'') = mstep (map_prod mk_db id tdb) st' in (V' @ V'', st''))" by (induct \ arbitrary: st) (auto split: prod.splits) lemma msteps_psnoc: "last_ts \ \ snd tdb \ msteps (psnoc \ tdb) st = (let (V', st') = msteps \ st; (V'', st'') = mstep (map_prod mk_db id tdb) st' in (V' @ V'', st''))" by transfer' (auto simp: msteps0_snoc split: list.splits prod.splits if_splits) definition monitor where "monitor \ \ = msteps_stateless \ (minit_safe \)" end lemma Suc_length_conv_snoc: "(Suc n = length xs) = (\y ys. xs = ys @ [y] \ length ys = n)" by (cases xs rule: rev_cases) auto lemma (in verimon) wf_mstate_msteps: "wf_mstate \ \ R st \ mem_restr R v \ \ \ \' \ X = msteps (pdrop (plen \) \') st \ wf_mstate \ \' R (snd X) \ ((i, v) \ flatten_verdicts (fst X)) = ((i, v) \ M \' - M \)" proof (induct "plen \' - plen \" arbitrary: X st \ \') case 0 from 0(1,4,5) have "\ = \'" "X = ([], st)" by (transfer; auto)+ with 0(2) show ?case unfolding flatten_verdicts_def by simp next case (Suc x) from Suc(2,5) obtain \'' tdb where "x = plen \'' - plen \" "\ \ \''" "\' = psnoc \'' tdb" "pdrop (plen \) (psnoc \'' tdb) = psnoc (pdrop (plen \) \'') tdb" "last_ts (pdrop (plen \) \'') \ snd tdb" "last_ts \'' \ snd tdb" "\'' \ psnoc \'' tdb" proof (atomize_elim, transfer, elim exE, goal_cases prefix) case (prefix _ _ \' _ \_tdb) then show ?case proof (cases \_tdb rule: rev_cases) case (snoc \ tdb) with prefix show ?thesis by (intro bexI[of _ "\' @ \"] exI[of _ tdb]) (force simp: sorted_append append_eq_Cons_conv split: list.splits if_splits)+ qed simp qed with Suc(1)[OF this(1) Suc.prems(1,2) this(2) refl] Suc.prems show ?case unfolding msteps_msteps_stateless[symmetric] by (auto simp: msteps_psnoc split_beta mstep_mverdicts dest: mono_monitor[THEN set_mp, rotated] intro!: wf_mstate_mstep) qed lemma (in verimon) wf_mstate_msteps_stateless: assumes "wf_mstate \ \ R st" "mem_restr R v" "\ \ \'" shows "(i, v) \ flatten_verdicts (msteps_stateless (pdrop (plen \) \') st) \ (i, v) \ M \' - M \" using wf_mstate_msteps[OF assms refl] unfolding msteps_msteps_stateless by simp lemma (in verimon) wf_mstate_msteps_stateless_UNIV: "wf_mstate \ \ UNIV st \ \ \ \' \ flatten_verdicts (msteps_stateless (pdrop (plen \) \') st) = M \' - M \" by (auto dest: wf_mstate_msteps_stateless[OF _ mem_restr_UNIV]) lemma (in verimon) mverdicts_Nil: "M pnil = {}" by (simp add: M_def pprogress_eq) context maux begin lemma minit_safe_minit: "mmonitorable \ \ minit_safe \ = minit \" unfolding minit_safe_def monitorable_formula_code by simp lemma wf_mstate_minit_safe: "mmonitorable \ \ wf_mstate \ pnil R (minit_safe \)" using wf_mstate_minit minit_safe_minit mmonitorable_def by metis end lemma (in verimon) monitor_mverdicts: "flatten_verdicts (monitor \ \) = M \" unfolding monitor_def using monitorable by (subst wf_mstate_msteps_stateless_UNIV[OF wf_mstate_minit_safe, simplified]) (auto simp: mmonitorable_def mverdicts_Nil) subsection \Collected correctness results\ context verimon begin text \We summarize the main results proved above. \begin{enumerate} \item The term @{term M} describes semantically the monitor's expected behaviour: \begin{itemize} \item @{thm[source] mono_monitor}: @{thm mono_monitor[no_vars]} \item @{thm[source] sound_monitor}: @{thm sound_monitor[no_vars]} \item @{thm[source] complete_monitor}: @{thm complete_monitor[no_vars]} \item @{thm[source] sliceable_M}: @{thm sliceable_M[no_vars]} \end{itemize} \item The executable monitor's online interface @{term minit_safe} and @{term mstep} preserves the invariant @{term wf_mstate} and produces the the verdicts according to @{term M}: \begin{itemize} \item @{thm[source] wf_mstate_minit_safe}: @{thm wf_mstate_minit_safe[no_vars]} \item @{thm[source] wf_mstate_mstep}: @{thm wf_mstate_mstep[no_vars]} \item @{thm[source] mstep_mverdicts}: @{thm mstep_mverdicts[no_vars]} \end{itemize} \item The executable monitor's offline interface @{term monitor} implements @{term M}: \begin{itemize} \item @{thm[source] monitor_mverdicts}: @{thm monitor_mverdicts[no_vars]} \end{itemize} \end{enumerate} \ end (*<*) end (*>*) diff --git a/thys/Prim_Dijkstra_Simple/Dijkstra_Impl.thy b/thys/Prim_Dijkstra_Simple/Dijkstra_Impl.thy --- a/thys/Prim_Dijkstra_Simple/Dijkstra_Impl.thy +++ b/thys/Prim_Dijkstra_Simple/Dijkstra_Impl.thy @@ -1,285 +1,285 @@ section \Implementation of Dijkstra's Algorithm\ theory Dijkstra_Impl imports Dijkstra_Abstract "Directed_Graph_Impl" "HOL-Library.While_Combinator" "Priority_Search_Trees.PST_RBT" "HOL-Data_Structures.RBT_Map" begin subsection \Implementation using ADT Interfaces\ locale Dijkstra_Impl_Adts = G: adt_finite_wgraph G_invar G_succ G_empty G_add G_\ + M: Map M_empty M_update M_delete M_lookup M_invar + Q: PrioMap Q_empty Q_update Q_delete Q_invar Q_lookup Q_is_empty Q_getmin for G_\ :: "'g \ ('v) wgraph" and G_invar G_succ G_empty G_add and M_empty M_update M_delete and M_lookup :: "'m \ 'v \ nat option" and M_invar and Q_empty Q_update Q_delete Q_invar and Q_lookup :: "'q \ 'v \ nat option" and Q_is_empty Q_getmin begin text \Simplifier setup\ lemmas [simp] = G.wgraph_specs lemmas [simp] = M.map_specs lemmas [simp] = Q.prio_map_specs end context PrioMap begin lemma map_getminE: assumes "getmin m = (k,p)" "invar m" "lookup m \ Map.empty" obtains "lookup m k = Some p" "\k' p'. lookup m k' = Some p' \ p\p'" using map_getmin[OF assms] by (auto simp: ran_def) end locale Dijkstra_Impl_Defs = Dijkstra_Impl_Adts where G_\ = G_\ + Dijkstra \G_\ g\ s for G_\ :: "'g \ ('v::linorder) wgraph" and g s locale Dijkstra_Impl = Dijkstra_Impl_Defs where G_\ = G_\ for G_\ :: "'g \ ('v::linorder) wgraph" + assumes G_invar[simp]: "G_invar g" begin lemma finite_all_dnodes[simp, intro!]: "finite all_dnodes" proof - have "all_dnodes \ Set.insert s (snd ` edges)" by (fastforce simp: all_dnodes_def edges_def image_iff) also have "finite \" by (auto simp: G.finite) finally (finite_subset) show ?thesis . qed lemma finite_unfinished_dnodes[simp, intro!]: "finite (unfinished_dnodes S)" using finite_subset[OF unfinished_nodes_subset] by auto lemma (in -) fold_refine: assumes "I s" assumes "\s x. I s \ x\set l \ I (f x s) \ \ (f x s) = f' x (\ s)" shows "I (fold f l s) \ \ (fold f l s) = fold f' l (\ s)" using assms by (induction l arbitrary: s) auto definition (in Dijkstra_Impl_Defs) "Q_relax_outgoing u du V Q = fold (\(d,v) Q. case Q_lookup Q v of None \ if M_lookup V v \ None then Q else Q_update v (du+d) Q | Some d' \ Q_update v (min (du+d) d') Q) ((G_succ g u)) Q" lemma Q_relax_outgoing[simp]: assumes [simp]: "Q_invar Q" shows "Q_invar (Q_relax_outgoing u du V Q) \ Q_lookup (Q_relax_outgoing u du V Q) = relax_outgoing' u du (M_lookup V) (Q_lookup Q)" apply (subst relax_outgoing''_refine[symmetric, where l="G_succ g u"]) apply simp unfolding Q_relax_outgoing_def relax_outgoing''_def apply (rule fold_refine[where I=Q_invar and \=Q_lookup]) by (auto split: option.split) definition (in Dijkstra_Impl_Defs) "D_invar_impl Q V \ Q_invar Q \ M_invar V \ D_invar' (Q_lookup Q) (M_lookup V)" definition (in Dijkstra_Impl_Defs) "Q_initQ \ Q_update s 0 Q_empty" lemma Q_init_Q[simp]: shows "Q_invar (Q_initQ)" "Q_lookup (Q_initQ) = initQ" by (auto simp: Q_initQ_def initQ_def) definition (in Dijkstra_Impl_Defs) "M_initV \ M_empty" lemma M_initS[simp]: "M_invar M_initV" "M_lookup M_initV = initV" unfolding M_initV_def initV_def by auto term Q_getmin definition (in Dijkstra_Impl_Defs) "dijkstra_loop \ while (\(Q,V). \ Q_is_empty Q) (\(Q,V). let (u,du) = Q_getmin Q; Q = Q_relax_outgoing u du V Q; Q = Q_delete u Q; V = M_update u du V in (Q,V) ) (Q_initQ,M_initV)" definition (in Dijkstra_Impl_Defs) "dijkstra \ snd dijkstra_loop" lemma transfer_preconditions: assumes "coupling Q V D S" shows "Q u = Some du \ D u = enat du \ u\S" using assms by (auto simp: coupling_def) lemma dijkstra_loop_invar_and_empty: shows "case dijkstra_loop of (Q,V) \ D_invar_impl Q V \ Q_is_empty Q" unfolding dijkstra_loop_def apply (rule while_rule[where P="case_prod D_invar_impl" and r="inv_image finite_psubset (unfinished_dnodes' o M_lookup o snd)"]) apply (all \(clarsimp split: prod.splits)?\) subgoal apply (simp add: D_invar_impl_def) apply (simp add: D_invar'_def) apply (intro exI conjI) apply (rule coupling_init) using initD_def initS_def invar_init by auto proof - fix Q V u du assume "\ Q_is_empty Q" "D_invar_impl Q V" "Q_getmin Q = (u, du)" hence "Q_lookup Q \ Map.empty" "D_invar' (Q_lookup Q) (M_lookup V)" and [simp]: "Q_invar Q" "M_invar V" and "Q_lookup Q u = Some du" "\k' p'. Q_lookup Q k' = Some p' \ du \ p'" by (auto simp: D_invar_impl_def elim: Q.map_getminE) then obtain D S where "D_invar D S" and COUPLING: "coupling (Q_lookup Q) (M_lookup V) D S" and ABS_PRE: "D u = enat du" "u\S" "\v. v \ S \ D u \ D v" by (auto simp: D_invar'_def transfer_preconditions less_eq_enat_def split: enat.splits) then interpret Dijkstra_Invar "G_\ g" s D S by simp have COUPLING': "coupling ((relax_outgoing' u du (M_lookup V) (Q_lookup Q))(u := None)) - (M_lookup V(u \ du)) + ((M_lookup V)(u \ du)) (relax_outgoing u D) (Set.insert u S)" using coupling_step[OF COUPLING \u\S\ \D u = enat du\] by auto show "D_invar_impl (Q_delete u (Q_relax_outgoing u du V Q)) (M_update u du V)" using maintain_D_invar[OF \u\S\] ABS_PRE using COUPLING' by (auto simp: D_invar_impl_def D_invar'_def) show "unfinished_dnodes' (M_lookup (M_update u du V)) \ unfinished_dnodes' (M_lookup V) \ finite (unfinished_dnodes' (M_lookup V))" using coupling_unfinished[OF COUPLING] coupling_unfinished[OF COUPLING'] using unfinished_nodes_decr[OF \u\S\] ABS_PRE by simp qed lemma dijkstra_correct: "M_invar dijkstra" "M_lookup dijkstra u = Some d \ \ s u = enat d" using dijkstra_loop_invar_and_empty unfolding dijkstra_def apply - apply (all \clarsimp simp: D_invar_impl_def\) apply (clarsimp simp: D_invar'_def) subgoal for Q V D S using Dijkstra_Invar.invar_finish_imp_correct[of "G_\ g" s D S u] apply (clarsimp simp: coupling_def) by (auto simp: domIff) done end subsection \Instantiation of ADTs and Code Generation\ global_interpretation G: wgraph_by_map RBT_Set.empty RBT_Map.update RBT_Map.delete Lookup2.lookup RBT_Map.M.invar defines G_empty = G.empty_graph and G_add_edge = G.add_edge and G_succ = G.succ by unfold_locales global_interpretation Dijkstra_Impl_Adts G.\ G.invar G.succ G.empty_graph G.add_edge RBT_Set.empty RBT_Map.update RBT_Map.delete Lookup2.lookup RBT_Map.M.invar PST_RBT.empty PST_RBT.update PST_RBT.delete PST_RBT.PM.invar Lookup2.lookup PST_RBT.rbt_is_empty pst_getmin .. global_interpretation D: Dijkstra_Impl_Defs G.invar G.succ G.empty_graph G.add_edge RBT_Set.empty RBT_Map.update RBT_Map.delete Lookup2.lookup RBT_Map.M.invar PST_RBT.empty PST_RBT.update PST_RBT.delete PST_RBT.PM.invar Lookup2.lookup PST_RBT.rbt_is_empty pst_getmin G.\ g s for g and s::"'v::linorder" defines dijkstra = D.dijkstra and dijkstra_loop = D.dijkstra_loop and Q_relax_outgoing = D.Q_relax_outgoing and M_initV = D.M_initV and Q_initQ = D.Q_initQ .. (* TODO: Why is this fix necessary? *) lemmas [code] = D.dijkstra_def D.dijkstra_loop_def context fixes g assumes [simp]: "G.invar g" begin interpretation AUX: Dijkstra_Impl G.invar G.succ G.empty_graph G.add_edge RBT_Set.empty RBT_Map.update RBT_Map.delete Lookup2.lookup RBT_Map.M.invar PST_RBT.empty PST_RBT.update PST_RBT.delete PST_RBT.PM.invar Lookup2.lookup PST_RBT.rbt_is_empty pst_getmin g s G.\ for s by unfold_locales simp_all lemmas dijkstra_correct = AUX.dijkstra_correct[folded dijkstra_def] end subsection \Combination with Graph Parser\ text \We combine the algorithm with a parser from lists to graphs\ global_interpretation G: wgraph_from_list_algo G.\ G.invar G.succ G.empty_graph G.add_edge defines from_list = G.from_list .. definition "dijkstra_list l s \ if valid_graph_rep l then Some (dijkstra (from_list l) s) else None" theorem dijkstra_list_correct: "case dijkstra_list l s of None \ \valid_graph_rep l | Some D \ valid_graph_rep l \ M.invar D \ (\u d. lookup D u = Some d \ WGraph.\ (wgraph_of_list l) s u = enat d)" unfolding dijkstra_list_def by (auto simp: dijkstra_correct G.from_list_correct) export_code dijkstra_list checking SML OCaml? Scala Haskell? value "dijkstra_list [(1::nat,2,7),(1,3,1),(3,2,2)] 1" value "dijkstra_list [(1::nat,2,7),(1,3,1),(3,2,2)] 3" end diff --git a/thys/ROBDD/Pointer_Map.thy b/thys/ROBDD/Pointer_Map.thy --- a/thys/ROBDD/Pointer_Map.thy +++ b/thys/ROBDD/Pointer_Map.thy @@ -1,162 +1,162 @@ section\Pointermap\ theory Pointer_Map imports Main begin text\ We need a datastructure that supports the following two operations: \begin{itemize} \item Given an element, it can construct a pointer (i.e., a small representation) of that element. It will always construct the same pointer for equal elements. \item Given a pointer, we can retrieve the element \end{itemize} \ record 'a pointermap = entries :: "'a list" getentry :: "'a \ nat option" definition "pointermap_sane m \ (distinct (entries m) \ (\n \ {.. (\p i. getentry m p = Some i \ entries m ! i = p \ i < length (entries m)))" definition "empty_pointermap \ \entries = [], getentry = \p. None \" lemma pointermap_empty_sane[simp, intro!]: "pointermap_sane empty_pointermap" unfolding empty_pointermap_def pointermap_sane_def by simp definition "pointermap_insert a m \ \entries = (entries m)@[a], getentry = (getentry m)(a \ length (entries m))\" definition "pm_pth m p \ entries m ! p" definition "pointermap_p_valid p m \ p < length (entries m)" definition "pointermap_getmk a m \ (case getentry m a of Some p \ (p,m) | None \ let u = pointermap_insert a m in (the (getentry u a), u))" lemma pointermap_sane_appendD: "pointermap_sane s \ m \ set (entries s) \ pointermap_sane (pointermap_insert m s)" unfolding pointermap_sane_def pointermap_insert_def proof(intro conjI[rotated],goal_cases) case 3 thus ?case by simp next case 2 { fix n have " \distinct (entries s) \ (\x. x \ {.. getentry s (entries s ! x) = Some x) \ (\p i. getentry s p = Some i \ entries s ! i = p \ i < length (entries s)); m \ set (entries s); - n \ {..entries = entries s @ [m], getentry = getentry s(m \ length (entries s))\)}; n < length (entries s)\ - \ getentry \entries = entries s @ [m], getentry = getentry s(m \ length (entries s))\ (entries \entries = entries s @ [m], getentry = getentry s(m \ length (entries s))\ ! n) = Some n" - "\distinct (entries s) \ (\x. x \ {.. getentry s (entries s ! x) = Some x) \ (\p i. getentry s p = Some i \ entries s ! i = p \ i < length (entries s)); m \ set (entries s); - n \ {..entries = entries s @ [m], getentry = getentry s(m \ length (entries s))\)}; \ n < length (entries s)\ - \ getentry \entries = entries s @ [m], getentry = getentry s(m \ length (entries s))\ (entries \entries = entries s @ [m], getentry = getentry s(m \ length (entries s))\ ! n) = Some n" + n \ {..entries = entries s @ [m], getentry = (getentry s)(m \ length (entries s))\)}; n < length (entries s)\ + \ getentry \entries = entries s @ [m], getentry = (getentry s)(m \ length (entries s))\ (entries \entries = entries s @ [m], getentry = (getentry s)(m \ length (entries s))\ ! n) = Some n" + "\distinct (entries s) \ (\x. x \ {.. (getentry s) (entries s ! x) = Some x) \ (\p i. getentry s p = Some i \ entries s ! i = p \ i < length (entries s)); m \ set (entries s); + n \ {..entries = entries s @ [m], getentry = (getentry s)(m \ length (entries s))\)}; \ n < length (entries s)\ + \ getentry \entries = entries s @ [m], getentry = (getentry s)(m \ length (entries s))\ (entries \entries = entries s @ [m], getentry = (getentry s)(m \ length (entries s))\ ! n) = Some n" proof(goal_cases) case 1 note goal1 = 1 from goal1(4) have sa: "\a. (entries s @ a) ! n = entries s ! n" by (simp add: nth_append) from goal1(1,4) have ih: "getentry s (entries s ! n) = Some n" by simp from goal1(2,4) have ne: "entries s ! n \ m" using nth_mem by fastforce from sa ih ne show ?case by simp next case 2 note goal2 = 2 from goal2(3,4) have ln: "n = length (entries s)" by simp hence sa: "\a. (entries s @ [a]) ! n = a" by simp from sa ln show ?case by simp qed } note h = this with 2 show ?case by blast (*apply(unfold Ball_def) apply(rule) apply(rule) apply(rename_tac n) apply(case_tac "n < length (entries s)") by(fact h)+*) next case 1 thus ?case by(clarsimp simp add: nth_append fun_upd_same Ball_def) force qed lemma luentries_noneD: "getentry s a = None \ pointermap_sane s \ a \ set (entries s)" unfolding pointermap_sane_def proof(rule ccontr, goal_cases) case 1 from 1(3) obtain n where "n < length (entries s)" "entries s ! n = a" unfolding in_set_conv_nth by blast with 1(2,1) show False by force qed lemma pm_pth_append: "pointermap_p_valid p m \ pm_pth (pointermap_insert a m) p = pm_pth m p" unfolding pointermap_p_valid_def pm_pth_def pointermap_insert_def by(simp add: nth_append) lemma pointermap_insert_in: "u = (pointermap_insert a m) \ pm_pth u (the (getentry u a)) = a" unfolding pointermap_insert_def pm_pth_def by(simp) lemma pointermap_insert_p_validI: "pointermap_p_valid p m \ pointermap_p_valid p (pointermap_insert a m)" unfolding pointermap_insert_def pointermap_p_valid_def by simp thm nth_eq_iff_index_eq lemma pth_eq_iff_index_eq: "pointermap_sane m \ pointermap_p_valid p1 m \ pointermap_p_valid p2 m \ (pm_pth m p1 = pm_pth m p2) \ (p1 = p2)" unfolding pointermap_sane_def pointermap_p_valid_def pm_pth_def using nth_eq_iff_index_eq by blast lemma pointermap_p_valid_updateI: "pointermap_sane m \ getentry m a = None \ u = pointermap_insert a m \ p = the (getentry u a) \ pointermap_p_valid p u" by(simp add: pointermap_sane_def pointermap_p_valid_def pointermap_insert_def) lemma pointermap_get_validI: "pointermap_sane m \ getentry m a = Some p \ pointermap_p_valid p m" by(simp add: pointermap_sane_def pointermap_p_valid_def) lemma pointermap_sane_getmkD: assumes sn: "pointermap_sane m" assumes res: "pointermap_getmk a m = (p,u)" shows "pointermap_sane u \ pointermap_p_valid p u" using sn res[symmetric] apply(cases "getentry m a") apply(simp_all add: pointermap_getmk_def Let_def split: option.split) apply(rule) apply(rule pointermap_sane_appendD) apply(clarify;fail)+ apply(rule luentries_noneD) apply(clarify;fail)+ apply(rule pointermap_p_valid_updateI[OF _ _ refl refl]) apply(clarify;fail)+ apply(erule pointermap_get_validI) by simp lemma pointermap_update_pthI: assumes sn: "pointermap_sane m" assumes res: "pointermap_getmk a m = (p,u)" shows "pm_pth u p = a" using assms apply(simp add: pointermap_getmk_def Let_def split: option.splits) apply(meson pointermap_insert_in) apply(clarsimp simp: pointermap_sane_def pm_pth_def) done lemma pointermap_p_valid_inv: assumes "pointermap_p_valid p m" assumes "pointermap_getmk a m = (x,u)" shows "pointermap_p_valid p u" using assms by(simp add: pointermap_getmk_def Let_def split: option.splits) (meson pointermap_insert_p_validI) lemma pointermap_p_pth_inv: assumes pv: "pointermap_p_valid p m" assumes u: "pointermap_getmk a m = (x,u)" shows "pm_pth u p = pm_pth m p" using pm_pth_append[OF pv] u by(clarsimp simp: pointermap_getmk_def Let_def split: option.splits) lemma pointermap_backward_valid: assumes puv: "pointermap_p_valid p u" assumes u: "pointermap_getmk a m = (x,u)" assumes ne: "x \ p" shows "pointermap_p_valid p m" (*using u unfolding pointermap_getmk_def apply(simp add: Let_def split: option.splits) prefer 2 using puv apply(simp) apply(clarify) apply(simp add: pointermap_insert_def) using puv apply(clarify) apply(simp add: pointermap_p_valid_def) using ne by linarith *) using assms by (auto simp: Let_def pointermap_getmk_def pointermap_p_valid_def pointermap_insert_def split: option.splits) end diff --git a/thys/Refine_Imperative_HOL/IICF/Impl/Heaps/IICF_Abs_Heapmap.thy b/thys/Refine_Imperative_HOL/IICF/Impl/Heaps/IICF_Abs_Heapmap.thy --- a/thys/Refine_Imperative_HOL/IICF/Impl/Heaps/IICF_Abs_Heapmap.thy +++ b/thys/Refine_Imperative_HOL/IICF/Impl/Heaps/IICF_Abs_Heapmap.thy @@ -1,1214 +1,1214 @@ section \Priority Maps implemented with List and Map\ theory IICF_Abs_Heapmap imports IICF_Abs_Heap "HOL-Library.Rewrite" "../../Intf/IICF_Prio_Map" begin type_synonym ('k,'v) ahm = "'k list \ ('k \ 'v)" subsection \Basic Setup\ text \First, we define a mapping to list-based heaps\ definition hmr_\ :: "('k,'v) ahm \ 'v heap" where "hmr_\ \ \(pq,m). map (the o m) pq" definition "hmr_invar \ \(pq,m). distinct pq \ dom m = set pq" definition "hmr_rel \ br hmr_\ hmr_invar" lemmas hmr_rel_defs = hmr_rel_def br_def hmr_\_def hmr_invar_def lemma hmr_empty_invar[simp]: "hmr_invar ([],Map.empty)" by (auto simp: hmr_invar_def) locale hmstruct = h: heapstruct prio for prio :: "'v \ 'b::linorder" begin text \Next, we define a mapping to priority maps.\ definition heapmap_\ :: "('k,'v) ahm \ ('k \ 'v)" where "heapmap_\ \ \(pq,m). m" definition heapmap_invar :: "('k,'v) ahm \ bool" where "heapmap_invar \ \hm. hmr_invar hm \ h.heap_invar (hmr_\ hm)" definition "heapmap_rel \ br heapmap_\ heapmap_invar" lemmas heapmap_rel_defs = heapmap_rel_def br_def heapmap_\_def heapmap_invar_def lemma [refine_dref_RELATES]: "RELATES hmr_rel" by (simp add: RELATES_def) lemma h_heap_invarI[simp]: "heapmap_invar hm \ h.heap_invar (hmr_\ hm)" by (simp add: heapmap_invar_def) lemma hmr_invarI[simp]: "heapmap_invar hm \ hmr_invar hm" unfolding heapmap_invar_def by blast lemma set_hmr_\[simp]: "hmr_invar hm \ set (hmr_\ hm) = ran (heapmap_\ hm)" apply (clarsimp simp: hmr_\_def hmr_invar_def heapmap_\_def eq_commute[of "dom _" "set _"] ran_def) apply force done lemma in_h_hmr_\_conv[simp]: "hmr_invar hm \ x \# h.\ (hmr_\ hm) \ x \ ran (heapmap_\ hm)" by (force simp: hmr_\_def hmr_invar_def heapmap_\_def in_multiset_in_set ran_is_image) subsection \Basic Operations\ (* length, val_of_op, update, butlast, append, empty *) text \In this section, we define the basic operations on heapmaps, and their relations to heaps and maps.\ subsubsection \Length\ text \Length of the list that represents the heap\ definition hm_length :: "('k,'v) ahm \ nat" where "hm_length \ \(pq,_). length pq" lemma hm_length_refine: "(hm_length, length) \ hmr_rel \ nat_rel" apply (intro fun_relI) unfolding hm_length_def by (auto simp: hmr_rel_defs) lemma hm_length_hmr_\[simp]: "length (hmr_\ hm) = hm_length hm" by (auto simp: hm_length_def hmr_\_def split: prod.splits) lemmas [refine] = hm_length_refine[param_fo] subsubsection \Valid\ text \Check whether index is valid\ definition "hm_valid hm i \ i>0 \ i\ hm_length hm" lemma hm_valid_refine: "(hm_valid,h.valid)\hmr_rel \ nat_rel \ bool_rel" apply (intro fun_relI) unfolding hm_valid_def h.valid_def by (parametricity add: hm_length_refine) lemma hm_valid_hmr_\[simp]: "h.valid (hmr_\ hm) = hm_valid hm" by (intro ext) (auto simp: h.valid_def hm_valid_def) subsubsection \Key-Of\ definition hm_key_of :: "('k,'v) ahm \ nat \ 'k" where "hm_key_of \ \(pq,m) i. pq!(i - 1)" definition hm_key_of_op :: "('k,'v) ahm \ nat \ 'k nres" where "hm_key_of_op \ \(pq,m) i. ASSERT (i>0) \ mop_list_get pq (i - 1)" lemma hm_key_of_op_unfold: shows "hm_key_of_op hm i = ASSERT (hm_valid hm i) \ RETURN (hm_key_of hm i)" unfolding hm_valid_def hm_length_def hm_key_of_op_def hm_key_of_def by (auto split: prod.splits simp: pw_eq_iff refine_pw_simps) lemma val_of_hmr_\[simp]: "hm_valid hm i \ h.val_of (hmr_\ hm) i = the (heapmap_\ hm (hm_key_of hm i))" by (auto simp: hmr_\_def h.val_of_def heapmap_\_def hm_key_of_def hm_valid_def hm_length_def split: prod.splits) lemma hm_\_key_ex[simp]: "\hmr_invar hm; hm_valid hm i\ \ (heapmap_\ hm (hm_key_of hm i) \ None)" unfolding heapmap_invar_def hmr_invar_def hm_valid_def heapmap_\_def hm_key_of_def hm_length_def by (auto split: prod.splits) subsubsection \Lookup\ abbreviation (input) hm_lookup where "hm_lookup \ heapmap_\" definition "hm_the_lookup_op hm k \ ASSERT (heapmap_\ hm k \ None \ hmr_invar hm) \ RETURN (the (heapmap_\ hm k))" subsubsection \Exchange\ text \Exchange two indices\ definition "hm_exch_op \ \(pq,m) i j. do { ASSERT (hm_valid (pq,m) i); ASSERT (hm_valid (pq,m) j); ASSERT (hmr_invar (pq,m)); pq \ mop_list_swap pq (i - 1) (j - 1); RETURN (pq,m) }" lemma hm_exch_op_invar: "hm_exch_op hm i j \\<^sub>n SPEC hmr_invar" unfolding hm_exch_op_def h.exch_op_def h.val_of_op_def h.update_op_def apply simp apply refine_vcg apply (auto simp: hm_valid_def map_swap hm_length_def hmr_rel_defs) done lemma hm_exch_op_refine: "(hm_exch_op,h.exch_op) \ hmr_rel \ nat_rel \ nat_rel \ \hmr_rel\nres_rel" apply (intro fun_relI nres_relI) unfolding hm_exch_op_def h.exch_op_def h.val_of_op_def h.update_op_def apply simp apply refine_vcg apply (auto simp: hm_valid_def map_swap hm_length_def hmr_rel_defs) done lemmas hm_exch_op_refine'[refine] = hm_exch_op_refine[param_fo, THEN nres_relD] definition hm_exch :: "('k,'v) ahm \ nat \ nat \ ('k,'v) ahm" where "hm_exch \ \(pq,m) i j. (swap pq (i-1) (j-1),m)" lemma hm_exch_op_\_correct: "hm_exch_op hm i j \\<^sub>n SPEC (\hm'. hm_valid hm i \ hm_valid hm j \ hm'=hm_exch hm i j )" unfolding hm_exch_op_def apply refine_vcg apply (vc_solve simp: hm_valid_def hm_length_def heapmap_\_def solve: asm_rl) apply (auto simp add: hm_key_of_def hm_exch_def swap_def) [] done lemma hm_exch_\[simp]: "heapmap_\ (hm_exch hm i j) = (heapmap_\ hm)" by (auto simp: heapmap_\_def hm_exch_def split: prod.splits) lemma hm_exch_valid[simp]: "hm_valid (hm_exch hm i j) = hm_valid hm" by (intro ext) (auto simp: hm_valid_def hm_length_def hm_exch_def split: prod.splits) lemma hm_exch_length[simp]: "hm_length (hm_exch hm i j) = hm_length hm" by (auto simp: hm_length_def hm_exch_def split: prod.splits) lemma hm_exch_same[simp]: "hm_exch hm i i = hm" by (auto simp: hm_exch_def split: prod.splits) lemma hm_key_of_exch_conv[simp]: "\hm_valid hm i; hm_valid hm j; hm_valid hm k\ \ hm_key_of (hm_exch hm i j) k = ( if k=i then hm_key_of hm j else if k=j then hm_key_of hm i else hm_key_of hm k )" unfolding hm_exch_def hm_valid_def hm_length_def hm_key_of_def by (auto split: prod.splits) lemma hm_key_of_exch_matching[simp]: "\hm_valid hm i; hm_valid hm j\ \ hm_key_of (hm_exch hm i j) i = hm_key_of hm j" "\hm_valid hm i; hm_valid hm j\ \ hm_key_of (hm_exch hm i j) j = hm_key_of hm i" by simp_all subsubsection \Index\ text \Obtaining the index of a key\ definition "hm_index \ \(pq,m) k. index pq k + 1" lemma hm_index_valid[simp]: "\hmr_invar hm; heapmap_\ hm k \ None\ \ hm_valid hm (hm_index hm k)" by (force simp: hm_valid_def heapmap_\_def hmr_invar_def hm_index_def hm_length_def Suc_le_eq) lemma hm_index_key_of[simp]: "\hmr_invar hm; heapmap_\ hm k \ None\ \ hm_key_of hm (hm_index hm k) = k" by (force simp: hm_valid_def heapmap_\_def hmr_invar_def hm_index_def hm_length_def hm_key_of_def Suc_le_eq) definition "hm_index_op \ \(pq,m) k. do { ASSERT (hmr_invar (pq,m) \ heapmap_\ (pq,m) k \ None); i \ mop_list_index pq k; RETURN (i+1) }" lemma hm_index_op_correct: assumes "hmr_invar hm" assumes "heapmap_\ hm k \ None" shows "hm_index_op hm k \ SPEC (\r. r= hm_index hm k)" using assms unfolding hm_index_op_def apply refine_vcg apply (auto simp: heapmap_\_def hmr_invar_def hm_index_def index_nth_id) done lemmas [refine_vcg] = hm_index_op_correct subsubsection \Update\ text \Updating the heap at an index\ definition hm_update_op :: "('k,'v) ahm \ nat \ 'v \ ('k,'v) ahm nres" where "hm_update_op \ \(pq,m) i v. do { ASSERT (hm_valid (pq,m) i \ hmr_invar (pq,m)); k \ mop_list_get pq (i - 1); RETURN (pq, m(k \ v)) }" lemma hm_update_op_invar: "hm_update_op hm k v \\<^sub>n SPEC hmr_invar" unfolding hm_update_op_def h.update_op_def apply refine_vcg by (auto simp: hmr_rel_defs map_distinct_upd_conv hm_valid_def hm_length_def) lemma hm_update_op_refine: "(hm_update_op, h.update_op) \ hmr_rel \ nat_rel \ Id \ \hmr_rel\nres_rel" apply (intro fun_relI nres_relI) unfolding hm_update_op_def h.update_op_def mop_list_get_alt mop_list_set_alt apply refine_vcg apply (auto simp: hmr_rel_defs map_distinct_upd_conv hm_valid_def hm_length_def) done lemmas [refine] = hm_update_op_refine[param_fo, THEN nres_relD] lemma hm_update_op_\_correct: assumes "hmr_invar hm" assumes "heapmap_\ hm k \ None" shows "hm_update_op hm (hm_index hm k) v \\<^sub>n SPEC (\hm'. heapmap_\ hm' = (heapmap_\ hm)(k\v))" using assms unfolding hm_update_op_def apply refine_vcg apply (force simp: heapmap_rel_defs hmr_rel_defs hm_index_def) done subsubsection \Butlast\ text \Remove last element\ definition hm_butlast_op :: "('k,'v) ahm \ ('k,'v) ahm nres" where "hm_butlast_op \ \(pq,m). do { ASSERT (hmr_invar (pq,m)); k \ mop_list_get pq (length pq - 1); pq \ mop_list_butlast pq; let m = m(k:=None); RETURN (pq,m) }" lemma hm_butlast_op_refine: "(hm_butlast_op, h.butlast_op) \ hmr_rel \ \hmr_rel\nres_rel" supply [simp del] = map_upd_eq_restrict apply (intro fun_relI nres_relI) unfolding hm_butlast_op_def h.butlast_op_def apply simp apply refine_vcg apply (clarsimp_all simp: hmr_rel_defs map_butlast distinct_butlast) apply (auto simp: neq_Nil_rev_conv) [] done lemmas [refine] = hm_butlast_op_refine[param_fo, THEN nres_relD] lemma hm_butlast_op_\_correct: "hm_butlast_op hm \\<^sub>n SPEC ( \hm'. heapmap_\ hm' = (heapmap_\ hm)( hm_key_of hm (hm_length hm) := None ))" unfolding hm_butlast_op_def apply refine_vcg apply (auto simp: heapmap_\_def hm_key_of_def hm_length_def) done subsubsection \Append\ text \Append new element at end of heap\ definition hm_append_op :: "('k,'v) ahm \ 'k \ 'v \ ('k,'v) ahm nres" where "hm_append_op \ \(pq,m) k v. do { ASSERT (k \ dom m); ASSERT (hmr_invar (pq,m)); pq \ mop_list_append pq k; let m = m (k \ v); RETURN (pq,m) }" lemma hm_append_op_invar: "hm_append_op hm k v \\<^sub>n SPEC hmr_invar" unfolding hm_append_op_def h.append_op_def apply refine_vcg unfolding heapmap_\_def hmr_rel_defs apply auto done lemma hm_append_op_refine: "\ heapmap_\ hm k = None; (hm,h)\hmr_rel \ \ (hm_append_op hm k v, h.append_op h v) \ \hmr_rel\nres_rel" apply (intro fun_relI nres_relI) unfolding hm_append_op_def h.append_op_def apply refine_vcg unfolding heapmap_\_def hmr_rel_defs apply auto done lemmas hm_append_op_refine'[refine] = hm_append_op_refine[param_fo, THEN nres_relD] lemma hm_append_op_\_correct: "hm_append_op hm k v \\<^sub>n SPEC (\hm'. heapmap_\ hm' = (heapmap_\ hm) (k \ v))" unfolding hm_append_op_def apply refine_vcg by (auto simp: heapmap_\_def) subsection \Auxiliary Operations\ text \Auxiliary operations on heapmaps, which are derived from the basic operations, but do not correspond to operations of the priority map interface\ text \We start with some setup\ lemma heapmap_hmr_relI: "(hm,h)\heapmap_rel \ (hm,hmr_\ hm) \ hmr_rel" by (auto simp: heapmap_rel_defs hmr_rel_defs) lemma heapmap_hmr_relI': "heapmap_invar hm \ (hm,hmr_\ hm) \ hmr_rel" by (auto simp: heapmap_rel_defs hmr_rel_defs) text \The basic principle how we prove correctness of our operations: Invariant preservation is shown by relating the operations to operations on heaps. Then, only correctness on the abstraction remains to be shown, assuming the operation does not fail. \ lemma heapmap_nres_relI': assumes "hm \ \hmr_rel h'" assumes "h' \ SPEC (h.heap_invar)" assumes "hm \\<^sub>n SPEC (\hm'. RETURN (heapmap_\ hm') \ h)" shows "hm \ \heapmap_rel h" using assms unfolding heapmap_rel_defs hmr_rel_def by (auto simp: pw_le_iff pw_leof_iff refine_pw_simps) lemma heapmap_nres_relI'': assumes "hm \ \hmr_rel h'" assumes "h' \ SPEC \" assumes "\h'. \ h' \ h.heap_invar h'" assumes "hm \\<^sub>n SPEC (\hm'. RETURN (heapmap_\ hm') \ h)" shows "hm \ \heapmap_rel h" apply (rule heapmap_nres_relI') apply fact apply (rule order_trans, fact) apply (clarsimp; fact) apply fact done subsubsection \Val-of\ text \Indexing into the heap\ definition hm_val_of_op :: "('k,'v) ahm \ nat \ 'v nres" where "hm_val_of_op \ \hm i. do { k \ hm_key_of_op hm i; v \ hm_the_lookup_op hm k; RETURN v }" lemma hm_val_of_op_refine: "(hm_val_of_op,h.val_of_op) \ (hmr_rel \ nat_rel \ \Id\nres_rel)" apply (intro fun_relI nres_relI) unfolding hm_val_of_op_def h.val_of_op_def hm_key_of_op_def hm_key_of_def hm_valid_def hm_length_def hm_the_lookup_op_def apply clarsimp apply (rule refine_IdD) apply refine_vcg apply (auto simp: hmr_rel_defs heapmap_\_def) done lemmas [refine] = hm_val_of_op_refine[param_fo, THEN nres_relD] subsubsection \Prio-of\ text \Priority of key\ definition "hm_prio_of_op h i \ do {v \ hm_val_of_op h i; RETURN (prio v)}" lemma hm_prio_of_op_refine: "(hm_prio_of_op, h.prio_of_op) \ hmr_rel \ nat_rel \ \Id\nres_rel" apply (intro fun_relI nres_relI) unfolding hm_prio_of_op_def h.prio_of_op_def apply refine_rcg by auto lemmas hm_prio_of_op_refine'[refine] = hm_prio_of_op_refine[param_fo, THEN nres_relD] subsubsection \Swim\ definition hm_swim_op :: "('k,'v) ahm \ nat \ ('k,'v) ahm nres" where "hm_swim_op h i \ do { RECT (\swim (h,i). do { ASSERT (hm_valid h i \ h.swim_invar (hmr_\ h) i); if hm_valid h (h.parent i) then do { ppi \ hm_prio_of_op h (h.parent i); pi \ hm_prio_of_op h i; if (\ppi \ pi) then do { h \ hm_exch_op h i (h.parent i); swim (h, h.parent i) } else RETURN h } else RETURN h }) (h,i) }" lemma hm_swim_op_refine: "(hm_swim_op, h.swim_op) \ hmr_rel \ nat_rel \ \hmr_rel\nres_rel" apply (intro fun_relI nres_relI) unfolding hm_swim_op_def h.swim_op_def apply refine_rcg apply refine_dref_type apply (clarsimp_all simp: hm_valid_refine[param_fo, THEN IdD]) apply (simp add: hmr_rel_def in_br_conv) done lemmas hm_swim_op_refine'[refine] = hm_swim_op_refine[param_fo, THEN nres_relD] lemma hm_swim_op_nofail_imp_valid: "nofail (hm_swim_op hm i) \ hm_valid hm i \ h.swim_invar (hmr_\ hm) i" unfolding hm_swim_op_def apply (subst (asm) RECT_unfold, refine_mono) by (auto simp: refine_pw_simps) lemma hm_swim_op_\_correct: "hm_swim_op hm i \\<^sub>n SPEC (\hm'. heapmap_\ hm' = heapmap_\ hm)" apply (rule leof_add_nofailI) apply (drule hm_swim_op_nofail_imp_valid) unfolding hm_swim_op_def apply (rule RECT_rule_leof[where pre="\(hm',i). hm_valid hm' i \ heapmap_\ hm' = heapmap_\ hm" and V = "inv_image less_than snd" ]) apply simp apply simp unfolding hm_prio_of_op_def hm_val_of_op_def hm_exch_op_def hm_key_of_op_def hm_the_lookup_op_def apply (refine_vcg) apply (vc_solve simp add: hm_valid_def hm_length_def) apply rprems apply (vc_solve simp: heapmap_\_def h.parent_def) done subsubsection \Sink\ definition hm_sink_op where "hm_sink_op h k \ RECT (\D (h,k). do { ASSERT (k>0 \ k\hm_length h); let len = hm_length h; if (2*k \ len) then do { let j = 2*k; pj \ hm_prio_of_op h j; j \ ( if j hm_prio_of_op h (Suc j); if pj>psj then RETURN (j+1) else RETURN j } else RETURN j); pj \ hm_prio_of_op h j; pk \ hm_prio_of_op h k; if (pk > pj) then do { h \ hm_exch_op h k j; D (h,j) } else RETURN h } else RETURN h }) (h,k)" lemma hm_sink_op_refine: "(hm_sink_op, h.sink_op) \ hmr_rel \ nat_rel \ \hmr_rel\nres_rel" apply (intro fun_relI nres_relI) unfolding hm_sink_op_def h.sink_op_opt_eq[symmetric] h.sink_op_opt_def apply refine_rcg apply refine_dref_type unfolding hmr_rel_def heapmap_rel_def apply (clarsimp_all simp: in_br_conv) done lemmas hm_sink_op_refine'[refine] = hm_sink_op_refine[param_fo, THEN nres_relD] lemma hm_sink_op_nofail_imp_valid: "nofail (hm_sink_op hm i) \ hm_valid hm i" unfolding hm_sink_op_def apply (subst (asm) RECT_unfold, refine_mono) by (auto simp: refine_pw_simps hm_valid_def) lemma hm_sink_op_\_correct: "hm_sink_op hm i \\<^sub>n SPEC (\hm'. heapmap_\ hm' = heapmap_\ hm)" apply (rule leof_add_nofailI) apply (drule hm_sink_op_nofail_imp_valid) unfolding hm_sink_op_def apply (rule RECT_rule_leof[where pre="\(hm',i). hm_valid hm' i \ heapmap_\ hm' = heapmap_\ hm \ hm_length hm' = hm_length hm" and V = "measure (\(hm',i). hm_length hm' - i)" ]) apply simp apply simp unfolding hm_prio_of_op_def hm_val_of_op_def hm_exch_op_def hm_key_of_op_def hm_the_lookup_op_def apply (refine_vcg) apply (vc_solve simp add: hm_valid_def hm_length_def) (* Takes long *) apply rprems apply (vc_solve simp: heapmap_\_def h.parent_def split: prod.splits) apply (auto) done subsubsection \Repair\ definition "hm_repair_op hm i \ do { hm \ hm_sink_op hm i; hm \ hm_swim_op hm i; RETURN hm }" lemma hm_repair_op_refine: "(hm_repair_op, h.repair_op) \ hmr_rel \ nat_rel \ \hmr_rel\nres_rel" apply (intro fun_relI nres_relI) unfolding hm_repair_op_def h.repair_op_def by refine_rcg lemmas hm_repair_op_refine'[refine] = hm_repair_op_refine[param_fo, THEN nres_relD] lemma hm_repair_op_\_correct: "hm_repair_op hm i \\<^sub>n SPEC (\hm'. heapmap_\ hm' = heapmap_\ hm)" unfolding hm_repair_op_def apply (refine_vcg hm_swim_op_\_correct[THEN leof_trans] hm_sink_op_\_correct[THEN leof_trans]) by auto subsection \Operations\ text \In this section, we define the operations that implement the priority-map interface\ subsubsection \Empty\ definition hm_empty_op :: "('k,'v) ahm nres" where "hm_empty_op \ RETURN ([],Map.empty)" lemma hm_empty_aref: "(hm_empty_op,RETURN op_map_empty) \ \heapmap_rel\nres_rel" unfolding hm_empty_op_def by (auto simp: heapmap_rel_defs hmr_rel_defs intro: nres_relI) subsubsection \Insert\ definition hm_insert_op :: "'k \ 'v \ ('k,'v) ahm \ ('k,'v) ahm nres" where "hm_insert_op \ \k v h. do { ASSERT (h.heap_invar (hmr_\ h)); h \ hm_append_op h k v; let l = hm_length h; h \ hm_swim_op h l; RETURN h }" lemma hm_insert_op_refine[refine]: "\ heapmap_\ hm k = None; (hm,h)\hmr_rel \ \ hm_insert_op k v hm \ \hmr_rel (h.insert_op v h)" unfolding hm_insert_op_def h.insert_op_def apply refine_rcg by (auto simp: hmr_rel_def br_def) lemma hm_insert_op_aref: "(hm_insert_op,mop_map_update_new) \ Id \ Id \ heapmap_rel \ \heapmap_rel\nres_rel" apply (intro fun_relI nres_relI) unfolding mop_map_update_new_alt apply (rule ASSERT_refine_right) apply (rule heapmap_nres_relI''[OF hm_insert_op_refine h.insert_op_correct]) apply (unfold heapmap_rel_def in_br_conv; clarsimp) apply (erule heapmap_hmr_relI) apply (unfold heapmap_rel_def in_br_conv; clarsimp) apply (unfold heapmap_rel_def in_br_conv; clarsimp) unfolding hm_insert_op_def apply (refine_vcg hm_append_op_\_correct[THEN leof_trans] hm_swim_op_\_correct[THEN leof_trans]) apply (unfold heapmap_rel_def in_br_conv; clarsimp) done subsubsection \Is-Empty\ lemma hmr_\_empty_iff[simp]: "hmr_invar hm \ hmr_\ hm = [] \ heapmap_\ hm = Map.empty" by (auto simp: hmr_\_def heapmap_invar_def heapmap_\_def hmr_invar_def split: prod.split) definition hm_is_empty_op :: "('k,'v) ahm \ bool nres" where "hm_is_empty_op \ \hm. do { ASSERT (hmr_invar hm); let l = hm_length hm; RETURN (l=0) }" lemma hm_is_empty_op_refine: "(hm_is_empty_op, h.is_empty_op) \ hmr_rel \ \bool_rel\nres_rel" apply (intro fun_relI nres_relI) unfolding hm_is_empty_op_def h.is_empty_op_def apply refine_rcg apply (auto simp: hmr_rel_defs) [] apply (parametricity add: hm_length_refine) done lemma hm_is_empty_op_aref: "(hm_is_empty_op, RETURN o op_map_is_empty) \ heapmap_rel \ \bool_rel\nres_rel" apply (intro fun_relI nres_relI) unfolding hm_is_empty_op_def apply refine_vcg apply (auto simp: hmr_rel_defs heapmap_rel_defs hm_length_def) done subsubsection \Lookup\ definition hm_lookup_op :: "'k \ ('k,'v) ahm \ 'v option nres" where "hm_lookup_op \ \k hm. ASSERT (heapmap_invar hm) \ RETURN (hm_lookup hm k)" lemma hm_lookup_op_aref: "(hm_lookup_op,RETURN oo op_map_lookup) \ Id \ heapmap_rel \ \\Id\option_rel\nres_rel" apply (intro fun_relI nres_relI) unfolding hm_lookup_op_def heapmap_rel_def in_br_conv apply refine_vcg apply simp_all done subsubsection \Contains-Key\ definition "hm_contains_key_op \ \k (pq,m). ASSERT (heapmap_invar (pq,m)) \ RETURN (k\dom m)" lemma hm_contains_key_op_aref: "(hm_contains_key_op,RETURN oo op_map_contains_key) \ Id \ heapmap_rel \ \bool_rel\nres_rel" apply (intro fun_relI nres_relI) unfolding hm_contains_key_op_def heapmap_rel_defs apply refine_vcg by (auto) subsubsection \Decrease-Key\ definition "hm_decrease_key_op \ \k v hm. do { ASSERT (heapmap_invar hm); ASSERT (heapmap_\ hm k \ None \ prio v \ prio (the (heapmap_\ hm k))); i \ hm_index_op hm k; hm \ hm_update_op hm i v; hm_swim_op hm i }" definition (in heapstruct) "decrease_key_op i v h \ do { ASSERT (valid h i \ prio v \ prio_of h i); h \ update_op h i v; swim_op h i }" lemma (in heapstruct) decrease_key_op_invar: "\heap_invar h; valid h i; prio v \ prio_of h i\ \ decrease_key_op i v h \ SPEC heap_invar" unfolding decrease_key_op_def apply refine_vcg by (auto simp: swim_invar_decr) lemma index_op_inline_refine: assumes "heapmap_invar hm" assumes "heapmap_\ hm k \ None" assumes "f (hm_index hm k) \ m" shows "do {i \ hm_index_op hm k; f i} \ m" using hm_index_op_correct[of hm k] assms by (auto simp: pw_le_iff refine_pw_simps) lemma hm_decrease_key_op_refine: "\(hm,h)\hmr_rel; (hm,m)\heapmap_rel; m k = Some v'\ \ hm_decrease_key_op k v hm \\hmr_rel (h.decrease_key_op (hm_index hm k) v h)" unfolding hm_decrease_key_op_def h.decrease_key_op_def (*apply (rewrite at "Let (hm_index hm k) _" Let_def)*) apply (refine_rcg index_op_inline_refine) unfolding hmr_rel_def heapmap_rel_def in_br_conv apply (clarsimp_all) done lemma hm_index_op_inline_leof: assumes "f (hm_index hm k) \\<^sub>n m" shows "do {i \ hm_index_op hm k; f i} \\<^sub>n m" using hm_index_op_correct[of hm k] assms unfolding hm_index_op_def by (auto simp: pw_le_iff pw_leof_iff refine_pw_simps split: prod.splits) lemma hm_decrease_key_op_\_correct: - "heapmap_invar hm \ hm_decrease_key_op k v hm \\<^sub>n SPEC (\hm'. heapmap_\ hm' = heapmap_\ hm(k\v))" + "heapmap_invar hm \ hm_decrease_key_op k v hm \\<^sub>n SPEC (\hm'. heapmap_\ hm' = (heapmap_\ hm)(k\v))" unfolding hm_decrease_key_op_def apply (refine_vcg hm_update_op_\_correct[THEN leof_trans] hm_swim_op_\_correct[THEN leof_trans] hm_index_op_inline_leof ) apply simp_all done lemma hm_decrease_key_op_aref: "(hm_decrease_key_op, PR_CONST (mop_pm_decrease_key prio)) \ Id \ Id \ heapmap_rel \ \heapmap_rel\nres_rel" unfolding PR_CONST_def apply (intro fun_relI nres_relI) apply (frule heapmap_hmr_relI) unfolding mop_pm_decrease_key_alt apply (rule ASSERT_refine_right; clarsimp) apply (rule heapmap_nres_relI') apply (rule hm_decrease_key_op_refine; assumption) unfolding heapmap_rel_def hmr_rel_def in_br_conv apply (rule h.decrease_key_op_invar; simp; fail ) apply (refine_vcg hm_decrease_key_op_\_correct[THEN leof_trans]; simp; fail) done subsubsection \Increase-Key\ definition "hm_increase_key_op \ \k v hm. do { ASSERT (heapmap_invar hm); ASSERT (heapmap_\ hm k \ None \ prio v \ prio (the (heapmap_\ hm k))); i \ hm_index_op hm k; hm \ hm_update_op hm i v; hm_sink_op hm i }" definition (in heapstruct) "increase_key_op i v h \ do { ASSERT (valid h i \ prio v \ prio_of h i); h \ update_op h i v; sink_op h i }" lemma (in heapstruct) increase_key_op_invar: "\heap_invar h; valid h i; prio v \ prio_of h i\ \ increase_key_op i v h \ SPEC heap_invar" unfolding increase_key_op_def apply refine_vcg by (auto simp: sink_invar_incr) lemma hm_increase_key_op_refine: "\(hm,h)\hmr_rel; (hm,m)\heapmap_rel; m k = Some v'\ \ hm_increase_key_op k v hm \\hmr_rel (h.increase_key_op (hm_index hm k) v h)" unfolding hm_increase_key_op_def h.increase_key_op_def (*apply (rewrite at "Let (hm_index hm k) _" Let_def)*) apply (refine_rcg index_op_inline_refine) unfolding hmr_rel_def heapmap_rel_def in_br_conv apply (clarsimp_all) done lemma hm_increase_key_op_\_correct: - "heapmap_invar hm \ hm_increase_key_op k v hm \\<^sub>n SPEC (\hm'. heapmap_\ hm' = heapmap_\ hm(k\v))" + "heapmap_invar hm \ hm_increase_key_op k v hm \\<^sub>n SPEC (\hm'. heapmap_\ hm' = (heapmap_\ hm)(k\v))" unfolding hm_increase_key_op_def apply (refine_vcg hm_update_op_\_correct[THEN leof_trans] hm_sink_op_\_correct[THEN leof_trans] hm_index_op_inline_leof) apply simp_all done lemma hm_increase_key_op_aref: "(hm_increase_key_op, PR_CONST (mop_pm_increase_key prio)) \ Id \ Id \ heapmap_rel \ \heapmap_rel\nres_rel" unfolding PR_CONST_def apply (intro fun_relI nres_relI) apply (frule heapmap_hmr_relI) unfolding mop_pm_increase_key_alt apply (rule ASSERT_refine_right; clarsimp) apply (rule heapmap_nres_relI') apply (rule hm_increase_key_op_refine; assumption) unfolding heapmap_rel_def hmr_rel_def in_br_conv apply (rule h.increase_key_op_invar; simp; fail ) apply (refine_vcg hm_increase_key_op_\_correct[THEN leof_trans]; simp) done subsubsection \Change-Key\ definition "hm_change_key_op \ \k v hm. do { ASSERT (heapmap_invar hm); ASSERT (heapmap_\ hm k \ None); i \ hm_index_op hm k; hm \ hm_update_op hm i v; hm_repair_op hm i }" definition (in heapstruct) "change_key_op i v h \ do { ASSERT (valid h i); h \ update_op h i v; repair_op h i }" lemma (in heapstruct) change_key_op_invar: "\heap_invar h; valid h i\ \ change_key_op i v h \ SPEC heap_invar" unfolding change_key_op_def apply (refine_vcg) apply hypsubst apply refine_vcg by (auto simp: sink_invar_incr) lemma hm_change_key_op_refine: "\(hm,h)\hmr_rel; (hm,m)\heapmap_rel; m k = Some v'\ \ hm_change_key_op k v hm \\hmr_rel (h.change_key_op (hm_index hm k) v h)" unfolding hm_change_key_op_def h.change_key_op_def (*apply (rewrite at "Let (hm_index hm k) _" Let_def)*) apply (refine_rcg index_op_inline_refine) unfolding hmr_rel_def heapmap_rel_def in_br_conv apply (clarsimp_all) done lemma hm_change_key_op_\_correct: - "heapmap_invar hm \ hm_change_key_op k v hm \\<^sub>n SPEC (\hm'. heapmap_\ hm' = heapmap_\ hm(k\v))" + "heapmap_invar hm \ hm_change_key_op k v hm \\<^sub>n SPEC (\hm'. heapmap_\ hm' = (heapmap_\ hm)(k\v))" unfolding hm_change_key_op_def apply (refine_vcg hm_update_op_\_correct[THEN leof_trans] hm_repair_op_\_correct[THEN leof_trans] hm_index_op_inline_leof) unfolding heapmap_rel_def in_br_conv apply simp apply simp done lemma hm_change_key_op_aref: "(hm_change_key_op, mop_map_update_ex) \ Id \ Id \ heapmap_rel \ \heapmap_rel\nres_rel" apply (intro fun_relI nres_relI) apply (frule heapmap_hmr_relI) unfolding mop_map_update_ex_alt apply (rule ASSERT_refine_right; clarsimp) apply (rule heapmap_nres_relI') apply (rule hm_change_key_op_refine; assumption) unfolding heapmap_rel_def hmr_rel_def in_br_conv apply (rule h.change_key_op_invar; simp; fail ) apply ((refine_vcg hm_change_key_op_\_correct[THEN leof_trans]; simp)) done subsubsection \Set\ text \Realized as generic algorithm!\ (* TODO: Implement as such! *) lemma (in -) op_pm_set_gen_impl: "RETURN ooo op_map_update = (\k v m. do { c \ RETURN (op_map_contains_key k m); if c then mop_map_update_ex k v m else mop_map_update_new k v m })" apply (intro ext) unfolding op_map_contains_key_def mop_map_update_ex_def mop_map_update_new_def by simp definition "hm_set_op k v hm \ do { c \ hm_contains_key_op k hm; if c then hm_change_key_op k v hm else hm_insert_op k v hm }" lemma hm_set_op_aref: "(hm_set_op, RETURN ooo op_map_update) \ Id \ Id \ heapmap_rel \ \heapmap_rel\nres_rel" unfolding op_pm_set_gen_impl apply (intro fun_relI nres_relI) unfolding hm_set_op_def o_def apply (refine_rcg hm_contains_key_op_aref[param_fo, unfolded o_def, THEN nres_relD] hm_change_key_op_aref[param_fo, THEN nres_relD] hm_insert_op_aref[param_fo, THEN nres_relD] ) by auto subsubsection \Pop-Min\ definition hm_pop_min_op :: "('k,'v) ahm \ (('k\'v) \ ('k,'v) ahm) nres" where "hm_pop_min_op hm \ do { ASSERT (heapmap_invar hm); ASSERT (hm_valid hm 1); k \ hm_key_of_op hm 1; v \ hm_the_lookup_op hm k; let l = hm_length hm; hm \ hm_exch_op hm 1 l; hm \ hm_butlast_op hm; if (l\1) then do { hm \ hm_sink_op hm 1; RETURN ((k,v),hm) } else RETURN ((k,v),hm) }" lemma hm_pop_min_op_refine: "(hm_pop_min_op, h.pop_min_op) \ hmr_rel \ \UNIV \\<^sub>r hmr_rel\nres_rel" apply (intro fun_relI nres_relI) unfolding hm_pop_min_op_def h.pop_min_op_def (* Project away stuff of second component *) unfolding ignore_snd_refine_conv hm_the_lookup_op_def hm_key_of_op_unfold apply (simp cong: if_cong add: Let_def) apply (simp add: unused_bind_conv h.val_of_op_def refine_pw_simps) (* Prove refinement *) apply refine_rcg unfolding hmr_rel_def in_br_conv apply (unfold heapmap_invar_def;simp) apply (auto simp: in_br_conv) done text \We demonstrate two different approaches for proving correctness here. The first approach uses the relation to plain heaps only to establish the invariant. The second approach also uses the relation to heaps to establish correctness of the result. The first approach seems to be more robust against badly set up simpsets, which may be the case in early stages of development. Assuming a working simpset, the second approach may be less work, and the proof may look more elegant. \ text_raw \\paragraph{First approach}\ text \Transfer heapmin-property to heapmap-domain\ lemma heapmap_min_prop: assumes INV: "heapmap_invar hm" assumes V': "heapmap_\ hm k = Some v'" assumes NE: "hm_valid hm (Suc 0)" shows "prio (the (heapmap_\ hm (hm_key_of hm (Suc 0)))) \ prio v'" proof - \ \Transform into the domain of heaps\ obtain pq m where [simp]: "hm=(pq,m)" by (cases hm) from NE have [simp]: "pq\[]" by (auto simp: hm_valid_def hm_length_def) have CNV_LHS: "prio (the (heapmap_\ hm (hm_key_of hm (Suc 0)))) = h.prio_of (hmr_\ hm) (Suc 0)" by (auto simp: heapmap_\_def hm_key_of_def hmr_\_def h.val_of_def) from INV have INV': "h.heap_invar (hmr_\ hm)" unfolding heapmap_invar_def by auto from V' INV obtain i where IDX: "h.valid (hmr_\ hm) i" and CNV_RHS: "prio v' = h.prio_of (hmr_\ hm) i" apply (clarsimp simp: heapmap_\_def heapmap_invar_def hmr_invar_def hmr_\_def h.valid_def h.val_of_def) by (metis (no_types, opaque_lifting) Suc_leI comp_apply diff_Suc_Suc diff_zero domI index_less_size_conv neq0_conv nth_index nth_map old.nat.distinct(2) option.sel) from h.heap_min_prop[OF INV' IDX] show ?thesis unfolding CNV_LHS CNV_RHS . qed text \With the above lemma, the correctness proof is straightforward\ lemma hm_pop_min_\_correct: "hm_pop_min_op hm \\<^sub>n SPEC (\((k,v),hm'). heapmap_\ hm k = Some v \ heapmap_\ hm' = (heapmap_\ hm)(k:=None) \ (\k' v'. heapmap_\ hm k' = Some v' \ prio v \ prio v'))" unfolding hm_pop_min_op_def hm_key_of_op_unfold hm_the_lookup_op_def apply (refine_vcg hm_exch_op_\_correct[THEN leof_trans] hm_butlast_op_\_correct[THEN leof_trans] hm_sink_op_\_correct[THEN leof_trans] ) apply (auto simp: heapmap_min_prop) done lemma heapmap_nres_rel_prodI: assumes "hmx \ \(UNIV \\<^sub>r hmr_rel) h'x" assumes "h'x \ SPEC (\(_,h'). h.heap_invar h')" assumes "hmx \\<^sub>n SPEC (\(r,hm'). RETURN (r,heapmap_\ hm') \ \(R\\<^sub>rId) hx)" shows "hmx \ \(R\\<^sub>rheapmap_rel) hx" using assms unfolding heapmap_rel_def hmr_rel_def br_def heapmap_invar_def apply (auto simp: pw_le_iff pw_leof_iff refine_pw_simps; blast) done lemma hm_pop_min_op_aref: "(hm_pop_min_op, PR_CONST (mop_pm_pop_min prio)) \ heapmap_rel \ \(Id\\<^sub>rId)\\<^sub>rheapmap_rel\nres_rel" unfolding PR_CONST_def apply (intro fun_relI nres_relI) apply (frule heapmap_hmr_relI) unfolding mop_pm_pop_min_alt apply (intro ASSERT_refine_right) apply (rule heapmap_nres_rel_prodI) apply (rule hm_pop_min_op_refine[param_fo, THEN nres_relD]; assumption) unfolding heapmap_rel_def hmr_rel_def in_br_conv apply (refine_vcg; simp) apply (refine_vcg hm_pop_min_\_correct[THEN leof_trans]; simp split: prod.splits) done text_raw \\paragraph{Second approach}\ (* Alternative approach: Also use knowledge about result in multiset domain. Obtaining property seems infeasible at first attempt! *) definition "hm_kv_of_op hm i \ do { ASSERT (hm_valid hm i \ hmr_invar hm); k \ hm_key_of_op hm i; v \ hm_the_lookup_op hm k; RETURN (k, v) }" definition "kvi_rel hm i \ {((k,v),v) | k v. hm_key_of hm i = k}" lemma hm_kv_op_refine[refine]: assumes "(hm,h)\hmr_rel" shows "hm_kv_of_op hm i \ \(kvi_rel hm i) (h.val_of_op h i)" unfolding hm_kv_of_op_def h.val_of_op_def kvi_rel_def hm_key_of_op_unfold hm_the_lookup_op_def apply simp apply refine_vcg using assms by (auto simp: hm_valid_def hm_length_def hmr_rel_defs heapmap_\_def hm_key_of_def split: prod.splits) definition hm_pop_min_op' :: "('k,'v) ahm \ (('k\'v) \ ('k,'v) ahm) nres" where "hm_pop_min_op' hm \ do { ASSERT (heapmap_invar hm); ASSERT (hm_valid hm 1); kv \ hm_kv_of_op hm 1; let l = hm_length hm; hm \ hm_exch_op hm 1 l; hm \ hm_butlast_op hm; if (l\1) then do { hm \ hm_sink_op hm 1; RETURN (kv,hm) } else RETURN (kv,hm) }" lemma hm_pop_min_op_refine': "\ (hm,h)\hmr_rel \ \ hm_pop_min_op' hm \ \(kvi_rel hm 1 \\<^sub>r hmr_rel) (h.pop_min_op h)" unfolding hm_pop_min_op'_def h.pop_min_op_def (* Project away stuff of second component *) unfolding ignore_snd_refine_conv (* Prove refinement *) apply refine_rcg unfolding hmr_rel_def heapmap_rel_def apply (unfold heapmap_invar_def; simp add: in_br_conv) apply (simp_all add: in_br_conv) done lemma heapmap_nres_rel_prodI': assumes "hmx \ \(S \\<^sub>r hmr_rel) h'x" assumes "h'x \ SPEC \" assumes "\h' r. \ (r,h') \ h.heap_invar h'" assumes "hmx \\<^sub>n SPEC (\(r,hm'). (\r'. (r,r')\S \ \ (r',hmr_\ hm')) \ hmr_invar hm' \ RETURN (r,heapmap_\ hm') \ \(R\\<^sub>rId) hx)" shows "hmx \ \(R\\<^sub>rheapmap_rel) hx" using assms unfolding heapmap_rel_def hmr_rel_def heapmap_invar_def apply (auto simp: pw_le_iff pw_leof_iff refine_pw_simps in_br_conv ) by meson lemma ex_in_kvi_rel_conv: "(\r'. (r,r')\kvi_rel hm i \ \ r') \ (fst r = hm_key_of hm i \ \ (snd r))" unfolding kvi_rel_def apply (cases r) apply auto done lemma hm_pop_min_aref': "(hm_pop_min_op', mop_pm_pop_min prio) \ heapmap_rel \ \(Id\\<^sub>rId) \\<^sub>r heapmap_rel\nres_rel" apply (intro fun_relI nres_relI) apply (frule heapmap_hmr_relI) unfolding mop_pm_pop_min_alt apply (intro ASSERT_refine_right) apply (rule heapmap_nres_rel_prodI') apply (erule hm_pop_min_op_refine') apply (unfold heapmap_rel_def hmr_rel_def in_br_conv) [] apply (rule h.pop_min_op_correct) apply simp apply simp apply simp apply (clarsimp simp: ex_in_kvi_rel_conv split: prod.splits) unfolding hm_pop_min_op'_def hm_kv_of_op_def hm_key_of_op_unfold hm_the_lookup_op_def apply (refine_vcg hm_exch_op_\_correct[THEN leof_trans] hm_butlast_op_\_correct[THEN leof_trans] hm_sink_op_\_correct[THEN leof_trans] ) unfolding heapmap_rel_def hmr_rel_def in_br_conv apply (auto intro: ranI) done subsubsection \Remove\ definition "hm_remove_op k hm \ do { ASSERT (heapmap_invar hm); ASSERT (k \ dom (heapmap_\ hm)); i \ hm_index_op hm k; let l = hm_length hm; hm \ hm_exch_op hm i l; hm \ hm_butlast_op hm; if i \ l then hm_repair_op hm i else RETURN hm }" definition (in heapstruct) "remove_op i h \ do { ASSERT (heap_invar h); ASSERT (valid h i); let l = length h; h \ exch_op h i l; h \ butlast_op h; if i \ l then repair_op h i else RETURN h }" lemma (in -) swap_empty_iff[iff]: "swap l i j = [] \ l=[]" by (auto simp: swap_def) lemma (in heapstruct) butlast_exch_last: "butlast (exch h i (length h)) = update (butlast h) i (last h)" unfolding exch_def update_def apply (cases h rule: rev_cases) apply (auto simp: swap_def butlast_list_update) done lemma (in heapstruct) remove_op_invar: "\ heap_invar h; valid h i \ \ remove_op i h \ SPEC heap_invar" unfolding remove_op_def apply refine_vcg apply (auto simp: valid_def) [] apply (auto simp: valid_def exch_def) [] apply (simp add: butlast_exch_last) apply refine_vcg apply auto [] apply auto [] apply (auto simp: valid_def) [] apply auto [] apply auto [] done lemma hm_remove_op_refine[refine]: "\ (hm,m)\heapmap_rel; (hm,h)\hmr_rel; heapmap_\ hm k \ None\ \ hm_remove_op k hm \ \hmr_rel (h.remove_op (hm_index hm k) h)" unfolding hm_remove_op_def h.remove_op_def heapmap_rel_def (*apply (rewrite at "Let (hm_index hm k) _" Let_def)*) apply (refine_rcg index_op_inline_refine) unfolding hmr_rel_def apply (auto simp: in_br_conv) done lemma hm_remove_op_\_correct: "hm_remove_op k hm \\<^sub>n SPEC (\hm'. heapmap_\ hm' = (heapmap_\ hm)(k:=None))" unfolding hm_remove_op_def apply (refine_vcg hm_exch_op_\_correct[THEN leof_trans] hm_butlast_op_\_correct[THEN leof_trans] hm_repair_op_\_correct[THEN leof_trans] hm_index_op_inline_leof ) apply (auto; fail) apply clarsimp apply (rewrite at "hm_index _ k = hm_length _" in asm eq_commute) apply (auto; fail) done lemma hm_remove_op_aref: "(hm_remove_op,mop_map_delete_ex) \ Id \ heapmap_rel \ \heapmap_rel\nres_rel" apply (intro fun_relI nres_relI) unfolding mop_map_delete_ex_alt apply (rule ASSERT_refine_right) apply (frule heapmap_hmr_relI) apply (rule heapmap_nres_relI') apply (rule hm_remove_op_refine; assumption?) apply (unfold heapmap_rel_def in_br_conv; auto) unfolding heapmap_rel_def hmr_rel_def in_br_conv apply (refine_vcg h.remove_op_invar; clarsimp; fail) apply (refine_vcg hm_remove_op_\_correct[THEN leof_trans]; simp; fail) done subsubsection \Peek-Min\ definition hm_peek_min_op :: "('k,'v) ahm \ ('k\'v) nres" where "hm_peek_min_op hm \ hm_kv_of_op hm 1" lemma hm_peek_min_op_aref: "(hm_peek_min_op, PR_CONST (mop_pm_peek_min prio)) \ heapmap_rel \ \Id\\<^sub>rId\nres_rel" unfolding PR_CONST_def apply (intro fun_relI nres_relI) proof - fix hm and m :: "'k \ 'v" assume A: "(hm,m)\heapmap_rel" from A have [simp]: "h.heap_invar (hmr_\ hm)" "hmr_invar hm" "m=heapmap_\ hm" unfolding heapmap_rel_def in_br_conv heapmap_invar_def by simp_all have "hm_peek_min_op hm \ \ (kvi_rel hm 1) (h.peek_min_op (hmr_\ hm))" unfolding hm_peek_min_op_def h.peek_min_op_def apply (refine_rcg hm_kv_op_refine) using A apply (simp add: heapmap_hmr_relI) done also have "\hmr_\ hm \ []\ \ (h.peek_min_op (hmr_\ hm)) \ SPEC (\v. v\ran (heapmap_\ hm) \ (\v'\ran (heapmap_\ hm). prio v \ prio v'))" apply refine_vcg by simp_all finally show "hm_peek_min_op hm \ \ (Id \\<^sub>r Id) (mop_pm_peek_min prio m)" unfolding mop_pm_peek_min_alt apply (simp add: pw_le_iff refine_pw_simps hm_peek_min_op_def hm_kv_of_op_def hm_key_of_op_unfold hm_the_lookup_op_def) apply (fastforce simp: kvi_rel_def ran_def) done qed end end diff --git a/thys/Transitive-Closure/RBT_Map_Set_Extension.thy b/thys/Transitive-Closure/RBT_Map_Set_Extension.thy --- a/thys/Transitive-Closure/RBT_Map_Set_Extension.thy +++ b/thys/Transitive-Closure/RBT_Map_Set_Extension.thy @@ -1,324 +1,324 @@ (* Title: Executable Transitive Closures of Finite Relations Author: Christian Sternagel René Thiemann Maintainer: Christian Sternagel and René Thiemann License: LGPL *) section \Accessing Values via Keys\ theory RBT_Map_Set_Extension imports Collections.RBTMapImpl Collections.RBTSetImpl Matrix.Utility begin text \ We provide two extensions of the red black tree implementation. The first extension provides two convenience methods on sets which are represented by red black trees: a check on subsets and the big union operator. The second extension is to provide two operations @{term elem_list_to_rm} and @{term rm_set_lookup} which can be used to index a set of values via keys. More precisely, given a list of values of type @{typ "'v"} and a key function of type @{typ "'v => 'k"}, @{term elem_list_to_rm} will generate a map of type @{typ "'k => 'v set"}. Then with @{term rs_set_lookup} we can efficiently access all values which match a given key. \ subsection \Subset and Union\ text \ For the subset operation @{term "r \ s"} we provide two implementations. The first one (@{term rs_subset}) traverses over @{term r} and then performs membership tests \\ s\. Its complexity is ${\cal O}(|r| \cdot log (|s|))$. The second one (@{term rs_subset_list}) generates sorted lists for both @{term r} and @{term s} and then linearly checks the subset condition. Its complexity is ${\cal O}(|r| + |s|)$. \ text \ As union operator we use the standard fold function. Note that the order of the union is important so that new sets are added to the big union. \ (* perhaps there is a smarter way to use two iterates at the same time, however, when writing this theory this feature was not detected in the RBTSetImpl theory *) definition rs_subset :: "('a :: linorder) rs \ 'a rs \ 'a option" where "rs_subset as bs = rs.iteratei as (\ maybe. case maybe of None \ True | Some _ \ False) (\ a _. if rs.memb a bs then None else Some a) None" lemma rs_subset [simp]: "rs_subset as bs = None \ rs.\ as \ rs.\ bs" proof - let ?abort = "\ maybe. case maybe of None \ True | Some _ \ False" let ?I = "\ aas maybe. maybe = None \ (\ a. a \ rs.\ as - aas \ a \ rs.\ bs)" let ?it = "rs_subset as bs" have "?I {} ?it \ (\ it \ rs.\ as. it \ {} \ \ ?abort ?it \ ?I it ?it)" unfolding rs_subset_def by (rule rs.iteratei_rule_P [where I="?I"]) (auto simp: rs.correct) then show ?thesis by auto qed definition rs_subset_list :: "('a :: linorder) rs \ 'a rs \ 'a option" where "rs_subset_list as bs = sorted_list_subset (rs.to_sorted_list as) (rs.to_sorted_list bs)" lemma rs_subset_list [simp]: "rs_subset_list as bs = None \ rs.\ as \ rs.\ bs" unfolding rs_subset_list_def sorted_list_subset[OF rs.to_sorted_list_correct(3)[OF rs.invar, of as] rs.to_sorted_list_correct(3)[OF rs.invar, of bs]] by (simp add: rs.to_sorted_list_correct) definition rs_Union :: "('q :: linorder) rs list \ 'q rs" where "rs_Union = foldl rs.union (rs.empty ())" lemma rs_Union [simp]: "rs.\ (rs_Union qs) = \ (rs.\ ` set qs)" proof - { fix start have "rs.\ (foldl rs.union start qs) = rs.\ start \ \ (rs.\ ` set qs)" by (induct qs arbitrary: start, auto simp: rs.correct) } from this[of "rs.empty ()"] show ?thesis unfolding rs_Union_def by (auto simp: rs.correct) qed subsection \Grouping Values via Keys\ text \ The functions to produce the index (@{term elem_list_to_rm}) and the lookup function (@{term rm_set_lookup}) are straight-forward, however it requires some tedious reasoning that they perform as they should. \ fun elem_list_to_rm :: "('d \ 'k :: linorder) \ 'd list \ ('k, 'd list) rm" where "elem_list_to_rm key [] = rm.empty ()" | "elem_list_to_rm key (d # ds) = (let rm = elem_list_to_rm key ds; k = key d in (case rm.\ rm k of None \ rm.update_dj k [d] rm | Some data \ rm.update k (d # data) rm))" definition "rm_set_lookup rm = (\ a. (case rm.\ rm a of None \ [] | Some rules \ rules))" lemma rm_to_list_empty [simp]: "rm.to_list (rm.empty ()) = []" proof - have "map_of (rm.to_list (rm.empty ())) = Map.empty" by (simp add: rm.correct) moreover have map_of_empty_iff: "\l. map_of l = Map.empty \ l = []" by (case_tac l) auto ultimately show ?thesis by metis qed locale rm_set = fixes rm :: "('k :: linorder, 'd list) rm" and key :: "'d \ 'k" and data :: "'d set" assumes rm_set_lookup: "\ k. set (rm_set_lookup rm k) = {d \ data. key d = k}" begin lemma data_lookup: "data = \ {set (rm_set_lookup rm k) | k. True}" (is "_ = ?R") proof - { fix d assume d: "d \ data" then have d: "d \ {d' \ data. key d' = key d}" by auto have "d \ ?R" by (rule UnionI[OF _ d], rule CollectI, rule exI[of _ "key d"], unfold rm_set_lookup[of "key d"], simp) } moreover { fix d assume "d \ ?R" from this[unfolded rm_set_lookup] have "d \ data" by auto } ultimately show ?thesis by blast qed lemma finite_data: "finite data" unfolding data_lookup proof show "finite {set (rm_set_lookup rm k) | k. True}" (is "finite ?L") proof - let ?rmset = "rm.\ rm" let ?M = "?rmset ` Map.dom ?rmset" let ?N = "((\ e. set (case e of None \ [] | Some ds \ ds)) ` ?M)" let ?K = "?N \ {{}}" from rm.finite[of rm] have fin: "finite ?K" by auto show ?thesis proof (rule finite_subset[OF _ fin], rule) fix ds assume "ds \ ?L" from this[unfolded rm_set_lookup_def] obtain fn where ds: "ds = set (case rm.\ rm fn of None \ [] | Some ds \ ds)" by auto show "ds \ ?K" proof (cases "rm.\ rm fn") case None then show ?thesis unfolding ds by auto next case (Some rules) from Some have fn: "fn \ Map.dom ?rmset" by auto have "ds \ ?N" unfolding ds by (rule, rule refl, rule, rule refl, rule fn) then show ?thesis by auto qed qed qed qed (force simp: rm_set_lookup_def) end interpretation elem_list_to_rm: rm_set "elem_list_to_rm key ds" key "set ds" proof fix k show "set (rm_set_lookup (elem_list_to_rm key ds) k) = {d \ set ds. key d = k}" proof (induct ds arbitrary: k) case Nil then show ?case unfolding rm_set_lookup_def by (simp add: rm.correct) next case (Cons d ds k) let ?el = "elem_list_to_rm key" let ?l = "\k ds. set (rm_set_lookup (?el ds) k)" let ?r = "\k ds. {d \ set ds. key d = k}" from Cons have ind: "\ k. ?l k ds = ?r k ds" by auto show "?l k (d # ds) = ?r k (d # ds)" proof (cases "rm.\ (?el ds) (key d)") case None from None ind[of "key d"] have r: "{da \ set ds. key da = key d} = {}" unfolding rm_set_lookup_def by auto from None have el: "?el (d # ds) = rm.update_dj (key d) [d] (?el ds)" by simp from None have ndom: "key d \ Map.dom (rm.\ (?el ds))" by auto have r: "?r k (d # ds) = ?r k ds \ {da. key da \ key d} \ {da . key da = k \ da = d}" (is "_ = ?r1 \ ?r2") using r by auto from ndom have l: "?l k (d # ds) = - set (case (rm.\ (elem_list_to_rm key ds)(key d \ [d])) k of None \ [] + set (case ((rm.\ (elem_list_to_rm key ds))(key d \ [d])) k of None \ [] | Some rules \ rules)" (is "_ = ?l") unfolding el rm_set_lookup_def by (simp add: rm.correct) { fix da assume "da \ ?r1 \ ?r2" then have "da \ ?l" proof assume "da \ ?r2" then have da: "da = d" and k: "key d = k" by auto show ?thesis unfolding da k by auto next assume "da \ ?r1" from this[unfolded ind[symmetric] rm_set_lookup_def] obtain das where rm: "rm.\ (?el ds) k = Some das" and da: "da \ set das" and k: "key da \ key d" by (cases "rm.\ (?el ds) k", auto) from ind[of k, unfolded rm_set_lookup_def] rm da k have k: "key d \ k" by auto - have rm: "(rm.\ (elem_list_to_rm key ds)(key d \ [d])) k = Some das" + have rm: "((rm.\ (elem_list_to_rm key ds))(key d \ [d])) k = Some das" unfolding rm[symmetric] using k by auto show ?thesis unfolding rm using da by auto qed } moreover { fix da assume l: "da \ ?l" let ?rm = "((rm.\ (elem_list_to_rm key ds))(key d \ [d])) k" from l obtain das where rm: "?rm = Some das" and da: "da \ set das" by (cases ?rm, auto) have "da \ ?r1 \ ?r2" proof (cases "k = key d") case True with rm da have da: "da = d" by auto then show ?thesis using True by auto next case False with rm have "rm.\ (?el ds) k = Some das" by auto from ind[of k, unfolded rm_set_lookup_def this] da False show ?thesis by auto qed } ultimately have "?l = ?r1 \ ?r2" by blast then show ?thesis unfolding l r . next case (Some das) from Some ind[of "key d"] have das: "{da \ set ds. key da = key d} = set das" unfolding rm_set_lookup_def by auto from Some have el: "?el (d # ds) = rm.update (key d) (d # das) (?el ds)" by simp from Some have dom: "key d \ Map.dom (rm.\ (?el ds))" by auto from dom have l: "?l k (d # ds) = - set (case (rm.\ (elem_list_to_rm key ds)(key d \ (d # das))) k of None \ [] + set (case ((rm.\ (elem_list_to_rm key ds))(key d \ (d # das))) k of None \ [] | Some rules \ rules)" (is "_ = ?l") unfolding el rm_set_lookup_def by (simp add: rm.correct) have r: "?r k (d # ds) = ?r k ds \ {da. key da = k \ da = d}" (is "_ = ?r1 \ ?r2") by auto { fix da assume "da \ ?r1 \ ?r2" then have "da \ ?l" proof assume "da \ ?r2" then have da: "da = d" and k: "key d = k" by auto show ?thesis unfolding da k by auto next assume "da \ ?r1" from this[unfolded ind[symmetric] rm_set_lookup_def] obtain das' where rm: "rm.\ (?el ds) k = Some das'" and da: "da \ set das'" by (cases "rm.\ (?el ds) k", auto) from ind[of k, unfolded rm_set_lookup_def rm] have das': "set das' = {d \ set ds. key d = k}" by auto show ?thesis proof (cases "k = key d") case True show ?thesis using das' das da unfolding True by simp next case False then show ?thesis using das' da rm by auto qed qed } moreover { fix da assume l: "da \ ?l" let ?rm = "((rm.\ (elem_list_to_rm key ds))(key d \ d # das)) k" from l obtain das' where rm: "?rm = Some das'" and da: "da \ set das'" by (cases ?rm, auto) have "da \ ?r1 \ ?r2" proof (cases "k = key d") case True with rm da das have da: "da \ set (d # das)" by auto then have "da = d \ da \ set das" by auto then have k: "key da = k" proof assume "da = d" then show ?thesis using True by simp next assume "da \ set das" with das True show ?thesis by auto qed from da k show ?thesis using das by auto next case False with rm have "rm.\ (?el ds) k = Some das'" by auto from ind[of k, unfolded rm_set_lookup_def this] da False show ?thesis by auto qed } ultimately have "?l = ?r1 \ ?r2" by blast then show ?thesis unfolding l r . qed qed qed end