diff --git a/thys/CAVA_LTL_Modelchecker/BoolProgs/BoolProgs_LTL_Conv.thy b/thys/CAVA_LTL_Modelchecker/BoolProgs/BoolProgs_LTL_Conv.thy --- a/thys/CAVA_LTL_Modelchecker/BoolProgs/BoolProgs_LTL_Conv.thy +++ b/thys/CAVA_LTL_Modelchecker/BoolProgs/BoolProgs_LTL_Conv.thy @@ -1,59 +1,49 @@ theory BoolProgs_LTL_Conv imports "BoolProgs_Extras" "~~/src/HOL/Library/Mapping" "../../LTL_to_GBA/LTL" begin -context begin interpretation LTL_Syntax . - fun b2l :: "bexp \ nat ltlc" where "b2l TT = true\<^sub>c" | "b2l FF = false\<^sub>c" | "b2l (bexp.V v) = prop\<^sub>c(v)" | "b2l (bexp.Not e) = not\<^sub>c (b2l e)" | "b2l (And e1 e2) = b2l e1 and\<^sub>c b2l e2" | "b2l (Or e1 e2) = b2l e1 or\<^sub>c b2l e2" -end - datatype propc = CProp String.literal | FProp "String.literal * integer" -context begin interpretation LTL_Syntax . - fun ltl_conv :: "const_map \ fun_map \ propc ltlc \ (nat ltlc + String.literal)" where "ltl_conv _ _ LTLcTrue = Inl LTLcTrue" | "ltl_conv _ _ LTLcFalse = Inl LTLcFalse" | "ltl_conv C _ (LTLcProp (CProp s)) = (case Mapping.lookup C s of Some c \ Inl (b2l c) | None \ Inr (STR ''Unknown constant: '' @@ s))" | "ltl_conv _ M (LTLcProp (FProp (s, arg))) = (case Mapping.lookup M s of Some f \ (Inl \ b2l \ f \ nat_of_integer) arg | None \ Inr (STR ''Unknown function: '' @@ s))" | "ltl_conv C M (LTLcNeg x) = (case ltl_conv C M x of Inl l \ Inl (LTLcNeg l) | Inr s \ Inr s)" | "ltl_conv C M (LTLcNext x) = (case ltl_conv C M x of Inl l \ Inl (LTLcNext l) | Inr s \ Inr s)" | "ltl_conv C M (LTLcFinal x) = (case ltl_conv C M x of Inl l \ Inl (LTLcFinal l) | Inr s \ Inr s)" | "ltl_conv C M (LTLcGlobal x) = (case ltl_conv C M x of Inl l \ Inl (LTLcGlobal l) | Inr s \ Inr s)" | "ltl_conv C M (LTLcAnd x y) = (case ltl_conv C M x of Inr s \ Inr s | Inl l \ (case ltl_conv C M y of Inl r \ Inl (LTLcAnd l r) | Inr s \ Inr s))" | "ltl_conv C M (LTLcOr x y) = (case ltl_conv C M x of Inr s \ Inr s | Inl l \ (case ltl_conv C M y of Inl r \ Inl (LTLcOr l r) | Inr s \ Inr s))" | "ltl_conv C M (LTLcImplies x y) = (case ltl_conv C M x of Inr s \ Inr s | Inl l \ (case ltl_conv C M y of Inl r \ Inl (LTLcImplies l r) | Inr s \ Inr s))" -| "ltl_conv C M (LTLcIff x y) = (case ltl_conv C M x of - Inr s \ Inr s - | Inl l \ (case ltl_conv C M y of Inl r \ Inl (LTLcIff l r) | Inr s \ Inr s))" | "ltl_conv C M (LTLcUntil x y) = (case ltl_conv C M x of Inr s \ Inr s | Inl l \ (case ltl_conv C M y of Inl r \ Inl (LTLcUntil l r) | Inr s \ Inr s))" | "ltl_conv C M (LTLcRelease x y) = (case ltl_conv C M x of Inr s \ Inr s | Inl l \ (case ltl_conv C M y of Inl r \ Inl (LTLcRelease l r) | Inr s \ Inr s))" end -end diff --git a/thys/CAVA_LTL_Modelchecker/CAVA_Impl.thy b/thys/CAVA_LTL_Modelchecker/CAVA_Impl.thy --- a/thys/CAVA_LTL_Modelchecker/CAVA_Impl.thy +++ b/thys/CAVA_LTL_Modelchecker/CAVA_Impl.thy @@ -1,875 +1,874 @@ section {* Actual Implementation of the CAVA Model Checker *} theory CAVA_Impl imports "CAVA_Abstract" "../CAVA_Automata/Automata_Impl" "../LTL_to_GBA/LTL_to_GBA_impl" (* LTL to BA *) "../Gabow_SCC/Gabow_GBG_Code" (* Gabow's Algorithm *) "Nested_DFS/NDFS_SI" (* Nested-DFS, standard-invars formalization *) "BoolProgs/BoolProgs" (* Boolean Programs *) "BoolProgs/Programs/BoolProgs_Programs" (* the actual programs *) "BoolProgs/BoolProgs_LTL_Conv" (* LTL parsing setup *) "../Promela/PromelaLTL" (* Promela *) "../Promela/PromelaLTLConv" (* LTL parsing setup *) "~~/src/HOL/Library/AList_Mapping" "../CAVA_Automata/CAVA_Base/CAVA_Code_Target" begin (*<*) subsection {* Exporting Graphs *} (* TODO: frv_export is going to be replaced by more explicit implementation of graphs. For the moment, we just keep it here. *) definition "frv_edge_set G \ g_E G \ (g_V G \ UNIV)" definition "frv_edge_set_aimpl G \ FOREACHi (\it r. r = g_E G \ ((g_V G - it) \ UNIV)) (g_V G) (\u r. do { let E = (\v. (u,v))`(succ_of_E (g_E G) u); ASSERT (E \ r = {}); RETURN (E \ r) }) {}" lemma frv_edge_set_aimpl_correct: "finite (g_V G) \ frv_edge_set_aimpl G \ SPEC (\r. r = frv_edge_set G)" unfolding frv_edge_set_aimpl_def frv_edge_set_def apply (refine_rcg refine_vcg) apply auto [] apply auto [] apply (auto simp: succ_of_E_def) [] apply auto [] done schematic_goal frv_edge_set_impl_aux: assumes [autoref_rules]: "(eq,op =)\R \ R \ bool_rel" assumes [relator_props]: "single_valued R" shows "(?c, frv_edge_set_aimpl) \ \Re,R\frgv_impl_rel_ext \ \\R \\<^sub>r R\list_set_rel\nres_rel" unfolding frv_edge_set_aimpl_def[abs_def] apply (autoref (trace,keep_goal)) done concrete_definition frv_edge_set_impl uses frv_edge_set_impl_aux lemmas [autoref_rules] = frv_edge_set_impl.refine[OF GEN_OP_D PREFER_sv_D] schematic_goal frv_edge_set_code_aux: "RETURN ?c \ frv_edge_set_impl eq G" unfolding frv_edge_set_impl_def by (refine_transfer (post)) concrete_definition frv_edge_set_code for eq G uses frv_edge_set_code_aux lemmas [refine_transfer] = frv_edge_set_code.refine lemma frv_edge_set_autoref[autoref_rules]: assumes EQ[unfolded autoref_tag_defs]: "GEN_OP eq op = (R \ R \ bool_rel)" assumes SV[unfolded autoref_tag_defs]: "PREFER single_valued R" shows "(frv_edge_set_code eq,frv_edge_set) \ \Re,R\frgv_impl_rel_ext \ \R \\<^sub>r R\list_set_rel" proof (intro fun_relI) fix Gi G assume Gr: "(Gi, G) \ \Re, R\frgv_impl_rel_ext" hence [simp]: "finite (g_V G)" unfolding frgv_impl_rel_ext_def g_impl_rel_ext_def gen_g_impl_rel_ext_def apply (simp add: list_set_rel_def br_def) apply fastforce done note frv_edge_set_code.refine also note frv_edge_set_impl.refine[OF EQ SV, THEN fun_relD, OF Gr, THEN nres_relD] also note frv_edge_set_aimpl_correct finally show "(frv_edge_set_code eq Gi, frv_edge_set G) \ \R \\<^sub>r R\list_set_rel" by (rule RETURN_ref_SPECD) simp_all qed definition "frv_export G \ do { nodes \ SPEC (\l. set l = g_V G \ distinct l); V0 \ SPEC (\l. set l = g_V0 G \ distinct l); E \ SPEC (\l. set l = frv_edge_set G \ distinct l); RETURN (nodes,V0,E) }" schematic_goal frv_export_impl_aux: fixes R :: "('vi \ 'v) set" notes [autoref_tyrel] = TYRELI[where R = "\R \\<^sub>r R\list_set_rel"] assumes EQ[autoref_rules]: "(eq,op =)\R \ R \ bool_rel" assumes SVR[relator_props]: "single_valued R" shows "(?c, frv_export) \ \Re,R\frgv_impl_rel_ext \ \\R\list_rel \\<^sub>r \R\list_rel \\<^sub>r \R \\<^sub>r R\list_rel\nres_rel" unfolding frv_export_def[abs_def] using [[autoref_trace_failed_id]] apply (autoref (keep_goal, trace)) done concrete_definition frv_export_impl for eq uses frv_export_impl_aux lemmas [autoref_rules] = frv_export_impl.refine[OF GEN_OP_D PREFER_sv_D] schematic_goal frv_export_code_aux: "RETURN ?c \ frv_export_impl eq G" unfolding frv_export_impl_def apply (refine_transfer (post)) done concrete_definition frv_export_code for eq G uses frv_export_code_aux lemmas [refine_transfer] = frv_export_code.refine (*>*) subsection {* Setup *} subsubsection {* LTL to GBA conversion *} text {* In the following, we set up the algorithms for LTL to GBA conversion. *} definition is_ltl_to_gba_algo :: "('a ltlc \ (nat, 'a \ bool, unit) igbav_impl_scheme) \ bool" -- "Predicate that must be satisfied by an LTL to GBA conversion" where "is_ltl_to_gba_algo algo \ (algo, ltl_to_gba_spec) \ ltl_rel \ \igbav_impl_rel_ext unit_rel nat_rel (\Id\fun_set_rel)\plain_nres_rel" definition gerth_ltl_to_gba -- "Conversion based on Gerth's Algorithm" where "gerth_ltl_to_gba \ - \ create_name_igba (ltln_rewrite (ltl_to_ltln (ltl_pushneg (ltlc_to_ltl \))))" + \ create_name_igba (rewrite_iter_fast (ltlc_to_ltln \))" lemma gerth_ltl_to_gba_refine: "gerth_ltl_to_gba \ \ \Id (ltl_to_gba_spec \)" apply simp unfolding ltl_to_gba_spec_def gerth_ltl_to_gba_def apply (rule order_trans[OF create_name_igba_correct]) apply (rule SPEC_rule) proof (safe del: equalityI) fix G :: "(nat, 'a set) igba_rec" assume "igba G" interpret igba G by fact assume 1: "finite (g_V G)" - assume 2: "\ \. accept \ \ \ \\<^sub>n ltln_rewrite (ltl_to_ltln (ltl_pushneg (ltlc_to_ltl \)))" + assume 2: "\ \. accept \ \ \ \\<^sub>n rewrite_iter_fast (ltlc_to_ltln \)" show "lang = ltlc_language \" unfolding lang_def ltlc_language_def using 2 - by (auto simp: ltln_rewrite__equiv ltlc_to_ltl_equiv) + using rewrite_iter_fast_sound ltlc_to_ltln_semantics by auto show "finite ((g_E G)\<^sup>* `` g_V0 G)" using 1 reachable_V by (auto intro: finite_subset) qed -definition "gerth_ltl_to_gba_code \ - \ create_name_igba_code - (ltln_rewrite (ltl_to_ltln (ltl_pushneg (ltlc_to_ltl \))))" +definition + "gerth_ltl_to_gba_code \ \ create_name_igba_code (rewrite_iter_fast (ltlc_to_ltln \))" lemma gerth_ltl_to_gba_code_refine: "is_ltl_to_gba_algo gerth_ltl_to_gba_code" proof - have "\\. RETURN (gerth_ltl_to_gba_code \) \ \(igbav_impl_rel_ext unit_rel nat_rel (\Id\fun_set_rel)) (gerth_ltl_to_gba \)" using assms unfolding gerth_ltl_to_gba_def[abs_def] gerth_ltl_to_gba_code_def[abs_def] apply (rule order_trans[OF create_name_igba_code.refine]) apply (rule create_name_igba_impl.refine[THEN fun_relD, THEN nres_relD]) apply simp apply simp done also note gerth_ltl_to_gba_refine finally show ?thesis unfolding is_ltl_to_gba_algo_def by (blast intro: plain_nres_relI) qed text {* We define a function that chooses between the possible conversion algorithms. (Currently there is only one) *} datatype config_l2b = CFG_L2B_GERTHS definition "ltl_to_gba_code cfg \ case cfg of CFG_L2B_GERTHS \ gerth_ltl_to_gba_code" lemma ltl_to_gba_code_refine: "is_ltl_to_gba_algo (ltl_to_gba_code cfg)" apply (cases cfg) unfolding ltl_to_gba_code_def using gerth_ltl_to_gba_code_refine apply simp done subsubsection {* Counterexample Search *} definition is_find_ce_algo :: "(('a, unit)igbg_impl_scheme \ 'a lasso option option) \ bool" -- "Predicate that must be satisfied by counterexample search algorithm" where "is_find_ce_algo algo \ (algo, find_ce_spec) \ igbg_impl_rel_ext unit_rel Id \ \\\\Id\lasso_run_rel\Relators.option_rel\Relators.option_rel\plain_nres_rel" definition gabow_find_ce_code -- "Emptiness check based on Gabow's SCC Algorithm" where "gabow_find_ce_code \ map_option (Some o lasso_of_prpl) o Gabow_GBG_Code.find_lasso_tr" lemma gabow_find_ce_code_refine: "is_find_ce_algo (gabow_find_ce_code :: ('a::hashable, unit) igbg_impl_scheme \ 'a lasso option option)" proof - have AUX_EQ: "\gbgi::('a, unit) igbg_impl_scheme. RETURN (gabow_find_ce_code gbgi) = (do { l \ RETURN (find_lasso_tr gbgi); RETURN (map_option (Some \ lasso_of_prpl) l) })" by (auto simp: gabow_find_ce_code_def) show ?thesis (* TODO: Clean up this proof! *) unfolding is_find_ce_algo_def apply (intro fun_relI plain_nres_relI) unfolding find_ce_spec_def AUX_EQ apply (refine_rcg) apply (rule order_trans[ OF bind_mono(1)[OF Gabow_GBG_Code.find_lasso_tr_correct order_refl]]) apply assumption apply (rule igb_fr_graphI) apply assumption apply assumption unfolding igb_graph.find_lasso_spec_def apply (simp add: pw_le_iff refine_pw_simps) apply (clarsimp simp: lasso_run_rel_sv option_rel_sv split: option.split_asm) apply (metis igb_graph.is_lasso_prpl_of_lasso igb_graph.accepted_lasso prod.exhaust) apply (simp add: option_rel_def lasso_run_rel_def br_def) by (metis igb_graph.is_lasso_prpl_conv igb_graph.lasso_accepted) qed definition ndfs_find_ce :: "('q::hashable,_) igb_graph_rec_scheme \ 'q lasso option option nres" -- "Emptiness check based on nested DFS" where "ndfs_find_ce G \ do { ASSERT (igb_graph G); let G = igb_graph.degeneralize G; l \ blue_dfs G; case l of None \ RETURN None | Some l \ do { ASSERT (snd l \ []); RETURN (Some (Some (map_lasso fst (lasso_of_prpl l)))) } }" lemma ndfs_find_ce_refine: "(G',G)\Id \ ndfs_find_ce G' \ \(\\\Id\lasso_run_rel\option_rel\option_rel) (find_ce_spec G)" apply simp unfolding find_ce_spec_def proof (refine_rcg SPEC_refine refine_vcg) assume [simp]: "igb_graph G" then interpret igb_graph G . assume fr: "finite ((g_E G)\<^sup>* `` g_V0 G)" have [simp]: "b_graph degeneralize" by (simp add: degen_invar) (*then interpret bg: b_graph degeneralize .*) show "ndfs_find_ce G \ (SPEC (\res. \res'. (res,res')\(\\\Id\lasso_run_rel\Relators.option_rel\Relators.option_rel) \ (case res' of None \ \r. \ is_acc_run r | Some None \ Ex is_acc_run | Some (Some x) \ is_acc_run x)))" unfolding ndfs_find_ce_def find_lasso_spec_def ce_correct_def apply (refine_rcg refine_vcg order_trans[OF blue_dfs_correct]) apply (clarsimp_all) apply (auto intro: degen_finite_reachable fr) [] apply (auto elim!: degen_acc_run_complete[where m="\_. ()"] dest!: degen.accepted_lasso[OF degen_finite_reachable[OF fr]] simp: degen.is_lasso_prpl_of_lasso[symmetric] prpl_of_lasso_def simp del: degen.is_lasso_prpl_of_lasso) [] apply (auto simp: b_graph.is_lasso_prpl_def graph.is_lasso_prpl_pre_def) [] apply (auto split: option.split simp: degen.is_lasso_prpl_conv lasso_run_rel_def br_def dest: degen.lasso_accepted degen_acc_run_sound ) [] done qed schematic_goal ndfs_find_ce_impl_aux: "(?c, ndfs_find_ce) \ igbg_impl_rel_ext Rm Id \ \ \\\unit_rel,Id::('a::hashable\_) set\lasso_rel_ext\option_rel\option_rel \ nres_rel" unfolding ndfs_find_ce_def[abs_def] using [[autoref_trace_failed_id]] apply (autoref (trace, keep_goal)) done concrete_definition ndfs_find_ce_impl uses ndfs_find_ce_impl_aux schematic_goal ndfs_find_ce_code_aux: "RETURN ?c \ ndfs_find_ce_impl G" unfolding ndfs_find_ce_impl_def by (refine_transfer (post)) concrete_definition ndfs_find_ce_code uses ndfs_find_ce_code_aux lemma ndfs_find_ce_code_refine: "is_find_ce_algo ndfs_find_ce_code" unfolding is_find_ce_algo_def proof (intro fun_relI plain_nres_relI) fix gbgi :: "('a,unit) igbg_impl_scheme" and gbg :: "'a igb_graph_rec" assume R: "(gbgi, gbg) \ igbg_impl_rel_ext unit_rel Id" note ndfs_find_ce_code.refine also note ndfs_find_ce_impl.refine[param_fo, THEN nres_relD, OF R] also note ndfs_find_ce_refine[OF IdI] finally show "RETURN (ndfs_find_ce_code gbgi) \ \ (\\\Id\lasso_run_rel\Relators.option_rel\Relators.option_rel) (find_ce_spec gbg)" by (simp add: lasso_rel_ext_id) qed text {* We define a function that chooses between the emptiness check algorithms *} datatype config_ce = CFG_CE_SCC_GABOW | CFG_CE_NDFS definition find_ce_code where "find_ce_code cfg \ case cfg of CFG_CE_SCC_GABOW \ gabow_find_ce_code | CFG_CE_NDFS \ ndfs_find_ce_code" lemma find_ce_code_refine: "is_find_ce_algo (find_ce_code cfg)" apply (cases cfg) unfolding find_ce_code_def using gabow_find_ce_code_refine ndfs_find_ce_code_refine apply (auto split: config_ce.split) done subsection {* System-Agnostic Model-Checker *} text {* In this section, we implement the part of the model-checker that does not depend on the language used to describe the system to be checked. *} subsubsection {* Default Implementation of Lazy Intersection *} locale cava_inter_impl_loc = igba_sys_prod_precond G S for S :: "('s, 'l set) sa_rec" and G :: "('q,'l set) igba_rec" + fixes Gi Si Rq Rs Rl eqq assumes [autoref_rules]: "(eqq,op =) \ Rq \ Rq \ bool_rel" assumes [autoref_rules]: "(Gi,G) \ igbav_impl_rel_ext unit_rel Rq Rl" assumes [autoref_rules]: "(Si,S) \ sa_impl_rel_ext unit_rel Rs Rl" begin lemma cava_inter_impl_loc_this: "cava_inter_impl_loc S G Gi Si Rq Rs Rl eqq" by unfold_locales (*<*) text {* TODO/FIXME: Some black-magic is going on here: The performance of the mc seems to depend on the ordering of states, so we do some adjustments of the ordering here. *} lemma prod_impl_aux_alt_cava_reorder: "prod = (\ g_V = Collect (\(q,s). q \ igba.V \ s \ sa.V), g_E = E_of_succ (\(q,s). if igba.L q (sa.L s) then (\(s,q). (q,s))`(LIST_SET_REV_TAG (succ_of_E (sa.E) s) \ (succ_of_E (igba.E) q)) else {} ), g_V0 = igba.V0 \ sa.V0, igbg_num_acc = igba.num_acc, igbg_acc = \(q,s). if s\sa.V then igba.acc q else {} \)" unfolding prod_def apply (auto simp: succ_of_E_def E_of_succ_def split: if_split_asm) done schematic_goal vf_prod_impl_aux_cava_reorder: shows "(?c, prod) \ igbg_impl_rel_ext unit_rel (Rq \\<^sub>r Rs)" unfolding prod_impl_aux_alt_cava_reorder[abs_def] using [[autoref_trace_failed_id]] apply (autoref (trace, keep_goal)) done lemma (in -) map_concat_map_map_opt: "map f (concat (map (\xa. map (f' xa) (l1 xa)) l2)) = List.maps (\xa. map (f o f' xa) (l1 xa)) l2" "((\(xd, xe). (xe, xd)) \ Pair xa) = (\x. (x,xa))" -- "Very specific optimization used in the next refinement" apply - apply (induction l2) apply (auto simp: List.maps_def) [2] apply auto [] done concrete_definition (in -) gbav_sys_prod_impl_cava_reorder for eqq Gi Si uses cava_inter_impl_loc.vf_prod_impl_aux_cava_reorder[ param_fo, unfolded map_concat_map_map_opt] lemmas [autoref_rules] = gbav_sys_prod_impl_cava_reorder.refine[OF cava_inter_impl_loc_this] context begin interpretation autoref_syn . (* HACK: Overwrite pattern that rewrites outer-level prod, such that local rules apply here. *) lemma [autoref_op_pat]: "prod \ OP prod" by simp end (*>*) definition dflt_inter :: "('q \ 's) igb_graph_rec \ ('q \ 's \ 's)" where "dflt_inter \ (prod, snd)" lemma dflt_inter_refine: shows "RETURN dflt_inter \ inter_spec S G" unfolding inter_spec_def dflt_inter_def apply (refine_rcg le_ASSERTI refine_vcg) proof clarsimp_all show "igb_graph prod" by (rule prod_invar) fix r assume "finite (igba.E\<^sup>* `` igba.V0)" "finite ((g_E S)\<^sup>* `` g_V0 S)" thus "finite ((g_E prod)\<^sup>* `` g_V0 prod)" using prod_finite_reachable by auto show "(\r'. prod.is_acc_run r' \ r = snd \ r') \ (sa.is_run r \ sa.L \ r \ igba.lang)" using gsp_correct1 gsp_correct2 apply (auto simp: comp_assoc) done qed schematic_goal dflt_inter_impl_aux: shows "(?c, dflt_inter) \ igbg_impl_rel_ext unit_rel (Rq \\<^sub>r Rs) \\<^sub>r (Rq \\<^sub>r Rs \ Rs)" unfolding dflt_inter_def[abs_def] using [[autoref_trace_failed_id]] apply (autoref (keep_goal)) done concrete_definition (in -) dflt_inter_impl for eqq Si Gi uses cava_inter_impl_loc.dflt_inter_impl_aux lemmas [autoref_rules] = dflt_inter_impl.refine[OF cava_inter_impl_loc_this] end definition [simp]: "op_dflt_inter \ cava_inter_impl_loc.dflt_inter" lemma [autoref_op_pat]: "cava_inter_impl_loc.dflt_inter \ op_dflt_inter" by simp context begin interpretation autoref_syn . lemma dflt_inter_autoref[autoref_rules]: fixes G :: "('q,'l set) igba_rec" fixes S :: "('s, 'l set) sa_rec" fixes Gi Si Rq Rs Rl eqq assumes "SIDE_PRECOND (igba G)" assumes "SIDE_PRECOND (sa S)" assumes "GEN_OP eqq op = (Rq \ Rq \ bool_rel)" assumes "(Gi,G) \ igbav_impl_rel_ext unit_rel Rq Rl" assumes "(Si,S) \ sa_impl_rel_ext unit_rel Rs Rl" shows "(dflt_inter_impl eqq Si Gi, (OP op_dflt_inter ::: sa_impl_rel_ext unit_rel Rs Rl \ igbav_impl_rel_ext unit_rel Rq Rl \ igbg_impl_rel_ext unit_rel (Rq \\<^sub>r Rs) \\<^sub>r (Rq \\<^sub>r Rs \ Rs))$S$G ) \ igbg_impl_rel_ext unit_rel (Rq \\<^sub>r Rs) \\<^sub>r (Rq \\<^sub>r Rs \ Rs)" proof - from assms interpret igba: igba G by simp from assms interpret sa: sa S by simp from assms interpret cava_inter_impl_loc S G Gi Si Rq Rs Rl eqq by unfold_locales simp_all show ?thesis apply simp apply (rule dflt_inter_impl.refine) apply unfold_locales done qed end lemma inter_spec_refineI: assumes "\ sa S; igba G \ \ m \ \R (inter_spec S G)" shows "m \ \R (inter_spec S G)" using assms unfolding inter_spec_def apply refine_rcg apply simp done lemma dflt_inter_impl_refine: fixes Rs :: "('si \ 's) set" fixes Rq :: "('qi \ 'q) set" fixes Rprop :: "('pi \ 'p) set" assumes [relator_props]: "single_valued Rs" "Range Rs = UNIV" "single_valued Rq" "Range Rq = UNIV" assumes EQ: "(eqq,op=) \ Rq \ Rq \ bool_rel" shows "(dflt_inter_impl eqq, inter_spec) \ sa_impl_rel_ext unit_rel Rs (\Rprop\fun_set_rel) \ igbav_impl_rel_ext unit_rel Rq (\Rprop\fun_set_rel) \ \igbg_impl_rel_ext unit_rel (Rq \\<^sub>r Rs) \\<^sub>r ((Rq \\<^sub>r Rs) \ Rs)\plain_nres_rel" apply (intro fun_relI plain_nres_relI) apply (rule inter_spec_refineI) proof - fix Si S Gi G assume R:"(Si,S) \ sa_impl_rel_ext unit_rel Rs (\Rprop\fun_set_rel)" "(Gi, G) \ igbav_impl_rel_ext unit_rel Rq (\Rprop\fun_set_rel)" assume "sa S" and "igba G" then interpret sa: sa S + igba: igba G . interpret cava_inter_impl_loc S G Gi Si Rq Rs "\Rprop\fun_set_rel" eqq apply unfold_locales using EQ R by simp_all note RETURN_refine[OF dflt_inter_impl.refine[OF cava_inter_impl_loc_this]] also note dflt_inter_refine finally show "RETURN (dflt_inter_impl eqq Si Gi) \ \ (igbg_impl_rel_ext unit_rel (Rq \\<^sub>r Rs) \\<^sub>r (Rq \\<^sub>r Rs \ Rs)) (inter_spec S G)" . qed subsubsection {* Definition of Model-Checker *} text {* In this section, we instantiate the parametrized model checker with the actual implementations. *} setup Locale_Code.open_block interpretation cava_sys_agn: impl_model_checker "sa_impl_rel_ext unit_rel Id (\Id\fun_set_rel)" "igbav_impl_rel_ext unit_rel Id (\Id\fun_set_rel)" "igbg_impl_rel_ext unit_rel (Id \\<^sub>r Id)" "\Id \\<^sub>r Id\lasso_run_rel" "\Id\lasso_run_rel" "ltl_to_gba_code" "\_::unit. dflt_inter_impl op =" "find_ce_code" "map_lasso" apply unfold_locales apply tagged_solver apply (rule ltl_to_gba_code_refine[unfolded is_ltl_to_gba_algo_def]) using dflt_inter_impl_refine[of Id Id "op =" Id] apply simp using find_ce_code_refine[unfolded is_find_ce_algo_def] apply simp apply assumption using map_lasso_run_refine[of Id Id] apply simp done setup Locale_Code.close_block definition "cava_sys_agn \ cava_sys_agn.impl_model_check" text {* The correctness theorem states correctness of the model checker wrt.\ a model given as system automata. In the following sections, we will then refine the model description to Boolean programs and Promela. *} theorem cava_sys_agn_correct: fixes sysi :: "('s::hashable, 'p::linorder \ bool, unit) sa_impl_scheme" and sys :: "('s, 'p set) sa_rec" and \ :: "'p ltlc" and cfg :: "config_l2b \ unit \ config_ce" assumes "(sysi, sys) \ sa_impl_rel_ext unit_rel Id (\Id\fun_set_rel)" and "sa sys" "finite ((g_E sys)\<^sup>* `` g_V0 sys)" shows "case cava_sys_agn cfg sysi \ of None \ sa.lang sys \ ltlc_language \ | Some None \ \ sa.lang sys \ ltlc_language \ | Some (Some L) \ graph_defs.is_run sys (run_of_lasso L) \ sa_L sys \ (run_of_lasso L) \ ltlc_language \" using cava_sys_agn.impl_model_check_correct[OF assms, of \ cfg] unfolding cava_sys_agn_def by (auto split: option.splits simp: lasso_run_rel_def br_def) subsection {* Model Checker for Boolean Programs *} definition bpc_to_sa :: "bprog \ BoolProgs.config \ (BoolProgs.config,nat set) sa_rec" -- "Conversion of a Boolean program to a system automata." where "bpc_to_sa bpc \ let (bp,c0)=bpc in \ g_V = UNIV, g_E = E_of_succ (set o BoolProgs.nexts bp), g_V0 = {c0}, sa_L = \c. bs_\ (snd c) \" definition bpc_to_sa_impl :: "bprog \ BoolProgs.config \ (BoolProgs.config,nat \ bool,unit) sa_impl_scheme" where "bpc_to_sa_impl bpc \ let (bp,c0)=bpc in \ gi_V = \_. True, gi_E = remdups o BoolProgs.nexts bp, gi_V0 = [c0], sai_L = \c i. bs_mem i (snd c) \" lemma bpc_to_sa_impl_refine: "(bpc_to_sa_impl bpc, bpc_to_sa bpc) \ sa_impl_rel_ext unit_rel Id (\nat_rel\fun_set_rel)" unfolding bpc_to_sa_impl_def bpc_to_sa_def unfolding sa_impl_rel_eext_def g_impl_rel_ext_def unfolding gen_sa_impl_rel_eext_def gen_g_impl_rel_ext_def apply (clarsimp split: prod.split) apply (intro conjI) apply (auto simp: fun_set_rel_def br_def) [] apply (rule E_of_succ_refine[param_fo]) apply (auto simp: list_set_rel_def br_def) [] apply (auto simp: list_set_rel_def br_def) [] apply (auto simp: fun_set_rel_def br_def) [] done lemma shows bpc_to_sa_invar: "sa (bpc_to_sa bpc)" and bpc_to_sa_fr: "finite ((g_E (bpc_to_sa bpc))\<^sup>* `` g_V0 (bpc_to_sa bpc))" proof - obtain bp c where [simp]: "bpc = (bp,c)" by (cases bpc) show "sa (bpc_to_sa bpc)" apply unfold_locales apply (simp add: bpc_to_sa_def) apply (simp add: bpc_to_sa_def) done show "finite ((g_E (bpc_to_sa bpc))\<^sup>* `` g_V0 (bpc_to_sa bpc))" apply (simp add: bpc_to_sa_def) apply (rule finite_subset[OF _ BoolProgs.reachable_configs_finite[of bp c]]) apply (rule rtrancl_reachable_induct) apply (auto intro: BoolProgs.reachable_configs.intros simp: E_of_succ_def) done qed interpretation bpc_to_sa: sa "bpc_to_sa bpc" using bpc_to_sa_invar . lemma bpc_to_sa_run_conv[simp]: "graph_defs.is_run (bpc_to_sa bpc) = bpc_is_run bpc" apply (rule ext) unfolding graph_defs.is_run_def unfolding bpc_to_sa_def bpc_is_run_def ipath_def E_of_succ_def by auto lemma bpc_to_sa_L_conv[simp]: "sa_L (bpc_to_sa bpc) = bpc_props" apply (rule ext) unfolding bpc_to_sa_def bpc_props_def apply (auto simp: E_of_succ_def split: prod.split) done lemma bpc_to_sa_lang_conv[simp]: "sa.lang (bpc_to_sa bpc) = bpc_lang bpc" unfolding bpc_to_sa.lang_def bpc_to_sa.accept_def[abs_def] bpc_lang_def by auto definition "cava_bpc cfg bpc \ \ cava_sys_agn cfg (bpc_to_sa_impl bpc) \" text {* Correctness theorem for the model checker on boolean programs. Note that the semantics of Boolean programs is given by @{const "bpc_lang"}. *} theorem cava_bpc_correct: "case cava_bpc cfg bpc \ of None \ bpc_lang bpc \ ltlc_language \ | Some None \ (\(bpc_lang bpc \ ltlc_language \)) | Some (Some ce) \ bpc_is_run bpc (run_of_lasso ce) \ bpc_props o run_of_lasso ce \ ltlc_language \" using cava_sys_agn_correct[OF bpc_to_sa_impl_refine bpc_to_sa_invar bpc_to_sa_fr, of bpc \ cfg] unfolding cava_bpc_def by (auto split: option.split simp: lasso_run_rel_def br_def) export_code cava_bpc checking SML subsection {* Model Checker for Promela Programs *} definition promela_to_sa :: "PromelaDatastructures.program \ APs \ gState \ (gState, nat set) sa_rec" -- "Conversion of a Promela model to a system automata." where "promela_to_sa promg \ let (prog,APs,g\<^sub>0)=promg in \ g_V = UNIV, g_E = E_of_succ (ls.\ o Promela.nexts_code prog), g_V0 = {g\<^sub>0}, sa_L = promela_props_ltl APs \" definition promela_to_sa_impl :: "PromelaDatastructures.program \ APs \ gState \ (gState, nat \ bool, unit) sa_impl_scheme" where "promela_to_sa_impl promg \ let (prog,APs,g\<^sub>0)=promg in \ gi_V = \_. True, gi_E = ls.to_list o Promela.nexts_code prog, gi_V0 = [g\<^sub>0], sai_L = propValid APs \" lemma promela_to_sa_impl_refine: shows "(promela_to_sa_impl promg, promela_to_sa promg) \ sa_impl_rel_ext unit_rel Id (\nat_rel\fun_set_rel)" unfolding promela_to_sa_impl_def promela_to_sa_def unfolding sa_impl_rel_eext_def g_impl_rel_ext_def unfolding gen_sa_impl_rel_eext_def gen_g_impl_rel_ext_def apply (clarsimp split: prod.split) apply (intro conjI) apply (auto simp: fun_set_rel_def br_def) [] apply (rule E_of_succ_refine[param_fo]) apply (auto simp: list_set_rel_def br_def ls.correct) [] apply (auto simp: list_set_rel_def br_def) [] apply (auto simp: fun_set_rel_def br_def in_set_member promela_props_ltl_def) [] done definition "cava_promela cfg ast \ \ let (promg,\\<^sub>i) = PromelaLTL.prepare cfg ast \ in cava_sys_agn (fst cfg) (promela_to_sa_impl promg) \\<^sub>i" text {* The next theorem states correctness of the Promela model checker. The correctness is specified for some AST. *} lemma cava_promela_correct: shows "case cava_promela cfg ast \ of None \ promela_language ast \ ltlc_language \ | Some None \ (\(promela_language ast \ ltlc_language \)) | Some (Some ce) \ promela_is_run ast (run_of_lasso ce) \ promela_props o run_of_lasso ce \ ltlc_language \" proof - obtain APs \\<^sub>i where conv: "PromelaLTL.ltl_convert \ = (APs,\\<^sub>i)" by (metis prod.exhaust) obtain prog g\<^sub>0 where ast: "Promela.setUp ast = (prog,g\<^sub>0)" by (metis prod.exhaust) let ?promg = "(prog,APs,g\<^sub>0)" have promela_to_sa_invar: "sa (promela_to_sa ?promg)" apply unfold_locales apply (simp add: promela_to_sa_def) apply (simp add: promela_to_sa_def) done have promela_to_sa_fr: "finite ((g_E (promela_to_sa ?promg))\<^sup>* `` g_V0 (promela_to_sa ?promg))" apply (simp add: promela_to_sa_def) apply (rule finite_subset[OF _ Promela.reachable_states_finite[of prog g\<^sub>0]]) apply (rule rtrancl_reachable_induct) apply (auto intro: Promela.reachable_states.intros simp: E_of_succ_def) [2] apply (fact setUp_program_inv[OF ast]) apply (fact setUp_gState_inv[OF ast]) done interpret promela_to_sa: sa "promela_to_sa ?promg" using promela_to_sa_invar . have promela_to_sa_run_conv[simp]: "graph_defs.is_run (promela_to_sa ?promg) = promela_is_run_ltl ?promg" apply (rule ext) unfolding graph_defs.is_run_def unfolding promela_to_sa_def promela_is_run_ltl_def promela_is_run'_def ipath_def E_of_succ_def by auto have promela_to_sa_L_conv[simp]: "sa_L (promela_to_sa ?promg) = promela_props_ltl APs" apply (rule ext) unfolding promela_to_sa_def promela_props_ltl_def[abs_def] by (auto simp: E_of_succ_def) have promela_to_sa_lang_conv[simp]: "sa.lang (promela_to_sa ?promg) = promela_language_ltl ?promg" unfolding promela_to_sa.lang_def promela_to_sa.accept_def[abs_def] promela_language_ltl_def by auto show ?thesis using cava_sys_agn_correct[OF promela_to_sa_impl_refine promela_to_sa_invar promela_to_sa_fr, of \\<^sub>i "fst cfg"] using promela_language_sub_iff[OF conv ast] promela_run_in_language_iff[OF conv] unfolding cava_promela_def PromelaLTL.prepare_def by (auto split: option.split prod.splits simp: lasso_run_rel_def br_def conv ast promela_is_run_ltl_def) qed export_code cava_promela checking SML subsection {* Extraction of SML Code *} definition "dflt_cfg \ (CFG_L2B_GERTHS,(),CFG_CE_SCC_GABOW)" export_code (* Cava MC *) cava_bpc cava_promela dflt_cfg CAVA_Impl.CFG_CE_NDFS (* BP *) BoolProgs.print_config chose_prog list_progs BoolProgs_LTL_Conv.ltl_conv BoolProgs_LTL_Conv.CProp BoolProgs_Programs.default_prog BoolProgs_Programs.keys_of_map BoolProgs_Programs.default_prog BoolProgs_Programs.keys_of_map (* Promela *) printProcesses Promela.printConfigFromAST lookupLTL PromelaLTLConv.ltl_conv (* stat printing *) frv_export_code LTL_to_GBA_impl.create_name_gba_code (* Lasso *) lasso_v0 lasso_va lasso_reach lasso_cycle (* Arith *) nat_of_integer int_of_integer integer_of_nat integer_of_int checking SML export_code (* Cava MC *) cava_bpc cava_promela dflt_cfg CAVA_Impl.CFG_CE_NDFS (* BP *) BoolProgs.print_config chose_prog list_progs BoolProgs_LTL_Conv.ltl_conv BoolProgs_LTL_Conv.CProp BoolProgs_Programs.default_prog BoolProgs_Programs.keys_of_map BoolProgs_Programs.default_prog BoolProgs_Programs.keys_of_map (* Promela *) printProcesses Promela.printConfigFromAST lookupLTL PromelaLTLConv.ltl_conv PromelaLTLConv.CProp PromelaLTLConv.Eq PromelaLTLConv.Ident (* stat printing *) frv_export_code LTL_to_GBA_impl.create_name_gba_code (* Lasso *) lasso_v0 lasso_va lasso_reach lasso_cycle (* Arith *) nat_of_integer int_of_integer integer_of_nat integer_of_int in SML file "code/CAVA_Export.sml" end diff --git a/thys/LTL_to_DRA/Impl/LTL_CAVA_Compat.thy b/thys/LTL_to_DRA/Impl/LTL_CAVA_Compat.thy --- a/thys/LTL_to_DRA/Impl/LTL_CAVA_Compat.thy +++ b/thys/LTL_to_DRA/Impl/LTL_CAVA_Compat.thy @@ -1,140 +1,37 @@ (* Author: Salomon Sickert + Author: _ + Author: _ + Author: _ License: BSD *) -section \CAVA LTL Compatibility Layer\ +section \LTL with syntactic sugar\ theory LTL_CAVA_Compat - imports Main "../LTL" + imports Main "~~/src/HOL/Library/Omega_Words_Fun" "../../../../afp/LTL/LTL" "../LTL_FGXU" begin -subsection \Original CAVA LTL Definitions\ - --- \The following definitions are adapted from the CAVA formalisation. The aim is to reuse the parsing infrastructure of CAVA\ - -datatype (ltlc_aprops: 'a) - ltlc = LTLcTrue - | LTLcFalse - | LTLcProp 'a - | LTLcNeg "'a ltlc" - | LTLcAnd "'a ltlc" "'a ltlc" - | LTLcOr "'a ltlc" "'a ltlc" - | LTLcImplies "'a ltlc" "'a ltlc" - | LTLcIff "'a ltlc" "'a ltlc" - | LTLcNext "'a ltlc" - | LTLcFinal "'a ltlc" - | LTLcGlobal "'a ltlc" - | LTLcUntil "'a ltlc" "'a ltlc" - | LTLcRelease "'a ltlc" "'a ltlc" - -notation - LTLcTrue ("true\<^sub>c") - and LTLcFalse ("false\<^sub>c") - and LTLcProp ("prop\<^sub>c'(_')") - and LTLcNeg ("not\<^sub>c _" [85] 85) - and LTLcAnd ("_ and\<^sub>c _" [82,82] 81) - and LTLcOr ("_ or\<^sub>c _" [81,81] 80) - and LTLcImplies ("_ implies\<^sub>c _" [81,81] 80) - and LTLcIff ("_ iff\<^sub>c _" [81,81] 80) - and LTLcNext ("X\<^sub>c _" [88] 87) - and LTLcFinal ("F\<^sub>c _" [88] 87) - and LTLcGlobal ("G\<^sub>c _" [88] 87) - and LTLcUntil ("_ U\<^sub>c _" [84,84] 83) - and LTLcRelease ("_ V\<^sub>c _" [83,83] 82) - -fun ltlc_semantics - :: "['a set word, 'a ltlc] \ bool" ("_ \\<^sub>c _" [80,80] 80) -where - "\ \\<^sub>c true\<^sub>c = True" - | "\ \\<^sub>c false\<^sub>c = False" - | "\ \\<^sub>c prop\<^sub>c(q) = (q\\ 0)" - | "\ \\<^sub>c not\<^sub>c \ = (\ \ \\<^sub>c \)" - | "\ \\<^sub>c \ and\<^sub>c \ = (\ \\<^sub>c \ \ \ \\<^sub>c \)" - | "\ \\<^sub>c \ or\<^sub>c \ = (\ \\<^sub>c \ \ \ \\<^sub>c \)" - | "\ \\<^sub>c \ implies\<^sub>c \ = (\ \\<^sub>c \ \ \ \\<^sub>c \)" - | "\ \\<^sub>c \ iff\<^sub>c \ = (\ \\<^sub>c \ \ \ \\<^sub>c \)" - | "\ \\<^sub>c X\<^sub>c \ = (suffix 1 \ \\<^sub>c \)" - | "\ \\<^sub>c F\<^sub>c \ = (\i. suffix i \ \\<^sub>c \)" - | "\ \\<^sub>c G\<^sub>c \ = (\i. suffix i \ \\<^sub>c \)" - | "\ \\<^sub>c \ U\<^sub>c \ = (\i. suffix i \ \\<^sub>c \ \ (\j \\<^sub>c \))" - | "\ \\<^sub>c \ V\<^sub>c \ = (\i. suffix i \ \\<^sub>c \ \ (\j \\<^sub>c \))" - subsection \Translation\ ---\The following infrastructure translates the CAVA LTL datatype to the LTL datatype used in this project\ +--\The following infrastructure translates the arbitrary LTL datatype to the LTL datatype used in this project\ -fun ltlc_to_ltl :: "bool \ 'a ltlc \ 'a ltl" +fun ltlc_to_ltl :: "'a ltln \ 'a ltl" where - "ltlc_to_ltl False true\<^sub>c = true" -| "ltlc_to_ltl False false\<^sub>c = false" -| "ltlc_to_ltl False prop\<^sub>c(q) = p(q)" -| "ltlc_to_ltl False (not\<^sub>c \) = ltlc_to_ltl True \" -| "ltlc_to_ltl False (\ and\<^sub>c \) = ltlc_to_ltl False \ and ltlc_to_ltl False \" -| "ltlc_to_ltl False (\ or\<^sub>c \) = ltlc_to_ltl False \ or ltlc_to_ltl False \" -| "ltlc_to_ltl False (\ implies\<^sub>c \) = ltlc_to_ltl True \ or ltlc_to_ltl False \" -| "ltlc_to_ltl False (\ iff\<^sub>c \) = (ltlc_to_ltl True \ or ltlc_to_ltl False \) and (ltlc_to_ltl False \ or ltlc_to_ltl True \)" -| "ltlc_to_ltl False (F\<^sub>c \) = F (ltlc_to_ltl False \)" -| "ltlc_to_ltl False (G\<^sub>c \) = G (ltlc_to_ltl False \)" -| "ltlc_to_ltl False (\ U\<^sub>c \) = ltlc_to_ltl False \ U ltlc_to_ltl False \" -| "ltlc_to_ltl False (\ V\<^sub>c \) = G (ltlc_to_ltl False \) or (ltlc_to_ltl False \ U (ltlc_to_ltl False \ and ltlc_to_ltl False \))" -| "ltlc_to_ltl True true\<^sub>c = false" -| "ltlc_to_ltl True false\<^sub>c = true" -| "ltlc_to_ltl True prop\<^sub>c(q) = np(q)" -| "ltlc_to_ltl True (not\<^sub>c \) = ltlc_to_ltl False \" -| "ltlc_to_ltl True (\ and\<^sub>c \) = ltlc_to_ltl True \ or ltlc_to_ltl True \" -| "ltlc_to_ltl True (\ or\<^sub>c \) = ltlc_to_ltl True \ and ltlc_to_ltl True \" -| "ltlc_to_ltl True (\ implies\<^sub>c \) = ltlc_to_ltl False \ and ltlc_to_ltl True \" -| "ltlc_to_ltl True (\ iff\<^sub>c \) = (ltlc_to_ltl True \ and ltlc_to_ltl False \) or (ltlc_to_ltl False \ and ltlc_to_ltl True \)" -| "ltlc_to_ltl True (F\<^sub>c \) = G (ltlc_to_ltl True \)" -| "ltlc_to_ltl True (G\<^sub>c \) = F (ltlc_to_ltl True \)" -| "ltlc_to_ltl True (\ U\<^sub>c \) = G (ltlc_to_ltl True \) or (ltlc_to_ltl True \ U (ltlc_to_ltl True \ and ltlc_to_ltl True \))" -| "ltlc_to_ltl True (\ V\<^sub>c \) = ltlc_to_ltl True \ U ltlc_to_ltl True \" -| "ltlc_to_ltl b (X\<^sub>c \) = X (ltlc_to_ltl b \)" - -lemma V_equiv: - "w \\<^sub>c \ V\<^sub>c \ \ w \\<^sub>c (\ U\<^sub>c (\ and\<^sub>c \)) or\<^sub>c G\<^sub>c \" -proof (cases "\i. \suffix i w \\<^sub>c \") - case True - def i \ "Least (\i. \suffix i w \\<^sub>c \)" - have "\j. j < i \ suffix j w \\<^sub>c \" and "\ suffix i w \\<^sub>c \" - using True LeastI not_less_Least unfolding i_def by metis+ (* Slow *) - hence "(\i. suffix i w \\<^sub>c \ \ (\j\<^sub>c \)) \ (\i. (suffix i w \\<^sub>c \ \ suffix i w \\<^sub>c \) \ (\j\<^sub>c \))" - by fastforce - moreover - hence "(\i. (suffix i w \\<^sub>c \ \ suffix i w \\<^sub>c \) \ (\j\<^sub>c \)) \ (\i. suffix i w \\<^sub>c \ \ (\j\<^sub>c \))" - using linorder_cases by blast - ultimately - show ?thesis - using True by auto -qed auto - -lemma push_neg_U: - "w \\<^sub>c (not\<^sub>c (\ U\<^sub>c \)) \ w \\<^sub>c not\<^sub>c \ U\<^sub>c (not\<^sub>c \ and\<^sub>c not\<^sub>c \) or\<^sub>c G\<^sub>c (not\<^sub>c \)" -proof - - have "w \\<^sub>c (not\<^sub>c (\ U\<^sub>c \)) \ w \\<^sub>c (not\<^sub>c \) V\<^sub>c (not\<^sub>c \)" - by simp - also - have "\ \ w \\<^sub>c not\<^sub>c \ U\<^sub>c (not\<^sub>c \ and\<^sub>c not\<^sub>c \) or\<^sub>c G\<^sub>c (not\<^sub>c \)" - unfolding V_equiv by simp - finally - show ?thesis - by blast -qed + "ltlc_to_ltl true\<^sub>n = true" +| "ltlc_to_ltl false\<^sub>n = false" +| "ltlc_to_ltl prop\<^sub>n(q) = p(q)" +| "ltlc_to_ltl nprop\<^sub>n(q) = np(q)" +| "ltlc_to_ltl (\ and\<^sub>n \) = ltlc_to_ltl \ and ltlc_to_ltl \" +| "ltlc_to_ltl (\ or\<^sub>n \) = ltlc_to_ltl \ or ltlc_to_ltl \" +| "ltlc_to_ltl (\ U\<^sub>n \) = (if \ = true\<^sub>n then F (ltlc_to_ltl \) else ltlc_to_ltl \ U ltlc_to_ltl \)" +| "ltlc_to_ltl (\ V\<^sub>n \) = (if \ = false\<^sub>n then G (ltlc_to_ltl \) else G (ltlc_to_ltl \) or (ltlc_to_ltl \ U (ltlc_to_ltl \ and ltlc_to_ltl \)))" +| "ltlc_to_ltl (X\<^sub>n \) = X (ltlc_to_ltl \)" lemma translation_correct: - "w \\<^sub>c \ \ w \ ltlc_to_ltl False \" - "\ w \\<^sub>c \ \ w \ ltlc_to_ltl True \" -proof (induction \ arbitrary: w) - case (LTLcUntil \ \) - case 2 - thus ?case - by (unfold ltlc_to_ltl.simps ltl_semantics.simps LTLcUntil[symmetric] ltlc_semantics.simps(4)[symmetric] push_neg_U) auto -next - case (LTLcRelease \ \) - case 1 - thus ?case - by (unfold ltlc_to_ltl.simps ltl_semantics.simps LTLcRelease[symmetric] V_equiv) auto -qed auto + "w \\<^sub>n \ \ w \ ltlc_to_ltl \" + apply (induction \ arbitrary: w) + apply (simp_all add: del: ltln_semantics.simps(9)) + unfolding ltln_Release_alterdef by auto -end +end \ No newline at end of file diff --git a/thys/LTL_to_DRA/LTL.thy b/thys/LTL_to_DRA/LTL.thy deleted file mode 100755 --- a/thys/LTL_to_DRA/LTL.thy +++ /dev/null @@ -1,851 +0,0 @@ -(* - Author: Salomon Sickert - License: BSD -*) - -section \LTL (in Negation-Normal-Form)\ - -theory LTL - imports Main "Aux/Words2" -begin - -text \Inspired/Based on schimpf/LTL\ - -subsection \Syntax\ - -datatype (vars: 'a) ltl = - LTLTrue - | LTLFalse - | LTLProp 'a - | LTLPropNeg 'a - | LTLAnd "'a ltl" "'a ltl" - | LTLOr "'a ltl" "'a ltl" - | LTLNext "'a ltl" - | LTLGlobal (theG: "'a ltl") - | LTLFinal "'a ltl" - | LTLUntil "'a ltl" "'a ltl" - -notation - LTLTrue ("true") - and LTLFalse ("false") - and LTLProp ("p'(_')") - and LTLPropNeg ("np'(_')" [86] 85) - and LTLAnd ("_ and _" [83,83] 82) - and LTLOr ("_ or _" [82,82] 81) - and LTLNext ("X _" [88] 87) - and LTLGlobal ("G _" [85] 84) - and LTLFinal ("F _" [84] 83) - and LTLUntil ("_ U _" [87,87] 86) - -subsection \Semantics\ - -fun ltl_semantics :: "['a set word, 'a ltl] \ bool" (infix "\" 80) -where - "w \ true = True" -| "w \ false = False" -| "w \ p(q) = (q \ w 0)" -| "w \ np(q) = (q \ w 0)" -| "w \ \ and \ = (w \ \ \ w \ \)" -| "w \ \ or \ = (w \ \ \ w \ \)" -| "w \ X \ = (suffix 1 w \ \)" -| "w \ G \ = (\k. suffix k w \ \)" -| "w \ F \ = (\k. suffix k w \ \)" -| "w \ \ U \ = (\k. suffix k w \ \ \ (\j < k. suffix j w \ \))" - -fun ltl_prop_entailment :: "['a ltl set, 'a ltl] \ bool" (infix "\\<^sub>P" 80) -where - "\ \\<^sub>P true = True" -| "\ \\<^sub>P false = False" -| "\ \\<^sub>P \ and \ = (\ \\<^sub>P \ \ \ \\<^sub>P \)" -| "\ \\<^sub>P \ or \ = (\ \\<^sub>P \ \ \ \\<^sub>P \)" -| "\ \\<^sub>P \ = (\ \ \)" - -subsubsection \Properties\ - -lemma LTL_G_one_step_unfolding: - "w \ G \ \ (w \ \ \ w \ X (G \))" - (is "?lhs \ ?rhs") -proof - assume ?lhs - hence "w \ \" - using suffix_0[of w] ltl_semantics.simps(8)[of w \] by metis - moreover - from `?lhs` have "w \ X (G \)" - by simp - ultimately - show ?rhs by simp -next - assume ?rhs - hence "w \ X (G \)" by simp - hence "\k. suffix (k + 1) w \ \" by force - hence "\k > 0. suffix k w \ \" - by (metis Suc_eq_plus1 gr0_implies_Suc) - moreover - from `?rhs` have "(suffix 0 w) \ \" by simp - ultimately - show ?lhs - using neq0_conv ltl_semantics.simps(8)[of w \] by blast -qed - -lemma LTL_F_one_step_unfolding: - "w \ F \ \ (w \ \ \ w \ X (F \))" - (is "?lhs \ ?rhs") -proof - assume ?lhs - then obtain k where "suffix k w \ \" by fastforce - thus ?rhs by (cases k) auto -next - assume ?rhs - thus ?lhs - using suffix_0[of w] suffix_suffix[of _ 1 w] by (metis ltl_semantics.simps(7) ltl_semantics.simps(9)) -qed - -lemma LTL_U_one_step_unfolding: - "w \ \ U \ \ (w \ \ \ (w \ \ \ w \ X (\ U \)))" - (is "?lhs \ ?rhs") -proof - assume ?lhs - then obtain k where "suffix k w \ \" and "\j \" - by force - thus ?rhs - by (cases k) auto -next - assume ?rhs - thus ?lhs - proof (cases "w \ \") - case False - hence "w \ \ \ w \ X (\ U \)" - using `?rhs` by blast - obtain k where "suffix k (suffix 1 w) \ \" and "\j \" - using False `?rhs` by force - moreover - { - fix j assume "j < 1 + k" - hence "suffix j w \ \" - using `w \ \ \ w \ X (\ U \)` `\j \`[unfolded suffix_suffix] - by (cases j) simp+ - } - ultimately - show ?thesis - by auto - qed force -qed - -lemma LTL_GF_infinitely_many_suffixes: - "w \ G (F \) = (\\<^sub>\i. suffix i w \ \)" - (is "?lhs = ?rhs") -proof - let ?S = "{i | i j. suffix (i + j) w \ \}" - let ?S' = "{i + j | i j. suffix (i + j) w \ \}" - - assume ?lhs - hence "infinite ?S" - by auto - moreover - have "\s \ ?S. \s' \ ?S'. s \ s'" - by fastforce - ultimately - have "infinite ?S'" - using infinite_nat_iff_unbounded_le le_trans by meson - moreover - have "?S' = {k | k. suffix k w \ \}" - using monoid_add_class.add.left_neutral by metis - ultimately - have "infinite {k | k. suffix k w \ \}" - by metis - thus ?rhs unfolding Inf_many_def by force -next - assume ?rhs - { - fix i - from `?rhs` obtain k where "i \ k" and "suffix k w \ \" - using INFM_nat_le[of "\n. suffix n w \ \"] by blast - then obtain j where "k = i + j" - using le_iff_add[of i k] by fast - hence "suffix j (suffix i w) \ \" - using `suffix k w \ \` suffix_suffix by fastforce - hence "suffix i w \ F \" by auto - } - thus ?lhs by auto -qed - -lemma LTL_FG_almost_all_suffixes: - "w \ F G \ = (\\<^sub>\i. suffix i w \ \)" - (is "?lhs = ?rhs") -proof - let ?S = "{k. \ suffix k w \ \}" - - assume ?lhs - then obtain i where "suffix i w \ G \" - by fastforce - hence "\j. j \ i \ (suffix j w \ \)" - using le_iff_add[of i] by auto - hence "\j. \suffix j w \ \ \ j < i" - using le_less_linear by blast - hence "?S \ {k. k < i}" - by blast - hence "finite ?S" - using finite_subset by fast - thus ?rhs - unfolding Alm_all_def Inf_many_def by presburger -next - assume ?rhs - obtain S where S_def: "S = {k. \ suffix k w \ \}" by blast - hence "finite S" - using `?rhs` unfolding Alm_all_def Inf_many_def by fast - then obtain i where "i = Max S" by blast - { - fix j - assume "i < j" - hence "j \ S" - using `i = Max S` Max.coboundedI[OF `finite S`] less_le_not_le by blast - hence "suffix j w \ \" using S_def by fast - } - hence "\j > i. (suffix j w \ \)" by simp - hence "suffix (Suc i) w \ G \" by auto - thus ?lhs - using ltl_semantics.simps(9)[of w "G \"] by blast -qed - -lemma LTL_FG_suffix: - "(suffix i w) \ F (G \) = w \ F (G \)" -proof - - have "(\m. \n\m. suffix n w \ \) = (\m. \n\m. suffix n (suffix i w) \ \)" (is "?l = ?r") - proof - assume ?r - then obtain m where "\n\m. suffix n (suffix i w) \ \" - by blast - hence "\n\i+m. suffix n w \ \" - unfolding suffix_suffix by (metis le_iff_add add_leE add_le_cancel_left) - thus ?l - by auto - qed (metis suffix_suffix trans_le_add2) - thus ?thesis - unfolding LTL_FG_almost_all_suffixes MOST_nat_le .. -qed - -lemma LTL_GF_suffix: - "(suffix i w) \ G (F \) = w \ G (F \)" -proof - - have "(\m. \n\m. suffix n w \ \) = (\m. \n\m. suffix n (suffix i w) \ \)" (is "?l = ?r") - proof - assume ?l - thus ?r - by (metis suffix_suffix add_leE add_le_cancel_left le_Suc_ex) - qed (metis suffix_suffix trans_le_add2) - thus ?thesis - unfolding LTL_GF_infinitely_many_suffixes INFM_nat_le .. -qed - -lemma LTL_suffix_G: - "w \ G \ \ suffix i w \ G \" - using suffix_0 suffix_suffix by (induction i) simp_all - -lemma LTL_prop_entailment_monotonI[intro]: - "S \\<^sub>P \ \ S \ S' \ S' \\<^sub>P \" - by (induction rule: ltl_prop_entailment.induct) auto - -lemma ltl_models_equiv_prop_entailment: - "w \ \ = {\. w \ \} \\<^sub>P \" - by (induction \) auto - -subsubsection \Limit Behaviour of the G-operator\ - -abbreviation Only_G -where - "Only_G S \ \x \ S. \y. x = G y" - -lemma ltl_G_stabilize: - assumes "finite \" - assumes "Only_G \" - obtains i where "\\ j. \ \ \ \ suffix i w \ \ = suffix (i + j) w \ \" -proof - - have "finite \ \ Only_G \ \ \i. \\ \ \. \j. suffix i w \ \ = suffix (i + j) w \ \" - proof (induction rule: finite_induct) - case (insert \ \) - then obtain i\<^sub>1 where "\\ j. \ \ \ \ suffix i\<^sub>1 w \ \ = suffix (i\<^sub>1 + j) w \ \" - by blast - moreover - from insert obtain \ where "\ = G \" - by blast - have "\i. \j. suffix i w \ G \ = suffix (i + j) w \ G \" - by (metis LTL_suffix_G plus_nat.add_0 suffix_0 suffix_suffix) - then obtain i\<^sub>2 where "\j. suffix i\<^sub>2 w \ \ = suffix (i\<^sub>2 + j) w \ \" - unfolding `\ = G \` by blast - ultimately - have "\\' j. \' \ \ \ {\} \ suffix (i\<^sub>1 + i\<^sub>2) w \ \' = suffix (i\<^sub>1 + i\<^sub>2 + j) w \ \'" - unfolding Un_iff singleton_iff by (metis add.commute add.left_commute) - thus ?case - by blast - qed simp - with assms obtain i where "\\ j. \ \ \ \ suffix i w \ \ = suffix (i + j) w \ \" - by blast - thus ?thesis - using that by blast -qed - -lemma ltl_G_stabilize_property: - assumes "finite \" - assumes "Only_G \" - assumes "\\ j. \ \ \ \ suffix i w \ \ = suffix (i + j) w \ \" - assumes "G \ \ \ \ {\. w \ F \}" - shows "suffix i w \ G \" -proof - - obtain j where "suffix j w \ G \" - using assms by fastforce - thus "suffix i w \ G \" - by (cases "i \ j", insert assms, unfold le_iff_add, blast, - metis (erased, lifting) LTL_suffix_G `suffix j w \ G \` le_add_diff_inverse nat_le_linear suffix_suffix) -qed - -subsection \Subformulae\ - -subsubsection \Propositions\ - -fun propos :: "'a ltl \'a ltl set" -where - "propos true = {}" -| "propos false = {}" -| "propos (\ and \) = propos \ \ propos \" -| "propos (\ or \) = propos \ \ propos \" -| "propos \ = {\}" - -fun nested_propos :: "'a ltl \'a ltl set" -where - "nested_propos true = {}" -| "nested_propos false = {}" -| "nested_propos (\ and \) = nested_propos \ \ nested_propos \" -| "nested_propos (\ or \) = nested_propos \ \ nested_propos \" -| "nested_propos (F \) = {F \} \ nested_propos \" -| "nested_propos (G \) = {G \} \ nested_propos \" -| "nested_propos (X \) = {X \} \ nested_propos \" -| "nested_propos (\ U \) = {\ U \} \ nested_propos \ \ nested_propos \" -| "nested_propos \ = {\}" - -lemma finite_propos: - "finite (propos \)" "finite (nested_propos \)" - by (induction \) simp+ - -lemma propos_subset: - "propos \ \ nested_propos \" - by (induction \) auto - -lemma LTL_prop_entailment_restrict_to_propos: - "S \\<^sub>P \ = (S \ propos \) \\<^sub>P \" - by (induction \) auto - -lemma propos_foldl: - assumes "\x y. propos (f x y) = propos x \ propos y" - shows "\{propos y |y. y = i \ y \ set xs} = propos (foldl f i xs)" -proof (induction xs rule: rev_induct) - case (snoc x xs) - have "\{propos y |y. y = i \ y \ set (xs@[x])} = \{propos y |y. y = i \ y \ set xs} \ propos x" - by auto - also - have "\ = propos (foldl f i xs) \ propos x" - using snoc by auto - also - have "\ = propos (foldl f i (xs@[x]))" - using assms by simp - finally - show ?case . -qed simp - -subsubsection \G-Subformulae\ - -text \Notation for paper: mathds{G}\ - -fun G_nested_propos :: "'a ltl \'a ltl set" ("\<^bold>G") -where - "\<^bold>G (\ and \) = \<^bold>G \ \ \<^bold>G \" -| "\<^bold>G (\ or \) = \<^bold>G \ \ \<^bold>G \" -| "\<^bold>G (F \) = \<^bold>G \" -| "\<^bold>G (G \) = \<^bold>G \ \ {G \}" -| "\<^bold>G (X \) = \<^bold>G \" -| "\<^bold>G (\ U \) = \<^bold>G \ \ \<^bold>G \" -| "\<^bold>G \ = {}" - -lemma G_nested_finite: - "finite (\<^bold>G \)" - by (induction \) auto - -lemma G_nested_propos_alt_def: - "\<^bold>G \ = nested_propos \ \ {\. (\x. \ = G x)}" - by (induction \) auto - -lemma G_nested_propos_Only_G: - "Only_G (\<^bold>G \)" - unfolding G_nested_propos_alt_def by blast - -lemma G_not_in_G: - "G \ \ \<^bold>G \" -proof - - have "\\. \ \ \<^bold>G \ \ size \ \ size \" - by (induction \) fastforce+ - thus ?thesis - by fastforce -qed - -lemma G_subset_G: - "\ \ \<^bold>G \ \ \<^bold>G \ \ \<^bold>G \" - "G \ \ \<^bold>G \ \ \<^bold>G \ \ \<^bold>G \" - by (induction \) auto - -lemma \_properties: - assumes "\ \ \<^bold>G \" - shows \_finite: "finite \" and \_elements: "Only_G \" - using assms G_nested_finite G_nested_propos_alt_def by (blast dest: finite_subset)+ - -subsection \Propositional Implication and Equivalence\ - -definition ltl_prop_implies :: "['a ltl, 'a ltl] \ bool" (infix "\\<^sub>P" 75) -where - "\ \\<^sub>P \ \ \\. \ \\<^sub>P \ \ \ \\<^sub>P \" - -definition ltl_prop_equiv :: "['a ltl, 'a ltl] \ bool" (infix "\\<^sub>P" 75) -where - "\ \\<^sub>P \ \ \\. \ \\<^sub>P \ \ \ \\<^sub>P \" - -lemma ltl_prop_implies_equiv: - "\ \\<^sub>P \ \ \ \\<^sub>P \ \ \ \\<^sub>P \" - unfolding ltl_prop_implies_def ltl_prop_equiv_def by meson - -lemma ltl_prop_equiv_equivp: - "equivp (op \\<^sub>P)" - by (blast intro: equivpI[of "op \\<^sub>P", unfolded transp_def symp_def reflp_def ltl_prop_equiv_def]) - -lemma [trans]: - "\ \\<^sub>P \ \ \ \\<^sub>P \ \ \ \\<^sub>P \" - using ltl_prop_equiv_equivp[THEN equivp_transp] . - -subsubsection \Quotient Type for Propositional Equivalence\ - -quotient_type 'a ltl_prop_equiv_quotient = "'a ltl" / "op \\<^sub>P" - morphisms Rep Abs - by (simp add: ltl_prop_equiv_equivp) - -type_synonym 'a ltl\<^sub>P = "'a ltl_prop_equiv_quotient" - -instantiation ltl_prop_equiv_quotient :: (type) equal begin - -lift_definition ltl_prop_equiv_quotient_eq_test :: "'a ltl\<^sub>P \ 'a ltl\<^sub>P \ bool" is "\x y. x \\<^sub>P y" - by (metis ltl_prop_equiv_quotient.abs_eq_iff) - -definition - eq: "equal_class.equal \ ltl_prop_equiv_quotient_eq_test" - -instance - by (standard; simp add: eq ltl_prop_equiv_quotient_eq_test.rep_eq, metis Quotient_ltl_prop_equiv_quotient Quotient_rel_rep) - -end - -lemma ltl\<^sub>P_abs_rep: "Abs (Rep \) = \" - by (meson Quotient3_abs_rep Quotient3_ltl_prop_equiv_quotient) - -lift_definition ltl_prop_entails_abs :: "'a ltl set \ 'a ltl\<^sub>P \ bool" ("_ \\\<^sub>P _") is "op \\<^sub>P" - by (simp add: ltl_prop_equiv_def) - -lift_definition ltl_prop_implies_abs :: "'a ltl\<^sub>P \ 'a ltl\<^sub>P \ bool" ("_ \\\<^sub>P _") is "op \\<^sub>P" - by (simp add: ltl_prop_equiv_def ltl_prop_implies_def) - -subsubsection \Propositional Equivalence implies LTL Equivalence\ - -lemma ltl_prop_implication_implies_ltl_implication: - "w \ \ \ \ \\<^sub>P \ \ w \ \" - unfolding ltl_prop_implies_def ltl_models_equiv_prop_entailment by simp - -lemma ltl_prop_equiv_implies_ltl_equiv: - "\ \\<^sub>P \ \ w \ \ = w \ \" - using ltl_prop_implication_implies_ltl_implication ltl_prop_implies_equiv by blast - -subsection \Substitution\ - -fun subst :: "'a ltl \ ('a ltl \ 'a ltl) \ 'a ltl" -where - "subst true m = true" -| "subst false m = false" -| "subst (\ and \) m = subst \ m and subst \ m" -| "subst (\ or \) m = subst \ m or subst \ m" -| "subst \ m = (case m \ of Some \' \ \' | None \ \)" - -text \Based on Uwe Schoening's Translation Lemma (Logic for CS, p. 54)\ - -lemma ltl_prop_equiv_subst_S: - "S \\<^sub>P subst \ m = ((S - dom m) \ {\ | \ \'. \ \ dom m \ m \ = Some \' \ S \\<^sub>P \'}) \\<^sub>P \" - by (induction \) (auto split: option.split) - -lemma subst_respects_ltl_prop_entailment: - "\ \\<^sub>P \ \ subst \ m \\<^sub>P subst \ m" - "\ \\<^sub>P \ \ subst \ m \\<^sub>P subst \ m" - unfolding ltl_prop_equiv_def ltl_prop_implies_def ltl_prop_equiv_subst_S by blast+ - -lemma subst_respects_ltl_prop_entailment_generalized: - "(\\. (\x. x \ S \ \ \\<^sub>P x) \ \ \\<^sub>P y) \ (\x. x \ S \ \ \\<^sub>P subst x m) \ \ \\<^sub>P subst y m" - unfolding ltl_prop_equiv_subst_S by simp - -lemma decomposable_function_subst: - "\f true = true; f false = false; \\ \. f (\ and \) = f \ and f \; \\ \. f (\ or \) = f \ or f \\ \ f \ = subst \ (\\. Some (f \))" - by (induction \) auto - -subsection \Additional Operators\ - -subsubsection \And\ - -lemma foldl_LTLAnd_prop_entailment: - "S \\<^sub>P foldl LTLAnd i xs = (S \\<^sub>P i \ (\y \ set xs. S \\<^sub>P y))" - by (induction xs arbitrary: i) auto - -fun And :: "'a ltl list \ 'a ltl" -where - "And [] = true" -| "And (x#xs) = foldl LTLAnd x xs" - -lemma And_prop_entailment: - "S \\<^sub>P And xs = (\x \ set xs. S \\<^sub>P x)" - using foldl_LTLAnd_prop_entailment by (cases xs) auto - -lemma And_propos: - "propos (And xs) = \{propos x| x. x \ set xs}" -proof (cases xs) - case Nil - thus ?thesis by simp -next - case (Cons x xs) - thus ?thesis - using propos_foldl[of LTLAnd x] by auto -qed - -lemma And_semantics: - "w \ And xs = (\x \ set xs. w \ x)" -proof - - have And_propos': "\x. x \ set xs \ propos x \ propos (And xs)" - using And_propos by blast - - have "w \ And xs = {\. \ \ propos (And xs) \ w \ \} \\<^sub>P (And xs)" - using ltl_models_equiv_prop_entailment LTL_prop_entailment_restrict_to_propos by blast - also - have "\ = (\x \ set xs. {\. \ \ propos (And xs) \ w \ \} \\<^sub>P x)" - using And_prop_entailment by auto - also - have "\ = (\x \ set xs. {\. \ \ propos x \ w \ \} \\<^sub>P x)" - using LTL_prop_entailment_restrict_to_propos And_propos' by blast - also - have "\ = (\x \ set xs. w \ x)" - using ltl_models_equiv_prop_entailment LTL_prop_entailment_restrict_to_propos by blast - finally - show ?thesis . -qed - -lemma And_append_syntactic: - "xs \ [] \ And (xs @ ys) = And ((And xs)#ys)" - by (induction xs rule: list_nonempty_induct) simp+ - -lemma And_append_S: - "S \\<^sub>P And (xs @ ys) = S \\<^sub>P And xs and And ys" - using And_prop_entailment[of S] by auto - -lemma And_append: - "And (xs @ ys) \\<^sub>P And xs and And ys" - unfolding ltl_prop_equiv_def using And_append_S by blast - -subsubsection \Lifted Variant\ - -lift_definition and_abs :: "'a ltl\<^sub>P \ 'a ltl\<^sub>P \ 'a ltl\<^sub>P" ("_ \and _") is "\x y. x and y" - unfolding ltl_prop_equiv_def by simp - -fun And_abs :: "'a ltl\<^sub>P list \ 'a ltl\<^sub>P" ("\And") -where - "\And xs = foldl and_abs (Abs true) xs" - -lemma foldl_LTLAnd_prop_entailment_abs: - "S \\\<^sub>P foldl and_abs i xs = (S \\\<^sub>P i \ (\y\set xs. S \\\<^sub>P y))" - by (induction xs arbitrary: i) - (simp_all add: and_abs_def ltl_prop_entails_abs.abs_eq, metis ltl_prop_entails_abs.rep_eq) - -lemma And_prop_entailment_abs: - "S \\\<^sub>P \And xs = (\x \ set xs. S \\\<^sub>P x)" - by (simp add: foldl_LTLAnd_prop_entailment_abs ltl_prop_entails_abs.abs_eq) - -lemma and_abs_conjunction: - "S \\\<^sub>P \ \and \ \ S \\\<^sub>P \ \ S \\\<^sub>P \" - by (metis and_abs.abs_eq ltl\<^sub>P_abs_rep ltl_prop_entailment.simps(3) ltl_prop_entails_abs.abs_eq) - -subsubsection \Or\ - -lemma foldl_LTLOr_prop_entailment: - "S \\<^sub>P foldl LTLOr i xs = (S \\<^sub>P i \ (\y \ set xs. S \\<^sub>P y))" - by (induction xs arbitrary: i) auto - -fun Or :: "'a ltl list \ 'a ltl" -where - "Or [] = false" -| "Or (x#xs) = foldl LTLOr x xs" - -lemma Or_prop_entailment: - "S \\<^sub>P Or xs = (\x \ set xs. S \\<^sub>P x)" - using foldl_LTLOr_prop_entailment by (cases xs) auto - -lemma Or_propos: - "propos (Or xs) = \{propos x| x. x \ set xs}" -proof (cases xs) - case Nil - thus ?thesis by simp -next - case (Cons x xs) - thus ?thesis - using propos_foldl[of LTLOr x] by auto -qed - -lemma Or_semantics: - "w \ Or xs = (\x \ set xs. w \ x)" -proof - - have Or_propos': "\x. x \ set xs \ propos x \ propos (Or xs)" - using Or_propos by blast - - have "w \ Or xs = {\. \ \ propos (Or xs) \ w \ \} \\<^sub>P (Or xs)" - using ltl_models_equiv_prop_entailment LTL_prop_entailment_restrict_to_propos by blast - also - have "\ = (\x \ set xs. {\. \ \ propos (Or xs) \ w \ \} \\<^sub>P x)" - using Or_prop_entailment by auto - also - have "\ = (\x \ set xs. {\. \ \ propos x \ w \ \} \\<^sub>P x)" - using LTL_prop_entailment_restrict_to_propos Or_propos' by blast - also - have "\ = (\x \ set xs. w \ x)" - using ltl_models_equiv_prop_entailment LTL_prop_entailment_restrict_to_propos by blast - finally - show ?thesis . -qed - -lemma Or_append_syntactic: - "xs \ [] \ Or (xs @ ys) = Or ((Or xs)#ys)" - by (induction xs rule: list_nonempty_induct) simp+ - -lemma Or_append_S: - "S \\<^sub>P Or (xs @ ys) = S \\<^sub>P Or xs or Or ys" - using Or_prop_entailment[of S] by auto - -lemma Or_append: - "Or (xs @ ys) \\<^sub>P Or xs or Or ys" - unfolding ltl_prop_equiv_def using Or_append_S by blast - -subsubsection \$eval_G$\ - --- \Partly evaluate a formula by only considering G-subformulae\ - -fun eval\<^sub>G -where - "eval\<^sub>G S (\ and \) = eval\<^sub>G S \ and eval\<^sub>G S \" -| "eval\<^sub>G S (\ or \) = eval\<^sub>G S \ or eval\<^sub>G S \" -| "eval\<^sub>G S (G \) = (if G \ \ S then true else false)" -| "eval\<^sub>G S \ = \" - --- \Syntactic Properties\ - -lemma eval\<^sub>G_And_map: - "eval\<^sub>G S (And xs) = And (map (eval\<^sub>G S) xs)" -proof (induction xs rule: rev_induct) - case (snoc x xs) - thus ?case - by (cases xs) simp+ -qed simp - -lemma eval\<^sub>G_Or_map: - "eval\<^sub>G S (Or xs) = Or (map (eval\<^sub>G S) xs)" -proof (induction xs rule: rev_induct) - case (snoc x xs) - thus ?case - by (cases xs) simp+ -qed simp - -lemma eval\<^sub>G_G_nested: - "\<^bold>G (eval\<^sub>G \ \) \ \<^bold>G \" - by (induction \) (simp_all, blast+) - -lemma eval\<^sub>G_subst: - "eval\<^sub>G S \ = subst \ (\\. Some (eval\<^sub>G S \))" - by (induction \) simp_all - --- \Semantic Properties\ - -lemma eval\<^sub>G_prop_entailment: - "S \\<^sub>P eval\<^sub>G S \ \ S \\<^sub>P \" - by (induction \, auto) - -lemma eval\<^sub>G_respectfulness: - "\ \\<^sub>P \ \ eval\<^sub>G S \ \\<^sub>P eval\<^sub>G S \" - "\ \\<^sub>P \ \ eval\<^sub>G S \ \\<^sub>P eval\<^sub>G S \" - using subst_respects_ltl_prop_entailment eval\<^sub>G_subst by metis+ - -lemma eval\<^sub>G_respectfulness_generalized: - "(\\. (\x. x \ S \ \ \\<^sub>P x) \ \ \\<^sub>P y) \ (\x. x \ S \ \ \\<^sub>P eval\<^sub>G P x) \ \ \\<^sub>P eval\<^sub>G P y" - using subst_respects_ltl_prop_entailment_generalized[of S y \] eval\<^sub>G_subst[of P] by metis - -lift_definition eval\<^sub>G_abs :: "'a ltl set \ 'a ltl\<^sub>P \ 'a ltl\<^sub>P" ("\eval\<^sub>G") is eval\<^sub>G - by (insert eval\<^sub>G_respectfulness(2)) - -subsection \Finite Quotient Set\ - -text \If we restrict formulas to a finite set of propositions, the set of quotients of these is finite\ - -lemma Rep_Abs_prop_entailment[simp]: - "A \\<^sub>P Rep (Abs \) = A \\<^sub>P \" - using Quotient3_ltl_prop_equiv_quotient[THEN rep_abs_rsp] - by (auto simp add: ltl_prop_equiv_def) - -fun sat_models :: "'a ltl_prop_equiv_quotient \ 'a ltl set set" -where - "sat_models a = {A. A \\<^sub>P Rep(a)}" - -lemma sat_models_invariant: - "A \ sat_models (Abs \) = A \\<^sub>P \" - using Rep_Abs_prop_entailment by auto - -lemma sat_models_inj: - "inj sat_models" - using Quotient3_ltl_prop_equiv_quotient[THEN Quotient3_rel_rep] - by (auto simp add: ltl_prop_equiv_def inj_on_def) - -lemma sat_models_finite_image: - assumes "finite P" - shows "finite (sat_models ` {Abs \ | \. nested_propos \ \ P})" -proof - - have "sat_models (Abs \) = {A \ B | A B. A \ P \ A \\<^sub>P \ \ B \ UNIV - P}" (is "?lhs = ?rhs") - if "nested_propos \ \ P" for \ - proof - have "\A B. A \ sat_models (Abs \) \ A \ B \ sat_models (Abs \)" - unfolding sat_models_invariant by blast - moreover - have "{A | A. A \ P \ A \\<^sub>P \} \ sat_models (Abs \)" - using sat_models_invariant by fast - ultimately - show "?rhs \ ?lhs" - by blast - next - have "propos \ \ P" - using that propos_subset by blast - have "A \ {A \ B | A B. A \ P \ A \\<^sub>P \ \ B \ UNIV - P}" - if "A \ sat_models (Abs \)" for A - proof (standard, goal_cases prems) - case prems - then have "A \\<^sub>P \" - using that sat_models_invariant by blast - then obtain C D where "C = (A \ P)" and "D = A - P" and "A = C \ D" - by blast - then have "C \\<^sub>P \" and "C \ P" and "D \ UNIV - P" - using `A \\<^sub>P \` LTL_prop_entailment_restrict_to_propos `propos \ \ P` by blast+ - then have "C \ D \ {A \ B | A B. A \ P \ A \\<^sub>P \ \ B \ UNIV - P}" - by blast - thus ?case - using `A = C \ D` by simp - qed - thus "?lhs \ ?rhs" - by blast - qed - hence Equal: "{sat_models (Abs \) | \. nested_propos \ \ P} = {{A \ B | A B. A \ P \ A \\<^sub>P \ \ B \ UNIV - P} | \. nested_propos \ \ P}" - by (metis (lifting, no_types)) - - have Finite: "finite {{A \ B | A B. A \ P \ A \\<^sub>P \ \ B \ UNIV - P} | \. nested_propos \ \ P}" - proof - - let ?map = "\P S. {A \ B | A B. A \ S \ B \ UNIV - P}" - obtain S' where S'_def: "S' = {{A \ B | A B. A \ P \ A \\<^sub>P \ \ B \ UNIV - P} | \. nested_propos \ \ P}" - by blast - obtain S where S_def: "S = {{A | A. A \ P \ A \\<^sub>P \} | \. nested_propos \ \ P}" - by blast - - -- \Prove S and ?map applied to it is finite\ - hence "S \ Pow (Pow P)" - by blast - hence "finite S" - using `finite P` finite_Pow_iff infinite_super by fast - hence "finite {?map P A | A. A \ S}" - by fastforce - - -- \Prove that S' can be embedded into S using ?map\ - - have "S' \ {?map P A | A. A \ S}" - proof - fix A - assume "A \ S'" - then obtain \ where "nested_propos \ \ P" - and "A = {A \ B | A B. A \ P \ A \\<^sub>P \ \ B \ UNIV - P}" - using S'_def by blast - then have "?map P {A | A. A \ P \ A \\<^sub>P \} = A" - by simp - moreover - have "{A | A. A \ P \ A \\<^sub>P \} \ S" - using `nested_propos \ \ P` S_def by auto - ultimately - show "A \ {?map P A | A. A \ S}" - by blast - qed - - show ?thesis - using rev_finite_subset[OF `finite {?map P A | A. A \ S}` `S' \ {?map P A | A. A \ S}`] - unfolding S'_def . - qed - - have Finite2: "finite {sat_models (Abs \) | \. nested_propos \ \ P}" - unfolding Equal using Finite by blast - have Equal2: "sat_models ` {Abs \ | \. nested_propos \ \ P} = {sat_models (Abs \) | \. nested_propos \ \ P}" - by blast - - show ?thesis - unfolding Equal2 using Finite2 by blast -qed - -lemma ltl_prop_equiv_quotient_restricted_to_P_finite: - assumes "finite P" - shows "finite {Abs \ | \. nested_propos \ \ P}" -proof - - have "inj_on sat_models {Abs \ |\. nested_propos \ \ P}" - using sat_models_inj subset_inj_on by auto - thus ?thesis - using finite_imageD[OF sat_models_finite_image[OF assms]] by fast -qed - -locale lift_ltl_transformer = - fixes - f :: "'a ltl \ 'b \ 'a ltl" - assumes - respectfulness: "\ \\<^sub>P \ \ f \ \ \\<^sub>P f \ \" - assumes - nested_propos_bounded: "nested_propos (f \ \) \ nested_propos \" -begin - -lift_definition f_abs :: "'a ltl\<^sub>P \ 'b \ 'a ltl\<^sub>P" is f - using respectfulness . - -lift_definition f_foldl_abs :: "'a ltl\<^sub>P \ 'b list \ 'a ltl\<^sub>P" is "foldl f" -proof - - fix \ \ :: "'a ltl" fix w :: "'b list" assume "\ \\<^sub>P \" - thus "foldl f \ w \\<^sub>P foldl f \ w" - using respectfulness by (induction w arbitrary: \ \) simp+ -qed - -lemma f_foldl_abs_alt_def: - "f_foldl_abs (Abs \) w = foldl f_abs (Abs \) w" - by (induction w rule: rev_induct) (unfold f_foldl_abs.abs_eq foldl.simps foldl_append, (metis f_abs.abs_eq)+) - -definition abs_reach :: "'a ltl_prop_equiv_quotient \ 'a ltl_prop_equiv_quotient set" -where - "abs_reach \ = {foldl f_abs \ w |w. True}" - -lemma finite_abs_reach: - "finite (abs_reach (Abs \))" -proof - - { - fix w - have "nested_propos (foldl f \ w) \ nested_propos \" - by (induction w arbitrary: \) (simp, metis foldl_Cons nested_propos_bounded subset_trans) - } - hence "abs_reach (Abs \) \ {Abs \ | \. nested_propos \ \ nested_propos \}" - unfolding abs_reach_def f_foldl_abs_alt_def[symmetric] f_foldl_abs.abs_eq by blast - thus ?thesis - using ltl_prop_equiv_quotient_restricted_to_P_finite finite_propos - by (blast dest: finite_subset) -qed - -end - -end diff --git a/thys/LTL_to_DRA/LTL_FGXU.thy b/thys/LTL_to_DRA/LTL_FGXU.thy new file mode 100755 --- /dev/null +++ b/thys/LTL_to_DRA/LTL_FGXU.thy @@ -0,0 +1,851 @@ +(* + Author: Salomon Sickert + License: BSD +*) + +section \LTL (in Negation-Normal-Form)\ + +theory LTL_FGXU + imports Main "Aux/Words2" +begin + +text \Inspired/Based on schimpf/LTL\ + +subsection \Syntax\ + +datatype (vars: 'a) ltl = + LTLTrue + | LTLFalse + | LTLProp 'a + | LTLPropNeg 'a + | LTLAnd "'a ltl" "'a ltl" + | LTLOr "'a ltl" "'a ltl" + | LTLNext "'a ltl" + | LTLGlobal (theG: "'a ltl") + | LTLFinal "'a ltl" + | LTLUntil "'a ltl" "'a ltl" + +notation + LTLTrue ("true") + and LTLFalse ("false") + and LTLProp ("p'(_')") + and LTLPropNeg ("np'(_')" [86] 85) + and LTLAnd ("_ and _" [83,83] 82) + and LTLOr ("_ or _" [82,82] 81) + and LTLNext ("X _" [88] 87) + and LTLGlobal ("G _" [85] 84) + and LTLFinal ("F _" [84] 83) + and LTLUntil ("_ U _" [87,87] 86) + +subsection \Semantics\ + +fun ltl_semantics :: "['a set word, 'a ltl] \ bool" (infix "\" 80) +where + "w \ true = True" +| "w \ false = False" +| "w \ p(q) = (q \ w 0)" +| "w \ np(q) = (q \ w 0)" +| "w \ \ and \ = (w \ \ \ w \ \)" +| "w \ \ or \ = (w \ \ \ w \ \)" +| "w \ X \ = (suffix 1 w \ \)" +| "w \ G \ = (\k. suffix k w \ \)" +| "w \ F \ = (\k. suffix k w \ \)" +| "w \ \ U \ = (\k. suffix k w \ \ \ (\j < k. suffix j w \ \))" + +fun ltl_prop_entailment :: "['a ltl set, 'a ltl] \ bool" (infix "\\<^sub>P" 80) +where + "\ \\<^sub>P true = True" +| "\ \\<^sub>P false = False" +| "\ \\<^sub>P \ and \ = (\ \\<^sub>P \ \ \ \\<^sub>P \)" +| "\ \\<^sub>P \ or \ = (\ \\<^sub>P \ \ \ \\<^sub>P \)" +| "\ \\<^sub>P \ = (\ \ \)" + +subsubsection \Properties\ + +lemma LTL_G_one_step_unfolding: + "w \ G \ \ (w \ \ \ w \ X (G \))" + (is "?lhs \ ?rhs") +proof + assume ?lhs + hence "w \ \" + using suffix_0[of w] ltl_semantics.simps(8)[of w \] by metis + moreover + from `?lhs` have "w \ X (G \)" + by simp + ultimately + show ?rhs by simp +next + assume ?rhs + hence "w \ X (G \)" by simp + hence "\k. suffix (k + 1) w \ \" by force + hence "\k > 0. suffix k w \ \" + by (metis Suc_eq_plus1 gr0_implies_Suc) + moreover + from `?rhs` have "(suffix 0 w) \ \" by simp + ultimately + show ?lhs + using neq0_conv ltl_semantics.simps(8)[of w \] by blast +qed + +lemma LTL_F_one_step_unfolding: + "w \ F \ \ (w \ \ \ w \ X (F \))" + (is "?lhs \ ?rhs") +proof + assume ?lhs + then obtain k where "suffix k w \ \" by fastforce + thus ?rhs by (cases k) auto +next + assume ?rhs + thus ?lhs + using suffix_0[of w] suffix_suffix[of _ 1 w] by (metis ltl_semantics.simps(7) ltl_semantics.simps(9)) +qed + +lemma LTL_U_one_step_unfolding: + "w \ \ U \ \ (w \ \ \ (w \ \ \ w \ X (\ U \)))" + (is "?lhs \ ?rhs") +proof + assume ?lhs + then obtain k where "suffix k w \ \" and "\j \" + by force + thus ?rhs + by (cases k) auto +next + assume ?rhs + thus ?lhs + proof (cases "w \ \") + case False + hence "w \ \ \ w \ X (\ U \)" + using `?rhs` by blast + obtain k where "suffix k (suffix 1 w) \ \" and "\j \" + using False `?rhs` by force + moreover + { + fix j assume "j < 1 + k" + hence "suffix j w \ \" + using `w \ \ \ w \ X (\ U \)` `\j \`[unfolded suffix_suffix] + by (cases j) simp+ + } + ultimately + show ?thesis + by auto + qed force +qed + +lemma LTL_GF_infinitely_many_suffixes: + "w \ G (F \) = (\\<^sub>\i. suffix i w \ \)" + (is "?lhs = ?rhs") +proof + let ?S = "{i | i j. suffix (i + j) w \ \}" + let ?S' = "{i + j | i j. suffix (i + j) w \ \}" + + assume ?lhs + hence "infinite ?S" + by auto + moreover + have "\s \ ?S. \s' \ ?S'. s \ s'" + by fastforce + ultimately + have "infinite ?S'" + using infinite_nat_iff_unbounded_le le_trans by meson + moreover + have "?S' = {k | k. suffix k w \ \}" + using monoid_add_class.add.left_neutral by metis + ultimately + have "infinite {k | k. suffix k w \ \}" + by metis + thus ?rhs unfolding Inf_many_def by force +next + assume ?rhs + { + fix i + from `?rhs` obtain k where "i \ k" and "suffix k w \ \" + using INFM_nat_le[of "\n. suffix n w \ \"] by blast + then obtain j where "k = i + j" + using le_iff_add[of i k] by fast + hence "suffix j (suffix i w) \ \" + using `suffix k w \ \` suffix_suffix by fastforce + hence "suffix i w \ F \" by auto + } + thus ?lhs by auto +qed + +lemma LTL_FG_almost_all_suffixes: + "w \ F G \ = (\\<^sub>\i. suffix i w \ \)" + (is "?lhs = ?rhs") +proof + let ?S = "{k. \ suffix k w \ \}" + + assume ?lhs + then obtain i where "suffix i w \ G \" + by fastforce + hence "\j. j \ i \ (suffix j w \ \)" + using le_iff_add[of i] by auto + hence "\j. \suffix j w \ \ \ j < i" + using le_less_linear by blast + hence "?S \ {k. k < i}" + by blast + hence "finite ?S" + using finite_subset by fast + thus ?rhs + unfolding Alm_all_def Inf_many_def by presburger +next + assume ?rhs + obtain S where S_def: "S = {k. \ suffix k w \ \}" by blast + hence "finite S" + using `?rhs` unfolding Alm_all_def Inf_many_def by fast + then obtain i where "i = Max S" by blast + { + fix j + assume "i < j" + hence "j \ S" + using `i = Max S` Max.coboundedI[OF `finite S`] less_le_not_le by blast + hence "suffix j w \ \" using S_def by fast + } + hence "\j > i. (suffix j w \ \)" by simp + hence "suffix (Suc i) w \ G \" by auto + thus ?lhs + using ltl_semantics.simps(9)[of w "G \"] by blast +qed + +lemma LTL_FG_suffix: + "(suffix i w) \ F (G \) = w \ F (G \)" +proof - + have "(\m. \n\m. suffix n w \ \) = (\m. \n\m. suffix n (suffix i w) \ \)" (is "?l = ?r") + proof + assume ?r + then obtain m where "\n\m. suffix n (suffix i w) \ \" + by blast + hence "\n\i+m. suffix n w \ \" + unfolding suffix_suffix by (metis le_iff_add add_leE add_le_cancel_left) + thus ?l + by auto + qed (metis suffix_suffix trans_le_add2) + thus ?thesis + unfolding LTL_FG_almost_all_suffixes MOST_nat_le .. +qed + +lemma LTL_GF_suffix: + "(suffix i w) \ G (F \) = w \ G (F \)" +proof - + have "(\m. \n\m. suffix n w \ \) = (\m. \n\m. suffix n (suffix i w) \ \)" (is "?l = ?r") + proof + assume ?l + thus ?r + by (metis suffix_suffix add_leE add_le_cancel_left le_Suc_ex) + qed (metis suffix_suffix trans_le_add2) + thus ?thesis + unfolding LTL_GF_infinitely_many_suffixes INFM_nat_le .. +qed + +lemma LTL_suffix_G: + "w \ G \ \ suffix i w \ G \" + using suffix_0 suffix_suffix by (induction i) simp_all + +lemma LTL_prop_entailment_monotonI[intro]: + "S \\<^sub>P \ \ S \ S' \ S' \\<^sub>P \" + by (induction rule: ltl_prop_entailment.induct) auto + +lemma ltl_models_equiv_prop_entailment: + "w \ \ = {\. w \ \} \\<^sub>P \" + by (induction \) auto + +subsubsection \Limit Behaviour of the G-operator\ + +abbreviation Only_G +where + "Only_G S \ \x \ S. \y. x = G y" + +lemma ltl_G_stabilize: + assumes "finite \" + assumes "Only_G \" + obtains i where "\\ j. \ \ \ \ suffix i w \ \ = suffix (i + j) w \ \" +proof - + have "finite \ \ Only_G \ \ \i. \\ \ \. \j. suffix i w \ \ = suffix (i + j) w \ \" + proof (induction rule: finite_induct) + case (insert \ \) + then obtain i\<^sub>1 where "\\ j. \ \ \ \ suffix i\<^sub>1 w \ \ = suffix (i\<^sub>1 + j) w \ \" + by blast + moreover + from insert obtain \ where "\ = G \" + by blast + have "\i. \j. suffix i w \ G \ = suffix (i + j) w \ G \" + by (metis LTL_suffix_G plus_nat.add_0 suffix_0 suffix_suffix) + then obtain i\<^sub>2 where "\j. suffix i\<^sub>2 w \ \ = suffix (i\<^sub>2 + j) w \ \" + unfolding `\ = G \` by blast + ultimately + have "\\' j. \' \ \ \ {\} \ suffix (i\<^sub>1 + i\<^sub>2) w \ \' = suffix (i\<^sub>1 + i\<^sub>2 + j) w \ \'" + unfolding Un_iff singleton_iff by (metis add.commute add.left_commute) + thus ?case + by blast + qed simp + with assms obtain i where "\\ j. \ \ \ \ suffix i w \ \ = suffix (i + j) w \ \" + by blast + thus ?thesis + using that by blast +qed + +lemma ltl_G_stabilize_property: + assumes "finite \" + assumes "Only_G \" + assumes "\\ j. \ \ \ \ suffix i w \ \ = suffix (i + j) w \ \" + assumes "G \ \ \ \ {\. w \ F \}" + shows "suffix i w \ G \" +proof - + obtain j where "suffix j w \ G \" + using assms by fastforce + thus "suffix i w \ G \" + by (cases "i \ j", insert assms, unfold le_iff_add, blast, + metis (erased, lifting) LTL_suffix_G `suffix j w \ G \` le_add_diff_inverse nat_le_linear suffix_suffix) +qed + +subsection \Subformulae\ + +subsubsection \Propositions\ + +fun propos :: "'a ltl \'a ltl set" +where + "propos true = {}" +| "propos false = {}" +| "propos (\ and \) = propos \ \ propos \" +| "propos (\ or \) = propos \ \ propos \" +| "propos \ = {\}" + +fun nested_propos :: "'a ltl \'a ltl set" +where + "nested_propos true = {}" +| "nested_propos false = {}" +| "nested_propos (\ and \) = nested_propos \ \ nested_propos \" +| "nested_propos (\ or \) = nested_propos \ \ nested_propos \" +| "nested_propos (F \) = {F \} \ nested_propos \" +| "nested_propos (G \) = {G \} \ nested_propos \" +| "nested_propos (X \) = {X \} \ nested_propos \" +| "nested_propos (\ U \) = {\ U \} \ nested_propos \ \ nested_propos \" +| "nested_propos \ = {\}" + +lemma finite_propos: + "finite (propos \)" "finite (nested_propos \)" + by (induction \) simp+ + +lemma propos_subset: + "propos \ \ nested_propos \" + by (induction \) auto + +lemma LTL_prop_entailment_restrict_to_propos: + "S \\<^sub>P \ = (S \ propos \) \\<^sub>P \" + by (induction \) auto + +lemma propos_foldl: + assumes "\x y. propos (f x y) = propos x \ propos y" + shows "\{propos y |y. y = i \ y \ set xs} = propos (foldl f i xs)" +proof (induction xs rule: rev_induct) + case (snoc x xs) + have "\{propos y |y. y = i \ y \ set (xs@[x])} = \{propos y |y. y = i \ y \ set xs} \ propos x" + by auto + also + have "\ = propos (foldl f i xs) \ propos x" + using snoc by auto + also + have "\ = propos (foldl f i (xs@[x]))" + using assms by simp + finally + show ?case . +qed simp + +subsubsection \G-Subformulae\ + +text \Notation for paper: mathds{G}\ + +fun G_nested_propos :: "'a ltl \'a ltl set" ("\<^bold>G") +where + "\<^bold>G (\ and \) = \<^bold>G \ \ \<^bold>G \" +| "\<^bold>G (\ or \) = \<^bold>G \ \ \<^bold>G \" +| "\<^bold>G (F \) = \<^bold>G \" +| "\<^bold>G (G \) = \<^bold>G \ \ {G \}" +| "\<^bold>G (X \) = \<^bold>G \" +| "\<^bold>G (\ U \) = \<^bold>G \ \ \<^bold>G \" +| "\<^bold>G \ = {}" + +lemma G_nested_finite: + "finite (\<^bold>G \)" + by (induction \) auto + +lemma G_nested_propos_alt_def: + "\<^bold>G \ = nested_propos \ \ {\. (\x. \ = G x)}" + by (induction \) auto + +lemma G_nested_propos_Only_G: + "Only_G (\<^bold>G \)" + unfolding G_nested_propos_alt_def by blast + +lemma G_not_in_G: + "G \ \ \<^bold>G \" +proof - + have "\\. \ \ \<^bold>G \ \ size \ \ size \" + by (induction \) fastforce+ + thus ?thesis + by fastforce +qed + +lemma G_subset_G: + "\ \ \<^bold>G \ \ \<^bold>G \ \ \<^bold>G \" + "G \ \ \<^bold>G \ \ \<^bold>G \ \ \<^bold>G \" + by (induction \) auto + +lemma \_properties: + assumes "\ \ \<^bold>G \" + shows \_finite: "finite \" and \_elements: "Only_G \" + using assms G_nested_finite G_nested_propos_alt_def by (blast dest: finite_subset)+ + +subsection \Propositional Implication and Equivalence\ + +definition ltl_prop_implies :: "['a ltl, 'a ltl] \ bool" (infix "\\<^sub>P" 75) +where + "\ \\<^sub>P \ \ \\. \ \\<^sub>P \ \ \ \\<^sub>P \" + +definition ltl_prop_equiv :: "['a ltl, 'a ltl] \ bool" (infix "\\<^sub>P" 75) +where + "\ \\<^sub>P \ \ \\. \ \\<^sub>P \ \ \ \\<^sub>P \" + +lemma ltl_prop_implies_equiv: + "\ \\<^sub>P \ \ \ \\<^sub>P \ \ \ \\<^sub>P \" + unfolding ltl_prop_implies_def ltl_prop_equiv_def by meson + +lemma ltl_prop_equiv_equivp: + "equivp (op \\<^sub>P)" + by (blast intro: equivpI[of "op \\<^sub>P", unfolded transp_def symp_def reflp_def ltl_prop_equiv_def]) + +lemma [trans]: + "\ \\<^sub>P \ \ \ \\<^sub>P \ \ \ \\<^sub>P \" + using ltl_prop_equiv_equivp[THEN equivp_transp] . + +subsubsection \Quotient Type for Propositional Equivalence\ + +quotient_type 'a ltl_prop_equiv_quotient = "'a ltl" / "op \\<^sub>P" + morphisms Rep Abs + by (simp add: ltl_prop_equiv_equivp) + +type_synonym 'a ltl\<^sub>P = "'a ltl_prop_equiv_quotient" + +instantiation ltl_prop_equiv_quotient :: (type) equal begin + +lift_definition ltl_prop_equiv_quotient_eq_test :: "'a ltl\<^sub>P \ 'a ltl\<^sub>P \ bool" is "\x y. x \\<^sub>P y" + by (metis ltl_prop_equiv_quotient.abs_eq_iff) + +definition + eq: "equal_class.equal \ ltl_prop_equiv_quotient_eq_test" + +instance + by (standard; simp add: eq ltl_prop_equiv_quotient_eq_test.rep_eq, metis Quotient_ltl_prop_equiv_quotient Quotient_rel_rep) + +end + +lemma ltl\<^sub>P_abs_rep: "Abs (Rep \) = \" + by (meson Quotient3_abs_rep Quotient3_ltl_prop_equiv_quotient) + +lift_definition ltl_prop_entails_abs :: "'a ltl set \ 'a ltl\<^sub>P \ bool" ("_ \\\<^sub>P _") is "op \\<^sub>P" + by (simp add: ltl_prop_equiv_def) + +lift_definition ltl_prop_implies_abs :: "'a ltl\<^sub>P \ 'a ltl\<^sub>P \ bool" ("_ \\\<^sub>P _") is "op \\<^sub>P" + by (simp add: ltl_prop_equiv_def ltl_prop_implies_def) + +subsubsection \Propositional Equivalence implies LTL Equivalence\ + +lemma ltl_prop_implication_implies_ltl_implication: + "w \ \ \ \ \\<^sub>P \ \ w \ \" + unfolding ltl_prop_implies_def ltl_models_equiv_prop_entailment by simp + +lemma ltl_prop_equiv_implies_ltl_equiv: + "\ \\<^sub>P \ \ w \ \ = w \ \" + using ltl_prop_implication_implies_ltl_implication ltl_prop_implies_equiv by blast + +subsection \Substitution\ + +fun subst :: "'a ltl \ ('a ltl \ 'a ltl) \ 'a ltl" +where + "subst true m = true" +| "subst false m = false" +| "subst (\ and \) m = subst \ m and subst \ m" +| "subst (\ or \) m = subst \ m or subst \ m" +| "subst \ m = (case m \ of Some \' \ \' | None \ \)" + +text \Based on Uwe Schoening's Translation Lemma (Logic for CS, p. 54)\ + +lemma ltl_prop_equiv_subst_S: + "S \\<^sub>P subst \ m = ((S - dom m) \ {\ | \ \'. \ \ dom m \ m \ = Some \' \ S \\<^sub>P \'}) \\<^sub>P \" + by (induction \) (auto split: option.split) + +lemma subst_respects_ltl_prop_entailment: + "\ \\<^sub>P \ \ subst \ m \\<^sub>P subst \ m" + "\ \\<^sub>P \ \ subst \ m \\<^sub>P subst \ m" + unfolding ltl_prop_equiv_def ltl_prop_implies_def ltl_prop_equiv_subst_S by blast+ + +lemma subst_respects_ltl_prop_entailment_generalized: + "(\\. (\x. x \ S \ \ \\<^sub>P x) \ \ \\<^sub>P y) \ (\x. x \ S \ \ \\<^sub>P subst x m) \ \ \\<^sub>P subst y m" + unfolding ltl_prop_equiv_subst_S by simp + +lemma decomposable_function_subst: + "\f true = true; f false = false; \\ \. f (\ and \) = f \ and f \; \\ \. f (\ or \) = f \ or f \\ \ f \ = subst \ (\\. Some (f \))" + by (induction \) auto + +subsection \Additional Operators\ + +subsubsection \And\ + +lemma foldl_LTLAnd_prop_entailment: + "S \\<^sub>P foldl LTLAnd i xs = (S \\<^sub>P i \ (\y \ set xs. S \\<^sub>P y))" + by (induction xs arbitrary: i) auto + +fun And :: "'a ltl list \ 'a ltl" +where + "And [] = true" +| "And (x#xs) = foldl LTLAnd x xs" + +lemma And_prop_entailment: + "S \\<^sub>P And xs = (\x \ set xs. S \\<^sub>P x)" + using foldl_LTLAnd_prop_entailment by (cases xs) auto + +lemma And_propos: + "propos (And xs) = \{propos x| x. x \ set xs}" +proof (cases xs) + case Nil + thus ?thesis by simp +next + case (Cons x xs) + thus ?thesis + using propos_foldl[of LTLAnd x] by auto +qed + +lemma And_semantics: + "w \ And xs = (\x \ set xs. w \ x)" +proof - + have And_propos': "\x. x \ set xs \ propos x \ propos (And xs)" + using And_propos by blast + + have "w \ And xs = {\. \ \ propos (And xs) \ w \ \} \\<^sub>P (And xs)" + using ltl_models_equiv_prop_entailment LTL_prop_entailment_restrict_to_propos by blast + also + have "\ = (\x \ set xs. {\. \ \ propos (And xs) \ w \ \} \\<^sub>P x)" + using And_prop_entailment by auto + also + have "\ = (\x \ set xs. {\. \ \ propos x \ w \ \} \\<^sub>P x)" + using LTL_prop_entailment_restrict_to_propos And_propos' by blast + also + have "\ = (\x \ set xs. w \ x)" + using ltl_models_equiv_prop_entailment LTL_prop_entailment_restrict_to_propos by blast + finally + show ?thesis . +qed + +lemma And_append_syntactic: + "xs \ [] \ And (xs @ ys) = And ((And xs)#ys)" + by (induction xs rule: list_nonempty_induct) simp+ + +lemma And_append_S: + "S \\<^sub>P And (xs @ ys) = S \\<^sub>P And xs and And ys" + using And_prop_entailment[of S] by auto + +lemma And_append: + "And (xs @ ys) \\<^sub>P And xs and And ys" + unfolding ltl_prop_equiv_def using And_append_S by blast + +subsubsection \Lifted Variant\ + +lift_definition and_abs :: "'a ltl\<^sub>P \ 'a ltl\<^sub>P \ 'a ltl\<^sub>P" ("_ \and _") is "\x y. x and y" + unfolding ltl_prop_equiv_def by simp + +fun And_abs :: "'a ltl\<^sub>P list \ 'a ltl\<^sub>P" ("\And") +where + "\And xs = foldl and_abs (Abs true) xs" + +lemma foldl_LTLAnd_prop_entailment_abs: + "S \\\<^sub>P foldl and_abs i xs = (S \\\<^sub>P i \ (\y\set xs. S \\\<^sub>P y))" + by (induction xs arbitrary: i) + (simp_all add: and_abs_def ltl_prop_entails_abs.abs_eq, metis ltl_prop_entails_abs.rep_eq) + +lemma And_prop_entailment_abs: + "S \\\<^sub>P \And xs = (\x \ set xs. S \\\<^sub>P x)" + by (simp add: foldl_LTLAnd_prop_entailment_abs ltl_prop_entails_abs.abs_eq) + +lemma and_abs_conjunction: + "S \\\<^sub>P \ \and \ \ S \\\<^sub>P \ \ S \\\<^sub>P \" + by (metis and_abs.abs_eq ltl\<^sub>P_abs_rep ltl_prop_entailment.simps(3) ltl_prop_entails_abs.abs_eq) + +subsubsection \Or\ + +lemma foldl_LTLOr_prop_entailment: + "S \\<^sub>P foldl LTLOr i xs = (S \\<^sub>P i \ (\y \ set xs. S \\<^sub>P y))" + by (induction xs arbitrary: i) auto + +fun Or :: "'a ltl list \ 'a ltl" +where + "Or [] = false" +| "Or (x#xs) = foldl LTLOr x xs" + +lemma Or_prop_entailment: + "S \\<^sub>P Or xs = (\x \ set xs. S \\<^sub>P x)" + using foldl_LTLOr_prop_entailment by (cases xs) auto + +lemma Or_propos: + "propos (Or xs) = \{propos x| x. x \ set xs}" +proof (cases xs) + case Nil + thus ?thesis by simp +next + case (Cons x xs) + thus ?thesis + using propos_foldl[of LTLOr x] by auto +qed + +lemma Or_semantics: + "w \ Or xs = (\x \ set xs. w \ x)" +proof - + have Or_propos': "\x. x \ set xs \ propos x \ propos (Or xs)" + using Or_propos by blast + + have "w \ Or xs = {\. \ \ propos (Or xs) \ w \ \} \\<^sub>P (Or xs)" + using ltl_models_equiv_prop_entailment LTL_prop_entailment_restrict_to_propos by blast + also + have "\ = (\x \ set xs. {\. \ \ propos (Or xs) \ w \ \} \\<^sub>P x)" + using Or_prop_entailment by auto + also + have "\ = (\x \ set xs. {\. \ \ propos x \ w \ \} \\<^sub>P x)" + using LTL_prop_entailment_restrict_to_propos Or_propos' by blast + also + have "\ = (\x \ set xs. w \ x)" + using ltl_models_equiv_prop_entailment LTL_prop_entailment_restrict_to_propos by blast + finally + show ?thesis . +qed + +lemma Or_append_syntactic: + "xs \ [] \ Or (xs @ ys) = Or ((Or xs)#ys)" + by (induction xs rule: list_nonempty_induct) simp+ + +lemma Or_append_S: + "S \\<^sub>P Or (xs @ ys) = S \\<^sub>P Or xs or Or ys" + using Or_prop_entailment[of S] by auto + +lemma Or_append: + "Or (xs @ ys) \\<^sub>P Or xs or Or ys" + unfolding ltl_prop_equiv_def using Or_append_S by blast + +subsubsection \$eval_G$\ + +-- \Partly evaluate a formula by only considering G-subformulae\ + +fun eval\<^sub>G +where + "eval\<^sub>G S (\ and \) = eval\<^sub>G S \ and eval\<^sub>G S \" +| "eval\<^sub>G S (\ or \) = eval\<^sub>G S \ or eval\<^sub>G S \" +| "eval\<^sub>G S (G \) = (if G \ \ S then true else false)" +| "eval\<^sub>G S \ = \" + +-- \Syntactic Properties\ + +lemma eval\<^sub>G_And_map: + "eval\<^sub>G S (And xs) = And (map (eval\<^sub>G S) xs)" +proof (induction xs rule: rev_induct) + case (snoc x xs) + thus ?case + by (cases xs) simp+ +qed simp + +lemma eval\<^sub>G_Or_map: + "eval\<^sub>G S (Or xs) = Or (map (eval\<^sub>G S) xs)" +proof (induction xs rule: rev_induct) + case (snoc x xs) + thus ?case + by (cases xs) simp+ +qed simp + +lemma eval\<^sub>G_G_nested: + "\<^bold>G (eval\<^sub>G \ \) \ \<^bold>G \" + by (induction \) (simp_all, blast+) + +lemma eval\<^sub>G_subst: + "eval\<^sub>G S \ = subst \ (\\. Some (eval\<^sub>G S \))" + by (induction \) simp_all + +-- \Semantic Properties\ + +lemma eval\<^sub>G_prop_entailment: + "S \\<^sub>P eval\<^sub>G S \ \ S \\<^sub>P \" + by (induction \, auto) + +lemma eval\<^sub>G_respectfulness: + "\ \\<^sub>P \ \ eval\<^sub>G S \ \\<^sub>P eval\<^sub>G S \" + "\ \\<^sub>P \ \ eval\<^sub>G S \ \\<^sub>P eval\<^sub>G S \" + using subst_respects_ltl_prop_entailment eval\<^sub>G_subst by metis+ + +lemma eval\<^sub>G_respectfulness_generalized: + "(\\. (\x. x \ S \ \ \\<^sub>P x) \ \ \\<^sub>P y) \ (\x. x \ S \ \ \\<^sub>P eval\<^sub>G P x) \ \ \\<^sub>P eval\<^sub>G P y" + using subst_respects_ltl_prop_entailment_generalized[of S y \] eval\<^sub>G_subst[of P] by metis + +lift_definition eval\<^sub>G_abs :: "'a ltl set \ 'a ltl\<^sub>P \ 'a ltl\<^sub>P" ("\eval\<^sub>G") is eval\<^sub>G + by (insert eval\<^sub>G_respectfulness(2)) + +subsection \Finite Quotient Set\ + +text \If we restrict formulas to a finite set of propositions, the set of quotients of these is finite\ + +lemma Rep_Abs_prop_entailment[simp]: + "A \\<^sub>P Rep (Abs \) = A \\<^sub>P \" + using Quotient3_ltl_prop_equiv_quotient[THEN rep_abs_rsp] + by (auto simp add: ltl_prop_equiv_def) + +fun sat_models :: "'a ltl_prop_equiv_quotient \ 'a ltl set set" +where + "sat_models a = {A. A \\<^sub>P Rep(a)}" + +lemma sat_models_invariant: + "A \ sat_models (Abs \) = A \\<^sub>P \" + using Rep_Abs_prop_entailment by auto + +lemma sat_models_inj: + "inj sat_models" + using Quotient3_ltl_prop_equiv_quotient[THEN Quotient3_rel_rep] + by (auto simp add: ltl_prop_equiv_def inj_on_def) + +lemma sat_models_finite_image: + assumes "finite P" + shows "finite (sat_models ` {Abs \ | \. nested_propos \ \ P})" +proof - + have "sat_models (Abs \) = {A \ B | A B. A \ P \ A \\<^sub>P \ \ B \ UNIV - P}" (is "?lhs = ?rhs") + if "nested_propos \ \ P" for \ + proof + have "\A B. A \ sat_models (Abs \) \ A \ B \ sat_models (Abs \)" + unfolding sat_models_invariant by blast + moreover + have "{A | A. A \ P \ A \\<^sub>P \} \ sat_models (Abs \)" + using sat_models_invariant by fast + ultimately + show "?rhs \ ?lhs" + by blast + next + have "propos \ \ P" + using that propos_subset by blast + have "A \ {A \ B | A B. A \ P \ A \\<^sub>P \ \ B \ UNIV - P}" + if "A \ sat_models (Abs \)" for A + proof (standard, goal_cases prems) + case prems + then have "A \\<^sub>P \" + using that sat_models_invariant by blast + then obtain C D where "C = (A \ P)" and "D = A - P" and "A = C \ D" + by blast + then have "C \\<^sub>P \" and "C \ P" and "D \ UNIV - P" + using `A \\<^sub>P \` LTL_prop_entailment_restrict_to_propos `propos \ \ P` by blast+ + then have "C \ D \ {A \ B | A B. A \ P \ A \\<^sub>P \ \ B \ UNIV - P}" + by blast + thus ?case + using `A = C \ D` by simp + qed + thus "?lhs \ ?rhs" + by blast + qed + hence Equal: "{sat_models (Abs \) | \. nested_propos \ \ P} = {{A \ B | A B. A \ P \ A \\<^sub>P \ \ B \ UNIV - P} | \. nested_propos \ \ P}" + by (metis (lifting, no_types)) + + have Finite: "finite {{A \ B | A B. A \ P \ A \\<^sub>P \ \ B \ UNIV - P} | \. nested_propos \ \ P}" + proof - + let ?map = "\P S. {A \ B | A B. A \ S \ B \ UNIV - P}" + obtain S' where S'_def: "S' = {{A \ B | A B. A \ P \ A \\<^sub>P \ \ B \ UNIV - P} | \. nested_propos \ \ P}" + by blast + obtain S where S_def: "S = {{A | A. A \ P \ A \\<^sub>P \} | \. nested_propos \ \ P}" + by blast + + -- \Prove S and ?map applied to it is finite\ + hence "S \ Pow (Pow P)" + by blast + hence "finite S" + using `finite P` finite_Pow_iff infinite_super by fast + hence "finite {?map P A | A. A \ S}" + by fastforce + + -- \Prove that S' can be embedded into S using ?map\ + + have "S' \ {?map P A | A. A \ S}" + proof + fix A + assume "A \ S'" + then obtain \ where "nested_propos \ \ P" + and "A = {A \ B | A B. A \ P \ A \\<^sub>P \ \ B \ UNIV - P}" + using S'_def by blast + then have "?map P {A | A. A \ P \ A \\<^sub>P \} = A" + by simp + moreover + have "{A | A. A \ P \ A \\<^sub>P \} \ S" + using `nested_propos \ \ P` S_def by auto + ultimately + show "A \ {?map P A | A. A \ S}" + by blast + qed + + show ?thesis + using rev_finite_subset[OF `finite {?map P A | A. A \ S}` `S' \ {?map P A | A. A \ S}`] + unfolding S'_def . + qed + + have Finite2: "finite {sat_models (Abs \) | \. nested_propos \ \ P}" + unfolding Equal using Finite by blast + have Equal2: "sat_models ` {Abs \ | \. nested_propos \ \ P} = {sat_models (Abs \) | \. nested_propos \ \ P}" + by blast + + show ?thesis + unfolding Equal2 using Finite2 by blast +qed + +lemma ltl_prop_equiv_quotient_restricted_to_P_finite: + assumes "finite P" + shows "finite {Abs \ | \. nested_propos \ \ P}" +proof - + have "inj_on sat_models {Abs \ |\. nested_propos \ \ P}" + using sat_models_inj subset_inj_on by auto + thus ?thesis + using finite_imageD[OF sat_models_finite_image[OF assms]] by fast +qed + +locale lift_ltl_transformer = + fixes + f :: "'a ltl \ 'b \ 'a ltl" + assumes + respectfulness: "\ \\<^sub>P \ \ f \ \ \\<^sub>P f \ \" + assumes + nested_propos_bounded: "nested_propos (f \ \) \ nested_propos \" +begin + +lift_definition f_abs :: "'a ltl\<^sub>P \ 'b \ 'a ltl\<^sub>P" is f + using respectfulness . + +lift_definition f_foldl_abs :: "'a ltl\<^sub>P \ 'b list \ 'a ltl\<^sub>P" is "foldl f" +proof - + fix \ \ :: "'a ltl" fix w :: "'b list" assume "\ \\<^sub>P \" + thus "foldl f \ w \\<^sub>P foldl f \ w" + using respectfulness by (induction w arbitrary: \ \) simp+ +qed + +lemma f_foldl_abs_alt_def: + "f_foldl_abs (Abs \) w = foldl f_abs (Abs \) w" + by (induction w rule: rev_induct) (unfold f_foldl_abs.abs_eq foldl.simps foldl_append, (metis f_abs.abs_eq)+) + +definition abs_reach :: "'a ltl_prop_equiv_quotient \ 'a ltl_prop_equiv_quotient set" +where + "abs_reach \ = {foldl f_abs \ w |w. True}" + +lemma finite_abs_reach: + "finite (abs_reach (Abs \))" +proof - + { + fix w + have "nested_propos (foldl f \ w) \ nested_propos \" + by (induction w arbitrary: \) (simp, metis foldl_Cons nested_propos_bounded subset_trans) + } + hence "abs_reach (Abs \) \ {Abs \ | \. nested_propos \ \ nested_propos \}" + unfolding abs_reach_def f_foldl_abs_alt_def[symmetric] f_foldl_abs.abs_eq by blast + thus ?thesis + using ltl_prop_equiv_quotient_restricted_to_P_finite finite_propos + by (blast dest: finite_subset) +qed + +end + +end diff --git a/thys/LTL_to_DRA/af.thy b/thys/LTL_to_DRA/af.thy --- a/thys/LTL_to_DRA/af.thy +++ b/thys/LTL_to_DRA/af.thy @@ -1,655 +1,655 @@ (* Author: Salomon Sickert License: BSD *) section \af - Unfolding Functions\ theory af - imports Main "Aux/Words2" LTL + imports Main "Aux/Words2" LTL_FGXU begin subsection \af\ fun af_letter :: "'a ltl \ 'a set \ 'a ltl" where "af_letter true \ = true" | "af_letter false \ = false" | "af_letter p(a) \ = (if a \ \ then true else false)" | "af_letter (np(a)) \ = (if a \ \ then true else false)" | "af_letter (\ and \) \ = (af_letter \ \) and (af_letter \ \)" | "af_letter (\ or \) \ = (af_letter \ \) or (af_letter \ \)" | "af_letter (X \) \ = \" | "af_letter (G \) \ = (G \) and (af_letter \ \)" | "af_letter (F \) \ = (F \) or (af_letter \ \)" | "af_letter (\ U \) \ = (\ U \ and (af_letter \ \)) or (af_letter \ \)" abbreviation af :: "'a ltl \ 'a set list \ 'a ltl" ("af") where "af \ w \ foldl af_letter \ w" lemma af_decompose: "af (\ and \) w = (af \ w) and (af \ w)" "af (\ or \) w = (af \ w) or (af \ w)" by (induction w rule: rev_induct) simp_all lemma af_simps[simp]: "af true w = true" "af false w = false" "af (X \) (x#xs) = af \ (xs)" by (induction w) simp_all lemma af_F: "af (F \) w = Or (F \ # map (af \) (suffixes w))" proof (induction w) case (Cons x xs) have "af (F \) (x # xs) = af (af_letter (F \) x) xs" by simp also have "\ = (af (F \) xs) or (af (af_letter (\) x) xs)" unfolding af_decompose[symmetric] by simp finally show ?case using Cons Or_append_syntactic by force qed simp lemma af_G: "af (G \) w = And (G \ # map (af \) (suffixes w))" proof (induction w) case (Cons x xs) have "af (G \) (x # xs) = af (af_letter (G \) x) xs" by simp also have "\ = (af (G \) xs) and (af (af_letter (\) x) xs)" unfolding af_decompose[symmetric] by simp finally show ?case using Cons Or_append_syntactic by force qed simp lemma af_U: "af (\ U \) (x#xs) = (af (\ U \) xs and af \ (x#xs)) or af \ (x#xs)" by (induction xs) (simp add: af_decompose)+ lemma af_respectfulness: "\ \\<^sub>P \ \ af_letter \ \ \\<^sub>P af_letter \ \" "\ \\<^sub>P \ \ af_letter \ \ \\<^sub>P af_letter \ \" proof - { fix \ have "af_letter \ \ = subst \ (\\. Some (af_letter \ \))" by (induction \) auto } thus "\ \\<^sub>P \ \ af_letter \ \ \\<^sub>P af_letter \ \" and "\ \\<^sub>P \ \ af_letter \ \ \\<^sub>P af_letter \ \" using subst_respects_ltl_prop_entailment by metis+ qed lemma af_respectfulness': "\ \\<^sub>P \ \ af \ w \\<^sub>P af \ w" "\ \\<^sub>P \ \ af \ w \\<^sub>P af \ w" by (induction w arbitrary: \ \) (insert af_respectfulness, fastforce+) lemma af_nested_propos: "nested_propos (af_letter \ \) \ nested_propos \" by (induction \) auto subsection \@{term af\<^sub>G}\ fun af_G_letter :: "'a ltl \ 'a set \ 'a ltl" where "af_G_letter true \ = true" | "af_G_letter false \ = false" | "af_G_letter p(a) \ = (if a \ \ then true else false)" | "af_G_letter (np(a)) \ = (if a \ \ then true else false)" | "af_G_letter (\ and \) \ = (af_G_letter \ \) and (af_G_letter \ \)" | "af_G_letter (\ or \) \ = (af_G_letter \ \) or (af_G_letter \ \)" | "af_G_letter (X \) \ = \" | "af_G_letter (G \) \ = (G \)" | "af_G_letter (F \) \ = (F \) or (af_G_letter \ \)" | "af_G_letter (\ U \) \ = (\ U \ and (af_G_letter \ \)) or (af_G_letter \ \)" abbreviation af\<^sub>G :: "'a ltl \ 'a set list \ 'a ltl" where "af\<^sub>G \ w \ (foldl af_G_letter \ w)" lemma af\<^sub>G_decompose: "af\<^sub>G (\ and \) w = (af\<^sub>G \ w) and (af\<^sub>G \ w)" "af\<^sub>G (\ or \) w = (af\<^sub>G \ w) or (af\<^sub>G \ w)" by (induction w rule: rev_induct) simp_all lemma af\<^sub>G_simps[simp]: "af\<^sub>G true w = true" "af\<^sub>G false w = false" "af\<^sub>G (G \) w = G \" "af\<^sub>G (X \) (x#xs) = af\<^sub>G \ (xs)" by (induction w) simp_all lemma af\<^sub>G_F: "af\<^sub>G (F \) w = Or (F \ # map (af\<^sub>G \) (suffixes w))" proof (induction w) case (Cons x xs) have "af\<^sub>G (F \) (x # xs) = af\<^sub>G (af_G_letter (F \) x) xs" by simp also have "\ = (af\<^sub>G (F \) xs) or (af\<^sub>G (af_G_letter (\) x) xs)" unfolding af\<^sub>G_decompose[symmetric] by simp finally show ?case using Cons Or_append_syntactic by force qed simp lemma af\<^sub>G_U: "af\<^sub>G (\ U \) (x#xs) = (af\<^sub>G (\ U \) xs and af\<^sub>G \ (x#xs)) or af\<^sub>G \ (x#xs)" by (simp add: af\<^sub>G_decompose) lemma af\<^sub>G_subsequence_U: "af\<^sub>G (\ U \) (w [0 \ Suc n]) = (af\<^sub>G (\ U \) (w [1 \ Suc n]) and af\<^sub>G \ (w [0 \ Suc n])) or af\<^sub>G \ (w [0 \ Suc n])" proof - have "\n. w [0 \ Suc n] = w 0 # w [1 \ Suc n]" using subsequence_append[of w 1] by (simp add: subsequence_def upt_conv_Cons) thus ?thesis using af\<^sub>G_U by metis qed lemma af_G_respectfulness: "\ \\<^sub>P \ \ af_G_letter \ \ \\<^sub>P af_G_letter \ \" "\ \\<^sub>P \ \ af_G_letter \ \ \\<^sub>P af_G_letter \ \" proof - { fix \ have "af_G_letter \ \ = subst \ (\\. Some (af_G_letter \ \))" by (induction \) auto } thus "\ \\<^sub>P \ \ af_G_letter \ \ \\<^sub>P af_G_letter \ \" and "\ \\<^sub>P \ \ af_G_letter \ \ \\<^sub>P af_G_letter \ \" using subst_respects_ltl_prop_entailment by metis+ qed lemma af_G_respectfulness': "\ \\<^sub>P \ \ af\<^sub>G \ w \\<^sub>P af\<^sub>G \ w" "\ \\<^sub>P \ \ af\<^sub>G \ w \\<^sub>P af\<^sub>G \ w" by (induction w arbitrary: \ \) (insert af_G_respectfulness, fastforce+) lemma af_G_nested_propos: "nested_propos (af_G_letter \ \) \ nested_propos \" by (induction \) auto lemma af_G_letter_sat_core: "Only_G \ \ \ \\<^sub>P \ \ \ \\<^sub>P af_G_letter \ \" by (induction \) (simp_all, blast+) lemma af\<^sub>G_sat_core: "Only_G \ \ \ \\<^sub>P \ \ \ \\<^sub>P af\<^sub>G \ w" using af_G_letter_sat_core by (induction w rule: rev_induct) (simp_all, blast) lemma af\<^sub>G_sat_core_generalized: "Only_G \ \ i \ j \ \ \\<^sub>P af\<^sub>G \ (w [0 \ i]) \ \ \\<^sub>P af\<^sub>G \ (w [0 \ j])" by (metis af\<^sub>G_sat_core foldl_append subsequence_append le_add_diff_inverse) lemma af\<^sub>G_eval\<^sub>G: "Only_G \ \ \ \\<^sub>P af\<^sub>G (eval\<^sub>G \ \) w \ \ \\<^sub>P eval\<^sub>G \ (af\<^sub>G \ w)" by (induction \) (simp_all add: eval\<^sub>G_prop_entailment af\<^sub>G_decompose) lemma af\<^sub>G_keeps_F_and_S: assumes "ys \ []" assumes "S \\<^sub>P af\<^sub>G \ ys" shows "S \\<^sub>P af\<^sub>G (F \) (xs @ ys)" proof - have "af\<^sub>G \ ys \ set (map (af\<^sub>G \) (suffixes (xs @ ys)))" using assms(1) unfolding suffixes_append map_append by (induction ys rule: List.list_nonempty_induct) auto thus ?thesis unfolding af\<^sub>G_F Or_prop_entailment using assms(2) by force qed subsection \G-Subformulae Simplification\ lemma G_af_simp[simp]: "\<^bold>G (af \ w) = \<^bold>G \" proof - { fix \ \ have "\<^bold>G (af_letter \ \) = \<^bold>G \" by (induction \) auto } thus ?thesis by (induction w arbitrary: \ rule: rev_induct) fastforce+ qed lemma G_af\<^sub>G_simp[simp]: "\<^bold>G (af\<^sub>G \ w) = \<^bold>G \" proof - { fix \ \ have "\<^bold>G (af_G_letter \ \) = \<^bold>G \" by (induction \) auto } thus ?thesis by (induction w arbitrary: \ rule: rev_induct) fastforce+ qed subsection \Relation between af and $af_G$\ lemma af_G_letter_free_F: "\<^bold>G \ = {} \ \<^bold>G (af_letter \ \) = {}" "\<^bold>G \ = {} \ \<^bold>G (af_G_letter \ \) = {}" by (induction \) auto lemma af_G_free: assumes "\<^bold>G \ = {}" shows "af \ w = af\<^sub>G \ w" using assms proof (induction w arbitrary: \) case (Cons x xs) hence "af (af_letter \ x) xs = af\<^sub>G (af_letter \ x) xs" using af_G_letter_free_F[OF Cons.prems, THEN Cons.IH] by blast moreover have "af_letter \ x = af_G_letter \ x" using Cons.prems by (induction \) auto ultimately show ?case by simp qed simp lemma af_equals_af\<^sub>G_base_cases: "af true w = af\<^sub>G true w" "af false w = af\<^sub>G false w" "af p(a) w = af\<^sub>G p(a) w" "af (np(a)) w = af\<^sub>G (np(a)) w" by (auto intro: af_G_free) lemma af_implies_af\<^sub>G: "S \\<^sub>P af \ w \ S \\<^sub>P af\<^sub>G \ w" proof (induction w arbitrary: S rule: rev_induct) case (snoc x xs) hence "S \\<^sub>P af_letter (af \ xs) x" by simp hence "S \\<^sub>P af_letter (af\<^sub>G \ xs) x" using af_respectfulness(1) snoc.IH unfolding ltl_prop_implies_def by blast moreover { fix \ have "\\. S \\<^sub>P af_letter \ \ \ S \\<^sub>P af_G_letter \ \" by (induction \) auto } ultimately show ?case using snoc.prems foldl_append by simp qed simp lemma af_implies_af\<^sub>G_2: "w \ af \ xs \ w \ af\<^sub>G \ xs" by (metis ltl_prop_implication_implies_ltl_implication af_implies_af\<^sub>G ltl_prop_implies_def) lemma af\<^sub>G_implies_af_eval\<^sub>G': assumes "S \\<^sub>P eval\<^sub>G \ (af\<^sub>G \ w)" assumes "\\. G \ \ \ \ S \\<^sub>P G \" assumes "\\ i. G \ \ \ \ i < length w \ S \\<^sub>P eval\<^sub>G \ (af\<^sub>G \ (drop i w))" shows "S \\<^sub>P af \ w" using assms proof (induction \ arbitrary: w) case (LTLGlobal \) hence "G \ \ \" unfolding af\<^sub>G_simps eval\<^sub>G.simps by (cases "G \ \ \") simp+ hence "S \\<^sub>P G \" using LTLGlobal by simp moreover { fix x assume "x \ set (map (af \) (suffixes w))" then obtain w' where "x = af \ w'" and "w' \ set (suffixes w)" by auto then obtain i where "w' = drop i w" and "i < length w" by (auto simp add: suffixes_alt_def subsequence_def) hence "S \\<^sub>P eval\<^sub>G \ (af\<^sub>G \ w')" using LTLGlobal.prems `G \ \ \` by simp hence "S \\<^sub>P x" using LTLGlobal(1)[OF `S \\<^sub>P eval\<^sub>G \ (af\<^sub>G \ w')`] LTLGlobal(3-4) drop_drop unfolding `x = af \ w'` `w' = drop i w` by simp } ultimately show ?case unfolding af_G eval\<^sub>G_And_map And_prop_entailment by simp next case (LTLFinal \) then obtain x where x_def: "x \ set (F \ # map (eval\<^sub>G \ o af\<^sub>G \) (suffixes w))" and "S \\<^sub>P x" unfolding Or_prop_entailment af\<^sub>G_F eval\<^sub>G_Or_map by force hence "\y \ set (F \ # map (af \) (suffixes w)). S \\<^sub>P y" proof (cases "x \ F \") case True then obtain w' where "S \\<^sub>P eval\<^sub>G \ (af\<^sub>G \ w')" and "w' \ set (suffixes w)" using x_def `S \\<^sub>P x` by auto hence "\\ i. G \ \ \ \ i < length w' \ S \\<^sub>P eval\<^sub>G \ (af\<^sub>G \ (drop i w'))" using LTLFinal.prems by (auto simp add: suffixes_alt_def subsequence_def) moreover have "\\. G \ \ \ \ S \\<^sub>P eval\<^sub>G \ (G \)" using LTLFinal by simp ultimately have "S \\<^sub>P af \ w'" using LTLFinal.IH[OF `S \\<^sub>P eval\<^sub>G \ (af\<^sub>G \ w')`] using assms(2) by blast thus ?thesis using `w' \ set (suffixes w)` by auto qed simp thus ?case unfolding af_F Or_prop_entailment eval\<^sub>G_Or_map by simp next case (LTLNext \) thus ?case proof (cases w) case (Cons x xs) { fix \ i assume "G \ \ \" and "Suc i < length (x#xs)" hence "S \\<^sub>P eval\<^sub>G \ (af\<^sub>G \ (drop (Suc i) (x#xs)))" using LTLNext.prems unfolding Cons by blast hence "S \\<^sub>P eval\<^sub>G \ (af\<^sub>G \ (drop i xs))" by simp } hence "\\ i. G \ \ \ \ i < length xs \ S \\<^sub>P eval\<^sub>G \ (af\<^sub>G \ (drop i xs))" by simp thus ?thesis using LTLNext Cons by simp qed simp next case (LTLUntil \ \) thus ?case proof (induction w) case (Cons x xs) { assume "S \\<^sub>P eval\<^sub>G \ (af\<^sub>G \ (x # xs))" moreover have "\\ i. G \ \ \ \ i < length (x#xs) \ S \\<^sub>P eval\<^sub>G \ (af\<^sub>G \ (drop i (x#xs)))" using Cons by simp ultimately have "S \\<^sub>P af \ (x # xs)" using Cons.prems by blast hence ?case unfolding af_U by simp } moreover { assume "S \\<^sub>P eval\<^sub>G \ (af\<^sub>G (\ U \) xs)" and "S \\<^sub>P eval\<^sub>G \ (af\<^sub>G \ (x # xs))" moreover have "\\ i. G \ \ \ \ i < length (x#xs) \ S \\<^sub>P eval\<^sub>G \ (af\<^sub>G \ (drop i (x#xs)))" using Cons by simp ultimately have "S \\<^sub>P af \ (x # xs)" and "S \\<^sub>P af (\ U \) xs" using Cons by (blast, force) hence ?case unfolding af_U by simp } ultimately show ?case using Cons(4) unfolding af\<^sub>G_U by auto qed simp next case (LTLProp a) thus ?case proof (cases w) case (Cons x xs) thus ?thesis using LTLProp by (cases "a \ x") simp+ qed simp next case (LTLPropNeg a) thus ?case proof (cases w) case (Cons x xs) thus ?thesis using LTLPropNeg by (cases "a \ x") simp+ qed simp qed (unfold af_equals_af\<^sub>G_base_cases af\<^sub>G_decompose af_decompose, auto) lemma af\<^sub>G_implies_af_eval\<^sub>G: assumes "S \\<^sub>P eval\<^sub>G \ (af\<^sub>G \ (w [0\j]))" assumes "\\. G \ \ \ \ S \\<^sub>P G \" assumes "\\ i. G \ \ \ \ i \ j \ S \\<^sub>P eval\<^sub>G \ (af\<^sub>G \ (w [i \ j]))" shows "S \\<^sub>P af \ (w [0\j])" using af\<^sub>G_implies_af_eval\<^sub>G'[OF assms(1-2), unfolded subsequence_length subsequence_drop] assms(3) by force subsection \Continuation\ -- \af fulfills the infinite continuation w' of a word after skipping some finite prefix w. Corresponds to Lemma 7 in arXiv: 1402.3388v2\ lemma af_ltl_continuation: "(w \ w') \ \ \ w' \ af \ w" proof (induction w arbitrary: \ w') case (Cons x xs) have "((x # xs) \ w') 0 = x" unfolding conc_def nth_Cons_0 by simp moreover have "suffix 1 ((x # xs) \ w') = xs \ w'" unfolding suffix_def conc_def by fastforce moreover { fix \ :: "'a ltl" have "\w. w \ \ \ suffix 1 w \ af_letter \ (w 0)" by (induction \) ((unfold LTL_F_one_step_unfolding LTL_G_one_step_unfolding LTL_U_one_step_unfolding)?, auto) } ultimately have "((x # xs) \ w') \ \ \ (xs \ w') \ af_letter \ x" by metis also have "\ \ w' \ af \ (x#xs)" using Cons.IH by simp finally show ?case . qed simp lemma af_ltl_continuation_suffix: "w \ \ \ suffix i w \ af \ (w[0 \ i])" using af_ltl_continuation prefix_suffix subsequence_def by metis lemma af_G_ltl_continuation: "\\ \ \<^bold>G \. w' \ \ = (w \ w') \ \ \ (w \ w') \ \ \ w' \ af\<^sub>G \ w" proof (induction w arbitrary: w' \) case (Cons x xs) { fix \ :: "'a ltl" fix w w' w'' assume "w'' \ G \ = ((w @ w') \ w'') \ G \" hence "w'' \ G \ = (w' \ w'') \ G \" and "(w' \ w'') \ G \ = ((w @ w') \ w'') \ G \" by (induction w' arbitrary: w) (metis LTL_suffix_G suffix_conc_length conc_conc)+ } note G_stable = this have A: "\\\\<^bold>G (af\<^sub>G \ [x]). w' \ \ = (xs \ w') \ \" using G_stable(1)[of w' _ "[x]"] Cons.prems unfolding G_af\<^sub>G_simp conc_conc append.simps unfolding G_nested_propos_alt_def by blast have B: "\\\\<^bold>G \. ([x] \ xs \ w') \ \ = (xs \ w') \ \" using G_stable(2)[of w' _ "[x]"] Cons.prems unfolding conc_conc append.simps unfolding G_nested_propos_alt_def by blast hence "([x] \ xs \ w') \ \ = (xs \ w') \ af\<^sub>G \ [x]" proof (induction \) case (LTLFinal \) thus ?case unfolding LTL_F_one_step_unfolding by (auto simp add: suffix_conc_length[of "[x]", simplified]) next case (LTLUntil \ \) thus ?case unfolding LTL_U_one_step_unfolding by (auto simp add: suffix_conc_length[of "[x]", simplified]) qed (auto simp add: conc_fst[of 0 "[x]"] suffix_conc_length[of "[x]", simplified]) also have "... = w' \ af\<^sub>G \ (x # xs)" using Cons.IH[of "af\<^sub>G \ [x]" w'] A by simp finally show ?case unfolding conc_conc by simp qed simp lemma af\<^sub>G_ltl_continuation_suffix: "\\ \ \<^bold>G \. w \ \ = (suffix i w) \ \ \ w \ \ \ suffix i w \ af\<^sub>G \ (w [0 \ i])" by (metis af_G_ltl_continuation[of \ "suffix i w"] prefix_suffix subsequence_def) subsection \Eager Unfolding @{term af} and @{term af\<^sub>G}\ fun Unf :: "'a ltl \ 'a ltl" where "Unf (F \) = F \ or Unf \" | "Unf (G \) = G \ and Unf \" | "Unf (\ U \) = (\ U \ and Unf \) or Unf \" | "Unf (\ and \) = Unf \ and Unf \" | "Unf (\ or \) = Unf \ or Unf \" | "Unf \ = \" fun Unf\<^sub>G :: "'a ltl \ 'a ltl" where "Unf\<^sub>G (F \) = F \ or Unf\<^sub>G \" | "Unf\<^sub>G (G \) = G \" | "Unf\<^sub>G (\ U \) = (\ U \ and Unf\<^sub>G \) or Unf\<^sub>G \" | "Unf\<^sub>G (\ and \) = Unf\<^sub>G \ and Unf\<^sub>G \" | "Unf\<^sub>G (\ or \) = Unf\<^sub>G \ or Unf\<^sub>G \" | "Unf\<^sub>G \ = \" fun step :: "'a ltl \ 'a set \ 'a ltl" where "step p(a) \ = (if a \ \ then true else false)" | "step (np(a)) \ = (if a \ \ then true else false)" | "step (X \) \ = \" | "step (\ and \) \ = step \ \ and step \ \" | "step (\ or \) \ = step \ \ or step \ \" | "step \ \ = \" fun af_letter_opt where "af_letter_opt \ \ = Unf (step \ \)" fun af_G_letter_opt where "af_G_letter_opt \ \ = Unf\<^sub>G (step \ \)" abbreviation af_opt :: "'a ltl \ 'a set list \ 'a ltl" ("af\<^sub>\") where "af\<^sub>\ \ w \ (foldl af_letter_opt \ w)" abbreviation af_G_opt :: "'a ltl \ 'a set list \ 'a ltl" ("af\<^sub>G\<^sub>\") where "af\<^sub>G\<^sub>\ \ w \ (foldl af_G_letter_opt \ w)" lemma af_letter_alt_def: "af_letter \ \ = step (Unf \) \" "af_G_letter \ \ = step (Unf\<^sub>G \) \" by (induction \) simp_all lemma af_to_af_opt: "Unf (af \ w) = af\<^sub>\ (Unf \) w" "Unf\<^sub>G (af\<^sub>G \ w) = af\<^sub>G\<^sub>\ (Unf\<^sub>G \) w" by (induction w arbitrary: \) (simp_all add: af_letter_alt_def) lemma af_equiv: "af \ (w @ [\]) = step (af\<^sub>\ (Unf \) w) \" using af_to_af_opt(1) by (metis af_letter_alt_def(1) foldl_Cons foldl_Nil foldl_append) lemma af_equiv': "af \ (w [0 \ Suc i]) = step (af\<^sub>\ (Unf \) (w [0 \ i])) (w i)" using af_equiv unfolding subsequence_def by auto subsection \Lifted Functions\ lemma respectfulness: "\ \\<^sub>P \ \ af_letter_opt \ \ \\<^sub>P af_letter_opt \ \" "\ \\<^sub>P \ \ af_letter_opt \ \ \\<^sub>P af_letter_opt \ \" "\ \\<^sub>P \ \ af_G_letter_opt \ \ \\<^sub>P af_G_letter_opt \ \" "\ \\<^sub>P \ \ af_G_letter_opt \ \ \\<^sub>P af_G_letter_opt \ \" "\ \\<^sub>P \ \ step \ \ \\<^sub>P step \ \" "\ \\<^sub>P \ \ step \ \ \\<^sub>P step \ \" "\ \\<^sub>P \ \ Unf \ \\<^sub>P Unf \" "\ \\<^sub>P \ \ Unf \ \\<^sub>P Unf \" "\ \\<^sub>P \ \ Unf\<^sub>G \ \\<^sub>P Unf\<^sub>G \" "\ \\<^sub>P \ \ Unf\<^sub>G \ \\<^sub>P Unf\<^sub>G \" using decomposable_function_subst[of "\\. af_letter_opt \ \", simplified] af_letter_opt.simps using decomposable_function_subst[of "\\. af_G_letter_opt \ \", simplified] af_G_letter_opt.simps using decomposable_function_subst[of "\\. step \ \", simplified] using decomposable_function_subst[of Unf, simplified] using decomposable_function_subst[of Unf\<^sub>G, simplified] using subst_respects_ltl_prop_entailment by metis+ lemma nested_propos: "nested_propos (step \ \) \ nested_propos \" "nested_propos (Unf \) \ nested_propos \" "nested_propos (Unf\<^sub>G \) \ nested_propos \" "nested_propos (af_letter_opt \ \) \ nested_propos \" "nested_propos (af_G_letter_opt \ \) \ nested_propos \" by (induction \) auto text \Lift functions and bind to new names\ interpretation af_abs: lift_ltl_transformer af_letter using lift_ltl_transformer_def af_respectfulness af_nested_propos by blast definition af_letter_abs ("\af") where "\af \ af_abs.f_abs" interpretation af_G_abs: lift_ltl_transformer af_G_letter using lift_ltl_transformer_def af_G_respectfulness af_G_nested_propos by blast definition af_G_letter_abs ("\af\<^sub>G") where "\af\<^sub>G \ af_G_abs.f_abs" interpretation af_abs_opt: lift_ltl_transformer af_letter_opt using lift_ltl_transformer_def respectfulness nested_propos by blast definition af_letter_abs_opt ("\af\<^sub>\") where "\af\<^sub>\ \ af_abs_opt.f_abs" interpretation af_G_abs_opt: lift_ltl_transformer af_G_letter_opt using lift_ltl_transformer_def respectfulness nested_propos by blast definition af_G_letter_abs_opt ("\af\<^sub>G\<^sub>\") where "\af\<^sub>G\<^sub>\ \ af_G_abs_opt.f_abs" lift_definition step_abs :: "'a ltl\<^sub>P \ 'a set \ 'a ltl\<^sub>P" ("\step") is step by (insert respectfulness) lift_definition Unf_abs :: "'a ltl\<^sub>P \ 'a ltl\<^sub>P" ("\Unf") is Unf by (insert respectfulness) lift_definition Unf\<^sub>G_abs :: "'a ltl\<^sub>P \ 'a ltl\<^sub>P" ("\Unf\<^sub>G") is Unf\<^sub>G by (insert respectfulness) subsubsection \Properties\ lemma af_G_letter_opt_sat_core: "Only_G \ \ \ \\<^sub>P \ \ \ \\<^sub>P af_G_letter_opt \ \" by (induction \) auto lemma af_G_letter_sat_core_lifted: "Only_G \ \ \ \\<^sub>P Rep \ \ \ \\<^sub>P Rep (af_G_letter_abs \ \)" by (metis af_G_letter_sat_core Quotient_ltl_prop_equiv_quotient[THEN Quotient_rep_abs] Quotient3_ltl_prop_equiv_quotient[THEN Quotient3_abs_rep] af_G_abs.f_abs.abs_eq ltl_prop_equiv_def af_G_letter_abs_def) lemma af_G_letter_opt_sat_core_lifted: "Only_G \ \ \ \\<^sub>P Rep \ \ \ \\<^sub>P Rep (\af\<^sub>G\<^sub>\ \ \)" unfolding af_G_letter_abs_opt_def by (metis af_G_letter_opt_sat_core Quotient_ltl_prop_equiv_quotient[THEN Quotient_rep_abs] Quotient3_ltl_prop_equiv_quotient[THEN Quotient3_abs_rep] af_G_abs_opt.f_abs.abs_eq ltl_prop_equiv_def) lemma af_G_letter_abs_opt_split: "\Unf\<^sub>G (\step \ \) = \af\<^sub>G\<^sub>\ \ \" unfolding af_G_letter_abs_opt_def step_abs_def comp_def af_G_abs_opt.f_abs_def using map_fun_apply Unf\<^sub>G_abs.abs_eq af_G_letter_opt.simps by auto lemma af_unfold: "\af = (\\ \. \step (\Unf \) \)" by (metis Unf_abs_def af_abs.f_abs.abs_eq af_letter_abs_def af_letter_alt_def(1) ltl\<^sub>P_abs_rep map_fun_apply step_abs.abs_eq) lemma af_opt_unfold: "\af\<^sub>\ = (\\ \. \Unf (\step \ \))" by (metis (no_types, lifting) Quotient3_abs_rep Quotient3_ltl_prop_equiv_quotient Unf_abs.abs_eq af_abs_opt.f_abs.abs_eq af_letter_abs_opt_def af_letter_opt.elims id_apply map_fun_apply step_abs_def) lemma af_abs_equiv: "foldl \af \ (xs @ [x]) = \step (foldl \af\<^sub>\ (\Unf \) xs) x" unfolding af_unfold af_opt_unfold by (induction xs arbitrary: x \ rule: rev_induct) simp+ lemma Rep_Abs_equiv: "Rep (Abs \) \\<^sub>P \" using Rep_Abs_prop_entailment unfolding ltl_prop_equiv_def by auto lemma Rep_step: "Rep (\step \ \) \\<^sub>P step (Rep \) \" by (metis Quotient3_abs_rep Quotient3_ltl_prop_equiv_quotient ltl_prop_equiv_quotient.abs_eq_iff step_abs.abs_eq) lemma step_\: "Only_G \ \ \ \\<^sub>P \ \ \ \\<^sub>P step \ \" by (induction \) auto lemma Unf\<^sub>G_\: "Only_G \ \ \ \\<^sub>P \ \ \ \\<^sub>P Unf\<^sub>G \" by (induction \) auto hide_fact (open) respectfulness nested_propos end diff --git a/thys/LTL_to_GBA/All_Of_LTL_to_GBA.thy b/thys/LTL_to_GBA/All_Of_LTL_to_GBA.thy --- a/thys/LTL_to_GBA/All_Of_LTL_to_GBA.thy +++ b/thys/LTL_to_GBA/All_Of_LTL_to_GBA.thy @@ -1,6 +1,6 @@ (*<*) theory All_Of_LTL_to_GBA -imports LTL LTL_Rewrite LTL_Stutter LTL_to_GBA LTL_to_GBA_impl +imports LTL_Stutter LTL_to_GBA LTL_to_GBA_impl begin end (*>*) diff --git a/thys/LTL_to_GBA/LTL.thy b/thys/LTL_to_GBA/LTL.thy deleted file mode 100644 --- a/thys/LTL_to_GBA/LTL.thy +++ /dev/null @@ -1,1335 +0,0 @@ -section \Linear Temporal Logic\ -(* Author: Alexander Schimpf *) -theory LTL -imports - "~~/src/HOL/Library/Omega_Words_Fun" "../Automatic_Refinement/Lib/Refine_Util" -begin - -subsection "LTL formulas" - -subsubsection \Syntax\ - -datatype (plugins del: size) - 'a ltl = LTLTrue - | LTLFalse - | LTLProp 'a - | LTLNeg "'a ltl" - | LTLAnd "'a ltl" "'a ltl" - | LTLOr "'a ltl" "'a ltl" - | LTLNext "'a ltl" - | LTLUntil "'a ltl" "'a ltl" - | LTLRelease "'a ltl" "'a ltl" - -instantiation ltl :: (type) size -begin - -text \The new datatype package would give a size of 1 to - @{const LTLProp}, which breaks some of the proofs below.\ -primrec size_ltl :: "'a ltl \ nat" where - "size LTLTrue = 0" - | "size LTLFalse = 0" - | "size (LTLProp _) = 0" - | "size (LTLNeg \) = size \ + 1" - | "size (LTLAnd \ \) = size \ + size \ + 1" - | "size (LTLOr \ \) = size \ + size \ + 1" - | "size (LTLNext \) = size \ + 1" - | "size (LTLUntil \ \) = size \ + size \ + 1" - | "size (LTLRelease \ \) = size \ + size \ + 1" - -instance .. - -end - -text \The following locale defines syntactic sugar for - parsing and printing LTL formulas in Isabelle\ -locale LTL_Syntax begin -notation - LTLTrue ("true") - and LTLFalse ("false") - and LTLProp ("prop'(_')") - and LTLNeg ("not _" [85] 85) - and LTLAnd ("_ and _" [82,82] 81) - and LTLOr ("_ or _" [81,81] 80) - and LTLNext ("X _" [88] 87) - and LTLUntil ("_ U _" [84,84] 83) - and LTLRelease ("_ V _" [83,83] 82) -end - -subsubsection \Semantics\ -text \We first provide an abstract semantics, that is parameterized with - the semantics of atomic propositions\ - -context begin interpretation LTL_Syntax . - -primrec ltl_semantics :: "'ap set word \ 'ap ltl \ bool" - ("_ \ _" [80,80] 80) - where - "\ \ true = True" - | "\ \ false = False" - | "\ \ prop(q) = (q \ (\ 0))" - | "\ \ not \ = (\ \ \ \)" - | "\ \ \ and \ = (\ \ \ \ \ \ \)" - | "\ \ \ or \ = (\ \ \ \ \ \ \)" - | "\ \ X \ = (suffix 1 \ \ \)" - | "\ \ \ U \ = (\i. suffix i \ \ \ \ (\j \ \))" - | "\ \ \ V \ = (\i. suffix i \ \ \ \ (\j \ \))" - -definition "ltl_language \ \ {\. \ \ \}" -end - -subsubsection \Explicit Syntactic Sugar\ -text \In this section, we provide a formulation of LTL with - explicit syntactic sugar deeply embedded. This formalization - serves as a reference semantics. -\ -datatype (ltlc_aprops: 'a) - ltlc = LTLcTrue - | LTLcFalse - | LTLcProp 'a - | LTLcNeg "'a ltlc" - | LTLcAnd "'a ltlc" "'a ltlc" - | LTLcOr "'a ltlc" "'a ltlc" - | LTLcImplies "'a ltlc" "'a ltlc" - | LTLcIff "'a ltlc" "'a ltlc" - | LTLcNext "'a ltlc" - | LTLcFinal "'a ltlc" - | LTLcGlobal "'a ltlc" - | LTLcUntil "'a ltlc" "'a ltlc" - | LTLcRelease "'a ltlc" "'a ltlc" - -context LTL_Syntax begin - notation - LTLcTrue ("true\<^sub>c") - and LTLcFalse ("false\<^sub>c") - and LTLcProp ("prop\<^sub>c'(_')") - and LTLcNeg ("not\<^sub>c _" [85] 85) - and LTLcAnd ("_ and\<^sub>c _" [82,82] 81) - and LTLcOr ("_ or\<^sub>c _" [81,81] 80) - and LTLcImplies ("_ implies\<^sub>c _" [81,81] 80) - and LTLcIff (" _ iff\<^sub>c _" [81,81] 80) - and LTLcNext ("X\<^sub>c _" [88] 87) - and LTLcFinal ("F\<^sub>c _" [88] 87) - and LTLcGlobal ("G\<^sub>c _" [88] 87) - and LTLcUntil ("_ U\<^sub>c _" [84,84] 83) - and LTLcRelease ("_ V\<^sub>c _" [83,83] 82) -end - -context begin interpretation LTL_Syntax . - -primrec ltlc_semantics - :: "['a set word, 'a ltlc] \ bool" ("_ \\<^sub>c _" [80,80] 80) -where - "\ \\<^sub>c true\<^sub>c = True" - | "\ \\<^sub>c false\<^sub>c = False" - | "\ \\<^sub>c prop\<^sub>c(q) = (q\\ 0)" - | "\ \\<^sub>c not\<^sub>c \ = (\ \ \\<^sub>c \)" - | "\ \\<^sub>c \ and\<^sub>c \ = (\ \\<^sub>c \ \ \ \\<^sub>c \)" - | "\ \\<^sub>c \ or\<^sub>c \ = (\ \\<^sub>c \ \ \ \\<^sub>c \)" - | "\ \\<^sub>c \ implies\<^sub>c \ = (\ \\<^sub>c \ \ \ \\<^sub>c \)" - | "\ \\<^sub>c \ iff\<^sub>c \ = (\ \\<^sub>c \ \ \ \\<^sub>c \)" - | "\ \\<^sub>c X\<^sub>c \ = (suffix 1 \ \\<^sub>c \)" - | "\ \\<^sub>c F\<^sub>c \ = (\i. suffix i \ \\<^sub>c \)" - | "\ \\<^sub>c G\<^sub>c \ = (\i. suffix i \ \\<^sub>c \)" - | "\ \\<^sub>c \ U\<^sub>c \ = (\i. suffix i \ \\<^sub>c \ \ (\j \\<^sub>c \))" - | "\ \\<^sub>c \ V\<^sub>c \ = (\i. suffix i \ \\<^sub>c \ \ (\j \\<^sub>c \))" - -definition "ltlc_language \ \ {\. \ \\<^sub>c \}" - -lemma ltlc_language_negate[simp]: - "ltlc_language (not\<^sub>c \) = - ltlc_language \" -unfolding ltlc_language_def -by auto - -lemma ltlc_semantics_sugar: - "\ \\<^sub>c \ implies\<^sub>c \ = \ \\<^sub>c (not\<^sub>c \ or\<^sub>c \)" - "\ \\<^sub>c \ iff\<^sub>c \ = \ \\<^sub>c ((not\<^sub>c \ or\<^sub>c \) and\<^sub>c (not\<^sub>c \ or\<^sub>c \))" - "\ \\<^sub>c F\<^sub>c \ = \ \\<^sub>c (true\<^sub>c U\<^sub>c \)" - "\ \\<^sub>c G\<^sub>c \ = \ \\<^sub>c (false\<^sub>c V\<^sub>c \)" -by auto - -definition "pw_eq_on S w w' \ \i. w i \ S = w' i \ S" - -lemma - pw_eq_on_refl[simp]: "pw_eq_on S w w" - and pw_eq_on_sym: "pw_eq_on S w w' \ pw_eq_on S w' w" - and pw_eq_on_trans[trans]: - "\pw_eq_on S w w'; pw_eq_on S w' w''\ \ pw_eq_on S w w''" - unfolding pw_eq_on_def by auto - -lemma ltlc_eq_on: "pw_eq_on (ltlc_aprops \) w w' \ w \\<^sub>c \ \ w' \\<^sub>c \" - unfolding pw_eq_on_def - apply (induction \ arbitrary: w w') - apply (simp_all add: suffix_def) - apply (auto) [] - apply ((rprems, (auto) []) | fo_rule arg_cong2 arg_cong | intro ext | simp)+ - done - -lemma map_ltlc_semantics_aux: - assumes "inj_on f APs" - assumes "\(range \) \ APs" - assumes "ltlc_aprops \ \ APs" - shows "\ \\<^sub>c \ \ (\i. f ` \ i) \\<^sub>c map_ltlc f \" - using assms(2,3) - apply (induct \ arbitrary: \) - using assms(1) - apply (simp_all add: suffix_def inj_on_Un) - apply (auto dest: inj_onD) [] - apply ((rprems, (auto) []) | fo_rule arg_cong2 arg_cong | intro ext | simp)+ - done - - -definition "map_aprops f APs \ { i. \p\APs. f p = Some i }" - -lemma map_ltlc_semantics: - assumes INJ: "inj_on f (dom f)" and DOM: "ltlc_aprops \ \ dom f" - shows "\ \\<^sub>c \ \ (map_aprops f o \) \\<^sub>c map_ltlc (the o f) \" -proof - - let ?\r = "\i. \ i \ ltlc_aprops \" - let ?\r' = "\i. \ i \ dom f" - - have 1: "\range ?\r \ ltlc_aprops \" by auto - - have INJ_the_dom: "inj_on (the o f) (dom f)" - using assms - by (auto simp: inj_on_def domIff) - note 2 = subset_inj_on[OF this DOM] - - have 3: "(\i. (the o f) ` ?\r' i) = map_aprops f o \" using DOM INJ - apply (auto intro!: ext simp: map_aprops_def domIff image_iff) - by (metis Int_iff domI option.sel) - - have "\ \\<^sub>c \ \ ?\r \\<^sub>c \" - apply (rule ltlc_eq_on) - apply (auto simp: pw_eq_on_def) - done - also from map_ltlc_semantics_aux[OF 2 1 subset_refl] - have "\ \ (\i. (the o f) ` ?\r i) \\<^sub>c map_ltlc (the o f) \" . - also have "\ \ (\i. (the o f) ` ?\r' i) \\<^sub>c map_ltlc (the o f) \" - apply (rule ltlc_eq_on) using DOM INJ - apply (auto simp: pw_eq_on_def ltlc.set_map domIff image_iff) - by (metis Int_iff contra_subsetD domD domI inj_on_eq_iff option.sel) - also note 3 - finally show ?thesis . -qed - -lemma map_ltlc_semantics_inv: - assumes INJ: "inj_on f (dom f)" and DOM: "ltlc_aprops \ \ dom f" - shows "\ \\<^sub>c map_ltlc (the o f) \ \ (\i. (the o f) -` \ i) \\<^sub>c \" - using map_ltlc_semantics[OF assms] - apply simp - apply (intro ltlc_eq_on) - apply (auto simp add: pw_eq_on_def ltlc.set_map map_aprops_def) - by (metis DOM comp_apply contra_subsetD domD option.sel vimage_eq) - -text \Conversion from LTL with common syntax to LTL\ - -fun ltlc_to_ltl :: "'a ltlc \ 'a ltl" -where - "ltlc_to_ltl true\<^sub>c = true" -| "ltlc_to_ltl false\<^sub>c = false" -| "ltlc_to_ltl prop\<^sub>c(q) = prop(q)" -| "ltlc_to_ltl (not\<^sub>c \) = not (ltlc_to_ltl \)" -| "ltlc_to_ltl (\ and\<^sub>c \) = ltlc_to_ltl \ and ltlc_to_ltl \" -| "ltlc_to_ltl (\ or\<^sub>c \) = ltlc_to_ltl \ or ltlc_to_ltl \" -| "ltlc_to_ltl (\ implies\<^sub>c \) = (not (ltlc_to_ltl \)) or (ltlc_to_ltl \)" -| "ltlc_to_ltl (\ iff\<^sub>c \) = (let \' = ltlc_to_ltl \ in - let \' = ltlc_to_ltl \ in - (not \' or \') and (not \' or \'))" -| "ltlc_to_ltl (X\<^sub>c \) = X (ltlc_to_ltl \)" -| "ltlc_to_ltl (F\<^sub>c \) = true U ltlc_to_ltl \" -| "ltlc_to_ltl (G\<^sub>c \) = false V ltlc_to_ltl \" -| "ltlc_to_ltl (\ U\<^sub>c \) = ltlc_to_ltl \ U ltlc_to_ltl \" -| "ltlc_to_ltl (\ V\<^sub>c \) = ltlc_to_ltl \ V ltlc_to_ltl \" - -lemma ltlc_to_ltl_equiv: - "\ \ (ltlc_to_ltl \) \ \ \\<^sub>c \" - apply (induct \ arbitrary:\) - apply (auto simp: Let_def) - done - -end - -subsection \Semantic Preserving Syntax Transformations\ - -context begin interpretation LTL_Syntax . - -lemma ltl_true_or_con[simp]: - "\ \ prop(p) or (not prop(p)) \ \ \ true" - by auto - -lemma ltl_false_true_con[simp]: - "\ \ not true \ \ \ false" - by auto - -text\ - The negation symbol can be passed through the next operator. -\ -lemma ltl_Next_Neg_con[simp]: - "\ \ X (not \) \ \ \ not X \" - by auto - -text\ - The connection between Until and Release -\ -lemma ltl_Release_Until_con: - "\ \ \ V \ \ (\ \ \ (not \) U (not \))" - by auto - - -text\Expand strategy\ - -lemma ltl_expand_Until: - "\ \ \ U \ \ (\ \ \ or (\ and (X (\ U \))))" (is "?lhs = ?rhs") -proof - assume ?lhs - then obtain i - where psi_is: "suffix i \ \ \" - and phi_is: "\j \ \" by auto - show ?rhs - proof(cases i) - case 0 - then show ?rhs using psi_is by auto - next - case (Suc k) - with phi_is have "\ \ \" by auto - moreover - have "\ \ X (\ U \)" - using psi_is phi_is Suc by auto - ultimately show ?rhs by auto - qed -next - assume rhs: ?rhs - show ?lhs - proof(cases "\ \ \") - case True - then have "suffix 0 \ \ \" by auto - moreover - have "\j<0. suffix j \ \ \" by auto - ultimately - have "\i. suffix i \ \ \ - \ (\j \ \)" by blast - then show ?lhs by auto - next - case False - then have phi_is: "\ \ \" - and "\ \ X (\ U \)" using rhs by auto - then obtain i - where psi_suc_is: "suffix (Suc i) \ \ \" - and phi_suc_is: "\j \ \" by auto - have sbgoal: "\j<(Suc i). suffix j \ \ \" - proof(clarify) - fix j - assume j_less: "j \ \" - proof (cases j) - assume "j=0" - then show ?thesis using phi_is by auto - next - fix k - assume "j=Suc k" - then show ?thesis using j_less phi_suc_is by auto - qed - qed - then show ?lhs using psi_suc_is phi_is by auto - qed -qed - -lemma ltl_expand_Release: - "\ \ \ V \ \ (\ \ \ and (\ or (X (\ V \))))" -proof - - from ltl_expand_Until[of \ "not \" "not \"] - show ?thesis by auto -qed - -text\Double negation structure of an LTL formula\ - -lemma [simp]: - "not ((\\. not not \) ^^ n) \ = ((\\. not not \) ^^ n) (not \)" - by (induct n) auto - -lemma ltl_double_neg_struct: - shows "\n \. \ = ((\\. not not \) ^^ n) \ \ (\\. \ \ not not \)" - (is "\n \. ?Q \ n \") -proof(cases "\\. \ \ not \") - case True - then have "?Q \ 0 \" by auto - then show ?thesis by blast -next - case False - then show ?thesis - proof(induct \) - case (LTLNeg \') - show ?case - proof(cases "\\. \' \ not \") - case True - with LTLNeg have "?Q (not \') 0 (not \')" by auto - then show ?thesis by blast - next - case False - with LTLNeg obtain n' \' where Q: "?Q \' n' \'" by auto - show ?thesis - proof(cases "\\''. \' = not \''") - case True - then obtain \'' - where "\' = not \''" by auto - with Q have "?Q (not \') (n'+1) \''" by auto - then show ?thesis by blast - next - case False - with Q have "?Q (not \') n' (not \')" by auto - then show ?thesis by blast - qed - qed - qed auto -qed - -lemma ltl_size_double_neg: - assumes "\ = ((\\. not not \) ^^ n) \" - shows "size \ \ size \" -using assms proof (induct n arbitrary:\ \) - case (Suc k) - obtain \ where \_eq: "\ = ((\\. not not \) ^^ k) \" by auto - then have "\ = not not \" using Suc by auto - moreover have "size \ \ size \" using Suc \_eq by auto - ultimately show ?case by auto -qed auto - - -text\Pushing negation to the top of a proposition\ - -fun - ltl_pushneg :: "'a ltl \ 'a ltl" - where - "ltl_pushneg true = true" - | "ltl_pushneg false = false" - | "ltl_pushneg prop(q) = prop(q)" - | "ltl_pushneg (not true) = false" - | "ltl_pushneg (not false) = true" - | "ltl_pushneg (not prop(q)) = not prop(q)" - | "ltl_pushneg (not (not \)) = ltl_pushneg \" - | "ltl_pushneg (not (\ and \)) = ltl_pushneg (not \) or ltl_pushneg (not \)" - | "ltl_pushneg (not (\ or \)) = ltl_pushneg (not \) and ltl_pushneg (not \)" - | "ltl_pushneg (not (X \)) = X ltl_pushneg (not \)" - | "ltl_pushneg (not (\ U \)) = ltl_pushneg (not \) V ltl_pushneg (not \)" - | "ltl_pushneg (not (\ V \)) = ltl_pushneg (not \) U ltl_pushneg (not \)" - | "ltl_pushneg (\ and \) = (ltl_pushneg \) and (ltl_pushneg \)" - | "ltl_pushneg (\ or \) = (ltl_pushneg \) or (ltl_pushneg \)" - | "ltl_pushneg (X \) = X (ltl_pushneg \)" - | "ltl_pushneg (\ U \) = (ltl_pushneg \) U (ltl_pushneg \)" - | "ltl_pushneg (\ V \) = (ltl_pushneg \) V (ltl_pushneg \)" - - -text\ - In fact, the @{text ltl_pushneg} function does not change the - semantics of the input formula. -\ - -lemma ltl_pushneg_neg: - shows "\ \ ltl_pushneg (not \) \ \ \ not ltl_pushneg \" - by (induct \ arbitrary: \) auto - -theorem ltl_pushneg_equiv[simp]: - "\ \ ltl_pushneg \ \ \ \ \" -proof (induct \ arbitrary: \) - case (LTLNeg \) - with ltl_pushneg_neg show ?case by auto -qed auto - - -text\ - We can now show that @{text ltl_pushneg} does what it should do. - Actually the negation occurs after the transformation - only on top of a proposition. -\ (* TODO (Alex): Formulation! Describe the 'what' in one or two sentences! *) - -lemma ltl_pushneg_double_neg: - shows "ltl_pushneg (((\\. not not \) ^^ n) \) = ltl_pushneg \" -by (induct n arbitrary:\) auto - -lemma ltl_pushneg_neg_struct: - assumes "ltl_pushneg \ = not \" - shows "\q. \ = prop(q)" -proof - - from ltl_double_neg_struct - obtain n \ where \_eq: "\ = ((\\. not not \) ^^ n) \" - and \_neq: "(\\. \ \ not not \)" by blast - with ltl_pushneg_double_neg have "ltl_pushneg \ = ltl_pushneg \" by auto - then show ?thesis using \_eq assms \_neq - proof(induct \) - case (LTLNeg f) then show ?case by (cases f) auto - qed auto -qed - -inductive subfrml -where - "subfrml \ (not \)" -| "subfrml \ (\ and \)" -| "subfrml \ (\ and \)" -| "subfrml \ (\ or \)" -| "subfrml \ (\ or \)" -| "subfrml \ (X \)" -| "subfrml \ (\ U \)" -| "subfrml \ (\ U \)" -| "subfrml \ (\ V \)" -| "subfrml \ (\ V \)" - -abbreviation is_subfrml ("_ is'_subformula'_of _") -where - "is_subfrml \ \ \ subfrml\<^sup>*\<^sup>* \ \" - -lemma subfrml_size: - assumes "subfrml \ \" - shows "size \ < size \" -using assms by (induct \) auto - -lemma subformula_size: - assumes "\ is_subformula_of \" - shows "size \ < size \ \ \ = \" -using assms proof(induct \) - case base then show ?case by auto -next - case (step \ \) - then have "size \ < size \" by (rule_tac subfrml_size) - then show ?case using step by auto -qed - - -lemma subformula_on_ltl_pushneg: - assumes "\ is_subformula_of (ltl_pushneg \)" - shows "\\. \ = ltl_pushneg \" -proof(cases "\ = ltl_pushneg \") - case True then show ?thesis by blast -next - case False then show ?thesis using assms - proof(induct \ rule:ltl_pushneg.induct) - case 1 then show ?case using subformula_size by force - next - case 2 then show ?case using subformula_size by force - next - case 3 then show ?case using subformula_size by force - next - case 4 then show ?case using subformula_size by force - next - case 5 then show ?case using subformula_size by force - next - case (6 q) - let ?frml = "not prop(q)" - from rtrancl_eq_or_trancl[to_pred, of subfrml] - have t_prm: "subfrml\<^sup>+\<^sup>+ \ ?frml" - using 6 by auto - obtain \ - where sf_prm: "subfrml \ \" - and rt_prm: "\ is_subformula_of ?frml" - using tranclpD[OF t_prm] by blast - show ?case - proof(cases "\ = ?frml") - assume "\ = ?frml" - with sf_prm have "\ = ltl_pushneg prop(q)" - by (cases \) auto - then show ?thesis by blast - next - assume "\ \ ?frml" - with rtranclpD[OF rt_prm] - tranclp_into_rtranclp[of subfrml] - subformula_size[of \ ?frml] - have "size \ = 0" by auto - with subfrml_size[OF sf_prm] - show ?thesis by auto - qed - next - case 7 then show ?case by auto - next - case (8 \ \) - let ?frml = "not (\ and \)" - from tranclD2[to_pred, of subfrml \ "ltl_pushneg ?frml"] - rtrancl_eq_or_trancl[to_pred, of subfrml] - obtain z - where "subfrml z (ltl_pushneg ?frml)" - and rt_prm: "\ is_subformula_of z" - using 8 by auto - then have z_is: - "z = ltl_pushneg (not \) \ - z = ltl_pushneg (not \)" by (cases z) auto - show ?case - proof(cases "\ = z") - assume "\ = z" - with z_is show ?thesis by auto - next - assume "\ \ z" - with rtranclpD[OF rt_prm] - tranclp_into_rtranclp - have "\ is_subformula_of z" by auto - then show ?thesis using 8 z_is by auto - qed - next - case (9 \ \) - let ?frml = "not (\ or \)" - from tranclD2[to_pred, of subfrml \ "ltl_pushneg ?frml"] - rtrancl_eq_or_trancl[to_pred, of subfrml] - obtain z - where "subfrml z (ltl_pushneg ?frml)" - and rt_prm: "\ is_subformula_of z" - using 9 by auto - then have z_is: - "z = ltl_pushneg (not \) \ - z = ltl_pushneg (not \)" by (cases z) auto - show ?case - proof(cases "\ = z") - assume "\ = z" - with z_is show ?thesis by auto - next - assume "\ \ z" - with rtranclpD[OF rt_prm] - tranclp_into_rtranclp - have "\ is_subformula_of z" by auto - then show ?thesis using 9 z_is by auto - qed - next - case (10 \) - let ?frml = "not (X \)" - from tranclD2[to_pred, of subfrml \ "ltl_pushneg ?frml"] - rtrancl_eq_or_trancl[to_pred, of subfrml] - obtain z - where "subfrml z (ltl_pushneg ?frml)" - and rt_prm: "\ is_subformula_of z" - using 10 by auto - then have z_is: "z = ltl_pushneg (not \)" - by (cases z) auto - show ?case - proof(cases "\ = z") - assume "\ = z" - with z_is show ?thesis by auto - next - assume "\ \ z" - with rtranclpD[OF rt_prm] - tranclp_into_rtranclp - have "\ is_subformula_of z" by auto - then show ?thesis using 10 z_is by auto - qed - next - case (11 \ \) - let ?frml = "not (\ U \)" - from tranclD2[to_pred, of subfrml \ "ltl_pushneg ?frml"] - rtrancl_eq_or_trancl[to_pred, of subfrml] - obtain z - where "subfrml z (ltl_pushneg ?frml)" - and rt_prm: "\ is_subformula_of z" - using 11 by auto - then have z_is: - "z = ltl_pushneg (not \) \ - z = ltl_pushneg (not \)" by (cases z) auto - show ?case - proof(cases "\ = z") - assume "\ = z" - with z_is show ?thesis by auto - next - assume "\ \ z" - with rtranclpD[OF rt_prm] - tranclp_into_rtranclp - have "\ is_subformula_of z" by auto - then show ?thesis using 11 z_is by auto - qed - next - case (12 \ \) - let ?frml = "not (\ V \)" - from tranclD2[to_pred, of subfrml \ "ltl_pushneg ?frml"] - rtrancl_eq_or_trancl[to_pred, of subfrml] - obtain z - where "subfrml z (ltl_pushneg ?frml)" - and rt_prm: "\ is_subformula_of z" - using 12 by auto - then have z_is: - "z = ltl_pushneg (not \) \ - z = ltl_pushneg (not \)" by (cases z) auto - show ?case - proof(cases "\ = z") - assume "\ = z" - with z_is show ?thesis by auto - next - assume "\ \ z" - with rtranclpD[OF rt_prm] - tranclp_into_rtranclp - have "\ is_subformula_of z" by auto - then show ?thesis using 12 z_is by auto - qed - next - case (13 \ \) - let ?frml = "\ and \" - from tranclD2[to_pred, of subfrml \ "ltl_pushneg ?frml"] - rtrancl_eq_or_trancl[to_pred, of subfrml] - obtain z - where "subfrml z (ltl_pushneg ?frml)" - and rt_prm: "\ is_subformula_of z" - using 13 by auto - then have z_is: - "z = ltl_pushneg \ \ - z = ltl_pushneg \" by (cases z) auto - show ?case - proof(cases "\ = z") - assume "\ = z" - with z_is show ?thesis by auto - next - assume "\ \ z" - with rtranclpD[OF rt_prm] - tranclp_into_rtranclp - have "\ is_subformula_of z" by auto - then show ?thesis using 13 z_is by auto - qed - next - case (14 \ \) - let ?frml = "\ or \" - from tranclD2[to_pred, of subfrml \ "ltl_pushneg ?frml"] - rtrancl_eq_or_trancl[to_pred, of subfrml] - obtain z - where "subfrml z (ltl_pushneg ?frml)" - and rt_prm: "\ is_subformula_of z" - using 14 by auto - then have z_is: - "z = ltl_pushneg \ \ - z = ltl_pushneg \" by (cases z) auto - show ?case - proof(cases "\ = z") - assume "\ = z" - with z_is show ?thesis by auto - next - assume "\ \ z" - with rtranclpD[OF rt_prm] - tranclp_into_rtranclp - have "\ is_subformula_of z" by auto - then show ?thesis using 14 z_is by auto - qed - next - case (15 \) - let ?frml = "X \" - from tranclD2[to_pred, of subfrml \ "ltl_pushneg ?frml"] - rtrancl_eq_or_trancl[to_pred, of subfrml] - obtain z - where "subfrml z (ltl_pushneg ?frml)" - and rt_prm: "\ is_subformula_of z" - using 15 by auto - then have z_is: "z = ltl_pushneg \" by (cases z) auto - show ?case - proof(cases "\ = z") - assume "\ = z" - with z_is show ?thesis by auto - next - assume "\ \ z" - with rtranclpD[OF rt_prm] - tranclp_into_rtranclp - have "\ is_subformula_of z" by auto - then show ?thesis using 15 z_is by auto - qed - next - case (16 \ \) - let ?frml = "\ U \" - from tranclD2[to_pred, of subfrml \ "ltl_pushneg ?frml"] - rtrancl_eq_or_trancl[to_pred, of subfrml] - obtain z - where "subfrml z (ltl_pushneg ?frml)" - and rt_prm: "\ is_subformula_of z" - using 16 by auto - then have z_is: - "z = ltl_pushneg \ \ - z = ltl_pushneg \" by (cases z) auto - show ?case - proof(cases "\ = z") - assume "\ = z" - with z_is show ?thesis by auto - next - assume "\ \ z" - with rtranclpD[OF rt_prm] - tranclp_into_rtranclp - have "\ is_subformula_of z" by auto - then show ?thesis using 16 z_is by auto - qed - next - case (17 \ \) - let ?frml = "\ V \" - from tranclD2[to_pred, of subfrml \ "ltl_pushneg ?frml"] - rtrancl_eq_or_trancl[to_pred, of subfrml] - obtain z - where "subfrml z (ltl_pushneg ?frml)" - and rt_prm: "\ is_subformula_of z" - using 17 by auto - then have z_is: - "z = ltl_pushneg \ \ - z = ltl_pushneg \" by (cases z) auto - show ?case - proof(cases "\ = z") - assume "\ = z" - with z_is show ?thesis by auto - next - assume "\ \ z" - with rtranclpD[OF rt_prm] - tranclp_into_rtranclp - have "\ is_subformula_of z" by auto - then show ?thesis using 17 z_is by auto - qed - qed -qed - - -text\ - The fact that after pushing the negation the structure of a - formula changes, is shown by the following theorem. - Indeed, after pushing the negation symbol inside a - formula, it occurs at most on top of a proposition. -\ - -theorem ltl_pushneg_struct: - assumes "(not \) is_subformula_of (ltl_pushneg \)" - shows "\q. \ = prop(q)" -proof - - from assms subformula_on_ltl_pushneg - obtain \ - where prm:"not \ = ltl_pushneg \" by blast - from ltl_pushneg_neg_struct[OF sym[OF prm]] - show ?thesis by auto -qed - - -text\ - Now we want to show that the size of the formula, - which is transformed by @{text ltl_pushneg}, does not increase - 'too much', i.e. there is no exponential blowup - produced by the transformation. - For that purpose we need an additional function, - which counts the literals of the derivation tree - of a formula. - The idea is, that, assuming the worst case, the - pushing of negation can only increase the size of a - formula by putting the negation symbol on top of every - proposition inside the formula. -\ - -fun leafcnt :: "'a ltl \ nat" -where - "leafcnt true = 1" -| "leafcnt false = 1" -| "leafcnt prop(q) = 1" -| "leafcnt (not \) = leafcnt \" -| "leafcnt (\ and \) = (leafcnt \) + (leafcnt \)" -| "leafcnt (\ or \) = (leafcnt \) + (leafcnt \)" -| "leafcnt (X \) = leafcnt \" -| "leafcnt (\ U \) = (leafcnt \) + (leafcnt \)" -| "leafcnt (\ V \) = (leafcnt \) + (leafcnt \)" - - -lemma leafcnt_double_neg_ident: - "leafcnt (((\\. not not \) ^^ n) \) = leafcnt \" -by (induct n arbitrary:\) auto - -lemma ltl_pushneg_help: - "\\. ltl_pushneg \ = ltl_pushneg \ - \ ((\\. \ = not \ \ (\\. \ \ not \)) \ (\\. \ \ not \)) - \ size \ \ size \ - \ leafcnt \ = leafcnt \" - (is "\\. ?P \ \ \ ?Q \ \ ?R \") -proof - - from ltl_double_neg_struct - obtain n \ where dblneg: "\ = ((\\. not not \) ^^ n) \" - and \_neq: "(\\. \ \ not not \)" - by blast - have "?Q \" - proof(rule ccontr) - assume *: "\ ?Q \" - show "False" - proof(cases "\\. \ = not \") - case True - then obtain \ where "\ = not \" by blast - with * dblneg \_neq show ?thesis by auto - next - case False - with * show ?thesis by auto - qed - qed - with dblneg \_neq - ltl_pushneg_double_neg[of n \] - leafcnt_double_neg_ident[of n \] - ltl_size_double_neg[OF dblneg] show ?thesis by auto -qed - - -lemma ltl_pushneg_size_lin_help: - assumes "\ = ltl_pushneg \" - shows "size \ + 1 \ size \ + leafcnt \" - using assms -proof (induct \ arbitrary: \) - case LTLTrue show ?case by (cases \) auto -next - case LTLFalse show ?case by (cases \) auto -next - case LTLProp show ?case by (cases \) auto -next - case (LTLNeg \') - with ltl_pushneg_neg_struct[of \ \'] obtain q where q: "\' = prop(q)" by auto - with ltl_pushneg_help[of \] obtain \' where \': - "ltl_pushneg \ = ltl_pushneg \'" - "(\\. \' = not \ \ (\\. \ \ not \)) \ (\\. \' \ not \)" - "size \' \ size \" - "leafcnt \' = leafcnt \" - by auto - show ?case - proof(cases "\\. \' = not \ \ (\\. \ \ not \)") - case True - then obtain \ where \'is:"\' = not \" and "\\. \ \ not \" by auto - with q \' show ?thesis using LTLNeg by (cases \) auto - next - case False - with \' LTLNeg show ?thesis by (cases \') force+ - qed -next - case (LTLAnd f g) - with ltl_pushneg_help[of \] - obtain \' where \': - "ltl_pushneg \ = ltl_pushneg \'" - "(\\. \' = not \ \ (\\. \ \ not \)) \ (\\. \' \ not \)" - "size \' \ size \" - "leafcnt \' = leafcnt \" - by auto - show ?case - proof(cases "\\. \' = not \ \ (\\. \ \ not \)") - case True - then obtain \ where \'is: "\' = not \" and "\\. \ \ not \" by auto - with LTLAnd \' have "size (f and g) \ size \ + leafcnt \" by (cases \) force+ - with \' \'is show ?thesis by auto - next - case False - with LTLAnd \' show ?thesis by (cases \') force+ - qed -next - case (LTLOr f g) - with ltl_pushneg_help[of \] - obtain \' where \': - "ltl_pushneg \ = ltl_pushneg \'" - "(\\. \' = not \ \ (\\. \ \ not \)) \ (\\. \' \ not \)" - "size \' \ size \" - "leafcnt \' = leafcnt \" by auto - show ?case - proof(cases "\\. \' = not \ \ (\\. \ \ not \)") - case True - then obtain \ where \'is: "\' = not \" and "\\. \ \ not \" by auto - with LTLOr \' have "size (f or g) \ size \ + leafcnt \" by (cases \) force+ - with \' \'is show ?thesis by auto - next - case False - with LTLOr \' show ?thesis by (cases \') force+ - qed -next - case (LTLNext f) - with ltl_pushneg_help[of \] - obtain \' where \': - "ltl_pushneg \ = ltl_pushneg \'" - "(\\. \' = not \ \ (\\. \ \ not \)) \ (\\. \' \ not \)" - "size \' \ size \" - "leafcnt \' = leafcnt \" by auto - show ?case - proof(cases "\\. \' = not \ \ (\\. \ \ not \)") - case True - then obtain \ where \'is: "\' = not \" and "\\. \ \ not \" by auto - with LTLNext \' show ?thesis by (cases \) force+ - next - case False - with LTLNext \' show ?thesis by (cases \') force+ - qed -next - case (LTLUntil f g) - with ltl_pushneg_help[of \] - obtain \' where \': - "ltl_pushneg \ = ltl_pushneg \'" - "(\\. \' = not \ \ (\\. \ \ not \)) \ (\\. \' \ not \)" - "size \' \ size \" - "leafcnt \' = leafcnt \" by auto - show ?case - proof(cases "\\. \' = not \ \ (\\. \ \ not \)") - case True - then obtain \ where \'is: "\' = not \" and "\\. \ \ not \" by auto - with LTLUntil \' have "size (f U g) \ size \ + leafcnt \" by (cases \) force+ - with LTLUntil \' \'is show ?thesis by auto - next - case False - with LTLUntil \' show ?thesis by (cases \') force+ - qed -next - case (LTLRelease f g) - with ltl_pushneg_help[of \] - obtain \' where \': - "ltl_pushneg \ = ltl_pushneg \'" - "(\\. \' = not \ \ (\\. \ \ not \)) \ (\\. \' \ not \)" - "size \' \ size \" - "leafcnt \' = leafcnt \" by auto - show ?case - proof(cases "\\. \' = not \ \ (\\. \ \ not \)") - case True - then obtain \ where \'is:"\' = not \" and "\\. \ \ not \" by auto - with LTLRelease \' have "size (f V g) \ size \ + leafcnt \" by (cases \) force+ - with LTLRelease \' \'is show ?thesis by auto - next - case False - with LTLRelease \' show ?thesis by (cases \') force+ - qed -qed - -theorem ltl_pushneg_size_lin: - "size (ltl_pushneg \) \ 2 * size \" -proof - - have "leafcnt \ \ size \ + 1" by (induct \) auto - with ltl_pushneg_size_lin_help[of _ \] - have "size (ltl_pushneg \) + 1 \ size \ + size \ + 1" by force - then show ?thesis by auto -qed - -end - -subsection \LTL formula in negation normal form (NNF)\ -text\ - We define a type of LTL formula in negation normal form (NNF) -\ - - -datatype - 'a ltln = LTLnTrue - | LTLnFalse - | LTLnProp 'a - | LTLnNProp 'a - | LTLnAnd "'a ltln" "'a ltln" - | LTLnOr "'a ltln" "'a ltln" - | LTLnNext "'a ltln" - | LTLnUntil "'a ltln" "'a ltln" - | LTLnRelease "'a ltln" "'a ltln" - -context LTL_Syntax begin - notation - LTLnTrue ("true\<^sub>n") - and LTLnFalse ("false\<^sub>n") - and LTLnProp ("prop\<^sub>n'(_')") - and LTLnNProp ("nprop\<^sub>n'(_')") - and LTLnAnd ("_ and\<^sub>n _" [82,82] 81) - and LTLnOr ("_ or\<^sub>n _" [84,84] 83) - and LTLnNext ("X\<^sub>n _" [88] 87) - and LTLnUntil ("_ U\<^sub>n _" [84,84] 83) - and LTLnRelease ("_ V\<^sub>n _" [84,84] 83) - - abbreviation ltln_eventuality :: "'a ltln \ 'a ltln" ("\\<^sub>n _" [88] 87) - where "ltln_eventuality \ \ true\<^sub>n U\<^sub>n \" - - abbreviation ltln_universality :: "'a ltln \ 'a ltln" ("\\<^sub>n _" [88] 87) - where "ltln_universality \ \ false\<^sub>n V\<^sub>n \" - -end - -context begin interpretation LTL_Syntax . - -primrec ltln_semantics :: "['a set word, 'a ltln] \ bool" -("_ \\<^sub>n _" [80,80] 80) - where - "\ \\<^sub>n true\<^sub>n = True" - | "\ \\<^sub>n false\<^sub>n = False" - | "\ \\<^sub>n prop\<^sub>n(q) = (q\\ 0)" - | "\ \\<^sub>n nprop\<^sub>n(q) = (q\\ 0)" - | "\ \\<^sub>n \ and\<^sub>n \ = (\ \\<^sub>n \ \ \ \\<^sub>n \)" - | "\ \\<^sub>n \ or\<^sub>n \ = (\ \\<^sub>n \ \ \ \\<^sub>n \)" - | "\ \\<^sub>n X\<^sub>n \ = (suffix 1 \ \\<^sub>n \)" - | "\ \\<^sub>n \ U\<^sub>n \ = (\i. suffix i \ \\<^sub>n \ \ (\j \\<^sub>n \))" - | "\ \\<^sub>n \ V\<^sub>n \ = (\i. suffix i \ \\<^sub>n \ \ (\j \\<^sub>n \))" - -definition "ltln_language \ \ {\. \ \\<^sub>n \}" - -text\ - Conversion from LTL to LTL in NNF -\ - -fun ltl_to_ltln :: "'a ltl \ 'a ltln" -where - "ltl_to_ltln true = true\<^sub>n" -| "ltl_to_ltln false = false\<^sub>n" -| "ltl_to_ltln prop(q) = prop\<^sub>n(q)" -| "ltl_to_ltln (not prop(q)) = nprop\<^sub>n(q)" -| "ltl_to_ltln (\ and \) = ltl_to_ltln \ and\<^sub>n ltl_to_ltln \" -| "ltl_to_ltln (\ or \) = ltl_to_ltln \ or\<^sub>n ltl_to_ltln \" -| "ltl_to_ltln (X \) = X\<^sub>n (ltl_to_ltln \)" -| "ltl_to_ltln (\ U \) = ltl_to_ltln \ U\<^sub>n ltl_to_ltln \" -| "ltl_to_ltln (\ V \) = ltl_to_ltln \ V\<^sub>n ltl_to_ltln \" - - -lemma ltl_to_ltln_on_ltl_pushneg_equiv: - assumes "\ = ltl_pushneg \" - shows "\ \ \ \ \ \\<^sub>n ltl_to_ltln \" - using assms -proof(induct \ arbitrary: \ \) - case LTLTrue show ?case by auto -next - case LTLFalse show ?case by auto -next - case LTLProp show ?case by auto -next - case (LTLNeg \) - with ltl_pushneg_neg_struct[of \ \] - obtain q - where "\ = prop(q)" - by auto - then show ?case by auto -next - case (LTLAnd f g \ \) - then have frml_eq: "ltl_pushneg \ = f and g" by auto - with subformula_on_ltl_pushneg[of _ \] - obtain \ - where "f = ltl_pushneg \" - by (auto intro: subfrml.intros) - moreover - from frml_eq subformula_on_ltl_pushneg[of _ \] - obtain \ - where "g = ltl_pushneg \" - by (auto intro: subfrml.intros) - ultimately - have "\ \ f = \ \\<^sub>n ltl_to_ltln f" - and "\ \ g = \ \\<^sub>n ltl_to_ltln g" - using LTLAnd by auto - then show ?case by auto -next - case (LTLOr f g \ \) - then have frml_eq: "ltl_pushneg \ = f or g" by auto - with subformula_on_ltl_pushneg[of _ \] - obtain \ - where "f = ltl_pushneg \" - by (auto intro: subfrml.intros) - moreover - from frml_eq subformula_on_ltl_pushneg[of _ \] - obtain \ - where "g = ltl_pushneg \" - by (auto intro: subfrml.intros) - ultimately - have "\ \ f = \ \\<^sub>n ltl_to_ltln f" - and "\ \ g = \ \\<^sub>n ltl_to_ltln g" - using LTLOr by auto - then show ?case by auto -next - case (LTLNext \ \ \) - then have "ltl_pushneg \ = X \" by auto - with subformula_on_ltl_pushneg[of \ \] - obtain \ - where "\ = ltl_pushneg \" - by (auto intro: subfrml.intros) - with LTLNext(1) have "suffix 1 \ \ \ = suffix 1 \ \\<^sub>n ltl_to_ltln \" . - then show ?case by auto -next - case (LTLUntil f g \ \) - then have frml_eq: "ltl_pushneg \ = f U g" by auto - with subformula_on_ltl_pushneg[of _ \] - obtain \ - where "f = ltl_pushneg \" - by (auto intro: subfrml.intros) - moreover - from frml_eq subformula_on_ltl_pushneg[of _ \] - obtain \ - where "g = ltl_pushneg \" - by (auto intro: subfrml.intros) - ultimately - have "\i. suffix i \ \ f = suffix i \ \\<^sub>n ltl_to_ltln f" - and "\i. suffix i \ \ g = suffix i \ \\<^sub>n ltl_to_ltln g" - using LTLUntil(1,2) by blast+ - then show ?case by auto -next - case (LTLRelease f g \ \) - then have frml_eq: "ltl_pushneg \ = f V g" by auto - with subformula_on_ltl_pushneg[of _ \] - obtain \ - where "f = ltl_pushneg \" - by (auto intro: subfrml.intros) - moreover - from frml_eq subformula_on_ltl_pushneg[of _ \] - obtain \ - where "g = ltl_pushneg \" - by (auto intro: subfrml.intros) - ultimately - have "\i. suffix i \ \ f = suffix i \ \\<^sub>n ltl_to_ltln f" - and "\i. suffix i \ \ g = suffix i \ \\<^sub>n ltl_to_ltln g" - using LTLRelease(1,2) by blast+ - then show ?case by auto -qed - - -lemma ltl_nnf_equiv[simp]: - "\ \\<^sub>n ltl_to_ltln (ltl_pushneg \) \ \ \ \" - using sym[OF ltl_pushneg_equiv] ltl_to_ltln_on_ltl_pushneg_equiv by blast - - -fun subfrmlsn :: "'a ltln \ 'a ltln set" -where - "subfrmlsn (\ and\<^sub>n \) = {\ and\<^sub>n \} \ subfrmlsn \ \ subfrmlsn \" -| "subfrmlsn (X\<^sub>n \) = {X\<^sub>n \} \ subfrmlsn \" -| "subfrmlsn (\ U\<^sub>n \) = {\ U\<^sub>n \} \ subfrmlsn \ \ subfrmlsn \" -| "subfrmlsn (\ V\<^sub>n \) = {\ V\<^sub>n \} \ subfrmlsn \ \ subfrmlsn \" -| "subfrmlsn (\ or\<^sub>n \) = {\ or\<^sub>n \} \ subfrmlsn \ \ subfrmlsn \" -| "subfrmlsn x = {x}" - -lemma subfrmlsn_id[simp]: "\ \ subfrmlsn \" by (induct \) auto -lemma subfrmlsn_finite: "finite (subfrmlsn \)" by (induct \) auto -lemma subfrmlsn_subset:"\ \ subfrmlsn \ \ subfrmlsn \ \ subfrmlsn \" -by (induct \ arbitrary:\) auto - -fun size_frmln :: "'a ltln \ nat" -where - "size_frmln (\ and\<^sub>n \) = size_frmln \ + size_frmln \ + 1" -| "size_frmln (X\<^sub>n \) = size_frmln \ + 1" -| "size_frmln (\ U\<^sub>n \) = size_frmln \ + size_frmln \ + 1" -| "size_frmln (\ V\<^sub>n \) = size_frmln \ + size_frmln \ + 1" -| "size_frmln (\ or\<^sub>n \) = size_frmln \ + size_frmln \ + 1" -| "size_frmln _ = 1" - -lemma size_frmln_gt_zero[simp]: "size_frmln \ > 0" by (induct \) auto - -abbreviation - "frmlset_sumn \ \ setsum size_frmln \" -- "FIXME: lemmas about this?" - -lemma frmlset_sumn_diff_less[intro!]: - assumes finS:"finite S" - and "A\{}" - and subset:"A\S" - shows "frmlset_sumn (S - A) < frmlset_sumn S" -proof - - have finA: "finite A" using assms by (rule_tac finite_subset) - then have "frmlset_sumn A > 0" using assms size_frmln_gt_zero by (induct rule:finite_induct) auto - moreover - have "frmlset_sumn A \ frmlset_sumn S" using assms size_frmln_gt_zero by (rule_tac setsum_mono2) auto - ultimately show ?thesis using setsum_diff_nat[OF finA, of S size_frmln] assms by auto -qed - -definition - "frmln_props \ \ {p. prop\<^sub>n(p) \ subfrmlsn \ \ nprop\<^sub>n(p) \ subfrmlsn \}" - -lemma ltln_expand_Until: - "\ \\<^sub>n \ U\<^sub>n \ = (\ \\<^sub>n \ or\<^sub>n (\ and\<^sub>n (X\<^sub>n (\ U\<^sub>n \))))" (is "?lhs = ?rhs") -proof - assume ?lhs - then obtain i - where psi_is: "suffix i \ \\<^sub>n \" - and phi_is: "\j \\<^sub>n \" by auto - show ?rhs - proof(cases i) - assume "i=0" - then show ?rhs using psi_is by auto - next - fix k - assume i_eq: "i = Suc k" - with phi_is have "\ \\<^sub>n \" by auto - moreover - have "\ \\<^sub>n X\<^sub>n (\ U\<^sub>n \)" - using psi_is phi_is i_eq by auto - ultimately show ?rhs by auto - qed -next - assume rhs: ?rhs - show ?lhs - proof(cases "\ \\<^sub>n \") - assume "\ \\<^sub>n \" - then have "suffix 0 \ \\<^sub>n \" by auto - moreover - have "\j<0. suffix j \ \\<^sub>n \" by auto - ultimately - have "\i. suffix i \ \\<^sub>n \ - \ (\j \\<^sub>n \)" by blast - then show ?lhs by auto - next - assume "\ \ \\<^sub>n \" - then have phi_is: "\ \\<^sub>n \" - and "\ \\<^sub>n X\<^sub>n (\ U\<^sub>n \)" using rhs by auto - then obtain i - where psi_suc_is: "suffix (Suc i) \ \\<^sub>n \" - and phi_suc_is: "\j \\<^sub>n \" by auto - have sbgoal: "\j<(Suc i). suffix j \ \\<^sub>n \" - proof(clarify) - fix j - assume j_less: "j \\<^sub>n \" - proof (cases j) - assume "j=0" - then show ?thesis using phi_is by auto - next - fix k - assume "j=Suc k" - then show ?thesis using j_less phi_suc_is by auto - qed - qed - then show ?lhs using psi_suc_is phi_is by auto - qed -qed - -lemma ltln_expand_Release: - "\ \\<^sub>n \ V\<^sub>n \ = (\ \\<^sub>n \ and\<^sub>n (\ or\<^sub>n (X\<^sub>n (\ V\<^sub>n \))))" (is "?lhs = ?rhs") -proof - assume lhs: ?lhs - then have psi_is: "\ \\<^sub>n \" by force - - have "\i. \\ \ \\<^sub>n \; \ suffix (Suc i) \ \\<^sub>n \\ - \ (\j \\<^sub>n \)" - proof - - fix i - assume phi_nis: "\ \ \\<^sub>n \" - and "\ suffix (Suc i) \ \\<^sub>n \" - then obtain j - where "j \\<^sub>n \" using lhs by auto - then have "j - 1 < i \ suffix (Suc (j - 1)) \ \\<^sub>n \" - using phi_nis by (cases j) auto - then show "\j \\<^sub>n \" by auto - qed - then show ?rhs using psi_is by auto -next - assume rhs: ?rhs - then have psi_is: "\ \\<^sub>n \" by auto - - show ?lhs - proof(cases "\ \\<^sub>n \") - assume "\ \\<^sub>n \" - then show ?thesis using psi_is by force - next - assume phi_nis: "\ \ \\<^sub>n \" - - then have "\i. suffix (Suc i) \ \\<^sub>n \ - \ (\j \\<^sub>n \)" - using rhs by auto - - have "\i. \ suffix i \ \\<^sub>n \ - \ (\j \\<^sub>n \)" - proof - - fix i - assume psi_suf_nis: "\ suffix i \ \\<^sub>n \" - show "\j \\<^sub>n \" - proof(cases "i") - assume "i=0" - with psi_suf_nis psi_is show ?thesis by auto - next - fix k - assume i_eq: "i=Suc k" - with psi_suf_nis rhs show ?thesis by force - qed - qed - then show ?thesis by auto - qed -qed - -lemma ltln_Release_alterdef: - "\ \\<^sub>n \ V\<^sub>n \ \ \ \\<^sub>n (\\<^sub>n \) or\<^sub>n (\ U\<^sub>n (\ and\<^sub>n \))" (is "?lhs = ?rhs") -proof - assume ?lhs - { assume "\ (\i. suffix i \ \\<^sub>n \)" - then obtain i where \_neq: "\ suffix i \ \\<^sub>n \" by auto - let ?k = "LEAST i. \ suffix i \ \\<^sub>n \" - from \_neq \?lhs\ have "\j \\<^sub>n \" by (metis not_less_Least) - moreover - have "\ suffix ?k \ \\<^sub>n \" by (rule LeastI, rule \_neq) - moreover then obtain j where "j \\<^sub>n \" using \?lhs\ by auto - ultimately have "\ \\<^sub>n \ U\<^sub>n (\ and\<^sub>n \)" by auto } - with \?lhs\ show ?rhs by auto -next - assume ?rhs - { assume "\ \ \\<^sub>n \\<^sub>n \" - with \?rhs\ obtain i where "suffix i \ \\<^sub>n \ and\<^sub>n \" and "\j \\<^sub>n \" by auto - then have ?lhs by (auto, metis nat_neq_iff) } - then show ?lhs using \?rhs\ by auto -qed - - -end -end diff --git a/thys/LTL_to_GBA/LTL_Rewrite.thy b/thys/LTL_to_GBA/LTL_Rewrite.thy deleted file mode 100644 --- a/thys/LTL_to_GBA/LTL_Rewrite.thy +++ /dev/null @@ -1,666 +0,0 @@ -section \Rewriting LTL formulas\ -(* Author: Alexander Schimpf *) -theory LTL_Rewrite -imports - LTL -begin - -context begin interpretation LTL_Syntax . - -inductive_set ltln_pure_eventual_frmls :: "'a ltln set" -where - "\\<^sub>n \ \ ltln_pure_eventual_frmls" -| "\ \ \ ltln_pure_eventual_frmls; \ \ ltln_pure_eventual_frmls \ - \ \ and\<^sub>n \ \ ltln_pure_eventual_frmls" -| "\ \ \ ltln_pure_eventual_frmls; \ \ ltln_pure_eventual_frmls \ - \ \ or\<^sub>n \ \ ltln_pure_eventual_frmls" -| "\ \ \ ltln_pure_eventual_frmls; \ \ ltln_pure_eventual_frmls \ - \ \ U\<^sub>n \ \ ltln_pure_eventual_frmls" -| "\ \ \ ltln_pure_eventual_frmls; \ \ ltln_pure_eventual_frmls \ - \ \ V\<^sub>n \ \ ltln_pure_eventual_frmls" - -theorem ltln_pure_eventual_frmls_equiv: - assumes "\ \ ltln_pure_eventual_frmls" - shows "\ \\<^sub>n \ U\<^sub>n \ \ \ \\<^sub>n \" - using assms -proof (induct \ arbitrary:\ \) - case 1 - then show ?case - by force -next - case prems: 2 show ?case using prems(2)[of \ \] prems(4)[of \ \] - by (auto, metis less_nat_zero_code suffix_0) -next - case prems: 3 show ?case - using prems(2)[of \ \] prems(4)[of \ \] by auto -next - case prems: (4 \ \) - let ?\ = "\ U\<^sub>n \" - { assume "\ \\<^sub>n \ U\<^sub>n ?\" - then obtain i where "suffix i \ \\<^sub>n ?\" and "\j \\<^sub>n \" - by auto - moreover with prems(4)[of "suffix i \" \] have "suffix i \ \\<^sub>n \" - by blast - ultimately have "\ \\<^sub>n ?\" using prems(4)[of \ \] prems(4)[of \ \] - by auto } - moreover - { assume "\ \\<^sub>n ?\" - with prems have "\ \\<^sub>n \ U\<^sub>n \" by auto - then obtain i where "suffix i \ \\<^sub>n \" and "\j \\<^sub>n \" - by auto - moreover with prems(4)[of "suffix i \" \] have "suffix i \ \\<^sub>n \ U\<^sub>n \" - by blast - ultimately have "\ \\<^sub>n \ U\<^sub>n ?\" by auto } - ultimately show ?case by fast -next - case prems: (5 \ \) - let ?\ = "\ V\<^sub>n \" - { assume "\ \\<^sub>n \ U\<^sub>n ?\" - then obtain i where - V_suf_i: "suffix i \ \\<^sub>n \ V\<^sub>n \" - and phi_all_less_i: "\j \\<^sub>n \" - unfolding ltln_Release_alterdef[symmetric] by auto - then have \_suf_i: "suffix i \ \\<^sub>n \" - by (metis ltln_expand_Release ltln_semantics.simps(5)) - have \_less_i: "\j \\<^sub>n \" - proof(clarify) - fix k - assume "k) \\<^sub>n \" - and "\j) \\<^sub>n \" - using V_suf_i phi_all_less_i \_suf_i by auto - then show "suffix k \ \\<^sub>n \" using prems(4)[of "suffix k \" \] - using ltln_semantics.simps(8) by blast - qed - have "\ \\<^sub>n ?\" - proof - - { fix i' - { assume "i' < i" - then have "suffix i' \ \\<^sub>n \" using \_less_i by auto } - moreover - { assume "i' \ i" - then obtain i'' where i'_eq: "i' = i + i''" and "i'\i''" - by (metis Nat.diff_le_self le_add_diff_inverse2 add.commute) - then have "suffix i' \ \\<^sub>n \ \ (\j \\<^sub>n \)" - using V_suf_i by auto } - ultimately have "suffix i' \ \\<^sub>n \ \ (\j \\<^sub>n \)" - by (metis linorder_not_less) } - then show ?thesis by auto - qed } - moreover - { assume "\ \\<^sub>n ?\" - then have "suffix 0 \ \\<^sub>n ?\ \ (\j<0. suffix j \ \\<^sub>n \)" by auto - then have "\ \\<^sub>n \ U\<^sub>n ?\" unfolding ltln_semantics.simps by blast } - ultimately show ?case by fast -qed - -corollary ltln_pure_eventual_frmls_equiv_diamond: - assumes "\ \ ltln_pure_eventual_frmls" - shows "\ \\<^sub>n \\<^sub>n \ \ \ \\<^sub>n \" -by (rule ltln_pure_eventual_frmls_equiv[OF assms]) - - -inductive_set ltln_pure_universal_frmls :: "'a ltln set" -where - "\\<^sub>n \ \ ltln_pure_universal_frmls" -| "\ \ \ ltln_pure_universal_frmls; \ \ ltln_pure_universal_frmls \ - \ \ and\<^sub>n \ \ ltln_pure_universal_frmls" -| "\ \ \ ltln_pure_universal_frmls; \ \ ltln_pure_universal_frmls \ - \ \ or\<^sub>n \ \ ltln_pure_universal_frmls" -| "\ \ \ ltln_pure_universal_frmls; \ \ ltln_pure_universal_frmls \ - \ \ U\<^sub>n \ \ ltln_pure_universal_frmls" -| "\ \ \ ltln_pure_universal_frmls; \ \ ltln_pure_universal_frmls \ - \ \ V\<^sub>n \ \ ltln_pure_universal_frmls" - -theorem ltln_pure_universal_frmls_equiv: - assumes "\ \ ltln_pure_universal_frmls" - shows "\ \\<^sub>n \ V\<^sub>n \ \ \ \\<^sub>n \" - using assms -proof (induct \ arbitrary:\ \) - case 1 - then show ?case by force -next - case prems: 2 - show ?case - using prems(2)[of \ \] prems(4)[of \ \] by auto -next - case prems: 3 show ?case - using prems(2)[of \ \] prems(4)[of \ \] - by (auto, metis less_nat_zero_code suffix_0) -next - case prems: (4 \ \) - let ?\ = "\ U\<^sub>n \" - { assume assm: "\ \\<^sub>n \ V\<^sub>n ?\" - { assume "\ \\<^sub>n \\<^sub>n ?\" - then have "suffix 0 \ \\<^sub>n ?\" - by (metis ltln_semantics.simps(9) ltln_semantics.simps(2)) - then have "\ \\<^sub>n ?\" by auto } - moreover - { assume "\ \ \\<^sub>n \\<^sub>n ?\" - then have "\ \\<^sub>n ?\ U\<^sub>n (\ and\<^sub>n ?\)" - using assm ltln_Release_alterdef[of \ \ ?\] - by (metis ltln_semantics.simps(6)) - then obtain i - where "suffix i \ \\<^sub>n \ and\<^sub>n ?\" and "\j \\<^sub>n ?\" - by auto - then have "\ \\<^sub>n ?\" - by (cases "i=0") (metis suffix_0 ltln_semantics.simps(5) - neq0_conv suffix_0)+ } - ultimately have "\ \\<^sub>n ?\" by fast } - moreover - { assume "\ \\<^sub>n ?\" - then obtain i where - \_suf_i: "suffix i \ \\<^sub>n \" - and \_less_i: "\j \\<^sub>n \" - by auto - with prems(4)[of "suffix i \"] \_suf_i - have "suffix i \ \\<^sub>n \\<^sub>n \ or\<^sub>n (\ U\<^sub>n (\ and\<^sub>n \))" - using ltln_Release_alterdef[of "suffix i \" \ \] by blast - moreover - { assume "suffix i \ \\<^sub>n \\<^sub>n \" - then have \_suf_i: "suffix i \ \\<^sub>n \ V\<^sub>n ?\" - by auto (metis ac_simps ltln_expand_Until ltln_semantics.simps) - from \_less_i have \_less_i: "\j \\<^sub>n ?\" - proof(clarify, goal_cases) - case (1 j) - then obtain j' where "i = j + j'" by (metis less_imp_add_positive) - then have "suffix j' (suffix j \) \\<^sub>n \" - and "\k) \\<^sub>n \" - using \_suf_i by auto (metis \_less_i \i = j + j'\ - add_less_cancel_right add.commute) - then show ?case by auto - qed - have "\ \\<^sub>n \ V\<^sub>n ?\" - proof - - { fix k - { assume "k < i" - with \_less_i have "suffix k \ \\<^sub>n ?\" by auto } - moreover - { assume "k \ i" - then obtain i' where "k = i + i'" by (metis le_iff_add) - with \_suf_i have "suffix k \ \\<^sub>n ?\ \ (\j \\<^sub>n \)" - by auto } - ultimately have "suffix k \ \\<^sub>n ?\ \ (\j \\<^sub>n \)" - by (metis linorder_not_less) } - then show ?thesis by auto - qed } - moreover - { assume "suffix i \ \\<^sub>n \ U\<^sub>n (\ and\<^sub>n \)" - then obtain k where "suffix (i+k) \ \\<^sub>n \" - and "suffix (i+k) \ \\<^sub>n \" - and "\j) \\<^sub>n \" by auto - then have "\j\i+k. suffix j \ \\<^sub>n ?\" - proof(clarify, goal_cases) - case prems: (1 j) - { assume "j < i" - then obtain j' where "i = j + j'" by (metis less_imp_add_positive) - with \_suf_i \_less_i have ?case by auto } - moreover - { assume "j \ i" - then obtain i' where "j = i + i'" by (metis le_iff_add) - with prems have ?case - by (auto, metis (full_types) add.comm_neutral add_Suc_right - le_neq_implies_less less_nat_zero_code) } - ultimately show ?case by (metis less_or_eq_imp_le linorder_neqE_nat) - qed - with \suffix (i+k) \ \\<^sub>n \\ have "\ \\<^sub>n \ V\<^sub>n ?\" - by (auto, metis linorder_not_less) } - ultimately have "\ \\<^sub>n \ V\<^sub>n ?\" by auto } - ultimately show ?case by fast -next - case prems: (5 \ \) - let ?\ = "\ V\<^sub>n \" - { assume "\ \\<^sub>n \ V\<^sub>n ?\" - moreover - { assume "\ \\<^sub>n \\<^sub>n ?\" - then have "\ \\<^sub>n ?\" - by (metis prems(4) ltln_semantics.simps(2) ltln_semantics.simps(9)) } - moreover - { assume "\ \\<^sub>n ?\ U\<^sub>n (\ and\<^sub>n ?\)" - then obtain i where "suffix i \ \\<^sub>n ?\" and "\j \\<^sub>n ?\" - by auto - then have "\ \\<^sub>n ?\" by (cases "i=0") (auto, metis suffix_0 suffix_suffix) } - ultimately have "\ \\<^sub>n ?\" using ltln_Release_alterdef[of \ \ ?\] by auto } - moreover - { assume "\ \\<^sub>n ?\" - { assume "\ \\<^sub>n \\<^sub>n \" - then have "\ \\<^sub>n \ V\<^sub>n ?\" by auto } - moreover - { assume assm: "\ \ \\<^sub>n \\<^sub>n \" - then have "\ \\<^sub>n \ U\<^sub>n (\ and\<^sub>n \)" - using ltln_Release_alterdef[of \ \ \] \\ \\<^sub>n ?\\ by auto - then have "\ \\<^sub>n \" using \\ \\<^sub>n ?\\ by (auto, metis calculation(1) prems(4)) - with prems(4)[of \ \] have "\ \\<^sub>n \ V\<^sub>n \" by auto - then have "\ \\<^sub>n \ U\<^sub>n (\ and\<^sub>n \)" - using assm ltln_Release_alterdef[of \ \ \] by auto - then obtain i - where "suffix i \ \\<^sub>n \ and\<^sub>n \" and "\j \\<^sub>n \" - by auto - moreover then have "\j\i. suffix j \ \\<^sub>n \ V\<^sub>n \" - by (metis \\ \\<^sub>n \\ assm prems(4)) - ultimately have "\ \\<^sub>n \ V\<^sub>n ?\" by (auto, metis linorder_not_le) } - ultimately have "\ \\<^sub>n \ V\<^sub>n ?\" by fast } - ultimately show ?case by fast -qed - - -text\Some simple rewrite rules\ - -fun ltln_rewrite_step :: "'a ltln \ 'a ltln" -where - "ltln_rewrite_step (_ U\<^sub>n true\<^sub>n) = true\<^sub>n" -| "ltln_rewrite_step (_ V\<^sub>n false\<^sub>n) = false\<^sub>n" -| "ltln_rewrite_step (true\<^sub>n U\<^sub>n (_ U\<^sub>n \)) = true\<^sub>n U\<^sub>n \" -| "ltln_rewrite_step (false\<^sub>n V\<^sub>n (_ V\<^sub>n \)) = false\<^sub>n V\<^sub>n \" -| "ltln_rewrite_step \ = (case \ of - \ U\<^sub>n \' \ - if \ = \' then \ - else if \' \ ltln_pure_eventual_frmls then \' - else \ - | \ V\<^sub>n \' \ - if \ = \' then \ - else if \' \ ltln_pure_universal_frmls then \' - else \ - | (\ U\<^sub>n \) and\<^sub>n (\ U\<^sub>n \') \ if \ = \' then (\ and\<^sub>n \) U\<^sub>n \ else \ - | (\ U\<^sub>n \) or\<^sub>n (\' U\<^sub>n \) \ if \ = \' then \ U\<^sub>n (\ or\<^sub>n \) else \ - | (\ V\<^sub>n \) and\<^sub>n (\' V\<^sub>n \) \ if \ = \' then \ V\<^sub>n (\ and\<^sub>n \) else \ - | (\ V\<^sub>n \) or\<^sub>n (\ V\<^sub>n \') \ if \ = \' then (\ or\<^sub>n \) V\<^sub>n \ else \ - | _ \ \) " - -lemma ltln_rewrite_step__size_less: - assumes "ltln_rewrite_step \ \ \" - shows "size (ltln_rewrite_step \) < size \" -proof (cases \) - case (LTLnUntil \ \) - with assms show ?thesis - by (cases \, cases \) (auto split:ltln.split, - metis+, cases \, auto split:ltln.split, metis+) -next - case (LTLnRelease \ \) - with assms show ?thesis - by (cases \, cases \) (auto split:ltln.split, - metis+, cases \, auto split:ltln.split, metis+) -qed (insert assms, auto split:ltln.split) - -lemma ltln_rewrite_step__size_leq: - "size (ltln_rewrite_step \) \ size \" - using ltln_rewrite_step__size_less[of \] - by (cases "ltln_rewrite_step \ = \") auto - - -theorem ltln_rewrite_step__equiv: - "\ \\<^sub>n ltln_rewrite_step \ \ \ \\<^sub>n \" -proof (cases \) - case \: (LTLnAnd \ \) - show ?thesis - proof (cases \) - case \: LTLnUntil - show ?thesis - proof (cases \) - case \: LTLnUntil - with \ \ \ show ?thesis by (auto, metis nat_neq_iff order_less_trans) - qed (insert \ \, auto) - qed (insert \, auto split:ltln.split) -next - case \: (LTLnOr \ \) - show ?thesis - proof (cases \) - case \: LTLnRelease - show ?thesis - proof (cases \) - case LTLnRelease - with \ \ show ?thesis by (auto, metis nat_neq_iff order_less_trans) - qed (insert \ \, auto) - qed (insert \, auto split:ltln.split) -next - case \: (LTLnUntil \ \) - show ?thesis - proof(cases "\ \ true\<^sub>n \ \ (\ = true\<^sub>n \ (\\' \'. \ = \' U\<^sub>n \')) \ \ \ \") - case True - with \ have "\ \\<^sub>n (if \ \ ltln_pure_eventual_frmls then \ else \) \ \ \\<^sub>n \" - using ltln_pure_eventual_frmls_equiv by auto - with \ True show ?thesis by (cases \) (cases \, auto split:ltln.split)+ - next - case False - with \ show ?thesis - by (cases \, auto split:ltln.split) (metis less_nat_zero_code neq0_conv - suffix_0 add.comm_neutral add_0)+ - qed -next - case \: (LTLnRelease \ \) - show ?thesis - proof(cases "\ \ false\<^sub>n \ \ (\ = false\<^sub>n \ (\\' \'. \ = \' V\<^sub>n \')) \ \ \ \") - case True - with \ have "\ \\<^sub>n (if \ \ ltln_pure_universal_frmls then \ else \) \ \ \\<^sub>n \" - using ltln_pure_universal_frmls_equiv by auto - with \ True show ?thesis by (cases \) (cases \, auto split:ltln.split)+ - next - case False - with \ show ?thesis - by (cases \, auto split:ltln.split) (metis less_nat_zero_code neq0_conv - suffix_0 add.comm_neutral add_0)+ - qed -qed (auto split:ltln.split) - -function ltln_rewrite_rec -where - "ltln_rewrite_rec \ = (case ltln_rewrite_step \ of - \ and\<^sub>n \ \ (ltln_rewrite_rec \) and\<^sub>n (ltln_rewrite_rec \) - | \ or\<^sub>n \ \ (ltln_rewrite_rec \) or\<^sub>n (ltln_rewrite_rec \) - | X\<^sub>n \ \ X\<^sub>n (ltln_rewrite_rec \) - | \ U\<^sub>n \ \ (ltln_rewrite_rec \) U\<^sub>n (ltln_rewrite_rec \) - | \ V\<^sub>n \ \ (ltln_rewrite_rec \) V\<^sub>n (ltln_rewrite_rec \) - | \ \ \)" -by pat_completeness auto -termination proof - - { - fix \ \ :: "'a ltln" and thesis - assume "ltln_rewrite_step \ = \" - thm ltln_rewrite_step__size_leq - moreover assume "\ltln_rewrite_step \ = \; - size (local.ltln_rewrite_step \) \ size \\ \ thesis" - ultimately have thesis using ltln_rewrite_step__size_leq[of \] - by blast - } note AUX=this - show ?thesis - apply (relation "inv_image less_than size") - apply simp_all - apply (erule AUX, simp)+ - done -qed - -declare ltln_rewrite_rec.simps [simp del] - -lemma ltln_rewrite_rec__size_less: - assumes "ltln_rewrite_rec \ \ \" - shows "size (ltln_rewrite_rec \) < size \" - using assms -proof (induct "ltln_rewrite_rec \" arbitrary: \) - case LTLnTrue - then show ?case - using ltln_rewrite_step__size_less[of \] ltln_rewrite_rec.simps[of \] - by (cases "ltln_rewrite_step \") auto -next - case LTLnFalse - then show ?case - using ltln_rewrite_step__size_less[of \] ltln_rewrite_rec.simps[of \] - by (cases "ltln_rewrite_step \") auto -next - case LTLnProp - then show ?case - using ltln_rewrite_step__size_less[of \] ltln_rewrite_rec.simps[of \] - by (cases "ltln_rewrite_step \") auto -next - case LTLnNProp - then show ?case - using ltln_rewrite_step__size_less[of \] ltln_rewrite_rec.simps[of \] - by (cases "ltln_rewrite_step \") auto -next - case (LTLnAnd \ \) - show ?case - proof (cases "ltln_rewrite_step \") - case \: (LTLnAnd \' \') - show ?thesis - using LTLnAnd ltln_rewrite_step__size_less[of \] - unfolding \ ltln_rewrite_rec.simps[of \] - by (cases "\' and\<^sub>n \' = \") fastforce+ - qed (insert LTLnAnd, auto simp add: ltln_rewrite_rec.simps[of \]) -next - case (LTLnOr \ \) - show ?case - proof (cases "ltln_rewrite_step \") - case \: (LTLnOr \' \') - show ?thesis - using LTLnOr ltln_rewrite_step__size_less[of \] - unfolding \ ltln_rewrite_rec.simps[of \] - by (cases "\' or\<^sub>n \' = \") fastforce+ - qed (insert LTLnOr, auto simp add: ltln_rewrite_rec.simps[of \]) -next - case (LTLnNext \) - show ?case - proof (cases "ltln_rewrite_step \") - case \: (LTLnNext \') - show ?thesis - using LTLnNext ltln_rewrite_step__size_less[of \] - unfolding \ ltln_rewrite_rec.simps[of \] - by (cases "X\<^sub>n \' = \") fastforce+ - qed (insert LTLnNext, auto simp add: ltln_rewrite_rec.simps[of \]) -next - case (LTLnUntil \ \) - show ?case - proof (cases "ltln_rewrite_step \") - case \: (LTLnUntil \' \') - show ?thesis - using LTLnUntil ltln_rewrite_step__size_less[of \] - unfolding \ ltln_rewrite_rec.simps[of \] - by (cases "\' U\<^sub>n \' = \") fastforce+ - qed (insert LTLnUntil, auto simp add: ltln_rewrite_rec.simps[of \]) -next - case (LTLnRelease \ \) - show ?case - proof (cases "ltln_rewrite_step \") - case \: (LTLnRelease \' \') - show ?thesis - using LTLnRelease ltln_rewrite_step__size_less[of \] - unfolding \ ltln_rewrite_rec.simps[of \] - by (cases "\' V\<^sub>n \' = \") fastforce+ - qed (insert LTLnRelease, auto simp add: ltln_rewrite_rec.simps[of \]) -qed - -lemma ltln_rewrite_rec__size_leq: - "size (ltln_rewrite_rec \) \ size \" -using ltln_rewrite_rec__size_less[of \] by (cases "ltln_rewrite_rec \ = \") auto - - -theorem ltln_rewrite_rec__equiv: "\ \\<^sub>n ltln_rewrite_rec \ \ \ \\<^sub>n \" -proof (induct "ltln_rewrite_rec \" arbitrary: \ \) - case LTLnTrue - then show ?case - using ltln_rewrite_step__equiv[of _ \] ltln_rewrite_rec.simps[of \] - by (cases "ltln_rewrite_step \") auto -next - case LTLnFalse - then show ?case - using ltln_rewrite_step__equiv[of _ \] ltln_rewrite_rec.simps[of \] - by (cases "ltln_rewrite_step \") auto -next - case LTLnProp - then show ?case - using ltln_rewrite_step__equiv[of _ \] ltln_rewrite_rec.simps[of \] - by (cases "ltln_rewrite_step \") auto -next - case LTLnNProp - then show ?case - using ltln_rewrite_step__equiv[of _ \] ltln_rewrite_rec.simps[of \] - by (cases "ltln_rewrite_step \") auto -next - case LTLnAnd - then show ?case - using ltln_rewrite_step__equiv[of _ \] ltln_rewrite_rec.simps[of \] - by (cases "ltln_rewrite_step \") auto -next - case LTLnOr - then show ?case - using ltln_rewrite_step__equiv[of _ \] ltln_rewrite_rec.simps[of \] - by (cases "ltln_rewrite_step \") auto -next - case LTLnNext - then show ?case - using ltln_rewrite_step__equiv[of _ \] ltln_rewrite_rec.simps[of \] - by (cases "ltln_rewrite_step \") auto -next - case LTLnUntil - then show ?case - using ltln_rewrite_step__equiv[of _ \] ltln_rewrite_rec.simps[of \] - by (cases "ltln_rewrite_step \") auto -next - case LTLnRelease - then show ?case - using ltln_rewrite_step__equiv[of _ \] ltln_rewrite_rec.simps[of \] - by (cases "ltln_rewrite_step \") auto -qed - - -function ltln_rewrite -where - "ltln_rewrite \ = - (let \ = ltln_rewrite_rec \ - in if \ \ \ then ltln_rewrite \ else \)" -by pat_completeness auto - -termination -proof (relation "inv_image less_than size", goal_cases) - case 1 - show ?case by auto -next - case (2 \) - then show ?case - using ltln_rewrite_rec__size_less[of \] by auto -qed - -declare ltln_rewrite.simps [simp del] - - -lemma ltln_rewrite__size_less: - assumes "ltln_rewrite \ \ \" - shows "size (ltln_rewrite \) < size \" - using assms -proof (induct \ rule: ltln_rewrite.induct) - case prems: (1 \) - show ?case - proof (cases "ltln_rewrite_rec \ = \") - case True - then show ?thesis - using prems(2) unfolding ltln_rewrite.simps[of \] by auto - next - case False - then have rw_\_eq: "ltln_rewrite \ = ltln_rewrite (ltln_rewrite_rec \)" - unfolding ltln_rewrite.simps[of \] by auto - from prems(1)[OF refl False, folded rw_\_eq] - and ltln_rewrite_rec__size_less[OF False] - show ?thesis - by (cases "ltln_rewrite \ = ltln_rewrite_rec \") auto - qed -qed - -lemma ltln_rewrite__size_leq: "size (ltln_rewrite \) \ size \" - using ltln_rewrite__size_less[of \] by (cases "ltln_rewrite \ = \") auto - - -lemma ltln_rewrite__equiv: "\ \\<^sub>n ltln_rewrite \ \ \ \\<^sub>n \" -proof (induct \ rule: ltln_rewrite.induct) - case prems: (1 \) - show ?case - using prems[OF refl] ltln_rewrite_rec__equiv[of \ \] - unfolding ltln_rewrite.simps[of \] - by (cases "ltln_rewrite_rec \ = \") auto -qed - - -fun ltln_pure_eventual_frmls_impl -where - "ltln_pure_eventual_frmls_impl (\\<^sub>n \) = True" -| "ltln_pure_eventual_frmls_impl (\ and\<^sub>n \) - = (ltln_pure_eventual_frmls_impl \ \ ltln_pure_eventual_frmls_impl \)" -| "ltln_pure_eventual_frmls_impl (\ or\<^sub>n \) - = (ltln_pure_eventual_frmls_impl \ \ ltln_pure_eventual_frmls_impl \)" -| "ltln_pure_eventual_frmls_impl (\ U\<^sub>n \) - = (ltln_pure_eventual_frmls_impl \ \ ltln_pure_eventual_frmls_impl \)" -| "ltln_pure_eventual_frmls_impl (\ V\<^sub>n \) - = (ltln_pure_eventual_frmls_impl \ \ ltln_pure_eventual_frmls_impl \)" -| "ltln_pure_eventual_frmls_impl _ = False" - -lemma ltln_pure_eventual_frmls_unfold[code_unfold]: - "\ \ ltln_pure_eventual_frmls \ ltln_pure_eventual_frmls_impl \" - (is "?lhs = ?rhs") -proof - assume ?lhs - then show ?rhs - proof (induct rule: ltln_pure_eventual_frmls.induct) - case (4 \) - then show ?case by (cases "\ = true\<^sub>n") (cases \, auto)+ - qed auto -next - assume ?rhs - then show ?lhs - proof (induct \) - case (LTLnUntil \ \) - then show ?case - by (cases "\ = true\<^sub>n") (cases \, - auto intro: ltln_pure_eventual_frmls.intros)+ - qed (auto intro: ltln_pure_eventual_frmls.intros) -qed - -fun ltln_pure_universal_frmls_impl -where - "ltln_pure_universal_frmls_impl (\\<^sub>n \) = True" -| "ltln_pure_universal_frmls_impl (\ and\<^sub>n \) - = (ltln_pure_universal_frmls_impl \ \ ltln_pure_universal_frmls_impl \)" -| "ltln_pure_universal_frmls_impl (\ or\<^sub>n \) - = (ltln_pure_universal_frmls_impl \ \ ltln_pure_universal_frmls_impl \)" -| "ltln_pure_universal_frmls_impl (\ U\<^sub>n \) - = (ltln_pure_universal_frmls_impl \ \ ltln_pure_universal_frmls_impl \)" -| "ltln_pure_universal_frmls_impl (\ V\<^sub>n \) - = (ltln_pure_universal_frmls_impl \ \ ltln_pure_universal_frmls_impl \)" -| "ltln_pure_universal_frmls_impl _ = False" - -lemma ltln_pure_universal_frmls_unfold[code_unfold]: - "\ \ ltln_pure_universal_frmls \ ltln_pure_universal_frmls_impl \" - (is "?lhs = ?rhs") -proof - assume ?lhs - then show ?rhs - proof (induct rule: ltln_pure_universal_frmls.induct) - case (5 \) - then show ?case by (cases "\ = false\<^sub>n") (cases \, auto)+ - qed auto -next - assume ?rhs - then show ?lhs - proof (induct \) - case (LTLnRelease \ \) - then show ?case - by (cases "\ = false\<^sub>n") (cases \, auto intro: ltln_pure_universal_frmls.intros)+ - qed (auto intro: ltln_pure_universal_frmls.intros) -qed - -definition - "ltln_rewrite_step_impl \ \ case \ of - \ U\<^sub>n \ \ if \ = true\<^sub>n then true\<^sub>n - else ( - case (\, \) of - (true\<^sub>n, _ U\<^sub>n \') \ true\<^sub>n U\<^sub>n \' - | _ \ if \ = \ then \ - else if \\ltln_pure_eventual_frmls then \ - else \) - | \ V\<^sub>n \ \ if \ = false\<^sub>n then false\<^sub>n - else ( - case (\, \) of - (false\<^sub>n, _ V\<^sub>n \') \ false\<^sub>n V\<^sub>n \' - | _ \ if \ = \ then \ - else if \\ltln_pure_universal_frmls then \ - else \) - | \1 and\<^sub>n \2 \ ( - case (\1, \2) of - (\ U\<^sub>n \, \ U\<^sub>n \') \ if \ = \' then (\ and\<^sub>n \) U\<^sub>n \ else \ - | (\ V\<^sub>n \, \' V\<^sub>n \) \ if \ = \' then \ V\<^sub>n (\ and\<^sub>n \) else \ - | _ \ \) - | \1 or\<^sub>n \2 \ ( - case (\1, \2) of - (\ U\<^sub>n \, \' U\<^sub>n \) \ if \ = \' then \ U\<^sub>n (\ or\<^sub>n \) else \ - | (\ V\<^sub>n \, \ V\<^sub>n \') \ if \ = \' then (\ or\<^sub>n \) V\<^sub>n \ else \ - | _ \ \) - | _ \ \" - -lemma ltln_rewrite_step_unfold[code_unfold]: - "ltln_rewrite_step = ltln_rewrite_step_impl" -proof - - { fix \ :: "'a ltln" - have "ltln_rewrite_step \ = ltln_rewrite_step_impl \" - unfolding ltln_rewrite_step_impl_def by (cases \) (auto split:ltln.split) } - then show ?thesis by auto -qed - -end - -end diff --git a/thys/LTL_to_GBA/LTL_Stutter.thy b/thys/LTL_to_GBA/LTL_Stutter.thy --- a/thys/LTL_to_GBA/LTL_Stutter.thy +++ b/thys/LTL_to_GBA/LTL_Stutter.thy @@ -1,114 +1,53 @@ section \Stutter Invariance of next-free LTL Formula\ -theory LTL_Stutter -imports LTL "../Stuttering_Equivalence/PLTL" -begin - text \This theory builds on the AFP-entry by Stephan Merz\ - - text \Get rid of overlapping notation\ - no_notation PLTL.holds_of ("_ \ _" [70,70] 40) - - hide_const (open) PLTL.false PLTL.atom PLTL.implies PLTL.next PLTL.until - hide_const (open) PLTL.not PLTL.true PLTL.or PLTL.and - hide_const (open) PLTL.eventually PLTL.always - - no_notation Samplers.suffix ("_[_ ..]" [80, 15] 80) - hide_const (open) Samplers.suffix - hide_fact (open) Samplers.suffix_def - - lemma PLTL_suffix_cnv[simp]: "Samplers.suffix = (\w i. suffix i w)" - apply (intro ext) - unfolding Samplers.suffix_def Omega_Words_Fun.suffix_def .. - - hide_const (open) PLTL.next_free - - primrec ltl_next_free :: "'a ltl \ bool" where - "ltl_next_free LTLTrue \ True" - | "ltl_next_free LTLFalse \ True" - | "ltl_next_free (LTLProp _) \ True" - | "ltl_next_free (LTLNeg \) \ ltl_next_free \" - | "ltl_next_free (LTLAnd \ \) \ ltl_next_free \ \ ltl_next_free \" - | "ltl_next_free (LTLOr \ \) \ ltl_next_free \ \ ltl_next_free \" - | "ltl_next_free (LTLNext _) \ False" - | "ltl_next_free (LTLUntil \ \) \ ltl_next_free \ \ ltl_next_free \" - | "ltl_next_free (LTLRelease \ \) \ ltl_next_free \ \ ltl_next_free \" - - text \Conversion between the two LTL formalizations\ - - primrec cnv :: "'a LTL.ltl \ 'a set PLTL.pltl" where - "cnv LTLTrue = PLTL.true" - | "cnv LTLFalse = PLTL.false" - | "cnv (LTLProp a) = PLTL.atom (op \ a)" - | "cnv (LTLNeg \) = PLTL.not (cnv \)" - | "cnv (LTLAnd \ \) = PLTL.and (cnv \) (cnv \)" - | "cnv (LTLOr \ \) = PLTL.or (cnv \) (cnv \)" - | "cnv (LTLNext \) = PLTL.next (cnv \)" - | "cnv (LTLUntil \ \) = PLTL.until (cnv \) (cnv \)" - | "cnv (LTLRelease \ \) - = PLTL.not (PLTL.until (PLTL.not (cnv \)) (PLTL.not (cnv \)))" - - lemma PLTL_holds_of_cnv[simp]: "PLTL.holds_of r (cnv \) \ ltl_semantics r \" - by (induction \ arbitrary: r) simp_all - - lemma PLTL_next_free_cnv[simp]: "PLTL.next_free (cnv \) \ ltl_next_free \" - by (induction \) auto - - theorem next_free_stutter_invariant: - assumes next_free: "ltl_next_free \" - assumes eq: "r \ r'" - shows "r \ \ \ r' \ \" - -- \A next free formula cannot distinguish between - stutter-equivalent runs.\ - proof - - { - fix r r' - assume eq: "r \ r'" and holds: "r \ \" - then have "holds_of r (cnv \)" by simp - - from next_free_stutter_invariant[of "cnv \"] next_free - have "PLTL.stutter_invariant (cnv \)" by simp - from stutter_invariantD[OF this eq] holds have "r' \ \" by simp - } note aux=this - - from aux[of r r'] aux[of r' r] eq stutter_equiv_sym[OF eq] show ?thesis - by blast - qed - +theory LTL_Stutter + imports "../../../afp/LTL/LTL" "../Stuttering_Equivalence/PLTL" +begin - context begin interpretation LTL_Syntax . - primrec ltlc_next_free :: "'a ltlc \ bool" - where - "ltlc_next_free true\<^sub>c = True" - | "ltlc_next_free false\<^sub>c = True" - | "ltlc_next_free (prop\<^sub>c(q)) = True" - | "ltlc_next_free (not\<^sub>c \) = ltlc_next_free \" - | "ltlc_next_free (\ and\<^sub>c \) = (ltlc_next_free \ \ ltlc_next_free \)" - | "ltlc_next_free (\ or\<^sub>c \) = (ltlc_next_free \ \ ltlc_next_free \)" - | "ltlc_next_free (\ implies\<^sub>c \) = (ltlc_next_free \ \ ltlc_next_free \)" - | "ltlc_next_free (\ iff\<^sub>c \) = (ltlc_next_free \ \ ltlc_next_free \)" - | "ltlc_next_free (X\<^sub>c \) = False" - | "ltlc_next_free (F\<^sub>c \) = ltlc_next_free \" - | "ltlc_next_free (G\<^sub>c \) = ltlc_next_free \" - | "ltlc_next_free (\ U\<^sub>c \) = (ltlc_next_free \ \ ltlc_next_free \)" - | "ltlc_next_free (\ V\<^sub>c \) = (ltlc_next_free \ \ ltlc_next_free \)" - end +text \This theory builds on the AFP-entry by Stephan Merz\ - lemma ltlc_to_ltl_next_free_iff: - "ltl_next_free (ltlc_to_ltl \) \ ltlc_next_free \" - by (induction \) (auto simp: Let_def) - +text \Get rid of overlapping notation\ - theorem ltlc_next_free_stutter_invariant: - assumes next_free: "ltlc_next_free \" - assumes eq: "r \ r'" - shows "r \\<^sub>c \ \ r' \\<^sub>c \" - -- \A next free formula cannot distinguish between - stutter-equivalent runs.\ - proof - - note NF' = next_free[folded ltlc_to_ltl_next_free_iff] - from next_free_stutter_invariant[OF NF' eq] show ?thesis - by (auto simp: ltlc_to_ltl_equiv) - qed +primrec ltlc_next_free :: "'a ltlc \ bool" + where + "ltlc_next_free true\<^sub>c = True" + | "ltlc_next_free false\<^sub>c = True" + | "ltlc_next_free (prop\<^sub>c(q)) = True" + | "ltlc_next_free (not\<^sub>c \) = ltlc_next_free \" + | "ltlc_next_free (\ and\<^sub>c \) = (ltlc_next_free \ \ ltlc_next_free \)" + | "ltlc_next_free (\ or\<^sub>c \) = (ltlc_next_free \ \ ltlc_next_free \)" + | "ltlc_next_free (\ implies\<^sub>c \) = (ltlc_next_free \ \ ltlc_next_free \)" + | "ltlc_next_free (X\<^sub>c \) = False" + | "ltlc_next_free (F\<^sub>c \) = ltlc_next_free \" + | "ltlc_next_free (G\<^sub>c \) = ltlc_next_free \" + | "ltlc_next_free (\ U\<^sub>c \) = (ltlc_next_free \ \ ltlc_next_free \)" + | "ltlc_next_free (\ V\<^sub>c \) = (ltlc_next_free \ \ ltlc_next_free \)" + +text \Conversion between the two LTL formalizations\ + +lemma PLTL_next_free_cnv[simp]: "PLTL.next_free (ltlc_to_pltl \) \ ltlc_next_free \" + by (induction \) auto + +theorem next_free_stutter_invariant: + assumes next_free: "ltlc_next_free \" + assumes eq: "r \ r'" + shows "r \\<^sub>c \ \ r' \\<^sub>c \" + -- \A next free formula cannot distinguish between + stutter-equivalent runs.\ +proof - + { + fix r r' + assume eq: "r \ r'" and holds: "r \\<^sub>c \" + then have "r \\<^sub>p (ltlc_to_pltl \)"by simp + + from next_free_stutter_invariant[of "ltlc_to_pltl \"] next_free + have "PLTL.stutter_invariant (ltlc_to_pltl \)" by simp + from stutter_invariantD[OF this eq] holds have "r' \\<^sub>c \" by simp + } note aux=this + + from aux[of r r'] aux[of r' r] eq stutter_equiv_sym[OF eq] show ?thesis + by blast +qed end diff --git a/thys/LTL_to_GBA/LTL_to_GBA.thy b/thys/LTL_to_GBA/LTL_to_GBA.thy --- a/thys/LTL_to_GBA/LTL_to_GBA.thy +++ b/thys/LTL_to_GBA/LTL_to_GBA.thy @@ -1,3235 +1,3231 @@ section \LTL to GBA translation\ (* Author: Alexander Schimpf Modified by Peter Lammich **) theory LTL_to_GBA imports "../CAVA_Automata/CAVA_Base/CAVA_Base" - LTL LTL_Rewrite + "../../../afp/LTL/LTL" "../../../afp/LTL/LTL_Rewrite" "../CAVA_Automata/Automata" begin subsection \Statistics\ code_printing code_module Gerth_Statistics \ (SML) \ structure Gerth_Statistics = struct val active = Unsynchronized.ref false val data = Unsynchronized.ref (0,0,0) fun is_active () = !active fun set_data num_states num_init num_acc = ( active := true; data := (num_states, num_init, num_acc) ) fun to_string () = let val (num_states, num_init, num_acc) = !data in "Num states: " ^ IntInf.toString (num_states) ^ "\n" ^ " Num initial: " ^ IntInf.toString num_init ^ "\n" ^ " Num acc-classes: " ^ IntInf.toString num_acc ^ "\n" end val _ = Statistics.register_stat ("Gerth LTL_to_GBA",is_active,to_string) end \ code_reserved SML Gerth_Statistics consts stat_set_data_int :: "integer \ integer \ integer \ unit" code_printing constant stat_set_data_int \ (SML) "Gerth'_Statistics.set'_data" definition "stat_set_data ns ni na \ stat_set_data_int (integer_of_nat ns) (integer_of_nat ni) (integer_of_nat na)" lemma [autoref_rules]: "(stat_set_data,stat_set_data) \ nat_rel \ nat_rel \ nat_rel \ unit_rel" by auto abbreviation "stat_set_data_nres ns ni na \ RETURN (stat_set_data ns ni na)" lemma discard_stat_refine[refine]: "m1\m2 \ stat_set_data_nres ns ni na \ m1 \ m2" by simp_all subsection \Preliminaries\ text \Some very special lemmas for reasoning about the nres-monad\ lemma SPEC_rule_nested2: "\m \ SPEC P; \r1 r2. P (r1, r2) \ g (r1, r2) \ SPEC P\ \ m \ SPEC (\r'. g r' \ SPEC P)" by (simp add: pw_le_iff) blast lemma SPEC_rule_param2: assumes "f x \ SPEC (P x)" and "\r1 r2. (P x) (r1, r2) \ (P x') (r1, r2)" shows "f x \ SPEC (P x')" using assms by (simp add: pw_le_iff) lemma SPEC_rule_weak: assumes "f x \ SPEC (Q x)" and "f x \ SPEC (P x)" and "\r1 r2. \(Q x) (r1, r2); (P x) (r1, r2)\ \ (P x') (r1, r2)" shows "f x \ SPEC (P x')" using assms by (simp add: pw_le_iff) lemma SPEC_rule_weak_nested2: "\f \ SPEC Q; f \ SPEC P; \r1 r2. \Q (r1, r2); P (r1, r2)\ \ g (r1, r2) \ SPEC P\ \ f \ SPEC (\r'. g r' \ SPEC P)" by (simp add: pw_le_iff) blast subsection \Creation of States\ text \ In this section, the first part of the algorithm, which creates the states of the automaton, is formalized. \ (* FIXME: Abstraktion über node_name *) type_synonym node_name = nat type_synonym 'a frml = "'a ltln" type_synonym 'a interprt = "'a set word" record 'a node = name :: node_name incoming :: "node_name set" new :: "'a frml set" old :: "'a frml set" "next" :: "'a frml set" context begin -interpretation LTL_Syntax . - fun frml_neg where "frml_neg true\<^sub>n = false\<^sub>n" | "frml_neg false\<^sub>n = true\<^sub>n" | "frml_neg prop\<^sub>n(q) = nprop\<^sub>n(q)" | "frml_neg nprop\<^sub>n(q) = prop\<^sub>n(q)" fun new1 where "new1 (\ and\<^sub>n \) = {\,\}" | "new1 (\ U\<^sub>n \) = {\}" | "new1 (\ V\<^sub>n \) = {\}" | "new1 (\ or\<^sub>n \) = {\}" | "new1 _ = {}" fun next1 where "next1 (X\<^sub>n \) = {\}" | "next1 (\ U\<^sub>n \) = {\ U\<^sub>n \}" | "next1 (\ V\<^sub>n \) = {\ V\<^sub>n \}" | "next1 _ = {}" fun new2 where "new2 (\ U\<^sub>n \) = {\}" | "new2 (\ V\<^sub>n \) = {\, \}" | "new2 (\ or\<^sub>n \) = {\}" | "new2 _ = {}" (* FIXME: Abstraction over concrete definition *) definition "expand_init \ 0" definition "expand_new_name \ Suc" lemma expand_new_name_expand_init: "expand_init < expand_new_name nm" by (auto simp add:expand_init_def expand_new_name_def) lemma expand_new_name_step[intro]: fixes n :: "'a node" shows "name n < expand_new_name (name n)" by (auto simp add:expand_new_name_def) lemma expand_new_name__less_zero[intro]: "0 < expand_new_name nm" proof - from expand_new_name_expand_init have "expand_init < expand_new_name nm" by auto then show ?thesis by (metis gr0I less_zeroE) qed abbreviation "upd_incoming_f n \ (\n'. if (old n' = old n \ next n' = next n) then n'\ incoming := incoming n \ incoming n' \ else n' )" definition "upd_incoming n ns \ ((upd_incoming_f n) ` ns)" lemma upd_incoming__elem: assumes "nd\upd_incoming n ns" shows "nd\ns \ (\nd'\ns. nd = nd'\ incoming := incoming n \ incoming nd' \ \ old nd' = old n \ next nd' = next n)" proof - obtain nd' where "nd'\ns" and nd_eq: "nd = upd_incoming_f n nd'" using assms unfolding upd_incoming_def by blast then show ?thesis by auto qed lemma upd_incoming__ident_node: assumes "nd\upd_incoming n ns" and "nd\ns" shows "incoming n \ incoming nd \ \ (old nd = old n \ next nd = next n)" proof (rule ccontr) assume "\ ?thesis" then have nsubset: "\ (incoming n \ incoming nd)" and old_next_eq: "old nd = old n \ next nd = next n" by auto moreover obtain nd' where "nd'\ns" and nd_eq: "nd = upd_incoming_f n nd'" using assms unfolding upd_incoming_def by blast { assume "old nd' = old n \ next nd' = next n" with nd_eq have "nd = nd'\incoming := incoming n \ incoming nd'\" by auto with nsubset have "False" by auto } moreover { assume "\ (old nd' = old n \ next nd' = next n)" with nd_eq old_next_eq have "False" by auto } ultimately show "False" by fast qed lemma upd_incoming__ident: assumes "\n\ns. P n" and "\n. \n\ns; P n\ \ (\v. P (n\ incoming := v \))" shows "\n\upd_incoming n ns. P n" proof fix nd v let ?f = "\n'. if (old n' = old n \ next n' = next n) then n'\ incoming := incoming n \ incoming n' \ else n'" assume "nd\upd_incoming n ns" then obtain nd' where "nd'\ns" and nd_eq: "nd = ?f nd'" unfolding upd_incoming_def by blast { assume "old nd' = old n \ next nd' = next n" then obtain v where "nd = nd'\ incoming := v \" using nd_eq by auto with assms \nd'\ns\ have "P nd" by auto } then show "P nd" using nd_eq \nd'\ns\ assms by auto qed lemma name_upd_incoming_f[simp]: "name (upd_incoming_f n x) = name x" by auto lemma name_upd_incoming[simp]: "name ` (upd_incoming n ns) = name ` ns" (is "?lhs = ?rhs") unfolding upd_incoming_def proof safe fix x assume "x\ns" then have "upd_incoming_f n x \ (\n'. local.upd_incoming_f n n') ` ns" by auto then have "name (upd_incoming_f n x) \ name ` (\n'. local.upd_incoming_f n n') ` ns" by blast then show "name x \ name ` (\n'. local.upd_incoming_f n n') ` ns" by simp qed auto abbreviation expand_body where "expand_body \ (\expand (n,ns). if new n = {} then ( if (\n'\ns. old n' = old n \ next n' = next n) then RETURN (name n, upd_incoming n ns) else expand ( \ name=expand_new_name (name n), incoming={name n}, new=next n, old={}, next={} \, {n} \ ns) ) else do { \ \ SPEC (\x. x\(new n)); let n = n\ new := new n - {\} \; if (\q. \ = prop\<^sub>n(q) \ \ = nprop\<^sub>n(q)) then (if (frml_neg \)\old n then RETURN (name n, ns) else expand (n\ old := {\} \ old n \, ns)) else if \ = true\<^sub>n then expand (n\ old := {\} \ old n \, ns) else if \ = false\<^sub>n then RETURN (name n, ns) else if (\\ \. (\ = \ and\<^sub>n \) \ (\ = X\<^sub>n \)) then expand ( n\ new := new1 \ \ new n, old := {\} \ old n, next := next1 \ \ next n \, ns) else do { (nm, nds) \ expand ( n\ new := new1 \ \ new n, old := {\} \ old n, next := next1 \ \ next n \, ns); expand (n\ name := nm, new := new2 \ \ new n, old := {\} \ old n \, nds) } } )" lemma expand_body_mono: "trimono expand_body" by refine_mono definition expand :: "('a node \ ('a node set)) \ (node_name \ 'a node set) nres" where "expand \ REC expand_body" lemma REC_rule_old: (* TODO: Adapt proofs below, have fun with subgoal 27 ...*) fixes x::"'x" assumes M: "trimono body" assumes I0: "\ x" assumes IS: "\f x. \ \x. \ x \ f x \ M x; \ x; f \ REC body \ \ body f x \ M x" shows "REC body x \ M x" by (rule REC_rule_arb[where pre="\_. \" and M="\_. M", OF assms]) lemma expand_rec_rule: assumes I0: "\ x" assumes IS: "\f x. \ \x. f x \ expand x; \x. \ x \ f x \ M x; \ x \ \ expand_body f x \ M x" shows "expand x \ M x" unfolding expand_def apply (rule REC_rule_old[where \="\"]) apply (rule expand_body_mono) apply (rule I0) apply (rule IS[unfolded expand_def]) apply (blast dest: le_funD) apply blast apply blast done abbreviation "expand_assm_incoming n_ns \ (\nm\incoming (fst n_ns). name (fst n_ns) > nm) \ 0 < name (fst n_ns) \ (\q\snd n_ns. name (fst n_ns) > name q \ (\nm\incoming q. name (fst n_ns) > nm))" abbreviation "expand_rslt_incoming nm_nds \ (\q\snd nm_nds. (fst nm_nds > name q \ (\nm'\incoming q. fst nm_nds > nm')))" abbreviation "expand_rslt_name n_ns nm_nds \ (name (fst n_ns) \ fst nm_nds \ name ` (snd n_ns) \ name ` (snd nm_nds)) \ name ` (snd nm_nds) = name ` (snd n_ns) \ name ` {nd\snd nm_nds. name nd \ name (fst n_ns)}" abbreviation "expand_name_ident nds \ (\q\nds. \!q'\nds. name q = name q')" abbreviation "expand_assm_exist \ n_ns \ {\. \\. \ U\<^sub>n \ \ old (fst n_ns) \ \ \\<^sub>n \} \ new (fst n_ns) \ old (fst n_ns) \ (\\\new (fst n_ns). \ \\<^sub>n \) \ (\\\old (fst n_ns). \ \\<^sub>n \) \ (\\\next (fst n_ns). \ \\<^sub>n X\<^sub>n \)" abbreviation "expand_rslt_exist__node_prop \ n nd \ incoming n \ incoming nd \ (\\\old nd. \ \\<^sub>n \) \ (\\\next nd. \ \\<^sub>n X\<^sub>n \) \ {\. \\. \ U\<^sub>n \ \ old nd \ \ \\<^sub>n \} \ old nd" abbreviation "expand_rslt_exist \ n_ns nm_nds \ (\nd\snd nm_nds. expand_rslt_exist__node_prop \ (fst n_ns) nd)" abbreviation "expand_rslt_exist_eq__node n nd \ name n = name nd \ old n = old nd \ next n = next nd \ incoming n \ incoming nd" abbreviation "expand_rslt_exist_eq n_ns nm_nds \ (\n\snd n_ns. \nd\snd nm_nds. expand_rslt_exist_eq__node n nd)" lemma expand_name_propag: assumes "expand_assm_incoming n_ns \ expand_name_ident (snd n_ns)" (is "?Q n_ns") shows "expand n_ns \ SPEC (\r. expand_rslt_incoming r \ expand_rslt_name n_ns r \ expand_name_ident (snd r))" (is "expand _ \ SPEC (?P n_ns)") using assms proof (rule_tac expand_rec_rule[where \="?Q"], simp, intro refine_vcg, goal_cases) case prems: (1 _ _ n ns) then have Q: "?Q (n, ns)" by fast let ?nds = "upd_incoming n ns" from prems have "\q\?nds. name q < name n" by (rule_tac upd_incoming__ident) auto moreover have "\q\?nds. \nm'\incoming q. nm' < name n" (is "\q\_. ?sg q") proof fix q assume q_in:"q\?nds" show "?sg q" proof (cases "q\ns") case True with prems show ?thesis by auto next case False with upd_incoming__elem[OF q_in] obtain nd' where nd'_def:"nd'\ns \ q = nd'\incoming := incoming n \ incoming nd'\" by blast { fix nm' assume "nm'\incoming q" moreover { assume "nm'\incoming n" with Q have "nm' < name n" by auto } moreover { assume "nm'\incoming nd'" with Q nd'_def have "nm' < name n" by auto } ultimately have "nm' < name n" using nd'_def by auto } then show ?thesis by fast qed qed moreover have "expand_name_ident ?nds" proof (rule upd_incoming__ident, goal_cases) case 1 show ?case proof fix q assume "q\ns" with Q have "\!q'\ns. name q = name q'" by auto then obtain q' where "q'\ns" and "name q = name q'" and q'_all: "\q''\ns. name q' = name q'' \ q' = q''" by auto let ?q' = "upd_incoming_f n q'" have P_a: "?q'\?nds \ name q = name ?q'" using \q'\ns\ \name q = name q'\ q'_all unfolding upd_incoming_def by auto have P_all: "\q''\?nds. name ?q' = name q'' \ ?q' = q''" proof(clarify) fix q'' assume "q''\?nds" and q''_name_eq: "name ?q' = name q''" { assume "q''\ns" with upd_incoming__elem[OF \q''\?nds\] obtain nd'' where "nd''\ns" and q''_is: "q'' = nd''\incoming := incoming n \ incoming nd''\ \ old nd'' = old n \ next nd'' = next n" by auto then have "name nd'' = name q'" using q''_name_eq by (cases "old q' = old n \ next q' = next n") auto with \nd''\ns\ q'_all have "nd'' = q'" by auto then have "?q' = q''" using q''_is by simp } moreover { assume "q''\ns" moreover have "name q' = name q''" using q''_name_eq by (cases "old q' = old n \ next q' = next n") auto moreover then have "incoming n \ incoming q'' \ incoming q'' = incoming n \ incoming q''" by auto ultimately have "?q' = q''" using upd_incoming__ident_node[OF \q''\?nds\] q'_all by auto } ultimately show "?q' = q''" by fast qed show "\!q'\upd_incoming n ns. name q = name q'" proof(rule ex1I[of _ ?q'], goal_cases) case 1 then show ?case using P_a by simp next case 2 then show ?case using P_all unfolding P_a[THEN conjunct2, THEN sym] by blast qed qed qed simp ultimately show ?case using prems by auto next case prems: (2 f x n ns) then have step: "\x. ?Q x \ f x \ SPEC (?P x)" by simp from prems have Q: "?Q (n, ns)" by auto show ?case unfolding \x = (n, ns)\ proof (rule_tac SPEC_rule_param2[where P = "?P"], rule_tac step, goal_cases) case 1 with expand_new_name_step[of n] show ?case using Q by (auto elim: order_less_trans[rotated]) next case prems': (2 _ nds) then have "name ` ns \ name ` nds" by auto with expand_new_name_step[of n] show ?case using prems' by auto qed next case 3 then show ?case by auto next case prems: (4 f) then have step: "\x. ?Q x \ f x \ SPEC (?P x)" by simp_all with prems show ?case by (rule_tac SPEC_rule_param2, rule_tac step) auto next case prems: (5 f) then have step: "\x. ?Q x \ f x \ SPEC (?P x)" by simp_all from prems show ?case by (rule_tac SPEC_rule_param2, rule_tac step) auto next case 6 then show ?case by auto next case prems: (7 f) then have step: "\x. ?Q x \ f x \ SPEC (?P x)" by simp_all from prems show ?case by (rule_tac SPEC_rule_param2, rule_tac step) auto next case prems: (8 f x n ns \) then have goal_assms: "\ \ new n \ (\\ \. \ = \ or\<^sub>n \ \ \ = \ U\<^sub>n \ \ \ = \ V\<^sub>n \)" by (cases \) auto from prems have Q: "?Q (n, ns)" and step: "\x. ?Q x \ f x \ SPEC (?P x)" by simp_all show ?case using goal_assms Q unfolding case_prod_unfold \x = (n, ns)\ proof (rule_tac SPEC_rule_nested2, goal_cases) case 1 then show ?case by (rule_tac SPEC_rule_param2, rule_tac step) auto next case prems: (2 nm nds) then have P_x: "(name n \ nm \ name ` ns \ name ` nds) \ name ` nds = name ` ns \ name ` {nd\nds. name nd \ name n}" (is "_ \ ?nodes_eq nds ns (name n)") by auto show ?case proof (rule_tac SPEC_rule_param2[where P = "?P"], goal_cases) case 1 with prems show ?case by (rule_tac step) auto next case prems': (2 nm' nds') then have "\q\nds'. name q < nm' \ (\nm''\incoming q. nm'' < nm')" by auto moreover have "?nodes_eq nds ns (name n)" and "?nodes_eq nds' nds nm" and "name n \ nm" using prems' P_x by auto then have "?nodes_eq nds' ns (name n)" by auto then have "expand_rslt_name (n, ns) (nm', nds')" using prems' P_x subset_trans[of "name ` ns" "name ` nds"] by auto ultimately show ?case using prems' by auto qed qed qed lemmas expand_name_propag__incoming = SPEC_rule_conjunct1[OF expand_name_propag] lemmas expand_name_propag__name = SPEC_rule_conjunct1[OF SPEC_rule_conjunct2[OF expand_name_propag]] lemmas expand_name_propag__name_ident = SPEC_rule_conjunct2[OF SPEC_rule_conjunct2[OF expand_name_propag]] lemma expand_rslt_exist_eq: shows "expand n_ns \ SPEC (expand_rslt_exist_eq n_ns)" (is "_ \ SPEC (?P n_ns)") using assms proof (rule_tac expand_rec_rule[where \="\_. True"], simp, intro refine_vcg, goal_cases) case prems: (1 f x n ns) let ?r = "(name n, upd_incoming n ns)" have "expand_rslt_exist_eq (n, ns) ?r" unfolding snd_conv proof fix n' assume "n'\ns" { assume "old n' = old n \ next n' = next n" with \n'\ns\ have "n'\ incoming := incoming n \ incoming n' \ \ upd_incoming n ns" unfolding upd_incoming_def by auto } moreover { assume "\ (old n' = old n \ next n' = next n)" with \n'\ns\ have "n' \ upd_incoming n ns" unfolding upd_incoming_def by auto } ultimately show "\nd\upd_incoming n ns. expand_rslt_exist_eq__node n' nd" by force qed with prems show ?case by auto next case prems: (2 f) then have step: "\x. f x \ SPEC (?P x)" by simp with prems show ?case by (rule_tac SPEC_rule_param2, rule_tac step) auto next case 3 then show ?case by auto next case prems: (4 f) then have step: "\x. f x \ SPEC (?P x)" by simp with prems show ?case by (rule_tac SPEC_rule_param2, rule_tac step) auto next case prems: (5 f) then have step: "\x. f x \ SPEC (?P x)" by simp with prems show ?case by (rule_tac SPEC_rule_param2, rule_tac step) auto next case 6 then show ?case by auto next case prems: (7 f) then have step: "\x. f x \ SPEC (?P x)" by simp with prems show ?case by (rule_tac SPEC_rule_param2, rule_tac step) auto next case prems: (8 f x n ns) then have step: "\x. f x \ SPEC (?P x)" by simp show ?case proof (rule_tac SPEC_rule_nested2, goal_cases) case 1 with prems show ?case by (rule_tac SPEC_rule_param2, rule_tac step) auto next case (2 nm nds) with prems have P_x: "?P (n, ns) (nm, nds)" by fast show ?case unfolding case_prod_unfold \x = (n, ns)\ proof (rule_tac SPEC_rule_param2[where P = "?P"], goal_cases) case 1 then show ?case by (rule_tac step) next case prems': (2 nm' nds') { fix n' assume "n'\ns" with P_x obtain nd where "nd\nds" and n'_split: "expand_rslt_exist_eq__node n' nd" by auto with prems' obtain nd' where "nd'\nds'" and "expand_rslt_exist_eq__node nd nd'" by auto then have "\nd'\nds'. expand_rslt_exist_eq__node n' nd'" using n'_split subset_trans[of "incoming n'"] by auto } then have "expand_rslt_exist_eq (n, ns) (nm', nds')" by auto with prems show ?case by auto qed qed qed lemma expand_prop_exist: "expand n_ns \ SPEC (\r. expand_assm_exist \ n_ns \ expand_rslt_exist \ n_ns r)" (is "_ \ SPEC (?P n_ns)") using assms proof (rule_tac expand_rec_rule[where \="\_. True"], simp, intro refine_vcg, goal_cases) case prems: (1 f x n ns) let ?nds = "upd_incoming n ns" let ?r = "(name n, ?nds)" { assume Q: "expand_assm_exist \ (n, ns)" note \\n'\ns. old n' = old n \ next n' = next n\ then obtain n' where "n'\ns" and assm_eq: "old n' = old n \ next n' = next n" by auto let ?nd = "n'\ incoming := incoming n \ incoming n'\" have "?nd \ ?nds" using \n'\ns\ assm_eq unfolding upd_incoming_def by auto moreover have "incoming n \ incoming ?nd" by auto moreover have "expand_rslt_exist__node_prop \ n ?nd" using Q assm_eq \new n = {}\ by simp ultimately have "expand_rslt_exist \ (n, ns) ?r" unfolding fst_conv snd_conv by blast } with prems show ?case by auto next case prems: (2 f x n ns) then have step: "\x. f x \ SPEC (?P x)" and f_sup: "\x. f x \ expand x" by auto show ?case unfolding \x = (n, ns)\ proof (rule_tac SPEC_rule_weak[where Q = "expand_rslt_exist_eq"], goal_cases) case 1 then show ?case by (rule_tac order_trans, rule_tac f_sup, rule_tac expand_rslt_exist_eq) next case 2 then show ?case by (rule_tac step) next case prems': (3 nm nds) then have "name ` ns \ name ` nds" by auto moreover { assume assm_ex: "expand_assm_exist \ (n, ns)" with prems' obtain nd where "nd\nds" and "expand_rslt_exist_eq__node n nd" by force+ then have "expand_rslt_exist__node_prop \ n nd" using assm_ex \new n = {}\ by auto then have "expand_rslt_exist \ (n, ns) (nm, nds)" using \nd\nds\ by auto } ultimately show ?case using expand_new_name_step[of n] prems' by auto qed next case prems: (3 f x n ns \) { assume "expand_assm_exist \ (n, ns)" with prems have "\ \\<^sub>n \" and "\ \\<^sub>n frml_neg \" by auto with prems have False by auto } with prems show ?case by auto next case prems: (4 f x n ns \) then have goal_assms: "\ \ new n \ (\q. \ = prop\<^sub>n(q) \ \ = nprop\<^sub>n(q))" and step: "\x. f x \ SPEC (?P x)" by simp_all show ?case using goal_assms unfolding \x = (n, ns)\ proof (rule_tac SPEC_rule_param2, rule_tac step, goal_cases) case prems': (1 nm nds) { assume "expand_assm_exist \ (n, ns)" with prems' have "expand_rslt_exist \ (n, ns) (nm, nds)" by auto } then show ?case by auto qed next case prems: (5 f x n ns \) then have goal_assms: "\ \ new n \ \ = true\<^sub>n" and step: "\x. f x \ SPEC (?P x)" by simp_all show ?case using goal_assms unfolding \x = (n, ns)\ proof (rule_tac SPEC_rule_param2, rule_tac step, goal_cases) case prems': (1 nm nds) { assume "expand_assm_exist \ (n, ns)" with prems' have "expand_rslt_exist \ (n, ns) (nm, nds)" by auto } then show ?case by auto qed next case prems: (6 f x n ns \) { assume "expand_assm_exist \ (n, ns)" with prems have "\ \\<^sub>n false\<^sub>n" by auto } with prems show ?case by auto next case prems: (7 f x n ns \) then have goal_assms: "\ \ new n \ (\\ \. \ = \ and\<^sub>n \ \ \ = X\<^sub>n \)" and step: "\x. f x \ SPEC (?P x)" by simp_all show ?case using goal_assms unfolding \x = (n, ns)\ proof (rule_tac SPEC_rule_param2, rule_tac step, goal_cases) case prems': (1 nm nds) { assume "expand_assm_exist \ (n, ns)" with prems' have "expand_rslt_exist \ (n, ns) (nm, nds)" by auto } then show ?case by auto qed next case prems: (8 f x n ns \) then have goal_assms: "\ \ new n \ (\\ \. \ = \ or\<^sub>n \ \ \ = \ U\<^sub>n \ \ \ = \ V\<^sub>n \)" by (cases \) auto from prems have step: "\x. f x \ SPEC (?P x)" and f_sup: "\x. f x \ expand x" by auto let ?x1 = "(n\new := new n - {\}, new := new1 \ \ new (n\new := new n - {\}\), old := {\} \ old n, next := next1 \ \ next n\, ns)" let ?new1_assm_sel = "\\. (case \ of \ U\<^sub>n \ => \ | \ V\<^sub>n \ \ \ | \ or\<^sub>n \ \ \)" { assume new1_assm: "\ (\ \\<^sub>n (?new1_assm_sel \))" then have ?case using goal_assms unfolding \x = (n, ns)\ proof (rule_tac SPEC_rule_nested2, goal_cases) case prems': 1 then show ?case proof (rule_tac SPEC_rule_param2, rule_tac step, goal_cases) case prems'': (1 nm nds) { assume "expand_assm_exist \ (n, ns)" with prems'' have "expand_assm_exist \ ?x1" unfolding fst_conv proof (cases \, goal_cases) case \: (8 \ \) then have "\ \\<^sub>n \ U\<^sub>n \" by fast then have "\ \\<^sub>n \" and "\ \\<^sub>n X\<^sub>n (\ U\<^sub>n \)" using \ ltln_expand_Until[of \ \ \] by auto with \ show ?case by auto next case \: (9 \ \) then have *: "\ \\<^sub>n \ V\<^sub>n \" by fast with \ have "\ \\<^sub>n \" and "\ \\<^sub>n X\<^sub>n (\ V\<^sub>n \)" using ltln_expand_Release[of \ \ \] by auto with \ * show ?case by auto qed auto with prems'' have "expand_rslt_exist \ (n, ns) (nm, nds)" by force } with prems'' show ?case by auto qed next case prems': (2 nm nds) then have P_x: "?P (n, ns) (nm, nds)" by fast show ?case unfolding case_prod_unfold proof (rule_tac SPEC_rule_weak[where P = "?P" and Q = "expand_rslt_exist_eq"], goal_cases) case 1 then show ?case by (rule_tac order_trans, rule_tac f_sup, rule_tac expand_rslt_exist_eq) next case 2 then show ?case by (rule_tac step) next case prems'': (3 nm' nds') { assume "expand_assm_exist \ (n, ns)" with P_x have "expand_rslt_exist \ (n, ns) (nm, nds)" by simp then obtain nd where nd: "nd\nds" "expand_rslt_exist__node_prop \ n nd" using goal_assms by auto with prems'' obtain nd' where "nd'\nds'" and "expand_rslt_exist_eq__node nd nd'" by force with nd have "expand_rslt_exist__node_prop \ n nd'" using subset_trans[of "incoming n" "incoming nd"] by auto then have "expand_rslt_exist \ (n,ns) (nm', nds')" using \nd'\nds'\ goal_assms by auto } then show ?case by fast qed qed } moreover { assume new1_assm: "\ \\<^sub>n (?new1_assm_sel \)" let ?x2f = "\(nm::node_name, nds::'a node set). ( n\new := new n - {\}, name := nm, new := new2 \ \ new (n\new := new n - {\}\), old := {\} \ old n\, nds)" have P_x: "f ?x1 \ SPEC (?P ?x1)" by (rule step) moreover { fix r :: "node_name \ 'a node set" let ?x2 = "?x2f r" assume assm: "(?P ?x1) r" have "f ?x2 \ SPEC (?P (n, ns))" unfolding case_prod_unfold fst_conv proof (rule_tac SPEC_rule_param2, rule_tac step, goal_cases) case prems': (1 nm' nds') { assume "expand_assm_exist \ (n, ns)" with new1_assm goal_assms have "expand_assm_exist \ ?x2" proof (cases r, cases \, goal_cases) case prems'': (9 _ _ \ \) then have *: "\ \\<^sub>n \ V\<^sub>n \" unfolding fst_conv by fast with ltln_expand_Release[of \ \ \] have "\ \\<^sub>n \" by auto with prems'' * show ?case by auto qed auto with prems' have "expand_rslt_exist \ ?x2 (nm', nds')" unfolding case_prod_unfold fst_conv snd_conv by fast then have "expand_rslt_exist \ (n, ns) (nm', nds')" by (cases r, auto) } then show ?case by simp qed } then have "SPEC (?P ?x1) \ SPEC (\r. (case r of (nm, nds) => f (?x2f (nm, nds))) \ SPEC (?P (n, ns)))" using goal_assms by (rule_tac SPEC_rule) force finally have ?case unfolding case_prod_unfold \x = (n, ns)\ by simp } ultimately show ?case by fast qed text \Termination proof\ definition expand\<^sub>T :: "('a node \ ('a node set)) \ (node_name \ 'a node set) nres" where "expand\<^sub>T n_ns \ REC\<^sub>T expand_body n_ns" abbreviation "old_next_pair n \ (old n, next n)" abbreviation "old_next_limit \ \ Pow (subfrmlsn \) \ Pow (subfrmlsn \)" lemma old_next_limit_finite: "finite (old_next_limit \)" using subfrmlsn_finite by auto definition "expand_ord \ \ inv_image (finite_psupset (old_next_limit \) <*lex*> less_than) (\(n, ns). (old_next_pair ` ns, frmlset_sumn (new n)))" lemma expand_ord_wf[simp]: "wf (expand_ord \)" using finite_psupset_wf[OF old_next_limit_finite] unfolding expand_ord_def by auto abbreviation "expand_inv_node \ n \ finite (new n) \ finite (old n) \ finite (next n) \ (new n) \ (old n) \ (next n) \ subfrmlsn \" abbreviation "expand_inv_result \ ns \ finite ns \ (\n'\ns. (new n') \ (old n') \ (next n') \ subfrmlsn \)" definition "expand_inv \ n_ns \ (case n_ns of (n, ns) \ expand_inv_node \ n \ expand_inv_result \ ns)" lemma new1_less_setsum: "frmlset_sumn (new1 \) < frmlset_sumn {\}" proof - have "frmlset_sumn {\,\} < Suc (size_frmln \ + size_frmln \)" for \ \ by (cases "\ = \") auto then show ?thesis by (cases \) auto qed lemma new2_less_setsum: "frmlset_sumn (new2 \) < frmlset_sumn {\}" proof - have "frmlset_sumn {\,\} < Suc (size_frmln \ + size_frmln \)" for \ \ by (cases "\ = \") auto then show ?thesis by (cases \) auto qed lemma new1_finite[intro]: "finite (new1 \)" by (cases \) auto lemma new1_subset_frmls: "\ \ new1 \ \ \ \ subfrmlsn \" by (cases \) auto lemma new2_finite[intro]: "finite (new2 \)" by (cases \) auto lemma new2_subset_frmls: "\ \ new2 \ \ \ \ subfrmlsn \" by (cases \) auto lemma next1_finite[intro]: "finite (next1 \)" by (cases \) auto lemma next1_subset_frmls: "\ \ next1 \ \ \ \ subfrmlsn \" by (cases \) auto lemma expand_inv_impl[intro!]: assumes "expand_inv \ (n, ns)" and newn: "\ \ new n" and "old_next_pair ` ns \ old_next_pair ` ns'" and "expand_inv_result \ ns'" and "(n' = n\ new := new n - {\}, old := {\} \ old n \) \ (n' = n\ new := new1 \ \ (new n - {\}), old := {\} \ old n, next := next1 \ \ next n \) \ (n' = n\ name := nm, new := new2 \ \ (new n - {\}), old := {\} \ old n \)" (is "?case1 \ ?case2 \ ?case3") shows "((n', ns'), (n, ns))\expand_ord \ \ expand_inv \ (n', ns')" (is "?concl1 \ ?concl2") proof from assms consider ?case1 | ?case2 | ?case3 by blast then show ?concl1 proof cases case n'def: 1 with assms show ?thesis unfolding expand_inv_def expand_ord_def finite_psupset_def by auto next case n'def: 2 have \innew: "\\ new n" and fin_new: "finite (new n)" using assms unfolding expand_inv_def by auto from \innew setsum_diff1_nat[of size_frmln "new n" \] have "frmlset_sumn (new n - {\}) = frmlset_sumn (new n) - size_frmln \" by simp moreover from fin_new setsum_Un_nat[OF new1_finite _, of "new n - {\}" size_frmln \] have "frmlset_sumn (new n') \ frmlset_sumn (new1 \) + frmlset_sumn (new n - {\})" unfolding n'def by auto moreover have sum_leq:"frmlset_sumn (new n) \ frmlset_sumn {\}" using \innew setsum_mono2[OF fin_new, of "{\}" size_frmln] size_frmln_gt_zero by simp ultimately have "frmlset_sumn (new n') < frmlset_sumn (new n)" using new1_less_setsum[of \] sum_leq by auto with assms show ?thesis unfolding expand_ord_def finite_psupset_def by auto next case n'def: 3 have \innew: "\\ new n" and fin_new: "finite (new n)" using assms unfolding expand_inv_def by auto from \innew setsum_diff1_nat[of size_frmln "new n" \] have "frmlset_sumn (new n - {\}) = frmlset_sumn (new n) - size_frmln \" by simp moreover from fin_new setsum_Un_nat[OF new2_finite _, of "new n - {\}" size_frmln \] have "frmlset_sumn (new n') \ frmlset_sumn (new2 \) + frmlset_sumn (new n - {\})" unfolding n'def by auto moreover have sum_leq:"frmlset_sumn (new n) \ frmlset_sumn {\}" using \innew setsum_mono2[OF fin_new, of "{\}" size_frmln] size_frmln_gt_zero by simp ultimately have "frmlset_sumn (new n') < frmlset_sumn (new n)" using new2_less_setsum[of \] sum_leq by auto with assms show ?thesis unfolding expand_ord_def finite_psupset_def by auto qed next have "new1 \ \ subfrmlsn \" and "new2 \ \ subfrmlsn \" and "next1 \ \ subfrmlsn \" using assms subfrmlsn_subset[OF new1_subset_frmls[of _ \]] subfrmlsn_subset[of \ \, OF set_rev_mp[of _ "new n"]] subfrmlsn_subset[OF new2_subset_frmls[of _ \]] subfrmlsn_subset[OF next1_subset_frmls[of _ \]] unfolding expand_inv_def (* This proof is merely a speed optimization. A single force+ does the job, but takes very long *) apply - apply (clarsimp split: prod.splits) apply (metis in_mono new1_subset_frmls) apply (clarsimp split: prod.splits) apply (metis new2_subset_frmls set_rev_mp) apply (clarsimp split: prod.splits) apply (metis in_mono next1_subset_frmls) done with assms show ?concl2 unfolding expand_inv_def by auto qed lemma expand_term_prop_help: assumes "((n', ns'), (n, ns))\expand_ord \ \ expand_inv \ (n', ns')" and assm_rule: "\expand_inv \ (n', ns'); ((n', ns'), n, ns) \ expand_ord \\ \ f (n', ns') \ SPEC P" shows "f (n', ns') \ SPEC P" using assms by (rule_tac assm_rule) auto lemma expand_inv_upd_incoming: assumes "expand_inv \ (n, ns)" shows "expand_inv_result \ (upd_incoming n ns)" using assms unfolding expand_inv_def upd_incoming_def by force lemma upd_incoming_eq_old_next_pair: "old_next_pair ` ns = old_next_pair ` (upd_incoming n ns)" (is "?A = ?B") proof show "?A \ ?B" proof fix x let ?f = "\n'. n'\incoming := incoming n \ incoming n'\" assume "x \ ?A" then obtain n' where "n' \ ns" and xeq: "x = (old n', next n')" by auto have "x \ (old_next_pair ` (\n'. n'\incoming := incoming n \ incoming n'\) ` (ns \ {n' \ ns. old n' = old n \ next n' = next n})) \ (old_next_pair ` (ns \ {n' \ ns. old n' \ old n \ next n' \ next n}))" proof (cases "old n' = old n \ next n' = next n") case True with \n' \ ns\ have "?f n' \ ?f ` (ns \ {n'\ns. old n' = old n \ next n' = next n})" (is "_ \ ?C") by auto then have "old_next_pair (?f n') \ old_next_pair ` ?C" by (rule_tac imageI) auto with xeq have "x\old_next_pair ` ?C" by auto then show ?thesis by blast next case False with \n' \ ns\ xeq have "x \ old_next_pair ` (ns \ {n'\ns. old n' \ old n \ next n' \ next n})" by auto then show ?thesis by blast qed then show "x \ ?B" using \x \ ?A\ unfolding upd_incoming_def by auto qed show "?B \ ?A" unfolding upd_incoming_def by (force intro:imageI) qed lemma expand_term_prop: "expand_inv \ n_ns \ expand\<^sub>T n_ns \ SPEC (\(_, nds). old_next_pair ` snd n_ns\old_next_pair `nds \ expand_inv_result \ nds)" (is "_ \ _ \ SPEC (?P n_ns)") unfolding expand\<^sub>T_def apply (rule_tac RECT_rule[where pre="\(n, ns). expand_inv \ (n, ns)" and V="expand_ord \"]) apply (refine_mono) apply simp apply simp proof (intro refine_vcg, goal_cases) case prems: (1 _ _ n ns) have "old_next_pair ` ns \ old_next_pair ` (upd_incoming n ns)" by (rule equalityD1[OF upd_incoming_eq_old_next_pair]) with prems show ?case using expand_inv_upd_incoming[of \ n ns] by auto next case prems: (2 expand x n ns) let ?n' = "\ name = expand_new_name (name n), incoming = {name n}, new = next n, old = {}, next = {}\" let ?ns' = "{n} \ ns" from prems have SPEC_sub:"SPEC (?P (?n', ?ns')) \ SPEC (?P x)" by (rule_tac SPEC_rule) auto from prems have *: "old_next_pair n\old_next_pair ` ns" by auto then have subset_next_pair: "old_next_pair ` ns \ old_next_pair ` (insert n ns)" by auto with prems * have "((?n', ?ns'), (n, ns))\expand_ord \" by (auto simp add: expand_ord_def expand_inv_def finite_psupset_def) moreover from prems have "expand_inv \ (?n', ?ns')" unfolding expand_inv_def by auto ultimately have "expand (?n', ?ns') \ SPEC (?P (?n', ?ns'))" using prems by fast with SPEC_sub show ?case by (rule_tac order_trans) fast+ next case 3 then show ?case by (auto simp add:expand_inv_def) next case 4 then show ?case apply (rule_tac expand_term_prop_help[OF expand_inv_impl]) apply (simp add: expand_inv_def)+ apply force done next case 5 then show ?case apply (rule_tac expand_term_prop_help[OF expand_inv_impl]) apply (simp add: expand_inv_def)+ apply force done next case 6 then show ?case by (simp add: expand_inv_def) next case 7 then show ?case apply (rule_tac expand_term_prop_help[OF expand_inv_impl]) apply (simp add: expand_inv_def)+ apply force done next case prems: (8 f x a b xa) let ?n' = "a\ new := new1 xa \ (new a - {xa}), old := insert xa (old a), next := next1 xa \ next a\" let ?n'' = "\nm. a\ name := nm, new := new2 xa \ (new a - {xa}), old := insert xa (old a)\" have step:"((?n', b), (a, b)) \ expand_ord \ \ expand_inv \ (?n', b)" using prems by (rule_tac expand_inv_impl) (auto simp add: expand_inv_def) with prems have assm1: "f (?n', b) \ SPEC (?P (a, b))" by auto moreover { fix nm::node_name and nds::"'a node set" assume assm1: "old_next_pair ` b \ old_next_pair ` nds" and assm2: "expand_inv_result \ nds" with prems step have "((?n'' nm, nds), (a, b)) \ expand_ord \ \ expand_inv \ (?n'' nm, nds)" by (rule_tac expand_inv_impl) auto with prems have "f (?n'' nm, nds) \ SPEC (?P (?n'' nm, nds))" by auto moreover have "SPEC (?P (?n'' nm, nds)) \ SPEC (?P (a, b))" using assm2 subset_trans[OF assm1] by auto ultimately have "f (?n'' nm, nds) \ SPEC (?P (a, b))" by (rule order_trans) } then have assm2: "SPEC (?P (a, b)) \ SPEC (\r. (case r of (nm, nds) \ f (?n'' nm, nds)) \ SPEC (?P (a, b)))" by (rule_tac SPEC_rule) auto from prems order_trans[OF assm1 assm2] show ?case by auto qed lemma expand_eq_expand\<^sub>T: assumes inv: "expand_inv \ n_ns" shows "expand\<^sub>T n_ns = expand n_ns" unfolding expand\<^sub>T_def expand_def apply (rule RECT_eq_REC) apply refine_mono unfolding expand\<^sub>T_def[symmetric] using expand_term_prop[OF inv] apply auto done lemma expand_nofail: assumes inv: "expand_inv \ n_ns" shows "nofail (expand\<^sub>T n_ns)" using expand_term_prop[OF inv] by (simp add: pw_le_iff) lemma [intro!]: "expand_inv \ ( \ name = expand_new_name expand_init, incoming = {expand_init}, new = {\}, old = {}, next = {} \, {})" by (auto simp add: expand_inv_def) definition create_graph :: "'a frml \ 'a node set nres" where "create_graph \ \ do { (_, nds) \ expand ( \ name = expand_new_name expand_init, incoming = {expand_init}, new = {\}, old = {}, next = {} \::'a node, {}::'a node set); RETURN nds }" (* "expand_eq_expand\<^sub>T" *) definition create_graph\<^sub>T :: "'a frml \ 'a node set nres" where "create_graph\<^sub>T \ \ do { (_, nds) \ expand\<^sub>T ( \ name = expand_new_name expand_init, incoming = {expand_init}, new = {\}, old = {}, next = {} \::'a node, {}::'a node set); RETURN nds }" lemma create_graph_eq_create_graph\<^sub>T: "create_graph \ = create_graph\<^sub>T \" unfolding create_graph\<^sub>T_def create_graph_def unfolding eq_iff proof (standard, goal_cases) case 1 then show ?case by refine_mono (unfold expand_def expand\<^sub>T_def, rule REC_le_RECT) next case 2 then show ?case by (refine_mono, rule expand_eq_expand\<^sub>T[unfolded eq_iff, THEN conjunct1]) auto qed lemma create_graph_finite: "create_graph \ \ SPEC finite" unfolding create_graph_def expand_def apply (intro refine_vcg) apply (rule_tac order_trans) apply (rule_tac REC_le_RECT) apply (fold expand\<^sub>T_def) apply (rule_tac order_trans[OF expand_term_prop]) apply auto[1] apply (rule_tac SPEC_rule) apply auto done lemma create_graph_nofail: "nofail (create_graph \)" by (rule_tac pwD1[OF create_graph_finite]) auto abbreviation "create_graph_rslt_exist \ nds \ \nd\nds. expand_init\incoming nd \ (\\\old nd. \ \\<^sub>n \) \ (\\\next nd. \ \\<^sub>n X\<^sub>n \) \ {\. \\. \ U\<^sub>n \ \ old nd \ \ \\<^sub>n \} \ old nd" lemma L4_7: assumes "\ \\<^sub>n \" shows "create_graph \ \ SPEC (create_graph_rslt_exist \)" using assms unfolding create_graph_def by (intro refine_vcg, rule_tac order_trans, rule_tac expand_prop_exist) ( auto simp add:expand_new_name_expand_init) lemma expand_incoming_name_exist: assumes "name (fst n_ns) > expand_init \ (\nm\incoming (fst n_ns). nm \ expand_init \ nm\name ` (snd n_ns)) \ expand_assm_incoming n_ns \ expand_name_ident (snd n_ns)" (is "?Q n_ns") and "\nd\snd n_ns. name nd > expand_init \ (\nm\incoming nd. nm \ expand_init \ nm\name ` (snd n_ns))" (is "?P (snd n_ns)") shows "expand n_ns \ SPEC (\nm_nds. ?P (snd nm_nds))" using assms apply (rule_tac expand_rec_rule[where \="\n_ns. ?Q n_ns \ ?P (snd n_ns)"]) apply simp apply (intro refine_vcg) proof goal_cases case (1 f x n ns) then show ?case proof (simp, clarify, goal_cases) case prems: (1 _ _ nd) { assume "nd\ns" with prems have ?case by auto } moreover { assume "nd\ns" with upd_incoming__elem[OF \nd\upd_incoming n ns\] obtain nd' where "nd'\ns" and "nd = nd'\incoming := incoming n \ incoming nd'\ \ old nd' = old n \ next nd' = next n" by auto with prems have ?case by auto } ultimately show ?case by fast qed next case (2 f x n ns) then have step: "\x. ?Q x \ ?P (snd x) \ f x \ SPEC (\x. ?P (snd x))" and QP: "?Q (n, ns) \ ?P ns" and f_sup: "\x. f x \ expand x" by auto show ?case unfolding \x = (n, ns)\ using QP expand_new_name_expand_init proof (rule_tac step, goal_cases) case prems: 1 then have name_less: "name n < expand_new_name (name n)" by auto moreover from prems name_less have "\nm\incoming n. nm < expand_new_name (name n)" by auto moreover from prems name_less have **: "\q\ns. name q < expand_new_name (name n) \ (\nm\incoming q. nm < expand_new_name (name n))" by force moreover from QP have "\!q'. (q' = n \ q' \ ns) \ name n = name q'" by force moreover have "\q\ns. \!q'. (q' = n \ q' \ ns) \ name q = name q' " using QP by auto ultimately show ?case using prems by simp qed next case 3 then show ?case by simp next case 4 then show ?case by simp next case 5 then show ?case by simp next case 6 then show ?case by simp next case 7 then show ?case by simp next case prems: (8 f x n ns \) then have goal_assms: "\ \ new n \ (\\ \. \ = \ or\<^sub>n \ \ \ = \ U\<^sub>n \ \ \ = \ V\<^sub>n \)" by (cases \) auto with prems have QP: "?Q (n, ns) \ ?P ns" and step: "\x. ?Q x \ ?P (snd x) \ f x \ SPEC (\x'. ?P (snd x'))" and f_sup: "\x. f x \ expand x" by auto let ?x = "(n\new := new n - {\}, new := new1 \ \ new (n\new := new n - {\}\), old := {\} \ old (n\new := new n - {\}\), next := next1 \ \ next (n\new := new n - {\}\)\, ns)" let ?props = "\x r. expand_rslt_incoming r \ expand_rslt_name x r \ expand_name_ident (snd r)" show ?case using goal_assms QP unfolding case_prod_unfold \x = (n, ns)\ proof (rule_tac SPEC_rule_weak_nested2[where Q = "?props ?x"], goal_cases) case 1 then show ?case by (rule_tac order_trans, rule_tac f_sup, rule_tac expand_name_propag) simp next case 2 then show ?case by (rule_tac SPEC_rule_param2[where P = "\x r. ?P (snd r)"], rule_tac step) auto next case (3 nm nds) then show ?case proof (rule_tac SPEC_rule_weak[where P = "\x r. ?P (snd r)" and Q = "\x r. expand_rslt_exist_eq x r \ ?props x r"], goal_cases) case 1 then show ?case by (rule_tac SPEC_rule_conjI, rule_tac order_trans, rule_tac f_sup, rule_tac expand_rslt_exist_eq, rule_tac order_trans, rule_tac f_sup, rule_tac expand_name_propag) force next case 2 then show ?case by (rule_tac SPEC_rule_param2[where P = "\x r. ?P (snd r)"], rule_tac step) force next case (3 nm' nds') then show ?case by simp qed qed qed lemma create_graph__incoming_name_exist: "create_graph \ \ SPEC (\nds. \nd\nds. expand_init < name nd \ (\nm\incoming nd. nm \ expand_init \ nm\name ` nds))" using assms unfolding create_graph_def by (intro refine_vcg, rule_tac order_trans, rule_tac expand_incoming_name_exist) ( auto simp add:expand_new_name_expand_init) abbreviation "expand_rslt_all__ex_equiv \ nd nds \ (\nd'\nds. name nd \ incoming nd' \ (\\\old nd'. suffix 1 \ \\<^sub>n \) \ (\\\next nd'. suffix 1 \ \\<^sub>n X\<^sub>n \) \ {\. \\. \ U\<^sub>n \ \ old nd' \ suffix 1 \ \\<^sub>n \} \ old nd')" abbreviation "expand_rslt_all \ n_ns nm_nds \ (\nd\snd nm_nds. name nd \ name ` (snd n_ns) \ (\\\old nd. \ \\<^sub>n \) \ (\\\next nd. \ \\<^sub>n X\<^sub>n \) \ expand_rslt_all__ex_equiv \ nd (snd nm_nds))" lemma expand_prop_all: assumes "expand_assm_incoming n_ns \ expand_name_ident (snd n_ns)" (is "?Q n_ns") shows "expand n_ns \ SPEC (expand_rslt_all \ n_ns)" (is "_ \ SPEC (?P n_ns)") using assms apply (rule_tac expand_rec_rule[where \="?Q"]) apply simp apply (intro refine_vcg) proof goal_cases case 1 then show ?case by (simp, rule_tac upd_incoming__ident) simp_all next case (2 f x n ns) then have step: "\x. ?Q x \ f x \ SPEC (?P x)" and Q: "?Q (n, ns)" and f_sup: "\x. f x \ expand x" by auto let ?x = "(\name = expand_new_name (name n), incoming = {name n}, new = next n, old = {}, next = {}\, {n} \ ns)" from Q have name_le: "name n < expand_new_name (name n)" by auto show ?case unfolding \x = (n, ns)\ proof (rule_tac SPEC_rule_weak[where Q = "\p r. (expand_assm_exist (suffix 1 \) ?x \ expand_rslt_exist (suffix 1 \) ?x r) \ expand_rslt_exist_eq p r \ (expand_name_ident (snd r))"], goal_cases) case 1 then show ?case proof (rule_tac SPEC_rule_conjI, rule_tac order_trans, rule_tac f_sup, rule_tac expand_prop_exist, rule_tac SPEC_rule_conjI, rule_tac order_trans, rule_tac f_sup, rule_tac expand_rslt_exist_eq, rule_tac order_trans, rule_tac f_sup, rule_tac expand_name_propag__name_ident, goal_cases) case 1 then show ?case using Q name_le by force qed next case 2 then show ?case using Q name_le by (rule_tac step) force next case prems: (3 nm nds) then obtain n' where "n'\nds" and eq_node: "expand_rslt_exist_eq__node n n'" by auto with prems have ex1_name: "\!q\nds. name n = name q" by auto then have nds_eq: "nds = {n'} \ {x \ nds. name n \ name x}" using eq_node \n'\nds\ by blast have name_notin: "name n \ name ` ns" using Q by auto from prems have P_x: "expand_rslt_all \ ?x (nm, nds)" by fast show ?case unfolding snd_conv proof clarify fix nd assume "nd \ nds" and name_img: "name nd \ name ` ns" and nd_old_equiv: "\\\old nd. \ \\<^sub>n \" and nd_next_equiv: "\\\next nd. \ \\<^sub>n X\<^sub>n \" show "expand_rslt_all__ex_equiv \ nd nds" proof (cases "name nd = name n") case True with nds_eq eq_node \nd\nds\ have "nd = n'" by auto with prems(1)[THEN conjunct1, simplified] nd_old_equiv nd_next_equiv eq_node show ?thesis by simp next case False with name_img \nd \ nds\ nd_old_equiv nd_next_equiv P_x show ?thesis by simp qed qed qed next case 3 then show ?case by auto next case prems: (4 f) then have step: "\x. ?Q x \ f x \ SPEC (?P x)" by simp from prems show ?case by (rule_tac SPEC_rule_param2, rule_tac step) auto next case prems: (5 f) then have step: "\x. ?Q x \ f x \ SPEC (?P x)" by simp from prems show ?case by (rule_tac SPEC_rule_param2, rule_tac step) auto next case 6 then show ?case by auto next case prems: (7 f) then have step: "\x. ?Q x \ f x \ SPEC (?P x)" by simp from prems show ?case by (rule_tac SPEC_rule_param2, rule_tac step) auto next case prems: (8 f x n ns \) then have goal_assms: "\ \ new n \ (\\ \. \ = \ or\<^sub>n \ \ \ = \ U\<^sub>n \ \ \ = \ V\<^sub>n \)" by (cases \) auto from prems have Q: "?Q (n, ns)" and step: "\x. ?Q x \ f x \ SPEC (?P x)" and f_sup: "\x. f x \ expand x" by auto let ?x = "(n\new := new n - {\}, new := new1 \ \ new (n\new := new n - {\}\), old := {\} \ old (n\new := new n - {\}\), next := next1 \ \ next (n\new := new n - {\}\) \, ns)" let ?props = "\x r. expand_rslt_incoming r \ expand_rslt_name x r \ expand_name_ident (snd r)" show ?case using goal_assms Q unfolding case_prod_unfold \x = (n, ns)\ proof (rule_tac SPEC_rule_weak_nested2[where Q = "?props ?x"], goal_cases) case 1 then show ?case by (rule_tac order_trans, rule_tac f_sup, rule_tac expand_name_propag) simp next case 2 then show ?case by (rule_tac SPEC_rule_param2[where P = "?P"], rule_tac step) auto next case prems: (3 nm nds) then show ?case proof (rule_tac SPEC_rule_weak[where P = "?P" and Q = "\x r. expand_rslt_exist_eq x r \ ?props x r"], goal_cases) case 1 then show ?case by (rule_tac SPEC_rule_conjI, rule_tac order_trans, rule_tac f_sup, rule_tac expand_rslt_exist_eq, rule_tac order_trans, rule_tac f_sup, rule_tac expand_name_propag) auto next case 2 then show ?case by (rule_tac SPEC_rule_param2[where P = "?P"], rule_tac step) auto next case prems': (3 nm' nds') then have P_x: "?P (n, ns) (nm, nds)" and P_x': "?P (n, nds) (nm', nds')" by simp_all show ?case unfolding snd_conv proof clarify fix nd assume "nd\nds'" and name_nd_notin: "name nd \ name ` ns" and old_equiv: "\\\old nd. \ \\<^sub>n \" and next_equiv: "\\\next nd. \ \\<^sub>n X\<^sub>n \" show "expand_rslt_all__ex_equiv \ nd nds'" proof (cases "name nd \ name ` nds") case True then obtain n' where "n' \ nds" and "name nd = name n'" by auto with prems' obtain nd' where "nd'\nds'" and nd'_eq: "expand_rslt_exist_eq__node n' nd'" by auto moreover from prems' have "\q\nds'. \!q'\nds'. name q = name q'" by simp ultimately have "nd' = nd" using \name nd = name n'\ \nd \ nds'\ by auto with nd'_eq have n'_eq: "expand_rslt_exist_eq__node n' nd" by simp then have "name n'\name ` ns" and "\\\old n'. \ \\<^sub>n \" and "\\\next n'. \ \\<^sub>n X\<^sub>n \" using name_nd_notin old_equiv next_equiv \n' \ nds\ by auto then have "expand_rslt_all__ex_equiv \ n' nds" (is "\nd'\nds. ?sthm n' nd'") using P_x \n' \ nds\ unfolding snd_conv by blast then obtain sucnd where sucnd: "sucnd\nds" and sthm: "?sthm n' sucnd" by blast moreover from prems' sucnd sthm obtain sucnd' where "sucnd'\nds'" and sucnd'_eq: "expand_rslt_exist_eq__node sucnd sucnd'" by auto ultimately have "?sthm n' sucnd'" by auto then show ?thesis using \sucnd' \ nds'\ unfolding \name nd = name n'\ by blast next case False with \nd \ nds'\ P_x' old_equiv next_equiv show ?thesis unfolding snd_conv by blast qed qed qed qed qed abbreviation "create_graph_rslt_all \ nds \ \q\nds. (\\\old q. \ \\<^sub>n \) \ (\\\next q. \ \\<^sub>n X\<^sub>n \) \ (\q'\nds. name q \ incoming q' \ (\\\old q'. suffix 1 \ \\<^sub>n \) \ (\\\next q'. suffix 1 \ \\<^sub>n X\<^sub>n \) \ {\. \\. \ U\<^sub>n \ \ old q' \ suffix 1 \ \\<^sub>n \} \ old q')" lemma L4_5: "create_graph \ \ SPEC (create_graph_rslt_all \)" unfolding create_graph_def by (intro refine_vcg, rule_tac order_trans, rule_tac expand_prop_all) (auto simp add:expand_new_name_expand_init) subsection \Creation of GBA\ text \This section formalizes the second part of the algorithm, that creates the actual generalized B\"uchi automata from the set of nodes.\ definition create_gba_from_nodes :: "'a frml \ 'a node set \ ('a node, 'a set) gba_rec" where "create_gba_from_nodes \ qs \ \ g_V = qs, g_E = {(q, q'). q\qs \ q'\qs \ name q\incoming q'}, g_V0 = {q\qs. expand_init\incoming q}, gbg_F = {{q\qs. \ U\<^sub>n \\old q \ \\old q}|\ \. \ U\<^sub>n \ \ subfrmlsn \}, gba_L = \q l. q\qs \ {p. prop\<^sub>n(p)\old q}\l \ {p. nprop\<^sub>n(p)\old q} \ l = {} \" end locale create_gba_from_nodes_precond = fixes \ :: "'a ltln" fixes qs :: "'a node set" assumes res: "inres (create_graph \) qs" begin lemma finite_qs[simp, intro!]: "finite qs" using res create_graph_finite by (auto simp add: pw_le_iff) lemma create_gba_from_nodes__invar: "gba (create_gba_from_nodes \ qs)" using [[simproc finite_Collect]] apply unfold_locales apply (auto intro!: finite_vimageI subfrmlsn_finite injI simp: create_gba_from_nodes_def) done sublocale gba: gba "create_gba_from_nodes \ qs" by (rule create_gba_from_nodes__invar) lemma create_gba_from_nodes__fin: "finite (g_V (create_gba_from_nodes \ qs))" unfolding create_gba_from_nodes_def by auto lemma create_gba_from_nodes__ipath: "ipath gba.E r \ (\i. r i \qs \ name (r i)\incoming (r (Suc i)))" unfolding create_gba_from_nodes_def ipath_def by auto lemma create_gba_from_nodes__is_run: "gba.is_run r \ expand_init \ incoming (r 0) \ (\i. r i\qs \ name (r i)\incoming (r (Suc i)))" unfolding gba.is_run_def apply (simp add: create_gba_from_nodes__ipath) apply (auto simp: create_gba_from_nodes_def) done context begin -interpretation LTL_Syntax . - abbreviation "auto_run_j j \ q \ (\\\old q. suffix j \ \\<^sub>n \) \ (\\\next q. suffix j \ \\<^sub>n X\<^sub>n \) \ {\. \\. \ U\<^sub>n \ \ old q \ suffix j \ \\<^sub>n \} \ old q" fun auto_run :: "['a interprt, 'a node set] \ 'a node word" where "auto_run \ nds 0 = (SOME q. q\nds \ expand_init \ incoming q \ auto_run_j 0 \ q)" | "auto_run \ nds (Suc k) = (SOME q'. q'\nds \ name (auto_run \ nds k) \ incoming q' \ auto_run_j (Suc k) \ q')" (* TODO: Remove. Only special instance of more generic principle lemma create_graph_comb: assumes "\x. inres (create_graph \) x \ P x" shows "create_graph \\SPEC P" using create_graph_nofail assms by (auto simp add: pw_le_iff refine_pw_simps) *) lemma run_propag_on_create_graph: assumes "ipath gba.E \" shows "\ k\qs \ name (\ k)\incoming (\ (k+1))" using assms by (auto simp: create_gba_from_nodes__ipath) lemma expand_false_propag: assumes "false\<^sub>n \ old (fst n_ns) \ (\nd\snd n_ns. false\<^sub>n \ old nd)" (is "?Q n_ns") shows "expand n_ns \ SPEC (\nm_nds. \nd\snd nm_nds. false\<^sub>n \ old nd)" using assms proof (rule_tac expand_rec_rule[where \="?Q"], simp, intro refine_vcg, goal_cases) case 1 then show ?case by (simp, rule_tac upd_incoming__ident) auto next case 8 then show ?case by (rule_tac SPEC_rule_nested2) auto qed auto lemma false_propag_on_create_graph: "create_graph \ \ SPEC (\nds. \nd\nds. false\<^sub>n \ old nd)" unfolding create_graph_def by (intro refine_vcg, rule_tac order_trans, rule_tac expand_false_propag) auto lemma expand_and_propag: assumes "\ and\<^sub>n \ \ old (fst n_ns) \ {\, \} \ old (fst n_ns) \ new (fst n_ns)" (is "?Q n_ns") and "\nd\snd n_ns. \ and\<^sub>n \ \ old nd \ {\, \} \ old nd" (is "?P (snd n_ns)") shows "expand n_ns \ SPEC (\nm_nds. ?P (snd nm_nds))" using assms proof (rule_tac expand_rec_rule[where \="\x. ?Q x \ ?P (snd x)"], simp, intro refine_vcg, goal_cases) case 1 then show ?case by (simp, rule_tac upd_incoming__ident) auto next case prems: (4 f x n ns) then have step: "\x. ?Q x \ ?P (snd x) \ f x \ SPEC (\x'. ?P (snd x'))" by simp with prems show ?case by (rule_tac step) auto next case prems: (5 f x n ns) then have step: "\x. ?Q x \ ?P (snd x) \ f x \ SPEC (\x'. ?P (snd x'))" by simp with prems show ?case by (rule_tac step) auto next case (6 f x n ns) then show ?case by auto next case prems: (7 f x n ns) then have step: "\x. ?Q x \ ?P (snd x) \ f x \ SPEC (\x'. ?P (snd x'))" by simp with prems show ?case by (rule_tac step) auto next case prems: (8 f x n ns \) then have goal_assms: "\ \ new n \ \ (\q. \ = prop\<^sub>n(q) \ \ = nprop\<^sub>n(q)) \ \ \ true\<^sub>n \ \ \ false\<^sub>n \ \ (\\ \. \ = \ and\<^sub>n \ \ \ = X\<^sub>n \) \ (\\ \. \ = \ or\<^sub>n \ \ \ = \ U\<^sub>n \ \ \ = \ V\<^sub>n \)" by (cases \) auto from prems have QP: "?Q (n, ns) \ ?P ns" and step: "\x. ?Q x \ ?P (snd x) \ f x \ SPEC (\x'. ?P (snd x'))" by simp_all show ?case using goal_assms QP unfolding case_prod_unfold \x = (n, ns)\ proof (rule_tac SPEC_rule_nested2, goal_cases) case 1 then show ?case by (rule_tac step) auto next case 2 then show ?case by (rule_tac step) auto qed qed auto lemma and_propag_on_create_graph: "create_graph \ \ SPEC (\nds. \nd\nds. \ and\<^sub>n \ \ old nd \ {\, \} \ old nd)" unfolding create_graph_def by (intro refine_vcg, rule_tac order_trans, rule_tac expand_and_propag) auto lemma expand_or_propag: assumes "\ or\<^sub>n \ \ old (fst n_ns) \ {\, \} \ (old (fst n_ns) \ new (fst n_ns)) \ {}" (is "?Q n_ns") and "\nd\snd n_ns. \ or\<^sub>n \ \ old nd \ {\, \} \ old nd \ {}" (is "?P (snd n_ns)") shows "expand n_ns \ SPEC (\nm_nds. ?P (snd nm_nds))" using assms proof (rule_tac expand_rec_rule[where \="\x. ?Q x \ ?P (snd x)"], simp, intro refine_vcg, goal_cases) case 1 then show ?case by (simp, rule_tac upd_incoming__ident) auto next case prems: (4 f x n ns) then have step: "\x. ?Q x \ ?P (snd x) \ f x \ SPEC (\x'. ?P (snd x'))" by simp with prems show ?case by (rule_tac step) auto next case prems: (5 f x n ns) then have step: "\x. ?Q x \ ?P (snd x) \ f x \ SPEC (\x'. ?P (snd x'))" by simp with prems show ?case by (rule_tac step) auto next case (6 f x n ns) then show ?case by auto next case prems: (7 f x n ns) then have step: "\x. ?Q x \ ?P (snd x) \ f x \ SPEC (\x'. ?P (snd x'))" by simp with prems show ?case by (rule_tac step) auto next case prems: (8 f x n ns \) then have goal_assms: "\ \ new n \ \ (\q. \ = prop\<^sub>n(q) \ \ = nprop\<^sub>n(q)) \ \ \ true\<^sub>n \ \ \ false\<^sub>n \ \ (\\ \. \ = \ and\<^sub>n \ \ \ = X\<^sub>n \) \ (\\ \. \ = \ or\<^sub>n \ \ \ = \ U\<^sub>n \ \ \ = \ V\<^sub>n \)" by (cases \) auto from prems have QP: "?Q (n, ns) \ ?P ns" and step: "\x. ?Q x \ ?P (snd x) \ f x \ SPEC (\x'. ?P (snd x'))" by simp_all show ?case using goal_assms QP unfolding case_prod_unfold \x = (n, ns)\ proof (rule_tac SPEC_rule_nested2, goal_cases) case 1 then show ?case by (rule_tac step) auto next case 2 then show ?case by (rule_tac step) auto qed qed auto lemma or_propag_on_create_graph: "create_graph \ \ SPEC (\nds. \nd\nds. \ or\<^sub>n \ \ old nd \ {\, \} \ old nd \ {})" unfolding create_graph_def by (intro refine_vcg, rule_tac order_trans, rule_tac expand_or_propag) auto abbreviation "next_propag__assm \ n_ns \ (X\<^sub>n \ \ old (fst n_ns) \ \ \ next (fst n_ns)) \ (\nd\snd n_ns. X\<^sub>n \ \ old nd \ name nd \ incoming (fst n_ns) \ \ \ old (fst n_ns) \ new (fst n_ns))" abbreviation "next_propag__rslt \ ns \ \nd\ns. \nd'\ns. X\<^sub>n \ \ old nd \ name nd \ incoming nd' \ \ \ old nd'" lemma expand_next_propag: fixes n_ns :: "_ \ 'a node set" assumes "next_propag__assm \ n_ns \ next_propag__rslt \ (snd n_ns) \ expand_assm_incoming n_ns \ expand_name_ident (snd n_ns)" (is "?Q n_ns") shows "expand n_ns \ SPEC (\r. next_propag__rslt \ (snd r))" (is "_ \ SPEC ?P") using assms proof (rule_tac expand_rec_rule[where \="?Q"], simp, intro refine_vcg, goal_cases) case (1 f x n ns) then show ?case proof (simp, rule_tac upd_incoming__ident, goal_cases) case prems: 1 { fix nd :: "'a node" and nd' :: "'a node" assume "nd\ns" and nd'_elem: "nd'\upd_incoming n ns" have "\ \ old nd'" if *: "X\<^sub>n \ \ old nd" and **: "name nd \ incoming nd'" proof (cases "nd'\ns") case True with prems * ** show ?thesis using \nd\ns\ by auto next case False with upd_incoming__elem[of nd' n ns] nd'_elem * ** obtain nd'' where "nd''\ns" and nd'_eq: "nd' = nd''\incoming := incoming n \ incoming nd''\" and old_eq: "old nd'' = old n" by auto have "\ \ old nd'" proof (cases "name nd \ incoming n") case True with prems * \nd\ns\ have "\ \ old n" by auto then show ?thesis using nd'_eq old_eq by simp next case False then have "name nd \ incoming nd''" using \name nd \incoming nd'\ nd'_eq by simp then show ?thesis unfolding nd'_eq using \nd\ns\ \nd''\ns\ * prems by auto qed then show ?thesis by auto qed } then show ?case by auto next case 2 then show ?case by simp qed next case prems: (2 f x n ns) then have step: "\x. ?Q x \ f x \ SPEC ?P" and f_sup: "\x. f x \ expand x" by auto from prems have Q: "?Q (n, ns)" by auto from Q have name_le: "name n < expand_new_name (name n)" by auto let ?x' = "(\name = expand_new_name (name n), incoming = {name n}, new = next n, old = {}, next = {}\, {n} \ ns)" have Q'1: "expand_assm_incoming ?x'\ expand_name_ident (snd ?x')" using \new n = {}\ Q[THEN conjunct2, THEN conjunct2] name_le by force have Q'2: "next_propag__assm \ ?x' \ next_propag__rslt \ (snd ?x')" using Q \new n = {}\ by auto show ?case using \new n = {}\ unfolding \x = (n, ns)\ proof (rule_tac SPEC_rule_weak[where Q = "\_ r. expand_name_ident (snd r)" and P = "\_. ?P"], goal_cases) case 1 then show ?case using Q'1 by (rule_tac order_trans, rule_tac f_sup, rule_tac expand_name_propag__name_ident) auto next case 2 then show ?case using Q'1 Q'2 by (rule_tac step) simp next case (3 nm nds) then show ?case by simp qed next case 3 then show ?case by auto next case prems: (4 f) then have step: "\x. ?Q x \ f x \ SPEC ?P" by simp from prems show ?case by (rule_tac SPEC_rule_param2, rule_tac step) auto next case prems: (5 f) then have step: "\x. ?Q x \ f x \ SPEC ?P" by simp from prems show ?case by (rule_tac SPEC_rule_param2, rule_tac step) auto next case 6 then show ?case by auto next case prems: (7 f) then have step: "\x. ?Q x \ f x \ SPEC ?P" by simp from prems show ?case by (rule_tac SPEC_rule_param2, rule_tac step) auto next case prems: (8 f x n ns \) then have goal_assms: "\ \ new n \ (\\ \. \ = \ or\<^sub>n \ \ \ = \ U\<^sub>n \ \ \ = \ V\<^sub>n \)" by (cases \) auto from prems have Q: "?Q (n, ns)" and step: "\x. ?Q x \ f x \ SPEC ?P" and f_sup: "\x. f x \ expand x" by auto let ?x = "(n\new := new n - {\}, new := new1 \ \ new (n\new := new n - {\}\), old := {\} \ old (n\new := new n - {\}\), next := next1 \ \ next (n\new := new n - {\}\) \, ns)" let ?props = "\x r. expand_rslt_exist_eq x r \ expand_rslt_incoming r \ expand_rslt_name x r \ expand_name_ident (snd r)" show ?case using goal_assms Q unfolding case_prod_unfold \x = (n, ns)\ proof (rule_tac SPEC_rule_weak_nested2[where Q = "?props ?x"], goal_cases) case 1 then show ?case by (rule_tac SPEC_rule_conjI, rule_tac order_trans, rule_tac f_sup, rule_tac expand_rslt_exist_eq, rule_tac order_trans, rule_tac f_sup, rule_tac expand_name_propag) simp+ next case 2 then show ?case by (rule_tac SPEC_rule_param2[where P = "\_. ?P"], rule_tac step) auto next case prems': (3 nm nds) let ?x' = "(n\new := new n - {\}, name := fst (nm, nds), new := new2 \ \ new (n\new := new n - {\}\), old := {\} \ old (n\new := new n - {\}\)\, nds)" from prems' show ?case proof (rule_tac step, goal_cases) case prems'': 1 then have "expand_assm_incoming ?x'" by auto moreover from prems'' have nds_ident: "expand_name_ident (snd ?x')" by simp moreover have "X\<^sub>n \ \ old (fst ?x') \ \\next (fst ?x')" using Q[THEN conjunct1] goal_assms by auto moreover from prems'' have "next_propag__rslt \ (snd ?x')" by simp moreover from prems'' have name_nds_eq: "name ` nds = name ` ns \ name ` {nd\nds. name nd \ name n}" by auto have "\nd\nds. (X\<^sub>n \ \ old nd \ name nd \ incoming (fst ?x')) \ \\old (fst ?x')\new (fst ?x')" (is "\nd\nds. ?assm (fst ?x') nd \ ?concl (fst ?x') nd") proof fix nd assume "nd\nds" { assume loc_assm: "name nd\name ` ns" then obtain n' where n': "n'\ns" "name n' = name nd" by auto moreover from prems'' n' obtain nd' where "nd'\nds" and n'_nd'_eq: "expand_rslt_exist_eq__node n' nd'" by auto ultimately have "nd = nd'" using nds_ident \nd\nds\ loc_assm by auto moreover from prems'' have "?assm n n' \ ?concl n n'" using \n'\ns\ by auto ultimately have "?assm (fst ?x') nd \ ?concl (fst ?x') nd" using n'_nd'_eq by auto } moreover { assume "name nd\name ` ns" with name_nds_eq \nd\nds\ have "name nd \ name n" by auto with prems'' have "\ (?assm (fst ?x') nd)" by auto } ultimately show "?assm (fst ?x') nd \ ?concl (fst ?x') nd" by auto qed ultimately show ?case by simp qed qed qed lemma next_propag_on_create_graph: "create_graph \ \ SPEC (\nds. \n\nds. \n'\nds. X\<^sub>n \\old n \ name n\incoming n' \ \\old n')" unfolding create_graph_def by (intro refine_vcg, rule_tac order_trans, rule_tac expand_next_propag) (auto simp add:expand_new_name_expand_init) abbreviation "release_propag__assm \ \ n_ns \ (\ V\<^sub>n \ \ old (fst n_ns) \ {\, \}\old (fst n_ns)\new (fst n_ns) \ (\\old (fst n_ns)\new (fst n_ns)) \ \ V\<^sub>n \ \ next (fst n_ns)) \ (\nd\snd n_ns. \ V\<^sub>n \ \ old nd \ name nd \ incoming (fst n_ns) \ {\, \}\old nd \ (\\old nd \ \ V\<^sub>n \ \ old (fst n_ns)\new (fst n_ns)))" abbreviation "release_propag__rslt \ \ ns \ \nd\ns. \nd'\ns. \ V\<^sub>n \ \ old nd \ name nd \ incoming nd' \ {\, \}\old nd \ (\\old nd \ \ V\<^sub>n \ \ old nd')" lemma expand_release_propag: fixes n_ns :: "_ \ 'a node set" assumes "release_propag__assm \ \ n_ns \ release_propag__rslt \ \ (snd n_ns) \ expand_assm_incoming n_ns \ expand_name_ident (snd n_ns)" (is "?Q n_ns") shows "expand n_ns \ SPEC (\r. release_propag__rslt \ \ (snd r))" (is "_ \ SPEC ?P") using assms proof (rule_tac expand_rec_rule[where \="?Q"], simp, intro refine_vcg, goal_cases) case (1 f x n ns) then show ?case proof (simp, rule_tac upd_incoming__ident, goal_cases) case prems: 1 { fix nd :: "'a node" and nd' :: "'a node" let ?V_prop = "\ V\<^sub>n \ \ old nd \ name nd \ incoming nd' \ {\, \} \ old nd \ \ \ old nd \ \ V\<^sub>n \ \ old nd'" assume "nd\ns" and nd'_elem: "nd'\upd_incoming n ns" { assume "nd'\ns" with prems have ?V_prop using \nd\ns\ by auto } moreover { assume "nd'\ns" and V_in_nd: "\ V\<^sub>n \ \ old nd" and "name nd \incoming nd'" with upd_incoming__elem[of nd' n ns] nd'_elem obtain nd'' where "nd''\ns" and nd'_eq: "nd' = nd''\incoming := incoming n \ incoming nd''\" and old_eq: "old nd'' = old n" by auto { assume "name nd \ incoming n" with prems V_in_nd \nd\ns\ have "{\, \} \ old nd \ \ \ old nd \ \ V\<^sub>n \ \ old n" by auto then have "{\, \} \ old nd \ \ \ old nd \ \ V\<^sub>n \ \ old nd'" using nd'_eq old_eq by simp } moreover { assume "name nd \ incoming n" then have "name nd \ incoming nd''" using \name nd \incoming nd'\ nd'_eq by simp then have "{\, \} \ old nd \ \ \ old nd \ \ V\<^sub>n \ \ old nd'" unfolding nd'_eq using prems \nd\ns\ \nd''\ns\ V_in_nd by auto } ultimately have "{\, \} \ old nd \ \ \ old nd \ \ V\<^sub>n \ \ old nd'" by fast } ultimately have ?V_prop by auto } then show ?case by auto next case 2 then show ?case by simp qed next case prems: (2 f x n ns) then have step: "\x. ?Q x \ f x \ SPEC ?P" and f_sup: "\x. f x \ expand x" by auto from prems have Q: "?Q (n, ns)" by auto from Q have name_le: "name n < expand_new_name (name n)" by auto let ?x' = "(\name = expand_new_name (name n), incoming = {name n}, new = next n, old = {}, next = {}\, {n} \ ns)" have Q'1: "expand_assm_incoming ?x'\ expand_name_ident (snd ?x')" using \new n = {}\ Q[THEN conjunct2, THEN conjunct2] name_le by force have Q'2: "release_propag__assm \ \ ?x' \ release_propag__rslt \ \ (snd ?x')" using Q \new n = {}\ by auto show ?case using \new n = {}\ unfolding \x = (n, ns)\ proof (rule_tac SPEC_rule_weak[where Q = "\_ r. expand_name_ident (snd r)" and P = "\_. ?P"], goal_cases) case 1 then show ?case using Q'1 by (rule_tac order_trans, rule_tac f_sup, rule_tac expand_name_propag__name_ident) auto next case 2 then show ?case using Q'1 Q'2 by (rule_tac step) simp next case (3 nm nds) then show ?case by simp qed next case 3 then show ?case by auto next case prems: (4 f x n ns \) then have goal_assms: "\ \ new n \ (\q. \ = prop\<^sub>n(q) \ \ = nprop\<^sub>n(q))" by simp from prems have step: "\x. ?Q x \ f x \ SPEC ?P" and Q: "?Q (n, ns)" by simp_all show ?case using Q goal_assms by (rule_tac SPEC_rule_param2, rule_tac step) auto next case prems: (5 f x n ns \) then have goal_assms: "\ \ new n \ \ = true\<^sub>n" by simp from prems have step: "\x. ?Q x \ f x \ SPEC ?P" and Q: "?Q (n, ns)" by simp_all show ?case using Q goal_assms by (rule_tac SPEC_rule_param2, rule_tac step) auto next case 6 then show ?case by auto next case prems: (7 f x n ns \) then have goal_assms: "\ \ new n \ (\\ \. \ = \ and\<^sub>n \ \ \ = X\<^sub>n \)" by simp from prems have step: "\x. ?Q x \ f x \ SPEC ?P" and Q: "?Q (n, ns)" by simp_all show ?case using Q goal_assms by (rule_tac SPEC_rule_param2, rule_tac step) auto next case prems: (8 f x n ns \) then have goal_assms: "\ \ new n \ (\\ \. \ = \ or\<^sub>n \ \ \ = \ U\<^sub>n \ \ \ = \ V\<^sub>n \)" by (cases \) auto from prems have Q: "?Q (n, ns)" and step: "\x. ?Q x \ f x \ SPEC ?P" and f_sup: "\x. f x \ expand x" by auto let ?x = "(n\new := new n - {\}, new := new1 \ \ new (n\new := new n - {\}\), old := {\} \ old (n\new := new n - {\}\), next := next1 \ \ next (n\new := new n - {\}\) \, ns)" let ?props = "\x r. expand_rslt_exist_eq x r \ expand_rslt_incoming r \ expand_rslt_name x r \ expand_name_ident (snd r)" show ?case using goal_assms Q unfolding case_prod_unfold \x = (n, ns)\ proof (rule_tac SPEC_rule_weak_nested2[where Q = "?props ?x"], goal_cases) case 1 then show ?case by (rule_tac SPEC_rule_conjI, rule_tac order_trans, rule_tac f_sup, rule_tac expand_rslt_exist_eq, rule_tac order_trans, rule_tac f_sup, rule_tac expand_name_propag) simp+ next case 2 then show ?case by (rule_tac SPEC_rule_param2[where P = "\_. ?P"], rule_tac step) auto next case prems': (3 nm nds) let ?x' = "(n\new := new n - {\}, name := fst (nm, nds), new := new2 \ \ new (n\new := new n - {\}\), old := {\} \ old (n\new := new n - {\}\)\, nds)" from prems' show ?case proof (rule_tac step, goal_cases) case prems'': 1 then have "expand_assm_incoming ?x'" by auto moreover from prems'' have nds_ident: "expand_name_ident (snd ?x')" by simp moreover have "(\ V\<^sub>n \ \ old (fst ?x') \ ({\, \}\old (fst ?x') \ new (fst ?x') \ (\\old (fst ?x')\new (fst ?x') \ \ V\<^sub>n \ \ next (fst ?x'))))" using Q[THEN conjunct1] goal_assms by auto moreover from prems'' have "release_propag__rslt \ \ (snd ?x')" by simp moreover from prems'' have name_nds_eq: "name ` nds = name ` ns \ name ` {nd\nds. name nd \ name n}" by auto have "\nd\nds. (\ V\<^sub>n \ \ old nd \ name nd \ incoming (fst ?x')) \ ({\, \}\old nd \ (\\old nd \ \ V\<^sub>n \ \old (fst ?x')\new (fst ?x')))" (is "\nd\nds. ?assm (fst ?x') nd \ ?concl (fst ?x') nd") proof fix nd assume "nd\nds" { assume loc_assm: "name nd\name ` ns" then obtain n' where n': "n'\ns" "name n' = name nd" by auto with prems'' obtain nd' where "nd'\nds" and n'_nd'_eq: "expand_rslt_exist_eq__node n' nd'" by auto with n' have "nd = nd'" using nds_ident \nd\nds\ loc_assm by auto moreover from prems'' have "?assm n n' \ ?concl n n'" using \n'\ns\ by auto ultimately have "?assm (fst ?x') nd \ ?concl (fst ?x') nd" using n'_nd'_eq by auto } moreover { assume "name nd\name ` ns" with name_nds_eq \nd\nds\ have "name nd \ name n" by auto with prems'' have "\ (?assm (fst ?x') nd)" by auto } ultimately show "?assm (fst ?x') nd \ ?concl (fst ?x') nd" by auto qed ultimately show ?case by simp qed qed qed lemma release_propag_on_create_graph: "create_graph \ \ SPEC (\nds. \n\nds. \n'\nds. \ V\<^sub>n \\old n \ name n\incoming n' \ ({\, \}\old n \ \\old n \ \ V\<^sub>n \\old n'))" unfolding create_graph_def by (intro refine_vcg, rule_tac order_trans, rule_tac expand_release_propag) (auto simp add:expand_new_name_expand_init) abbreviation "until_propag__assm f g n_ns \ (f U\<^sub>n g \ old (fst n_ns) \ (g\old (fst n_ns)\new (fst n_ns) \ (f\old (fst n_ns)\new (fst n_ns) \ f U\<^sub>n g \ next (fst n_ns)))) \ (\nd\snd n_ns. f U\<^sub>n g \ old nd \ name nd \ incoming (fst n_ns) \ (g\old nd \ (f\old nd \ f U\<^sub>n g \old (fst n_ns)\new (fst n_ns))))" abbreviation "until_propag__rslt f g ns \ \n\ns. \nd\ns. f U\<^sub>n g \ old n \ name n \ incoming nd \ (g \ old n \ (f\old n \ f U\<^sub>n g \ old nd))" lemma expand_until_propag: fixes n_ns :: "_ \ 'a node set" assumes "until_propag__assm \ \ n_ns \ until_propag__rslt \ \ (snd n_ns) \ expand_assm_incoming n_ns \ expand_name_ident (snd n_ns)" (is "?Q n_ns") shows "expand n_ns \ SPEC (\r. until_propag__rslt \ \ (snd r))" (is "_ \ SPEC ?P") using assms proof (rule_tac expand_rec_rule[where \="?Q"], simp, intro refine_vcg, goal_cases) case prems: (1 f x n ns) then show ?case proof (simp, rule_tac upd_incoming__ident, goal_cases) case prems': 1 { fix nd :: "'a node" and nd' :: "'a node" let ?U_prop = "\ U\<^sub>n \ \ old nd \ name nd \ incoming nd' \ \ \ old nd \ \ \ old nd \ \ U\<^sub>n \ \ old nd'" assume "nd\ns" and nd'_elem: "nd'\upd_incoming n ns" { assume "nd'\ns" with prems' have ?U_prop using \nd\ns\ by auto } moreover { assume "nd'\ns" and U_in_nd: "\ U\<^sub>n \ \ old nd" and "name nd \incoming nd'" with upd_incoming__elem[of nd' n ns] nd'_elem obtain nd'' where "nd''\ns" and nd'_eq: "nd' = nd''\incoming := incoming n \ incoming nd''\" and old_eq: "old nd'' = old n" by auto { assume "name nd \ incoming n" with prems' U_in_nd \nd\ns\ have "\ \ old nd \ \ \ old nd \ \ U\<^sub>n \ \ old n" by auto then have "\ \ old nd \ \ \ old nd \ \ U\<^sub>n \ \ old nd'" using nd'_eq old_eq by simp } moreover { assume "name nd \ incoming n" then have "name nd \ incoming nd''" using \name nd \incoming nd'\ nd'_eq by simp then have "\ \ old nd \ \ \ old nd \ \ U\<^sub>n \ \ old nd'" unfolding nd'_eq using prems' \nd\ns\ \nd''\ns\ U_in_nd by auto } ultimately have "\ \ old nd \ \ \ old nd \ \ U\<^sub>n \ \ old nd'" by fast } ultimately have ?U_prop by auto } then show ?case by auto next case 2 then show ?case by simp qed next case prems: (2 f x n ns) then have step: "\x. ?Q x \ f x \ SPEC ?P" and f_sup: "\x. f x \ expand x" by auto from prems have Q: "?Q (n, ns)" by auto from Q have name_le: "name n < expand_new_name (name n)" by auto let ?x' = "(\name = expand_new_name (name n), incoming = {name n}, new = next n, old = {}, next = {}\, {n} \ ns)" have Q'1: "expand_assm_incoming ?x'\ expand_name_ident (snd ?x')" using \new n = {}\ Q[THEN conjunct2, THEN conjunct2] name_le by force have Q'2: "until_propag__assm \ \ ?x' \ until_propag__rslt \ \ (snd ?x')" using Q \new n = {}\ by auto show ?case using \new n = {}\ unfolding \x = (n, ns)\ proof (rule_tac SPEC_rule_weak[where Q = "\_ r. expand_name_ident (snd r)" and P = "\_. ?P"], goal_cases) case 1 then show ?case using Q'1 by (rule_tac order_trans, rule_tac f_sup, rule_tac expand_name_propag__name_ident) auto next case 2 then show ?case using Q'1 Q'2 by (rule_tac step) simp next case (3 nm nds) then show ?case by simp qed next case 3 then show ?case by auto next case prems: (4 f) then have step: "\x. ?Q x \ f x \ SPEC ?P" by simp_all from prems show ?case by (rule_tac SPEC_rule_param2, rule_tac step) auto next case prems: (5 f) then have step: "\x. ?Q x \ f x \ SPEC ?P" by simp_all from prems show ?case by (rule_tac SPEC_rule_param2, rule_tac step) auto next case 6 then show ?case by auto next case prems: (7 f) then have step: "\x. ?Q x \ f x \ SPEC ?P" by simp_all from prems show ?case by (rule_tac SPEC_rule_param2, rule_tac step) auto next case prems: (8 f x n ns \) then have goal_assms: "\ \ new n \ (\\ \. \ = \ or\<^sub>n \ \ \ = \ U\<^sub>n \ \ \ = \ V\<^sub>n \)" by (cases \) auto from prems have Q: "?Q (n, ns)" and step: "\x. ?Q x \ f x \ SPEC ?P" and f_sup: "\x. f x \ expand x" by auto let ?x = "(n\new := new n - {\}, new := new1 \ \ new (n\new := new n - {\}\), old := {\} \ old (n\new := new n - {\}\), next := next1 \ \ next (n\new := new n - {\}\) \, ns)" let ?props = "\x r. expand_rslt_exist_eq x r \ expand_rslt_incoming r \ expand_rslt_name x r \ expand_name_ident (snd r)" show ?case using goal_assms Q unfolding case_prod_unfold \x = (n, ns)\ proof (rule_tac SPEC_rule_weak_nested2[where Q = "?props ?x"], goal_cases) case 1 then show ?case by (rule_tac SPEC_rule_conjI, rule_tac order_trans, rule_tac f_sup, rule_tac expand_rslt_exist_eq, rule_tac order_trans, rule_tac f_sup, rule_tac expand_name_propag) simp+ next case 2 then show ?case by (rule_tac SPEC_rule_param2[where P = "\_. ?P"], rule_tac step) auto next case prems: (3 nm nds) let ?x' = "(n\new := new n - {\}, name := fst (nm, nds), new := new2 \ \ new (n\new := new n - {\}\), old := {\} \ old (n\new := new n - {\}\)\, nds)" from prems show ?case proof (rule_tac step, goal_cases) case prems': 1 then have "expand_assm_incoming ?x'" by auto moreover from prems' have nds_ident: "expand_name_ident (snd ?x')" by simp moreover have "(\ U\<^sub>n \ \ old (fst ?x') \ (\\old (fst ?x')\new (fst ?x') \ (\\old (fst ?x')\new (fst ?x') \ \ U\<^sub>n \ \ next (fst ?x'))))" using Q[THEN conjunct1] goal_assms by auto moreover from prems' have "until_propag__rslt \ \ (snd ?x')" by simp moreover from prems' have name_nds_eq: "name ` nds = name ` ns \ name ` {nd\nds. name nd \ name n}" by auto have "\nd\nds. (\ U\<^sub>n \ \ old nd \ name nd \ incoming (fst ?x')) \ (\\old nd \ (\\old nd \ \ U\<^sub>n \ \old (fst ?x')\new (fst ?x')))" (is "\nd\nds. ?assm (fst ?x') nd \ ?concl (fst ?x') nd") proof fix nd assume "nd\nds" { assume loc_assm: "name nd\name ` ns" then obtain n' where n': "n'\ns" "name n' = name nd" by auto moreover from prems' n' obtain nd' where "nd'\nds" and n'_nd'_eq: "expand_rslt_exist_eq__node n' nd'" by auto ultimately have "nd = nd'" using nds_ident \nd\nds\ loc_assm by auto moreover from prems' have "?assm n n' \ ?concl n n'" using \n'\ns\ by auto ultimately have "?assm (fst ?x') nd \ ?concl (fst ?x') nd" using n'_nd'_eq by auto } moreover { assume "name nd\name ` ns" with name_nds_eq \nd\nds\ have "name nd \ name n" by auto with prems' have "\ (?assm (fst ?x') nd)" by auto } ultimately show "?assm (fst ?x') nd \ ?concl (fst ?x') nd" by auto qed ultimately show ?case by simp qed qed qed lemma until_propag_on_create_graph: "create_graph \ \ SPEC (\nds. \n\nds. \n'\nds. \ U\<^sub>n \\old n \ name n\incoming n' \ (\\old n \ \\old n \ \ U\<^sub>n \\old n'))" unfolding create_graph_def by (intro refine_vcg, rule_tac order_trans, rule_tac expand_until_propag) (auto simp add:expand_new_name_expand_init) definition all_subfrmls :: "'a node \ 'a frml set" where "all_subfrmls n \ \(subfrmlsn ` (new n \ old n \ next n))" lemma all_subfrmls__UnionD: assumes "(\x\A. subfrmlsn x) \ B" and "x\A" and "y\subfrmlsn x" shows "y\B" proof - note subfrmlsn_id[of x] also have "subfrmlsn x \ (\x\A. subfrmlsn x)" using assms by auto finally show ?thesis using assms by auto qed lemma expand_all_subfrmls_propag: assumes "all_subfrmls (fst n_ns) \ B \ (\nd\snd n_ns. all_subfrmls nd \ B)" (is "?Q n_ns") shows "expand n_ns \ SPEC (\r. \nd\snd r. all_subfrmls nd \ B)" (is "_ \ SPEC ?P") using assms proof (rule_tac expand_rec_rule[where \="?Q"], simp, intro refine_vcg, goal_cases) case 1 then show ?case proof (simp, rule_tac upd_incoming__ident, goal_cases) case 1 then show ?case by auto next case 2 then show ?case by (simp add: all_subfrmls_def) qed next case 2 then show ?case by (auto simp add: all_subfrmls_def) next case 3 then show ?case by (auto simp add: all_subfrmls_def) next case prems: (4 f) then have step: "\x. ?Q x \ f x \ SPEC ?P" by simp_all from prems show ?case by (rule_tac step) (auto simp add: all_subfrmls_def) next case prems: (5 f _ n ns \) then have goal_assms: "\ \ new n \ \ = true\<^sub>n" by simp from prems have step: "\x. ?Q x \ f x \ SPEC ?P" and Q: "?Q (n, ns)" by simp_all show ?case using Q goal_assms by (rule_tac step) (auto dest: all_subfrmls__UnionD simp add: all_subfrmls_def) next case 6 then show ?case by auto next case prems: (7 f x n ns \) then have goal_assms: "\ \ new n \ (\\ \. \ = \ and\<^sub>n \ \ \ = X\<^sub>n \)" by simp from prems have step: "\x. ?Q x \ f x \ SPEC ?P" and Q: "?Q (n, ns)" by simp_all show ?case using Q goal_assms by (rule_tac step) (auto dest: all_subfrmls__UnionD simp add: all_subfrmls_def) next case prems: (8 f x n ns \) then have goal_assms: "\ \ new n \ \ (\q. \ = prop\<^sub>n(q) \ \ = nprop\<^sub>n(q)) \ \ \ true\<^sub>n \ \ \ false\<^sub>n \ \ (\\ \. \ = \ and\<^sub>n \ \ \ = X\<^sub>n \) \ (\\ \. \ = \ or\<^sub>n \ \ \ = \ U\<^sub>n \ \ \ = \ V\<^sub>n \)" by (cases \) auto from prems have Q: "?Q (n, ns)" and step: "\x. ?Q x \ f x \ SPEC ?P" by simp_all show ?case using goal_assms Q unfolding case_prod_unfold \x = (n, ns)\ proof (rule_tac SPEC_rule_nested2, goal_cases) case 1 then show ?case by (rule_tac step) (auto dest: all_subfrmls__UnionD simp: all_subfrmls_def) next case 2 then show ?case by (rule_tac step) (auto dest: all_subfrmls__UnionD simp: all_subfrmls_def) qed qed lemma old_propag_on_create_graph: "create_graph \ \ SPEC (\nds. \n\nds. old n \ subfrmlsn \)" unfolding create_graph_def by (intro refine_vcg, rule_tac order_trans, rule_tac expand_all_subfrmls_propag[where B = "subfrmlsn \"]) (force simp add:all_subfrmls_def expand_new_name_expand_init)+ lemma L4_2__aux: assumes run: "ipath gba.E \" and "\ U\<^sub>n \ \ old (\ 0)" and "\j. (\i, \ U\<^sub>n \} \ old (\ i)) \ \ \ old (\ j)" shows "\i. {\, \ U\<^sub>n \} \ old (\ i) \ \ \ old (\ i)" proof - have "\i, \ U\<^sub>n \} \ old (\ i)" (is "?sbthm j") for j proof (induct j) show "?sbthm 0" by auto next fix k assume step: "?sbthm k" then have \_k_prop: "\ \ old (\ k) \ \ k\qs \ \ (Suc k)\qs \ name (\ k) \ incoming (\ (Suc k))" using assms run_propag_on_create_graph[OF run] by auto with inres_SPEC[OF res until_propag_on_create_graph[where \ = \ and \ = \]] have "{\, \ U\<^sub>n \} \ old (\ k)" (is "?subsetthm") proof (cases k) assume "k = 0" with assms \_k_prop inres_SPEC[OF res until_propag_on_create_graph[where \ = \ and \ = \]] show ?subsetthm by auto next fix l assume "k = Suc l" then have "{\, \ U\<^sub>n \}\old (\ l) \ \\old (\ l) \ \ l\qs \ \ k\qs \ name (\ l)\incoming (\ k)" using step assms run_propag_on_create_graph[OF run] by auto with inres_SPEC[OF res until_propag_on_create_graph[where \ = \ and \ = \]] have "\ U\<^sub>n \\old (\ k)" by auto with \_k_prop inres_SPEC[OF res until_propag_on_create_graph[where \ = \ and \ = \]] show ?subsetthm by auto qed with step show "?sbthm (Suc k)" by (metis less_SucE) qed with assms show ?thesis by auto qed lemma L4_2a: assumes "ipath gba.E \" and "\ U\<^sub>n \ \ old (\ 0)" shows "(\i. {\, \ U\<^sub>n \} \ old (\ i) \ \ \ old (\ i)) \ (\j. (\i, \ U\<^sub>n \} \ old (\ i)) \ \ \ old (\ j))" (is "?A \ ?B") proof (rule disjCI) assume "\ ?B" then show ?A using assms by (rule_tac L4_2__aux) blast+ qed lemma L4_2b: assumes run: "ipath gba.E \" and "\ U\<^sub>n \ \ old (\ 0)" and ACC: "gba.is_acc \" shows "\j. (\i, \ U\<^sub>n \} \ old (\ i)) \ \ \ old (\ j)" proof (rule ccontr) assume "\ ?thesis" then have contr: "\i. {\, \ U\<^sub>n \}\old(\ i) \ \\old(\ i)" using assms L4_2a[of \ \ \] by blast def S \ "{q \ qs. \ U\<^sub>n \ \ old q \ \ \ old q}" from assms inres_SPEC[OF res old_propag_on_create_graph] create_gba_from_nodes__ipath have "\ U\<^sub>n \ \ subfrmlsn \" by (metis assms(2) set_mp) then have "S\gbg_F(create_gba_from_nodes \ qs)" unfolding S_def create_gba_from_nodes_def by auto with ACC have 1: "\\<^sub>\i. \ i \ S" unfolding gba.is_acc_def by blast from INFM_EX[OF 1] obtain k where "\ k \ qs" and "\ U\<^sub>n \ \ old (\ k) \ \ \ old (\ k)" unfolding S_def by auto moreover have "{\, \ U\<^sub>n \}\old(\ k) \ \\old(\ k)" using contr by auto ultimately show False by auto qed lemma L4_2c: assumes run: "ipath gba.E \" and "\ V\<^sub>n \ \ old (\ 0)" shows "\i. \ \ old (\ i) \ (\j \ old (\ j))" proof - have "{\, \ V\<^sub>n \} \ old (\ i) \ (\j \ old (\ j))" (is "?thm i") for i proof (induct i) case 0 have "\ 0\qs \ \ 1\qs \ name (\ 0) \ incoming (\ 1)" using create_gba_from_nodes__ipath assms by auto then show ?case using assms inres_SPEC[OF res release_propag_on_create_graph, of \ \] by auto next case (Suc k) note \?thm k\ moreover { assume "{\, \ V\<^sub>n \} \ old (\ k)" moreover have "\ k\qs \ \ (Suc k)\qs \ name (\ k) \ incoming (\ (Suc k))" using create_gba_from_nodes__ipath assms by auto ultimately have "\ \ old (\ k) \ \ V\<^sub>n \ \ old (\ (Suc k))" using assms inres_SPEC[OF res release_propag_on_create_graph, of \ \] by auto moreover { assume "\ \ old (\ k)" then have ?case by blast } moreover { assume "\ V\<^sub>n \ \ old (\ (Suc k))" moreover have "\ (Suc k)\qs \ \ (Suc (Suc k))\qs \ name (\ (Suc k)) \ incoming (\ (Suc (Suc k)))" using create_gba_from_nodes__ipath assms by auto ultimately have ?case using assms inres_SPEC[OF res release_propag_on_create_graph, of \ \] by auto } ultimately have ?case by fast } moreover { assume "\j \ old (\ j)" then have ?case by auto } ultimately show ?case by auto qed then show ?thesis by auto qed lemma L4_8': assumes "ipath gba.E \" (is "?inf_run \") and "gba.is_acc \" (is "?gbarel_accept \") and "\i. gba.L (\ i) (\ i)" (is "?lgbarel_accept \ \") and "\ \ old (\ 0)" shows "\ \\<^sub>n \" using assms proof (induct \ arbitrary: \ \) case LTLnTrue show ?case by auto next case LTLnFalse then show ?case using inres_SPEC[OF res false_propag_on_create_graph] create_gba_from_nodes__ipath by (metis) next case (LTLnProp p) then show ?case unfolding create_gba_from_nodes_def by auto next case (LTLnNProp p) then show ?case unfolding create_gba_from_nodes_def by auto next case (LTLnAnd \ \) then show ?case using inres_SPEC[OF res and_propag_on_create_graph, of \ \] create_gba_from_nodes__ipath by (metis insert_subset ltln_semantics.simps(5)) next case (LTLnOr \ \) then have "\ \ old (\ 0) \ \ \ old (\ 0)" using inres_SPEC[OF res or_propag_on_create_graph, of \ \] create_gba_from_nodes__ipath by (metis (full_types) Int_empty_left Int_insert_left_if0) moreover have "\ \\<^sub>n \" if "\ \ old (\ 0)" using LTLnOr that by auto moreover have "\ \\<^sub>n \" if "\ \ old (\ 0)" using LTLnOr that by auto ultimately show ?case by auto next case (LTLnNext \) with create_gba_from_nodes__ipath[of \] have "\ 0 \ qs \ \ 1 \ qs \ name (\ 0) \ incoming (\ 1)" by auto with inres_SPEC[OF res next_propag_on_create_graph, of \] have "\\old (suffix 1 \ 0)" using LTLnNext by auto moreover have "?inf_run (suffix 1 \)" and "?gbarel_accept (suffix 1 \)" and "?lgbarel_accept (suffix 1 \) (suffix 1 \ )" using LTLnNext create_gba_from_nodes__ipath apply - apply (metis ipath_suffix) apply (auto simp del: suffix_nth) [] (* FIXME: "\a. suffix i \ a" is unfolded, but "suffix i \" is not! *) apply auto done ultimately show ?case using LTLnNext by simp next case (LTLnUntil \ \) then have "\j. (\i, \ U\<^sub>n \} \ old (\ i)) \ \ \ old (\ j)" using L4_2b by auto then obtain j where \_pre: "\i, \ U\<^sub>n \} \ old (\ i)" and "\ \ old (suffix j \ 0)" by auto moreover have "?inf_run (suffix j \)" and "?gbarel_accept (suffix j \)" and "?lgbarel_accept (suffix j \) (suffix j \)" unfolding limit_suffix using LTLnUntil create_gba_from_nodes__ipath apply - apply (metis ipath_suffix) apply (auto simp del: suffix_nth) [] (* FIXME: "\a. suffix i \ a" is unfolded, but "suffix i \" is not! *) apply auto done ultimately have "suffix j \ \\<^sub>n \" using LTLnUntil by simp moreover { fix i assume "i < j" have "?inf_run (suffix i \)" and "?gbarel_accept (suffix i \)" and "?lgbarel_accept (suffix i \) (suffix i \ )" unfolding limit_suffix using LTLnUntil create_gba_from_nodes__ipath apply - apply (metis ipath_suffix) apply (auto simp del: suffix_nth) [] (* FIXME: "\a. suffix i \ a" is unfolded, but "suffix i \" is not! *) apply auto done moreover have "\\old (suffix i \ 0)" using \_pre \i by auto ultimately have "suffix i \ \\<^sub>n \" using LTLnUntil by simp } ultimately show ?case by auto next case (LTLnRelease \ \) { fix i assume "\ \ old (\ i) \ (\j \ old (\ j))" moreover { assume *: "\ \ old (\ i)" have "?inf_run (suffix i \)" and "?gbarel_accept (suffix i \)" and "?lgbarel_accept (suffix i \) (suffix i \ )" unfolding limit_suffix using LTLnRelease create_gba_from_nodes__ipath apply - apply (metis ipath_suffix) apply (auto simp del: suffix_nth) [] (* FIXME: "\a. suffix i \ a" is unfolded, but "suffix i \" is not! *) apply auto done with * have "suffix i \ \\<^sub>n \" using LTLnRelease by auto } moreover { assume "\j \ old (\ j)" then obtain j where "j \ old (\ j)" by auto moreover have "?inf_run (suffix j \)" and "?gbarel_accept (suffix j \)" and "?lgbarel_accept (suffix j \) (suffix j \ )" unfolding limit_suffix using LTLnRelease create_gba_from_nodes__ipath apply - apply (metis ipath_suffix) apply (auto simp del: suffix_nth) [] (* FIXME: "\a. suffix i \ a" is unfolded, but "suffix i \" is not! *) apply auto done ultimately have "suffix j \ \\<^sub>n \" using LTLnRelease by auto then have "\j \\<^sub>n \" using \j by auto } ultimately have "suffix i \ \\<^sub>n \ \ (\j \\<^sub>n \)" by auto } then show ?case using LTLnRelease L4_2c by auto qed lemma L4_8: assumes "gba.is_acc_run \" and "\i. gba.L (\ i) (\ i)" and "\ \ old (\ 0)" shows "\ \\<^sub>n \" using assms unfolding gba.is_acc_run_def gba.is_run_def using L4_8' by blast lemma expand_expand_init_propag: assumes "(\nm\incoming n'. nm < name n') \ name n' \ name (fst n_ns) \ (incoming n' \ incoming (fst n_ns) \ {} \ new n' \ old (fst n_ns) \ new (fst n_ns)) " (is "?Q n_ns") and "\nd\snd n_ns. \nm\incoming n'. nm\incoming nd \ new n' \ old nd" (is "?P (snd n_ns)") shows "expand n_ns \ SPEC (\r. name n'\fst r \ ?P (snd r))" using assms proof (rule_tac expand_rec_rule[where \="\x. ?Q x \ ?P (snd x)"], simp, intro refine_vcg, goal_cases) case prems: (1 f x n ns) then have goal_assms: "new n = {} \ ?Q (n, ns) \ ?P ns" by simp { fix nd nm assume "nd\upd_incoming n ns" and "nm\incoming n'" and "nm\incoming nd" { assume "nd\ns" with goal_assms \nm\incoming n'\ \nm\incoming nd\ have "new n' \ old nd" by auto } moreover { assume "nd\ns" with upd_incoming__elem[OF \nd\upd_incoming n ns\] obtain nd' where "nd'\ns" and nd_eq: "nd = nd'\incoming := incoming n \ incoming nd'\" and old_next_eq: "old nd' = old n \ next nd' = next n" by auto { assume "nm\incoming nd'" with goal_assms \nm\incoming n'\ \nd'\ns\ have "new n' \ old nd" unfolding nd_eq by auto } moreover { assume "nm\incoming n" with nd_eq old_next_eq goal_assms \nm\incoming n'\ have "new n' \ old nd" by auto } ultimately have "new n' \ old nd" using \nm\incoming nd\ nd_eq by auto } ultimately have "new n' \ old nd" by fast } with prems show ?case by auto next case prems: (2 f x n ns) then have goal_assms: "new n = {} \ ?Q (n, ns) \ ?P ns" and step: "\x. ?Q x \ ?P (snd x) \ f x \ SPEC (\r. name n' \ fst r \ ?P (snd r))" by simp_all then show ?case proof (rule_tac step, goal_cases) case prems': 1 have expand_new_name_less: "name n < expand_new_name (name n)" by auto moreover have "name n \ incoming n'" using goal_assms by auto ultimately show ?case using prems' by auto qed next case 3 then show ?case by auto next case prems: (4 f x n ns) then have step: "\x. ?Q x \ ?P (snd x) \ f x \ SPEC (\r. name n' \ fst r \ ?P (snd r))" by simp from prems show ?case by (rule_tac step) auto next case prems: (5 f x n ns) then have step: "\x. ?Q x \ ?P (snd x) \ f x \ SPEC (\r. name n' \ fst r \ ?P (snd r))" by simp from prems show ?case by (rule_tac step) auto next case 6 then show ?case by auto next case prems: (7 f x n ns) then have step: "\x. ?Q x \ ?P (snd x) \ f x \ SPEC (\r. name n' \ fst r \ ?P (snd r))" by simp from prems show ?case by (rule_tac step) auto next case prems: (8 f x n ns \) then have goal_assms: "\ \ new n \ \ (\q. \ = prop\<^sub>n(q) \ \ = nprop\<^sub>n(q)) \ \ \ true\<^sub>n \ \ \ false\<^sub>n \ \ (\\ \. \ = \ and\<^sub>n \ \ \ = X\<^sub>n \) \ (\\ \. \ = \ or\<^sub>n \ \ \ = \ U\<^sub>n \ \ \ = \ V\<^sub>n \)" by (cases \) auto from prems have QP: "?Q (n, ns) \ ?P ns" and step: "\x. ?Q x \ ?P (snd x) \ f x \ SPEC (\r. name n' \ fst r \ ?P (snd r))" by simp_all show ?case using goal_assms QP unfolding case_prod_unfold \x = (n, ns)\ proof (rule_tac SPEC_rule_nested2, goal_cases) case 1 then show ?case by (rule_tac step) auto next case 2 then show ?case by (rule_tac step) auto qed qed lemma expand_init_propag_on_create_graph: "create_graph \ \ SPEC(\nds. \nd\nds. expand_init\incoming nd \ \ \ old nd)" unfolding create_graph_def by (intro refine_vcg, rule_tac order_trans, rule_tac expand_expand_init_propag[where n' = "\ name = expand_new_name expand_init, incoming = {expand_init}, new = {\}, old = {}, next = {} \"]) (auto simp add:expand_new_name_expand_init) lemma L4_6: assumes "q\gba.V0" shows "\\old q" using assms inres_SPEC[OF res expand_init_propag_on_create_graph] unfolding create_gba_from_nodes_def by auto lemma L4_9: assumes acc: "gba.accept \" shows "\ \\<^sub>n \" proof - from acc obtain \ where accept: "gba.is_acc_run \ \ (\i. gba.L (\ i) (\ i))" unfolding gba.accept_def by auto then have "\ 0\gba.V0" unfolding gba.is_acc_run_def gba.is_run_def by simp with L4_6 have "\\old (\ 0)" by auto with L4_8 accept show ?thesis by auto qed lemma L4_10: assumes "\ \\<^sub>n \" shows "gba.accept \" proof - def \ \ "auto_run \ qs" let ?G = "create_graph \" have \_prop_0: "(\ 0)\qs \ expand_init\incoming(\ 0) \ auto_run_j 0 \ (\ 0)" (is "?sbthm") using inres_SPEC[OF res L4_7[OF \\ \\<^sub>n \\]] unfolding \_def auto_run.simps by (rule_tac someI_ex, simp) blast have \_valid: "\j. \ j \ qs \ auto_run_j j \ (\ j)" (is "\j. ?\_valid j") proof fix j show "?\_valid j" proof(induct j) from \_prop_0 show "?\_valid 0" by fast next fix k assume goal_assms: "\ k \ qs \ auto_run_j k \ (\ k)" with inres_SPEC[OF res L4_5, of "suffix k \"] have sbthm: "\q'. q'\qs \ name (\ k)\incoming q' \ auto_run_j (Suc k) \ q'" (is "\q'. ?sbthm q'") by auto have "?sbthm (\ (Suc k))" using someI_ex[OF sbthm] unfolding \_def auto_run.simps by blast then show "?\_valid (Suc k)" by fast qed qed have \_prop_Suc: "\k. \ k\ qs \ \ (Suc k)\qs \ name (\ k)\incoming (\ (Suc k)) \ auto_run_j (Suc k) \ (\ (Suc k))" proof - fix k from \_valid have "\ k \ qs" and "auto_run_j k \ (\ k)" by blast+ with inres_SPEC[OF res L4_5, of "suffix k \"] have sbthm: "\q'. q'\qs \ name (\ k)\incoming q' \ auto_run_j (Suc k) \ q'" (is "\q'. ?sbthm q'") by auto show "\ k\ qs \ ?sbthm (\ (Suc k))" using \\ k\qs\ someI_ex[OF sbthm] unfolding \_def auto_run.simps by blast qed have \_vnaccpt: "\k \ \. \ U\<^sub>n \ \ old (\ k) \ \ (\i. {\, \ U\<^sub>n \} \ old (\ (k+i)) \ \ \ old (\ (k+i)))" proof clarify fix k \ \ assume U_in: "\ U\<^sub>n \ \ old (\ k)" and cntr_prm: "\i. {\, \ U\<^sub>n \} \ old (\ (k+i)) \ \ \ old (\ (k+i))" have "suffix k \ \\<^sub>n \ U\<^sub>n \" using U_in \_valid by force then obtain i where "suffix (k+i) \ \\<^sub>n \" and "\j \\<^sub>n \" by auto moreover have "\ U\<^sub>n \ \ old (\ (k+i)) \ \ \ old (\ (k+i))" using cntr_prm by auto ultimately show False using \_valid by force qed have \_exec: "gba.is_run \" using \_prop_0 \_prop_Suc \_valid unfolding gba.is_run_def ipath_def unfolding create_gba_from_nodes_def by auto moreover have \_vaccpt: "\k \ \. \ U\<^sub>n \ \ old (\ k) \ (\j. (\i, \ U\<^sub>n \} \ old (\ (k+i))) \ \ \ old (\ (k+j)))" proof(clarify) fix k \ \ assume U_in: "\ U\<^sub>n \ \ old (\ k)" then have "\ (\i. {\, \ U\<^sub>n \} \ old (suffix k \ i) \ \ \ old (suffix k \ i))" using \_vnaccpt[THEN allE, of k] by auto moreover have "suffix k \ 0 \ qs" using \_valid by auto ultimately show "\j. (\i, \ U\<^sub>n \} \ old (\ (k+i))) \ \ \ old (\ (k+j))" apply - apply (rule make_pos_rule'[OF L4_2a]) apply (fold suffix_def) apply (rule ipath_suffix) using \_exec[unfolded gba.is_run_def] apply simp using U_in apply simp apply simp done qed have "gba.is_acc \" unfolding gba.is_acc_def proof fix S assume "S\gba.F" then obtain \ \ where S_eq: "S = {q \ qs. \ U\<^sub>n \ \ old q \ \ \ old q}" and "\ U\<^sub>n \ \ subfrmlsn \" by (auto simp add: create_gba_from_nodes_def) have range_subset: "range \ \ qs" proof fix q assume "q\range \" with full_SetCompr_eq[of \] obtain k where "q = \ k" by auto then show "q \ qs" using \_valid by auto qed with limit_nonempty[of \] limit_in_range[of \] finite_subset[OF range_subset] inres_SPEC[OF res create_graph_finite] obtain q where q_in_limit: "q \ limit \" and q_is_node: "q\qs" by auto show "\\<^sub>\i. \ i \ S" proof (cases "\ U\<^sub>n \ \ old q") case False with S_eq q_in_limit q_is_node show ?thesis by (auto simp: limit_iff_frequent intro: INFM_mono) next case True obtain k where q_eq: "q = \ k" using q_in_limit unfolding limit_iff_frequent by (metis (lifting, full_types) INFM_nat_le) have "\\<^sub>\ k. \ \ old (\ k)" unfolding INFM_nat proof (rule ccontr) assume "\ (\m. \n>m. \ \ old (\ n))" then obtain m where "\n>m. \ \ old (\ n)" by auto moreover from q_eq q_in_limit limit_iff_frequent[of q \] INFM_nat[of "\n. \ n = q"] obtain n where "mn_eq: "\ n = \ k" by auto moreover obtain j where "\ \ old (\ (n+j))" using \_vaccpt \\ U\<^sub>n \ \ old q\ unfolding q_eq by (fold \n_eq) force ultimately show False by auto qed then have "\\<^sub>\ k. \ k \ qs \ \ \ old (\ k)" using \_valid by (auto intro: INF_mono) then show "\\<^sub>\ k. \ k \ S" unfolding S_eq by (rule INFM_mono) simp qed qed moreover have "gba.L (\ i) (\ i)" for i proof - from \_valid have [simp]: "\ i \ qs" by auto have "\\\old (\ i). suffix i \ \\<^sub>n \" using \_valid by auto then show ?thesis unfolding create_gba_from_nodes_def by auto qed ultimately show ?thesis unfolding gba.accept_def gba.is_acc_run_def by blast qed end end lemma create_graph__name_ident: "create_graph \ \ SPEC (\nds. \q\nds. \!q'\nds. name q = name q')" unfolding create_graph_def by (intro refine_vcg, rule_tac order_trans, rule_tac expand_name_propag__name_ident) (auto simp add:expand_new_name_expand_init) corollary create_graph__name_inj: "create_graph \ \ SPEC (\nds. inj_on name nds)" by (rule order_trans[OF create_graph__name_ident]) (auto intro: inj_onI) definition "create_gba \ \ do { nds \ create_graph\<^sub>T \; RETURN (create_gba_from_nodes \ nds) }" lemma create_graph_precond: "create_graph \ \ SPEC (create_gba_from_nodes_precond \)" apply (clarsimp simp: pw_le_iff create_graph_nofail) apply unfold_locales apply simp done lemma create_gba__invar: "create_gba \ \ SPEC gba" unfolding create_gba_def create_graph_eq_create_graph\<^sub>T[symmetric] apply (refine_rcg refine_vcg order_trans[OF create_graph_precond]) by (rule create_gba_from_nodes_precond.create_gba_from_nodes__invar) lemma create_gba_acc: shows "create_gba \ \ SPEC(\\. \\. gba.accept \ \ \ \ \\<^sub>n \)" unfolding create_gba_def create_graph_eq_create_graph\<^sub>T[symmetric] apply (refine_rcg refine_vcg order_trans[OF create_graph_precond]) using create_gba_from_nodes_precond.L4_9 using create_gba_from_nodes_precond.L4_10 by blast lemma create_gba__name_inj: shows "create_gba \ \ SPEC(\\. (inj_on name (g_V \)))" unfolding create_gba_def create_graph_eq_create_graph\<^sub>T[symmetric] apply (refine_rcg refine_vcg order_trans[OF create_graph__name_inj]) apply (auto simp: create_gba_from_nodes_def) done lemma create_gba__fin: shows "create_gba \ \ SPEC(\\. (finite (g_V \)))" unfolding create_gba_def create_graph_eq_create_graph\<^sub>T[symmetric] apply (refine_rcg refine_vcg order_trans[OF create_graph_finite]) apply (auto simp: create_gba_from_nodes_def) done lemma create_graph_old_finite: "create_graph \ \ SPEC (\nds. \nd\nds. finite (old nd))" proof - show ?thesis unfolding create_graph_def create_graph_eq_create_graph\<^sub>T[symmetric] unfolding expand_def apply (intro refine_vcg) apply (rule_tac order_trans) apply (rule_tac REC_le_RECT) apply (fold expand\<^sub>T_def) apply (rule_tac order_trans[OF expand_term_prop]) apply auto[1] apply (rule_tac SPEC_rule) apply auto by (metis infinite_super subfrmlsn_finite) qed lemma create_gba__old_fin: shows "create_gba \ \ SPEC(\\. \nd\g_V \. finite (old nd))" unfolding create_gba_def create_graph_eq_create_graph\<^sub>T[symmetric] apply (refine_rcg refine_vcg order_trans[OF create_graph_old_finite]) apply (simp add: create_gba_from_nodes_def) done lemma create_gba__incoming_exists: shows "create_gba \ \ SPEC(\\. \nd\g_V \. incoming nd \ insert expand_init (name ` (g_V \)))" unfolding create_gba_def create_graph_eq_create_graph\<^sub>T[symmetric] apply (refine_rcg refine_vcg order_trans[OF create_graph__incoming_name_exist]) apply (auto simp add: create_gba_from_nodes_def) done lemma create_gba__no_init: shows "create_gba \ \ SPEC(\\. expand_init \ name ` (g_V \))" unfolding create_gba_def create_graph_eq_create_graph\<^sub>T[symmetric] apply (refine_rcg refine_vcg order_trans[OF create_graph__incoming_name_exist]) apply (auto simp add: create_gba_from_nodes_def) done definition "nds_invars nds \ inj_on name nds \ finite nds \ expand_init \ name`nds \ (\nd\nds. finite (old nd) \ incoming nd \ insert expand_init (name ` nds))" lemma create_gba_nds_invars: "create_gba \ \ SPEC (\\. nds_invars (g_V \))" using create_gba__name_inj[of \] create_gba__fin[of \] create_gba__old_fin[of \] create_gba__incoming_exists[of \] create_gba__no_init[of \] unfolding nds_invars_def by (simp add: pw_le_iff) theorem T4_1: "create_gba \ \ SPEC( \\. gba \ \ finite (g_V \) \ (\\. gba.accept \ \ \ \ \\<^sub>n \) \ (nds_invars (g_V \)))" using create_gba__invar create_gba__fin create_gba_acc create_gba_nds_invars apply (simp add: pw_le_iff) apply blast done definition "create_name_gba \ \ do { G \ create_gba \; ASSERT (nds_invars (g_V G)); RETURN (gba_rename name G) }" theorem create_name_gba_correct: "create_name_gba \ \ SPEC(\\. gba \ \ finite (g_V \) \ (\\. gba.accept \ \ \ \ \\<^sub>n \))" unfolding create_name_gba_def apply (refine_rcg refine_vcg order_trans[OF T4_1]) apply (simp_all add: nds_invars_def gba_rename_correct) done definition create_name_igba :: "'a::linorder ltln \ _" where "create_name_igba \ \ do { A \ create_name_gba \; A' \ gba_to_idx A; stat_set_data_nres (card (g_V A)) (card (g_V0 A')) (igbg_num_acc A'); RETURN A' }" lemma create_name_igba_correct: "create_name_igba \ \ SPEC (\G. igba G \ finite (g_V G) \ (\\. igba.accept G \ \ \ \\<^sub>n \))" unfolding create_name_igba_def apply (refine_rcg order_trans[OF create_name_gba_correct] order_trans[OF gba.gba_to_idx_ext_correct] refine_vcg) apply clarsimp_all proof - fix G :: "(nat, 'a set) gba_rec" fix A :: "nat set" assume 1: "gba G" assume 2: "finite (g_V G)" "A \ gbg_F G" interpret gba G using 1 . show "finite A" using finite_V_Fe 2 . qed (* For presentation in paper*) context notes [refine_vcg] = order_trans[OF create_name_gba_correct] begin lemma "create_name_igba \ \ SPEC (\G. igba G \ (\\. igba.accept G \ \ \ \\<^sub>n \))" unfolding create_name_igba_def proof (refine_rcg refine_vcg, clarsimp_all) fix G :: "(nat, 'a set) gba_rec" assume "gba G" then interpret gba G . note [refine_vcg] = order_trans[OF gba_to_idx_ext_correct] assume "\\. gba.accept G \ = \ \\<^sub>n \" "finite (g_V G)" then show "gba_to_idx G \ SPEC (\G'. igba G' \ (\\. igba.accept G' \ = \ \\<^sub>n \))" by (refine_rcg refine_vcg) (auto intro: finite_V_Fe) qed end end diff --git a/thys/LTL_to_GBA/LTL_to_GBA_impl.thy b/thys/LTL_to_GBA/LTL_to_GBA_impl.thy --- a/thys/LTL_to_GBA/LTL_to_GBA_impl.thy +++ b/thys/LTL_to_GBA/LTL_to_GBA_impl.thy @@ -1,1161 +1,1163 @@ section \Refinement to Efficient Code\ (* Author: Peter Lammich Inspired by previous version of Alexander Schimpf. *) theory LTL_to_GBA_impl imports LTL_to_GBA "../Deriving/Comparator_Generator/Compare_Instances" "../CAVA_Automata/Automata_Impl" "../CAVA_Automata/CAVA_Base/CAVA_Code_Target" begin subsection \Parametricity Setup Boilerplate\ subsubsection \LTL Formulas\ derive linorder ltln inductive_set ltln_rel for R where "(LTLnTrue,LTLnTrue) \ ltln_rel R" | "(LTLnFalse,LTLnFalse) \ ltln_rel R" | "(a,a')\R \ (LTLnProp a,LTLnProp a') \ ltln_rel R" | "(a,a')\R \ (LTLnNProp a,LTLnNProp a') \ ltln_rel R" | "\(a,a')\ltln_rel R; (b,b')\ltln_rel R\ \ (LTLnAnd a b,LTLnAnd a' b') \ ltln_rel R" | "\(a,a')\ltln_rel R; (b,b')\ltln_rel R\ \ (LTLnOr a b,LTLnOr a' b') \ ltln_rel R" | "\(a,a')\ltln_rel R\ \ (LTLnNext a,LTLnNext a') \ ltln_rel R" | "\(a,a')\ltln_rel R; (b,b')\ltln_rel R\ \ (LTLnUntil a b,LTLnUntil a' b') \ ltln_rel R" | "\(a,a')\ltln_rel R; (b,b')\ltln_rel R\ \ (LTLnRelease a b,LTLnRelease a' b') \ ltln_rel R" lemmas ltln_rel_induct[induct set] = ltln_rel.induct[unfolded relAPP_def[of ltln_rel, symmetric]] lemmas ltln_rel_cases[cases set] = ltln_rel.cases[unfolded relAPP_def[of ltln_rel, symmetric]] lemmas ltln_rel_intros = ltln_rel.intros[unfolded relAPP_def[of ltln_rel, symmetric]] inductive_simps ltln_rel_left_simps[unfolded relAPP_def[of ltln_rel, symmetric]]: "(LTLnTrue,z) \ ltln_rel R" "(LTLnFalse,z) \ ltln_rel R" "(LTLnProp p, z) \ ltln_rel R" "(LTLnNProp p, z) \ ltln_rel R" "(LTLnAnd a b, z) \ ltln_rel R" "(LTLnOr a b, z) \ ltln_rel R" "(LTLnNext a, z) \ ltln_rel R" "(LTLnUntil a b, z) \ ltln_rel R" "(LTLnRelease a b, z) \ ltln_rel R" lemma ltln_rel_sv[relator_props]: assumes SV: "single_valued R" shows "single_valued (\R\ltln_rel)" proof (intro single_valuedI allI impI) fix x y z assume "(x, y) \ \R\ltln_rel" "(x, z) \ \R\ltln_rel" then show "y=z" apply (induction arbitrary: z) apply (simp (no_asm_use) only: ltln_rel_left_simps | blast intro: single_valuedD[OF SV])+ done qed lemma ltln_rel_id[relator_props]: "\ R = Id \ \ \R\ltln_rel = Id" proof (intro equalityI subsetI, clarsimp_all) fix a b assume "(a,b)\\Id\ltln_rel" then show "a=b" by induction auto next fix a show "(a,a)\\Id\ltln_rel" by (induction a) (auto intro: ltln_rel_intros) qed lemma ltln_rel_id_simp[simp]: "\Id\ltln_rel = Id" by (rule ltln_rel_id) simp consts i_ltln :: "interface \ interface" lemmas [autoref_rel_intf] = REL_INTFI[of ltln_rel i_ltln] thm ltln_rel_intros[no_vars] lemma ltln_con_param[param, autoref_rules]: "(LTLnTrue, LTLnTrue) \ \R\ltln_rel" "(LTLnFalse, LTLnFalse) \ \R\ltln_rel" "(LTLnProp, LTLnProp) \ R \ \R\ltln_rel" "(LTLnNProp, LTLnNProp) \ R \ \R\ltln_rel" "(LTLnAnd, LTLnAnd) \ \R\ltln_rel \ \R\ltln_rel \ \R\ltln_rel" "(LTLnOr, LTLnOr) \ \R\ltln_rel \ \R\ltln_rel \ \R\ltln_rel" "(LTLnNext, LTLnNext) \ \R\ltln_rel \ \R\ltln_rel" "(LTLnUntil, LTLnUntil) \ \R\ltln_rel \ \R\ltln_rel \ \R\ltln_rel" "(LTLnRelease, LTLnRelease) \ \R\ltln_rel \ \R\ltln_rel \ \R\ltln_rel" by (auto intro: ltln_rel_intros) lemma case_ltln_param[param, autoref_rules]: "(case_ltln,case_ltln) \ Rv \ Rv \ (R \ Rv) \ (R \ Rv) \ (\R\ltln_rel \ \R\ltln_rel \ Rv) \ (\R\ltln_rel \ \R\ltln_rel \ Rv) \ (\R\ltln_rel \ Rv) \ (\R\ltln_rel \ \R\ltln_rel \ Rv) \ (\R\ltln_rel \ \R\ltln_rel \ Rv) \ \R\ltln_rel \ Rv" apply (clarsimp) apply (case_tac ai, simp_all add: ltln_rel_left_simps) apply (clarsimp_all) apply parametricity+ done lemma rec_ltln_param[param, autoref_rules]: "(rec_ltln,rec_ltln) \ Rv \ Rv \ (R \ Rv) \ (R \ Rv) \ (\R\ltln_rel \ \R\ltln_rel \ Rv \ Rv \ Rv) \ (\R\ltln_rel \ \R\ltln_rel \ Rv \ Rv \ Rv) \ (\R\ltln_rel \ Rv \ Rv) \ (\R\ltln_rel \ \R\ltln_rel \ Rv \ Rv \ Rv) \ (\R\ltln_rel \ \R\ltln_rel \ Rv \ Rv \ Rv) \ \R\ltln_rel \ Rv" proof (clarsimp, goal_cases) case prems: 1 from prems(10) show ?case apply (induction) using prems(1-9) apply simp_all apply parametricity+ done qed lemma case_ltln_mono[refine_mono]: assumes "\ = LTLnTrue \ a\a'" assumes "\ = LTLnFalse \ b\b'" assumes "\p. \ = LTLnProp p \ c p\c' p" assumes "\p. \ = LTLnNProp p \ d p\d' p" assumes "\\ \. \ = LTLnAnd \ \ \ e \ \\e' \ \" assumes "\\ \. \ = LTLnOr \ \ \ f \ \\f' \ \" assumes "\\. \ = LTLnNext \ \ g \\g' \" assumes "\\ \. \ = LTLnUntil \ \ \ h \ \\h' \ \" assumes "\\ \. \ = LTLnRelease \ \ \ i \ \\i' \ \" shows "case_ltln a b c d e f g h i \ \ case_ltln a' b' c' d' e' f' g' h' i' \" using assms apply (cases \) apply simp_all done primrec ltln_eq where "ltln_eq eq LTLnTrue f \ (case f of LTLnTrue \ True | _ \ False)" | "ltln_eq eq LTLnFalse f \ (case f of LTLnFalse \ True | _ \ False)" | "ltln_eq eq (LTLnProp p) f \ (case f of LTLnProp p' \ eq p p' | _ \ False)" | "ltln_eq eq (LTLnNProp p) f \ (case f of LTLnNProp p' \ eq p p' | _ \ False)" | "ltln_eq eq (LTLnAnd \ \) f \ (case f of LTLnAnd \' \' \ ltln_eq eq \ \' \ ltln_eq eq \ \' | _ \ False)" | "ltln_eq eq (LTLnOr \ \) f \ (case f of LTLnOr \' \' \ ltln_eq eq \ \' \ ltln_eq eq \ \' | _ \ False)" | "ltln_eq eq (LTLnNext \) f \ (case f of LTLnNext \' \ ltln_eq eq \ \' | _ \ False)" | "ltln_eq eq (LTLnUntil \ \) f \ (case f of LTLnUntil \' \' \ ltln_eq eq \ \' \ ltln_eq eq \ \' | _ \ False)" | "ltln_eq eq (LTLnRelease \ \) f \ (case f of LTLnRelease \' \' \ ltln_eq eq \ \' \ ltln_eq eq \ \' | _ \ False)" lemma ltln_eq_autoref[autoref_rules]: assumes EQP: "(eq,op=) \ R \ R \ bool_rel" shows "(ltln_eq eq, op=) \ \R\ltln_rel \ \R\ltln_rel \ bool_rel" proof (intro fun_relI) fix \' \ \' \ assume "(\',\)\\R\ltln_rel" and "(\',\)\\R\ltln_rel" then show "(ltln_eq eq \' \', \=\)\bool_rel" apply (induction arbitrary: \' \) apply (erule ltln_rel_cases, simp_all) [] apply (erule ltln_rel_cases, simp_all) [] apply (erule ltln_rel_cases, simp_all add: EQP[THEN fun_relD, THEN fun_relD, THEN IdD]) [] apply (erule ltln_rel_cases, simp_all add: EQP[THEN fun_relD, THEN fun_relD, THEN IdD]) [] apply (rotate_tac -1) apply (erule ltln_rel_cases, simp_all) [] apply (rotate_tac -1) apply (erule ltln_rel_cases, simp_all) [] apply (rotate_tac -1) apply (erule ltln_rel_cases, simp_all) [] apply (rotate_tac -1) apply (erule ltln_rel_cases, simp_all) [] apply (rotate_tac -1) apply (erule ltln_rel_cases, simp_all) [] done qed lemma ltln_dflt_cmp[autoref_rules_raw]: assumes "PREFER_id R" shows "(dflt_cmp op \ op <, dflt_cmp op \ op <) \ \R\ltln_rel \ \R\ltln_rel \ comp_res_rel" using assms by simp type_synonym node_name_impl = node_name abbreviation (input) "node_name_rel \ Id :: (node_name_impl\node_name) set" lemma case_ltln_gtransfer: assumes "\ ai \ a" "\ bi \ b" "\a. \ (ci a) \ c a" "\a. \ (di a) \ d a" "\ltln1 ltln2. \ (ei ltln1 ltln2) \ e ltln1 ltln2" "\ltln1 ltln2. \ (fi ltln1 ltln2) \ f ltln1 ltln2" "\ltln. \ (gi ltln) \ g ltln" "\ltln1 ltln2. \ (hi ltln1 ltln2) \ h ltln1 ltln2" - "\ltln1 ltln2. \ (ii ltln1 ltln2) \ i ltln1 ltln2" - shows "\ (case_ltln ai bi ci di ei fi gi hi ii \) + "\ltln1 ltln2. \ (iiv ltln1 ltln2) \ i ltln1 ltln2" + shows "\ (case_ltln ai bi ci di ei fi gi hi iiv \) \ (case_ltln a b c d e f g h i \)" apply (cases \) apply (auto intro: assms) done lemmas [refine_transfer] = case_ltln_gtransfer[where \=nres_of] case_ltln_gtransfer[where \=RETURN] lemma [refine_transfer]: assumes "ai \ dSUCCEED" "bi \ dSUCCEED" "\a. ci a \ dSUCCEED" "\a. di a \ dSUCCEED" "\ltln1 ltln2. ei ltln1 ltln2 \ dSUCCEED" "\ltln1 ltln2. fi ltln1 ltln2 \ dSUCCEED" "\ltln. gi ltln \ dSUCCEED" "\ltln1 ltln2. hi ltln1 ltln2 \ dSUCCEED" - "\ltln1 ltln2. ii ltln1 ltln2 \ dSUCCEED" - shows "case_ltln ai bi ci di ei fi gi hi ii \ \ dSUCCEED" + "\ltln1 ltln2. iiv ltln1 ltln2 \ dSUCCEED" + shows "case_ltln ai bi ci di ei fi gi hi iiv \ \ dSUCCEED" apply (cases \) apply (simp_all add: assms) done +hide_const (open) Re + subsubsection \Nodes\ record 'a node_impl = name_impl :: node_name_impl incoming_impl :: "(node_name_impl,unit) RBT_Impl.rbt" new_impl :: "'a frml list" old_impl :: "'a frml list" next_impl :: "'a frml list" + + definition node_rel where node_rel_def_internal: "node_rel Re R \ {( \ name_impl = namei, incoming_impl = inci, new_impl = newi, old_impl = oldi, next_impl = nexti, \ = morei \, \ name = name, incoming = inc, new=new, old=old, next = next, \ = more \ ) | namei name inci inc newi new oldi old nexti next morei more. (namei,name)\node_name_rel \ (inci,inc)\\node_name_rel\dflt_rs_rel \ (newi,new)\\\R\ltln_rel\lss.rel \ (oldi,old)\\\R\ltln_rel\lss.rel \ (nexti,next)\\\R\ltln_rel\lss.rel \ (morei,more)\Re }" lemma node_rel_def: "\Re,R\node_rel = {( \ name_impl = namei, incoming_impl = inci, new_impl = newi, old_impl = oldi, next_impl = nexti, \ = morei \, \ name = name, incoming = inc, new=new, old=old, next = next, \ = more \ ) | namei name inci inc newi new oldi old nexti next morei more. (namei,name)\node_name_rel \ (inci,inc)\\node_name_rel\dflt_rs_rel \ (newi,new)\\\R\ltln_rel\lss.rel \ (oldi,old)\\\R\ltln_rel\lss.rel \ (nexti,next)\\\R\ltln_rel\lss.rel \ (morei,more)\Re }" by (simp add: node_rel_def_internal relAPP_def) lemma node_rel_sv[relator_props]: "single_valued Re \ single_valued R \ single_valued (\Re,R\node_rel)" apply (rule single_valuedI) apply (simp add: node_rel_def) apply (auto dest: single_valuedD lss.rel_sv[OF ltln_rel_sv] map2set_rel_sv[OF ahm_rel_sv] dest: single_valuedD[ OF map2set_rel_sv[OF rbt_map_rel_sv[OF single_valued_Id single_valued_Id]] ]) done consts i_node :: "interface \ interface \ interface" lemmas [autoref_rel_intf] = REL_INTFI[of node_rel i_node] lemma [autoref_rules]: "(node_impl_ext, node_ext) \ node_name_rel \ \node_name_rel\dflt_rs_rel \ \\R\ltln_rel\lss.rel \ \\R\ltln_rel\lss.rel \ \\R\ltln_rel\lss.rel \ Re \ \Re,R\node_rel" unfolding node_rel_def by auto term node.incoming_update lemma [autoref_rules]: "(node_impl.name_impl_update,node.name_update) \ (node_name_rel \ node_name_rel) \ \Re,R\node_rel \ \Re,R\node_rel" "(node_impl.incoming_impl_update,node.incoming_update) \ (\node_name_rel\dflt_rs_rel \ \node_name_rel\dflt_rs_rel) \ \Re,R\node_rel \ \Re,R\node_rel" "(node_impl.new_impl_update,node.new_update) \ (\\R\ltln_rel\lss.rel \ \\R\ltln_rel\lss.rel) \ \Re,R\node_rel \ \Re,R\node_rel" "(node_impl.old_impl_update,node.old_update) \ (\\R\ltln_rel\lss.rel \ \\R\ltln_rel\lss.rel) \ \Re,R\node_rel \ \Re,R\node_rel" "(node_impl.next_impl_update,node.next_update) \ (\\R\ltln_rel\lss.rel \ \\R\ltln_rel\lss.rel) \ \Re,R\node_rel \ \Re,R\node_rel" "(node_impl.more_update,node.more_update) \ (Re \ Re) \ \Re,R\node_rel \ \Re,R\node_rel" unfolding node_rel_def by (auto dest: fun_relD) term name lemma [autoref_rules]: "(node_impl.name_impl,node.name)\\Re,R\node_rel \ node_name_rel" "(node_impl.incoming_impl,node.incoming) \ \Re,R\node_rel \ \node_name_rel\dflt_rs_rel" "(node_impl.new_impl,node.new)\\Re,R\node_rel \ \\R\ltln_rel\lss.rel" "(node_impl.old_impl,node.old)\\Re,R\node_rel \ \\R\ltln_rel\lss.rel" "(node_impl.next_impl,node.next)\\Re,R\node_rel \ \\R\ltln_rel\lss.rel" "(node_impl.more, node.more)\\Re,R\node_rel \ Re" unfolding node_rel_def by auto subsection \Massaging the Abstract Algorithm\ text \ In a first step, we do some refinement steps on the abstract data structures, with the goal to make the algorithm more efficient. \ subsubsection \Creation of the Nodes\ text \ In the expand-algorithm, we replace nested conditionals by case-distinctions, and slightly stratify the code. \ +hide_const (open) exp + abbreviation (input) "expand2 exp n ns \ n1 nx1 n2 \ do { (nm, nds) \ exp ( n\ new := insert n1 (new n), old := insert \ (old n), next := nx1 \ next n \, ns); exp (n\ name := nm, new := n2 \ new n, old := {\} \ old n \, nds) }" -context begin interpretation LTL_Syntax . definition "expand_aimpl \ REC\<^sub>T (\expand (n,ns). if new n = {} then ( if (\n'\ns. old n' = old n \ next n' = next n) then RETURN (name n, upd_incoming n ns) else do { ASSERT (n \ ns); ASSERT (name n \ name`ns); expand (\ name=expand_new_name (name n), incoming={name n}, new=next n, old={}, next={} \, insert n ns) } ) else do { \ \ SPEC (\x. x\(new n)); let n = n\ new := new n - {\} \; case \ of prop\<^sub>n(q) \ if nprop\<^sub>n(q)\old n then RETURN (name n, ns) else expand (n\ old := {\} \ old n \, ns) | nprop\<^sub>n(q) \ if prop\<^sub>n(q)\old n then RETURN (name n, ns) else expand (n\ old := {\} \ old n \, ns) | true\<^sub>n \ expand (n\ old := {\} \ old n \, ns) | false\<^sub>n \ RETURN (name n, ns) | \ and\<^sub>n \ \ expand (n\ new := insert \ (insert \ (new n)), old := {\} \ old n, next := next n \, ns) | X\<^sub>n \ \ expand (n\ new := new n, old := {\} \ old n, next := insert \ (next n) \, ns) | \ or\<^sub>n \ \ expand2 expand n ns \ \ {} {\} | \ U\<^sub>n \ \ expand2 expand n ns \ \ {\} {\} | \ V\<^sub>n \ \ expand2 expand n ns \ \ {\} {\,\} (*| _ \ do { (nm, nds) \ expand ( n\ new := new1 \ \ new n, old := {\} \ old n, next := next1 \ \ next n \, ns); expand (n\ name := nm, new := new2 \ \ new n, old := {\} \ old n \, nds) }*) } )" -end lemma expand_aimpl_refine: fixes n_ns :: "('a node \ _)" defines "R \ Id \ {(_,(n,ns)). \n'\ns. n > name n'}" defines "R' \ Id \ {(_,(n,ns)). \n'\ns. name n > name n'}" assumes [refine]: "(n_ns',n_ns)\R'" shows "expand_aimpl n_ns' \ \R (expand\<^sub>T n_ns)" using [[goals_limit = 1]] proof - have [relator_props]: "single_valued R" by (auto simp add: R_def intro: single_valuedI) have [relator_props]: "single_valued R'" by (auto simp add: R'_def intro: single_valuedI) { fix n :: "'a node" and ns and n' ns' assume "((n', ns'), (n, ns)) \ R'" then have "(RETURN (name n', ns') \ \ R (RETURN (name n, ns)))" by (auto simp: R_def R'_def pw_le_iff refine_pw_simps) } note aux = this show ?thesis unfolding expand_aimpl_def expand\<^sub>T_def apply refine_rcg apply (simp add: R_def R'_def) apply (simp add: R_def R'_def) apply (auto simp add: R_def R'_def upd_incoming_def) [] apply (auto simp add: R_def R'_def upd_incoming_def) [] apply (auto simp add: R_def R'_def upd_incoming_def) [] apply rprems apply (auto simp: R'_def expand_new_name_def) [] apply (simp add: R'_def) apply (auto split: ltln.split) [] apply (fastforce simp: R_def R'_def) [] apply (fastforce simp: R_def R'_def) [] apply (auto simp: R_def R'_def) [] apply (auto simp: R_def R'_def) [] apply (auto simp: R_def R'_def) [] apply (auto simp: R_def R'_def) [] apply (auto simp: R_def R'_def) [] apply (auto simp: R_def R'_def) [] apply (auto simp: R_def R'_def) [] apply (auto simp: R_def R'_def) [] apply (auto simp: R_def R'_def) [] apply (refine_rcg, rprems, (fastforce simp: R_def R'_def)+) [] apply (fastforce simp: R'_def) [] apply (refine_rcg, rprems, (fastforce simp: R_def R'_def)+) [] apply (refine_rcg, rprems, (fastforce simp: R_def R'_def)+) [] done qed thm create_graph_def definition "create_graph_aimpl \ = do { (_, nds) \ expand_aimpl (\name = expand_new_name expand_init, incoming = {expand_init}, new = {\}, old = {}, next = {}\, {}); RETURN nds }" lemma create_graph_aimpl_refine: "create_graph_aimpl \ \ \Id (create_graph\<^sub>T \)" unfolding create_graph_aimpl_def create_graph\<^sub>T_def apply (refine_rcg expand_aimpl_refine) apply auto done subsubsection \Creation of GBA from Nodes\ text \ We summarize creation of the GBA and renaming of the nodes into one step \ lemma create_name_gba_alt: "create_name_gba \ = do { nds \ create_graph\<^sub>T \; ASSERT (nds_invars nds); RETURN (gba_rename_ext (\_. ()) name (create_gba_from_nodes \ nds)) }" proof - have [simp]: "\nds. g_V (create_gba_from_nodes \ nds) = nds" by (auto simp: create_gba_from_nodes_def) show ?thesis unfolding create_name_gba_def create_gba_def by simp qed text \In the following, we implement the componenents of the renamed GBA separately. \ text \\paragraph{Successor Function}\ definition "build_succ nds = FOREACH nds (\q' s. FOREACH (incoming q') (\qn s. if qn=expand_init then RETURN s else RETURN (s(qn \ insert (name q') (the_default {} (s qn)))) ) s ) Map.empty" lemma build_succ_aux1: assumes [simp]: "finite nds" assumes [simp]: "\q. q\nds \ finite (incoming q)" shows "build_succ nds \ SPEC (\r. r = (\qn. dflt_None_set {qn'. \q'. q'\nds \ qn' = name q' \ qn\incoming q' \ qn\expand_init }))" unfolding build_succ_def apply (refine_rcg refine_vcg FOREACH_rule[where I="\it s. s = (\qn. dflt_None_set {qn'. \q'. q'\nds-it \ qn' = name q' \ qn\incoming q' \ qn\expand_init })"]) apply (simp_all add: dflt_None_set_def) [2] apply (rename_tac q' it s) apply (rule_tac I="\it2 s. s = (\qn. dflt_None_set ( {qn'. \q'. q'\nds-it \ qn' = name q' \ qn\incoming q' \ qn\expand_init } \ {qn' . qn'=name q' \ qn\incoming q' - it2 \ qn\expand_init} ))" in FOREACH_rule) apply auto [] apply (simp_all add: dflt_None_set_def) apply (refine_rcg refine_vcg) apply (auto simp: dflt_None_set_def intro!: ext) [] apply (rule ext, (auto) [])+ done lemma build_succ_aux2: assumes NINIT: "expand_init \ name`nds" assumes CL: "\nd. nd\nds \ incoming nd \ insert expand_init (name`nds)" shows "(\qn. dflt_None_set {qn'. \q'. q'\nds \ qn' = name q' \ qn\incoming q' \ qn\expand_init }) = (\qn. dflt_None_set (succ_of_E (rename_E name {(q, q'). q \ nds \ q' \ nds \ name q \ incoming q'}) qn))" (is "(\qn. dflt_None_set (?L qn)) = (\qn. dflt_None_set (?R qn))") apply (intro ext) apply (fo_rule arg_cong) proof (intro ext equalityI subsetI) fix qn x assume "x\?R qn" then show "x\?L qn" using NINIT by (force simp: succ_of_E_def) next fix qn x assume XL: "x\?L qn" show "x\?R qn" proof (cases "qn = expand_init") case False from XL obtain q' where A: "q'\nds" "qn\incoming q'" and [simp]: "x=name q'" by auto from False obtain q where B: "q\nds" and [simp]: "qn = name q" using CL A by auto from A B show "x\?R qn" by (auto simp: succ_of_E_def image_def) next case [simp]: True from XL show "x\?R qn" by simp qed qed lemma build_succ_correct: assumes NINIT: "expand_init \ name`nds" assumes FIN: "finite nds" assumes CL: "\nd. nd\nds \ incoming nd \ insert expand_init (name`nds)" shows "build_succ nds \ SPEC (\r. E_of_succ (\qn. the_default {} (r qn)) = rename_E (\u. name u) {(q, q'). q \ nds \ q' \ nds \ name q \ incoming q'})" proof - from FIN CL have FIN': "\q. q\nds \ finite (incoming q)" by (metis finite_imageI finite_insert infinite_super) note build_succ_aux1[OF FIN FIN'] also note build_succ_aux2[OF NINIT CL] finally show ?thesis by (rule order_trans) auto qed text \\paragraph{ Accepting Sets}\ -context begin interpretation LTL_Syntax . primrec until_frmlsn :: "'a frml \ ('a frml \ 'a frml) set" where "until_frmlsn (\ and\<^sub>n \) = (until_frmlsn \) \ (until_frmlsn \)" | "until_frmlsn (X\<^sub>n \) = until_frmlsn \" | "until_frmlsn (\ U\<^sub>n \) = insert (\, \) ((until_frmlsn \) \ (until_frmlsn \))" | "until_frmlsn (\ V\<^sub>n \) = (until_frmlsn \) \ (until_frmlsn \)" | "until_frmlsn (\ or\<^sub>n \) = (until_frmlsn \) \ (until_frmlsn \)" | "until_frmlsn (true\<^sub>n) = {}" | "until_frmlsn (false\<^sub>n) = {}" | "until_frmlsn (prop\<^sub>n(_)) = {}" | "until_frmlsn (nprop\<^sub>n(_)) = {}" -end - lemma until_frmlsn_correct: "until_frmlsn \ = {(\, \). LTLnUntil \ \ \ subfrmlsn \}" by (induct \) auto definition "build_F nds \ \ (\(\,\). name ` {q \ nds. (LTLnUntil \ \ \ old q \ \ \ old q)}) ` until_frmlsn \" lemma build_F_correct: "build_F nds \ = {name ` A |A. \\ \. A = {q \ nds. LTLnUntil \ \ \ old q \ \ \ old q} \ LTLnUntil \ \ \ subfrmlsn \}" proof - have "{name ` A |A. \\ \. A = {q \ nds. LTLnUntil \ \ \ old q \ \ \ old q} \ LTLnUntil \ \ \ subfrmlsn \} = (\(\,\). name`{q\nds. LTLnUntil \ \ \ old q \ \ \ old q}) ` {(\, \). LTLnUntil \ \ \ subfrmlsn \}" by auto also have "\ = (\(\,\). name`{q\nds. LTLnUntil \ \ \ old q \ \ \ old q}) ` until_frmlsn \" unfolding until_frmlsn_correct .. finally show ?thesis unfolding build_F_def by simp qed text \\paragraph{ Labeling Function }\ definition "pn_props ps \ FOREACHi (\it (P,N). P = {p. LTLnProp p \ ps - it} \ N = {p. LTLnNProp p \ ps - it}) ps (\p (P,N). case p of LTLnProp p \ RETURN (insert p P,N) | LTLnNProp p \ RETURN (P, insert p N) | _ \ RETURN (P,N) ) ({},{})" lemma pn_props_correct: assumes [simp]: "finite ps" shows "pn_props ps \ SPEC(\r. r = ({p. LTLnProp p \ ps}, {p. LTLnNProp p \ ps}))" unfolding pn_props_def apply (refine_rcg refine_vcg) apply (auto split: ltln.split) done definition "pn_map nds \ FOREACH nds (\nd m. do { PN \ pn_props (old nd); RETURN (m(name nd \ PN)) }) Map.empty" lemma pn_map_correct: assumes [simp]: "finite nds" assumes FIN': "\nd. nd\nds \ finite (old nd)" assumes INJ: "inj_on name nds" shows "pn_map nds \ SPEC (\r. \qn. case r qn of None \ qn \ name`nds | Some (P,N) \ qn \ name`nds \ P = {p. LTLnProp p \ old (the_inv_into nds name qn)} \ N = {p. LTLnNProp p \ old (the_inv_into nds name qn)} )" unfolding pn_map_def apply (refine_rcg refine_vcg FOREACH_rule[where I="\it r. \qn. case r qn of None \ qn \ name`(nds - it) | Some (P,N) \ qn \ name`(nds - it) \ P = {p. LTLnProp p \ old (the_inv_into nds name qn)} \ N = {p. LTLnNProp p \ old (the_inv_into nds name qn)}"] order_trans[OF pn_props_correct] ) apply simp_all apply (blast dest: set_mp[THEN FIN']) [] apply (force split: option.splits simp: the_inv_into_f_f[OF INJ] it_step_insert_iff) [] apply (fastforce split: option.splits) [] done text \\paragraph{ Assembling the Implementation }\ definition "cr_rename_gba nds \ \ do { let V = name ` nds; let V0 = name ` {q \ nds. expand_init \ incoming q}; succmap \ build_succ nds; let E = E_of_succ (the_default {} o succmap); let F = build_F nds \; pnm \ pn_map nds; let L = (\qn l. case pnm qn of None \ False | Some (P,N) \ (\p\P. p\(l:::\<^sub>r\Id\fun_set_rel)) \ (\p\N. p\l) ); RETURN (\ g_V = V, g_E=E, g_V0=V0, gbg_F = F, gba_L = L \) }" lemma cr_rename_gba_refine: assumes INV: "nds_invars nds" assumes REL[simplified]: "(nds',nds)\Id" "(\',\)\Id" shows "cr_rename_gba nds' \' \ \Id (RETURN (gba_rename_ext (\_. ()) name (create_gba_from_nodes \ nds)))" unfolding RETURN_SPEC_conv proof (rule Id_SPEC_refine) from INV have NINIT: "expand_init \ name`nds" and FIN: "finite nds" and FIN': "\nd. nd\nds \ finite (old nd)" and CL: "\nd. nd\nds \ incoming nd \ insert expand_init (name`nds)" and INJ: "inj_on name nds" unfolding nds_invars_def by auto show "cr_rename_gba nds' \' \ SPEC (\x. x = gba_rename_ext (\_. ()) name (create_gba_from_nodes \ nds))" unfolding REL unfolding cr_rename_gba_def apply (refine_rcg refine_vcg order_trans[OF build_succ_correct[OF NINIT FIN CL]] order_trans[OF pn_map_correct[OF FIN FIN' INJ]] ) unfolding gba_rename_ecnv_def gb_rename_ecnv_def fr_rename_ext_def create_gba_from_nodes_def apply simp apply (intro conjI) apply (simp add: comp_def) apply (simp add: build_F_correct) apply (intro ext) apply (drule_tac x=qn in spec) apply (auto simp: the_inv_into_f_f[OF INJ] split: option.split prod.split) done qed definition "create_name_gba_aimpl \ \ do { nds \ create_graph_aimpl \; ASSERT (nds_invars nds); cr_rename_gba nds \ }" lemma create_name_gba_aimpl_refine: "create_name_gba_aimpl \ \ \Id (create_name_gba \)" unfolding create_name_gba_aimpl_def create_name_gba_alt apply (refine_rcg create_graph_aimpl_refine cr_rename_gba_refine) by auto subsection \Refinement to Efficient Data Structures\ subsubsection \Creation of GBA from Nodes\ schematic_goal until_frmlsn_impl_aux: assumes [relator_props, simp]: "R=Id" shows "(?c,until_frmlsn) \ \(R::(_\_::linorder) set)\ltln_rel \ \\R\ltln_rel \\<^sub>r \R\ltln_rel\dflt_rs_rel" unfolding until_frmlsn_def apply (autoref (keep_goal, trace)) done concrete_definition until_frmlsn_impl uses until_frmlsn_impl_aux lemmas [autoref_rules] = until_frmlsn_impl.refine[OF PREFER_id_D] schematic_goal build_succ_impl_aux: shows "(?c,build_succ) \ \\Rm,R\node_rel\list_set_rel \ \\nat_rel,\nat_rel\list_set_rel\iam_map_rel\nres_rel" unfolding build_succ_def[abs_def] expand_init_def using [[autoref_trace_failed_id]] apply (autoref (keep_goal, trace)) done concrete_definition build_succ_impl uses build_succ_impl_aux lemmas [autoref_rules] = build_succ_impl.refine (* TODO: Post-processing should be on by default! *) schematic_goal build_succ_code_aux: "RETURN ?c \ build_succ_impl x" unfolding build_succ_impl_def apply (refine_transfer (post)) done concrete_definition build_succ_code uses build_succ_code_aux lemmas [refine_transfer] = build_succ_code.refine schematic_goal build_F_impl_aux: assumes [relator_props]: "R = Id" shows "(?c,build_F) \ \\Rm,R\node_rel\list_set_rel \ \R\ltln_rel \ \\nat_rel\list_set_rel\list_set_rel" unfolding build_F_def[abs_def] using [[autoref_trace_failed_id]] apply (autoref (keep_goal, trace)) done concrete_definition build_F_impl uses build_F_impl_aux lemmas [autoref_rules] = build_F_impl.refine[OF PREFER_id_D] schematic_goal pn_map_impl_aux: shows "(?c,pn_map) \ \\Rm,Id\node_rel\list_set_rel \ \\nat_rel,\Id\list_set_rel \\<^sub>r \Id\list_set_rel\iam_map_rel\nres_rel" unfolding pn_map_def[abs_def] pn_props_def using [[autoref_trace_failed_id]] apply (autoref (keep_goal, trace)) done concrete_definition pn_map_impl uses pn_map_impl_aux lemma pn_map_impl_autoref[autoref_rules]: assumes "PREFER_id R" shows "(pn_map_impl,pn_map) \ \\Rm,R\node_rel\list_set_rel \ \\nat_rel,\R\list_set_rel \\<^sub>r \R\list_set_rel\iam_map_rel\nres_rel" using assms pn_map_impl.refine by simp schematic_goal pn_map_code_aux: "RETURN ?c \ pn_map_impl x" unfolding pn_map_impl_def apply (refine_transfer (post)) done concrete_definition pn_map_code uses pn_map_code_aux lemmas [refine_transfer] = pn_map_code.refine thm autoref_tyrel schematic_goal cr_rename_gba_impl_aux: assumes ID[relator_props]: "R=Id" notes [autoref_tyrel del] = TYRELI[of "\nat_rel\dflt_rs_rel"] shows "(?c,cr_rename_gba) \ \\Rm,R\node_rel\list_set_rel \ \R\ltln_rel \ (?R::(?'c \ _) set)" unfolding ID unfolding cr_rename_gba_def[abs_def] expand_init_def comp_def using [[autoref_trace_failed_id]] apply (autoref (keep_goal, trace)) done concrete_definition cr_rename_gba_impl uses cr_rename_gba_impl_aux thm cr_rename_gba_impl.refine lemma cr_rename_gba_autoref[autoref_rules]: assumes "PREFER_id R" shows "(cr_rename_gba_impl, cr_rename_gba) \ \\Rm, R\node_rel\list_set_rel \ \R\ltln_rel \ \gbav_impl_rel_ext unit_rel nat_rel (\R\fun_set_rel)\nres_rel" using assms cr_rename_gba_impl.refine[of R Rm] by simp schematic_goal cr_rename_gba_code_aux: "RETURN ?c \ cr_rename_gba_impl x y" unfolding cr_rename_gba_impl_def apply (refine_transfer (post)) done concrete_definition cr_rename_gba_code uses cr_rename_gba_code_aux lemmas [refine_transfer] = cr_rename_gba_code.refine subsubsection \Creation of Graph\ text \ The implementation of the node-set. The relation enforces that there are no different nodes with the same name. This effectively establishes an additional invariant, made explicit by an assertion in the refined program. This invariant allows for a more efficient implementation. \ definition ls_nds_rel_def_internal: "ls_nds_rel R \ \R\list_set_rel \ {(_,s). inj_on name s}" lemma ls_nds_rel_def: "\R\ls_nds_rel = \R\list_set_rel \ {(_,s). inj_on name s}" by (simp add: relAPP_def ls_nds_rel_def_internal) lemmas [autoref_rel_intf] = REL_INTFI[of ls_nds_rel i_set] lemma ls_nds_rel_sv[relator_props]: assumes "single_valued R" shows "single_valued (\R\ls_nds_rel)" using list_set_rel_sv[OF assms] unfolding ls_nds_rel_def by (metis inf.cobounded1 single_valued_subset) context begin interpretation autoref_syn . lemma lsnds_empty_autoref[autoref_rules]: assumes "PREFER_id R" shows "([],{})\\R\ls_nds_rel" using assms apply (simp add: ls_nds_rel_def) by autoref lemma lsnds_insert_autoref[autoref_rules]: assumes "SIDE_PRECOND (name n \ name`ns)" assumes "(n',n)\R" assumes "(ns',ns)\\R\ls_nds_rel" shows "(n'#ns',(OP insert ::: R \ \R\ls_nds_rel \ \R\ls_nds_rel)$n$ns) \ \R\ls_nds_rel" using assms unfolding ls_nds_rel_def apply simp proof (elim conjE, rule conjI) assume [autoref_rules]: "(n', n) \ R" "(ns', ns) \ \R\list_set_rel" assume "name n \ name ` ns" and "inj_on name ns" then have "n \ ns" by (auto) then show "(n' # ns', insert n ns) \ \R\list_set_rel" by autoref qed auto lemma ls_nds_image_autoref_aux: assumes [autoref_rules]: "(fi,f) \ Ra \ Rb" assumes "(l,s) \ \Ra\ls_nds_rel" assumes [simp]: "\x. name (f x) = name x" shows "(map fi l, f`s) \ \Rb\ls_nds_rel" proof - from assms have [autoref_rules]: "(l,s)\\Ra\list_set_rel" and INJ: "(inj_on name s)" by (auto simp add: ls_nds_rel_def) have [simp]: "inj_on f s" by (rule inj_onI) (metis INJ assms(3) inj_on_eq_iff) have "(map fi l, f`s) \ \Rb\list_set_rel" by (autoref (keep_goal)) moreover have "inj_on name (f`s)" apply (rule inj_onI) apply (auto simp: image_iff dest: inj_onD[OF INJ]) done ultimately show ?thesis by (auto simp: ls_nds_rel_def) qed lemma ls_nds_image_autoref[autoref_rules]: assumes "(fi,f) \ Ra \ Rb" assumes "SIDE_PRECOND (\x. name (f x) = name x)" shows "(map fi, (OP op ` ::: (Ra\Rb) \ \Ra\ls_nds_rel \ \Rb\ls_nds_rel)$f) \ \Ra\ls_nds_rel \ \Rb\ls_nds_rel" using assms unfolding autoref_tag_defs using ls_nds_image_autoref_aux by blast lemma list_set_autoref_to_list[autoref_ga_rules]: shows "is_set_to_sorted_list (\_ _. True) R ls_nds_rel id" unfolding is_set_to_list_def is_set_to_sorted_list_def ls_nds_rel_def it_to_sorted_list_def list_set_rel_def br_def by auto end +hide_const (open) Im context begin interpretation autoref_syn . lemma [autoref_itype]: "upd_incoming ::\<^sub>i \Im, I\\<^sub>ii_node \\<^sub>i \\Im', I\\<^sub>ii_node\\<^sub>ii_set \\<^sub>i \\Im', I\\<^sub>ii_node\\<^sub>ii_set" by simp end term upd_incoming schematic_goal upd_incoming_impl_aux: assumes "REL_IS_ID R" shows "(?c, upd_incoming)\\Rm1,R\node_rel \ \\Rm2,R\node_rel\ls_nds_rel \ \\Rm2,R\node_rel\ls_nds_rel" using assms apply simp unfolding upd_incoming_def[abs_def] using [[autoref_trace_failed_id]] apply (autoref (keep_goal)) done concrete_definition upd_incoming_impl uses upd_incoming_impl_aux lemmas [autoref_rules] = upd_incoming_impl.refine[OF PREFER_D[of REL_IS_ID]] schematic_goal expand_impl_aux: "(?c, expand_aimpl) \ \unit_rel,Id\node_rel \\<^sub>r \\unit_rel,Id\node_rel\ls_nds_rel \ \nat_rel \\<^sub>r \\unit_rel,Id\node_rel\ls_nds_rel\nres_rel" unfolding expand_aimpl_def[abs_def] expand_new_name_def using [[autoref_trace_failed_id, autoref_trace_intf_unif]] apply (autoref (trace,keep_goal)) done concrete_definition expand_impl uses expand_impl_aux context begin interpretation autoref_syn . lemma [autoref_itype]: "expand\<^sub>T ::\<^sub>i \\i_unit, I\\<^sub>ii_node, \\i_unit, I\\<^sub>ii_node\\<^sub>ii_set\\<^sub>ii_prod \\<^sub>i \\i_nat,\\i_unit, I\\<^sub>ii_node\\<^sub>ii_set\\<^sub>ii_prod\\<^sub>ii_nres" by simp lemma expand_autoref[autoref_rules]: assumes ID: "PREFER_id R" assumes A: "(n_ns', n_ns) \ \unit_rel,R\node_rel \\<^sub>r \\unit_rel,R\node_rel\list_set_rel" assumes B: "SIDE_PRECOND ( let (n,ns)=n_ns in inj_on name ns \ (\n'\ns. name n > name n') )" shows "(expand_impl n_ns', (OP expand_aimpl ::: \unit_rel,R\node_rel \\<^sub>r \\unit_rel,R\node_rel\list_set_rel \ \nat_rel \\<^sub>r \\unit_rel,R\node_rel\list_set_rel\nres_rel)$n_ns) \ \nat_rel \\<^sub>r \\unit_rel,R\node_rel\list_set_rel\nres_rel" proof simp from ID A B have 1: "(n_ns, n_ns) \ Id \ {(_,(n,ns)). \n'\ns. name n > name n'}" and 2: "(n_ns', n_ns) \ \unit_rel,Id\node_rel \\<^sub>r \\unit_rel,Id\node_rel\ls_nds_rel" unfolding ls_nds_rel_def apply - apply auto [] apply (cases n_ns') apply auto [] done have [simp]: "single_valued (nat_rel \\<^sub>r \\unit_rel, Id\node_rel\list_set_rel)" by tagged_solver have [relator_props]: "\R. single_valued (Id \ R)" by (metis IntE single_valuedD single_valuedI single_valued_Id) have [simp]: "single_valued ((nat_rel \\<^sub>r \\unit_rel, Id\node_rel\ls_nds_rel) O (Id \ {(_, n, ns). \n'\ns. name n' < n}))" by (tagged_solver) from expand_impl.refine[THEN fun_relD, OF 2, THEN nres_relD] show "(expand_impl n_ns', expand_aimpl n_ns) \ \nat_rel \\<^sub>r \\unit_rel, R\node_rel\list_set_rel\nres_rel" apply - apply (rule nres_relI) using ID apply (simp add: pw_le_iff refine_pw_simps) apply (fastforce simp: ls_nds_rel_def) done qed end schematic_goal expand_code_aux: "RETURN ?c \ expand_impl n_ns" unfolding expand_impl_def by (refine_transfer the_resI) concrete_definition expand_code uses expand_code_aux prepare_code_thms expand_code_def lemmas [refine_transfer] = expand_code.refine schematic_goal create_graph_impl_aux: assumes ID: "R=Id" shows "(?c, create_graph_aimpl) \ \R\ltln_rel \ \\\unit_rel,R\node_rel\list_set_rel\nres_rel" unfolding ID unfolding create_graph_aimpl_def[abs_def] expand_init_def expand_new_name_def using [[autoref_trace_failed_id]] apply (autoref (keep_goal)) done concrete_definition create_graph_impl uses create_graph_impl_aux lemmas [autoref_rules] = create_graph_impl.refine[OF PREFER_id_D] schematic_goal create_graph_code_aux: "RETURN ?c \ create_graph_impl \" unfolding create_graph_impl_def by refine_transfer concrete_definition create_graph_code uses create_graph_code_aux lemmas [refine_transfer] = create_graph_code.refine schematic_goal create_name_gba_impl_aux: "(?c, (create_name_gba_aimpl:: 'a::linorder ltln \ _)) \ \Id\ltln_rel \ (?R::(?'c\_) set)" unfolding create_name_gba_aimpl_def[abs_def] using [[autoref_trace_failed_id]] apply (autoref (keep_goal, trace)) done concrete_definition create_name_gba_impl uses create_name_gba_impl_aux lemma create_name_gba_autoref[autoref_rules]: assumes "PREFER_id R" shows "(create_name_gba_impl, create_name_gba) \ \R\ltln_rel \ \gbav_impl_rel_ext unit_rel nat_rel (\R\fun_set_rel)\nres_rel" (is "_:_\\?R\nres_rel") proof (intro fun_relI nres_relI) fix \ \' assume A: "(\,\')\\R\ltln_rel" from assms have RID[simp]: "R=Id" by simp note create_name_gba_impl.refine[THEN fun_relD,THEN nres_relD, OF A[unfolded RID]] also note create_name_gba_aimpl_refine finally show "create_name_gba_impl \ \ \ ?R (create_name_gba \')" by simp qed schematic_goal create_name_gba_code_aux: "RETURN ?c \ create_name_gba_impl \" unfolding create_name_gba_impl_def by (refine_transfer (post)) concrete_definition create_name_gba_code uses create_name_gba_code_aux lemmas [refine_transfer] = create_name_gba_code.refine schematic_goal create_name_igba_impl_aux: assumes RID: "R=Id" shows "(?c,create_name_igba)\ \R\ltln_rel \ \igbav_impl_rel_ext unit_rel nat_rel (\R\fun_set_rel)\nres_rel" unfolding RID unfolding create_name_igba_def[abs_def] using [[autoref_trace_failed_id]] apply (autoref (trace, keep_goal)) done concrete_definition create_name_igba_impl uses create_name_igba_impl_aux lemmas [autoref_rules] = create_name_igba_impl.refine[OF PREFER_id_D] schematic_goal create_name_igba_code_aux: "RETURN ?c \ create_name_igba_impl \" unfolding create_name_igba_impl_def by (refine_transfer (post)) concrete_definition create_name_igba_code uses create_name_igba_code_aux lemmas [refine_transfer] = create_name_igba_code.refine export_code create_name_igba_code checking SML end diff --git a/thys/LTL_to_GBA/ROOT b/thys/LTL_to_GBA/ROOT --- a/thys/LTL_to_GBA/ROOT +++ b/thys/LTL_to_GBA/ROOT @@ -1,19 +1,17 @@ chapter AFP session LTL_to_GBA (AFP) = CAVA_Automata + options [timeout = 900] theories [document = false] "../Stuttering_Equivalence/PLTL" theories - LTL - LTL_Rewrite LTL_Stutter LTL_to_GBA LTL_to_GBA_impl theories [document = false] All_Of_LTL_to_GBA document_files "root.tex" "intro.tex" "root.bib" diff --git a/thys/Promela/PromelaLTL.thy b/thys/Promela/PromelaLTL.thy --- a/thys/Promela/PromelaLTL.thy +++ b/thys/Promela/PromelaLTL.thy @@ -1,322 +1,321 @@ section "LTL integration" theory PromelaLTL imports "Promela" "../LTL_to_GBA/LTL" begin text {* We have a semantic engine for Promela. But we need to have an integration with LTL -- more specificly, we must know when a proposition is true in a global state. This is achieved in this theory. *} subsection {* LTL optimization *} text {* For efficiency reasons, we do not store the whole @{typ expr} on the labels of a system automaton, but @{typ nat} instead. This index then is used to look up the corresponding @{typ expr}. *} type_synonym APs = "expr iarray" primrec ltlc_aps_list' :: "'a ltlc \ 'a list \ 'a list" where "ltlc_aps_list' LTLcTrue l = l" | "ltlc_aps_list' LTLcFalse l = l" | "ltlc_aps_list' (LTLcProp p) l = (if List.member l p then l else p#l)" | "ltlc_aps_list' (LTLcNeg x) l = ltlc_aps_list' x l" | "ltlc_aps_list' (LTLcNext x) l = ltlc_aps_list' x l" | "ltlc_aps_list' (LTLcFinal x) l = ltlc_aps_list' x l" | "ltlc_aps_list' (LTLcGlobal x) l = ltlc_aps_list' x l" | "ltlc_aps_list' (LTLcAnd x y) l = ltlc_aps_list' y (ltlc_aps_list' x l)" | "ltlc_aps_list' (LTLcOr x y) l = ltlc_aps_list' y (ltlc_aps_list' x l)" | "ltlc_aps_list' (LTLcImplies x y) l = ltlc_aps_list' y (ltlc_aps_list' x l)" -| "ltlc_aps_list' (LTLcIff x y) l = ltlc_aps_list' y (ltlc_aps_list' x l)" | "ltlc_aps_list' (LTLcUntil x y) l = ltlc_aps_list' y (ltlc_aps_list' x l)" | "ltlc_aps_list' (LTLcRelease x y) l = ltlc_aps_list' y (ltlc_aps_list' x l)" lemma ltlc_aps_list'_correct: - "set (ltlc_aps_list' \ l) = ltlc_aprops \ \ set l" + "set (ltlc_aps_list' \ l) = ltlc_props \ \ set l" by (induct \ arbitrary: l) (auto simp add: in_set_member) lemma ltlc_aps_list'_distinct: "distinct l \ distinct (ltlc_aps_list' \ l)" by (induct \ arbitrary: l) (auto simp add: in_set_member) definition ltlc_aps_list :: "'a ltlc \ 'a list" where "ltlc_aps_list \ = ltlc_aps_list' \ []" lemma ltlc_aps_list_correct: - "set (ltlc_aps_list \) = ltlc_aprops \" + "set (ltlc_aps_list \) = ltlc_props \" unfolding ltlc_aps_list_def by (force simp: ltlc_aps_list'_correct) lemma ltlc_aps_list_distinct: "distinct (ltlc_aps_list \)" unfolding ltlc_aps_list_def by (auto intro: ltlc_aps_list'_distinct) primrec idx' :: "nat \ 'a list \ 'a \ nat option" where "idx' _ [] _ = None" | "idx' ctr (x#xs) y = (if x = y then Some ctr else idx' (ctr+1) xs y)" definition "idx = idx' 0" lemma idx'_correct: assumes "distinct xs" shows "idx' ctr xs y = Some n \ n \ ctr \ n < length xs + ctr \ xs ! (n-ctr) = y" using assms proof (induction xs arbitrary: n ctr) case (Cons x xs) show ?case proof (cases "x=y") case True with Cons.prems have *: "y \ set xs" by auto { assume A: "(y#xs)!(n-ctr) = y" and less: "ctr \ n" and length: "n < length (y#xs) + ctr" have "n = ctr" proof (rule ccontr) assume "n \ ctr" with less have "n-ctr > 0" by auto moreover from `n\ctr` length have "n-ctr < length(y#xs)" by auto ultimately have "(y#xs)!(n-ctr) \ set xs" by simp with A * show False by auto qed } with True Cons show ?thesis by auto next case False from Cons have "distinct xs" by simp with Cons.IH False have "idx' (Suc ctr) xs y = Some n \ Suc ctr \ n \ n < length xs + Suc ctr \ xs ! (n - Suc ctr) = y" by simp with False show ?thesis apply - apply (rule iffI) apply clarsimp_all done qed qed simp lemma idx_correct: assumes "distinct xs" shows "idx xs y = Some n \ n < length xs \ xs ! n = y" using idx'_correct[OF assms] by (simp add: idx_def) lemma idx_dom: assumes "distinct xs" shows "dom (idx xs) = set xs" by (auto simp add: idx_correct assms in_set_conv_nth) lemma idx_image_self: assumes "distinct xs" shows "(the \ idx xs) ` set xs = {.. set xs" with in_set_conv_nth obtain n where n: "n < length xs" "xs ! n = x" by metis with idx_correct[OF assms] have "idx xs x = Some n" by simp hence "the (idx xs x) = n" by simp with n show "(the \ idx xs) x < length xs" by simp next fix n assume "n < length xs" moreover with nth_mem have "xs ! n \ set xs" by simp then obtain x where "xs ! n = x" "x \ set xs" by simp_all ultimately have "idx xs x = Some n" by (simp add: idx_correct[OF assms]) hence "the (idx xs x) = n" by simp thus "n \ (the \ idx xs) ` set xs" using `x \ set xs` by auto qed lemma idx_ran: assumes "distinct xs" shows "ran (idx xs) = {.. APs \ nat ltlc" where "ltl_convert \ = ( let APs = ltlc_aps_list \; \\<^sub>i = map_ltlc (the \ idx APs) \ in (IArray APs, \\<^sub>i))" lemma ltl_convert_correct: assumes "ltl_convert \ = (APs, \\<^sub>i)" - shows "ltlc_aprops \ = set (IArray.list_of APs)" (is "?P1") - and "ltlc_aprops \\<^sub>i = {.. = set (IArray.list_of APs)" (is "?P1") + and "ltlc_props \\<^sub>i = {..\<^sub>i = map_ltlc (the \ idx (IArray.list_of APs)) \" (is "?P3") and "distinct (IArray.list_of APs)" proof - let ?APs = "IArray.list_of APs" from assms have APs_def: "?APs = ltlc_aps_list \" unfolding ltl_convert_def by auto with ltlc_aps_list_correct show APs_set: ?P1 by metis from assms show ?P3 unfolding ltl_convert_def by auto - from assms have "ltlc_aprops \\<^sub>i = (the \ idx ?APs) ` ltlc_aprops \" + from assms have "ltlc_props \\<^sub>i = (the \ idx ?APs) ` ltlc_props \" unfolding ltl_convert_def by (auto simp add: ltlc.set_map) moreover from APs_def ltlc_aps_list_distinct show "distinct ?APs" by simp note idx_image_self[OF this] moreover note APs_set ultimately show ?P2 by simp qed definition prepare :: "_ \ (program \ unit) \ ast \ expr ltlc \ (program \ APs \ gState) \ nat ltlc" where "prepare cfg ast \ \ let (prog,g\<^sub>0) = Promela.setUp ast; (APs,\\<^sub>i) = PromelaLTL.ltl_convert \ in ((prog, APs, g\<^sub>0), \\<^sub>i)" lemma prepare_instrument[code]: "prepare cfg ast \ \ let (_,printF) = cfg; _ = PromelaStatistics.start (); (prog,g\<^sub>0) = Promela.setUp ast; _ = printF prog; (APs,\\<^sub>i) = PromelaLTL.ltl_convert \; _ = PromelaStatistics.stop_timer () in ((prog, APs, g\<^sub>0), \\<^sub>i)" by (simp add: prepare_def) export_code prepare checking SML subsection {* Language of a Promela program *} definition propValid :: "APs \ gState \ nat \ bool" where "propValid APs g i \ i < IArray.length APs \ exprArith g emptyProc (APs!!i) \ 0" definition promela_E :: "program \ (gState \ gState) set" -- "Transition relation of a promela program" where "promela_E prog \ {(g,g'). g' \ ls.\ (nexts_code prog g)}" definition promela_E_ltl :: "program \ APs \ (gState \ gState) set" where "promela_E_ltl = promela_E \ fst" definition promela_is_run' :: "program \ gState \ gState word \ bool" -- "Predicate defining runs of promela programs" where "promela_is_run' progg r \ let (prog,g\<^sub>0)=progg in r 0 = g\<^sub>0 \ (\i. r (Suc i) \ ls.\ (nexts_code prog (r i)))" abbreviation "promela_is_run \ promela_is_run' \ setUp" definition promela_is_run_ltl :: "program \ APs \ gState \ gState word \ bool" where "promela_is_run_ltl promg r \ let (prog,APs,g) = promg in promela_is_run' (prog,g) r" definition promela_props :: "gState \ expr set" where "promela_props g = {e. exprArith g emptyProc e \ 0}" definition promela_props_ltl :: "APs \ gState \ nat set" where "promela_props_ltl APs g \ Collect (propValid APs g)" definition promela_language :: "ast \ expr set word set" where "promela_language ast \ {promela_props \ r | r. promela_is_run ast r}" definition promela_language_ltl :: "program \ APs \ gState \ nat set word set" where "promela_language_ltl promg \ let (prog,APs,g) = promg in {promela_props_ltl APs \ r | r. promela_is_run_ltl promg r}" lemma promela_props_ltl_map_aprops: assumes "ltl_convert \ = (APs,\\<^sub>i)" shows "promela_props_ltl APs = - map_aprops (idx (IArray.list_of APs)) \ promela_props" + map_props (idx (IArray.list_of APs)) \ promela_props" proof - let ?APs = "IArray.list_of APs" let ?idx = "idx ?APs" from ltl_convert_correct assms have D: "distinct ?APs" by simp show ?thesis proof (intro ext set_eqI iffI) fix g i assume "i \ promela_props_ltl APs g" hence "propValid APs g i" by (simp add: promela_props_ltl_def) hence l: "i < IArray.length APs" "exprArith g emptyProc (APs!!i) \ 0" by (simp_all add: propValid_def) hence "APs!!i \ promela_props g" by (simp add: promela_props_def) moreover from idx_correct l D have "?idx (APs!!i) = Some i" by fastforce - ultimately show "i \ (map_aprops ?idx \ promela_props) g" - unfolding o_def map_aprops_def + ultimately show "i \ (map_props ?idx \ promela_props) g" + unfolding o_def map_props_def by blast next fix g i - assume "i \ (map_aprops ?idx \ promela_props) g" + assume "i \ (map_props ?idx \ promela_props) g" then obtain p where p_def: "p \ promela_props g" "?idx p = Some i" - unfolding map_aprops_def o_def + unfolding map_props_def o_def by auto hence expr: "exprArith g emptyProc p \ 0" by (simp add: promela_props_def) from D p_def have "i < IArray.length APs" "APs !! i = p" using idx_correct by fastforce+ with expr have "propValid APs g i" by (simp add: propValid_def) thus "i \ promela_props_ltl APs g" by (simp add: promela_props_ltl_def) qed qed lemma promela_run_in_language_iff: assumes conv: "ltl_convert \ = (APs,\\<^sub>i)" shows "promela_props \ \ \ ltlc_language \ \ promela_props_ltl APs \ \ \ ltlc_language \\<^sub>i" (is "?L \ ?R") proof - let ?APs = "IArray.list_of APs" from conv have D: "distinct ?APs" by (simp add: ltl_convert_correct) - with conv have APs: "ltlc_aprops \ \ dom (idx ?APs)" + with conv have APs: "ltlc_props \ \ dom (idx ?APs)" by (simp add: idx_dom ltl_convert_correct) note map_semantics = map_ltlc_semantics[OF idx_inj_on_dom[OF D] APs] promela_props_ltl_map_aprops[OF conv] ltl_convert_correct[OF conv] have "?L \ (promela_props \ \) \\<^sub>c \" by (simp add: ltlc_language_def) also have "... \ (promela_props_ltl APs \ \) \\<^sub>c \\<^sub>i" using map_semantics by (simp add: o_assoc) also have "... \ ?R" by (simp add: ltlc_language_def) finally show ?thesis . qed lemma promela_language_sub_iff: assumes conv: "ltl_convert \ = (APs,\\<^sub>i)" and setUp: "setUp ast = (prog,g)" shows "promela_language_ltl (prog,APs,g) \ ltlc_language \\<^sub>i \ promela_language ast \ ltlc_language \" using promela_run_in_language_iff[OF conv] setUp by (auto simp add: promela_language_ltl_def promela_language_def promela_is_run_ltl_def) (* from PromelaDatastructures *) hide_const (open) abort abort' abortv err err' errv usc usc' warn the_warn with_warn hide_const (open) idx idx' end diff --git a/thys/Stuttering_Equivalence/PLTL.thy b/thys/Stuttering_Equivalence/PLTL.thy --- a/thys/Stuttering_Equivalence/PLTL.thy +++ b/thys/Stuttering_Equivalence/PLTL.thy @@ -1,879 +1,797 @@ theory PLTL -imports StutterEquivalence + imports Main "../../../afp/LTL/LTL" Samplers StutterEquivalence begin section {* Stuttering Invariant PLTL Formulas *} text {* We define the syntax and semantics of propositional linear-time temporal logic PLTL and show that its next-free fragment is invariant to finite stuttering. *} -subsection {* Syntax and semantics of PLTL *} - -text {* - PLTL formulas are built from atomic formulas, propositional connectives, - and the temporal operators ``next'' and ``until''. The following data - type definition is parameterized by the type of states over which - formulas are evaluated. -*} - -datatype 'a pltl = - "false" - | "atom" "'a \ bool" - | "implies" "'a pltl" "'a pltl" - | "next" "'a pltl" - | "until" "'a pltl" "'a pltl" - -text {* - The semantics of PLTL formulas is defined inductively w.r.t. - @{text "\"}-sequences of states. -*} - -fun holds_of :: "(nat \ 'a) \ 'a pltl \ bool" ("_ \ _" [70,70] 40) where - "(\ \ false) = False" -| "(\ \ atom p) = p (\ 0)" -| "(\ \ implies \ \) = ((\ \ \) \ (\ \ \))" -| "(\ \ next \) = (\[1..] \ \)" -| "(\ \ until \ \) = (\k. \[k..] \ \ \ (\i[i..] \ \))" - -text {* - Further connectives of PLTL can be defined as abbreviations. -*} - -definition not where "not \ = implies \ false" +notation pltl_semantics ("_ \ _" [80,80] 80) + and LTLpFalse ("false") + and LTLpImplies ("implies") + and LTLpAtom ("atom") + and LTLpUntil ("until") + and LTLpNext ("next") + and pltl_atoms ("atoms") -definition true where "true = not false" - -definition or where "or \ \ = implies (not \) \" - -definition "and" where "and \ \ = not (or (not \) (not \))" - -definition eventually where "eventually \ = until true \" - -definition always where "always \ = not (eventually (not \))" - -text {* - These definitions give rise to the expected semantics. -*} - -lemma holds_of_not [simp]: "(\ \ not \) = (\(\ \ \))" - by (simp add: not_def) - -lemma holds_of_true [simp]: "\ \ true" - by (simp add: true_def) - -lemma holds_of_or [simp]: "(\ \ or \ \) = (\ \ \ \ \ \ \)" - by (auto simp add: or_def) - -lemma holds_of_and [simp]: "(\ \ and \ \) = (\ \ \ \ \ \ \)" - by (simp add: and_def) - -lemma holds_of_eventually [simp]: "(\ \ eventually \) = (\k. \[k..] \ \)" - by (simp add: eventually_def) - -lemma holds_of_always [simp]: "(\ \ always \) = (\k. \[k..] \ \)" - by (simp add: always_def) - -text {* - We also define finite conjunctions and disjunctions. -*} +subsection {* Finite Conjunctions and Disjunctions *} (* It would be tempting to define these operators as follows: definition OR where "OR \ = Finite_Set.fold or false \" definition AND where "AND \ = Finite_Set.fold and true \" However, this would only work if "or" and "and" were left-commutative, which they are not (syntactically). We must therefore resort to the more general fold_graph predicate, effectively picking a conjunction (or disjunction) in arbitrary order. An alternative would be to define these generalized operators over lists of formulas, but working with sets is more natural in the following. *) -definition OR where "OR \ \ SOME \. fold_graph or false \ \" +definition OR where "OR \ \ SOME \. fold_graph LTLpOr LTLpFalse \ \" -definition AND where "AND \ \ SOME \. fold_graph and true \ \" +definition AND where "AND \ \ SOME \. fold_graph LTLpAnd LTLpTrue \ \" -lemma fold_graph_OR: "finite \ \ fold_graph or false \ (OR \)" +lemma fold_graph_OR: "finite \ \ fold_graph LTLpOr LTLpFalse \ (OR \)" unfolding OR_def by (rule someI2_ex[OF finite_imp_fold_graph]) -lemma fold_graph_AND: "finite \ \ fold_graph and true \ (AND \)" +lemma fold_graph_AND: "finite \ \ fold_graph LTLpAnd LTLpTrue \ (AND \)" unfolding AND_def by (rule someI2_ex[OF finite_imp_fold_graph]) lemma holds_of_OR [simp]: assumes fin: "finite (\::'a pltl set)" shows "(\ \ OR \) = (\\\\. \ \ \)" proof - { fix \::"'a pltl" - assume "fold_graph or false \ \" + assume "fold_graph LTLpOr LTLpFalse \ \" hence "(\ \ \) = (\\\\. \ \ \)" by (rule fold_graph.induct) auto } with fold_graph_OR[OF fin] show ?thesis by simp qed lemma holds_of_AND [simp]: assumes fin: "finite (\::'a pltl set)" shows "(\ \ AND \) = (\\\\. \ \ \)" proof - { fix \::"'a pltl" - assume "fold_graph and true \ \" + assume "fold_graph LTLpAnd LTLpTrue \ \" hence "(\ \ \) = (\\\\. \ \ \)" by (rule fold_graph.induct) auto } with fold_graph_AND[OF fin] show ?thesis by simp qed - subsection {* Next-Free PLTL Formulas *} text {* A PLTL formula is called \emph{next-free} if it does not contain any - subformula @{text "tnext f"}. + subformula. *} -fun "next_free" where +fun next_free :: "'a pltl \ bool" +where "next_free false = True" | "next_free (atom p) = True" | "next_free (implies \ \) = (next_free \ \ next_free \)" | "next_free (next \) = False" | "next_free (until \ \) = (next_free \ \ next_free \)" -lemma next_free_not [simp]: "next_free (not \) = next_free \" - by (simp add: not_def) - -lemma next_free_true [simp]: "next_free true" - by (simp add: true_def) +lemma next_free_not [simp]: + "next_free (not\<^sub>p \) = next_free \" + by (simp add: LTLpNot_def) -lemma next_free_or [simp]: "next_free (or \ \) = (next_free \ \ next_free \)" - by (simp add: or_def) +lemma next_free_true [simp]: + "next_free true\<^sub>p" + by (simp add: LTLpTrue_def) -lemma next_free_and [simp]: "next_free (and \ \) = (next_free \ \ next_free \)" - by (simp add: and_def) +lemma next_free_or [simp]: + "next_free (\ or\<^sub>p \) = (next_free \ \ next_free \)" + by (simp add: LTLpOr_def) -lemma next_free_eventually [simp]: "next_free (eventually \) = next_free \" - by (simp add: eventually_def) +lemma next_free_and [simp]: "next_free (\ and\<^sub>p \) = (next_free \ \ next_free \)" + by (simp add: LTLpAnd_def) -lemma next_free_always [simp]: "next_free (always \) = next_free \" - by (simp add: always_def) +lemma next_free_eventually [simp]: + "next_free (F\<^sub>p \) = next_free \" + by (simp add: LTLpEventually_def) + +lemma next_free_always [simp]: + "next_free (G\<^sub>p \) = next_free \" + by (simp add: LTLpAlways_def) lemma next_free_OR [simp]: assumes fin: "finite (\::'a pltl set)" shows "next_free (OR \) = (\\\\. next_free \)" proof - { fix \::"'a pltl" - assume "fold_graph or false \ \" + assume "fold_graph LTLpOr LTLpFalse \ \" hence "next_free \ = (\\\\. next_free \)" by (rule fold_graph.induct) auto } with fold_graph_OR[OF fin] show ?thesis by simp qed lemma next_free_AND [simp]: assumes fin: "finite (\::'a pltl set)" shows "next_free (AND \) = (\\\\. next_free \)" proof - { fix \::"'a pltl" - assume "fold_graph and true \ \" + assume "fold_graph LTLpAnd LTLpTrue \ \" hence "next_free \ = (\\\\. next_free \)" by (rule fold_graph.induct) auto } with fold_graph_AND[OF fin] show ?thesis by simp qed subsection {* Stuttering Invariance of PLTL Without ``Next'' *} text {* A PLTL formula is \emph{stuttering invariant} if for any stuttering equivalent state sequences @{text "\ \ \"}, the formula holds of @{text "\"} iff it holds of @{text "\"}. *} definition stutter_invariant where - "stutter_invariant \ = (\\ \. \ \ \ \ (\ \ \) = (\ \ \))" + "stutter_invariant \ = (\\ \. (\ \ \) \ (\ \ \) = (\ \ \))" text {* Since stuttering equivalence is symmetric, it is enough to show an implication in the above definition instead of an equivalence. *} lemma stutter_invariantI [intro!]: assumes "\\ \. \\ \ \; \ \ \\ \ \ \ \" shows "stutter_invariant \" proof - { fix \ \ assume st: "\ \ \" and f: "\ \ \" hence "\ \ \" by (rule assms) } moreover { fix \ \ assume st: "\ \ \" and f: "\ \ \" from st have "\ \ \" by (rule stutter_equiv_sym) from this f have "\ \ \" by (rule assms) } ultimately show ?thesis by (auto simp: stutter_invariant_def) qed lemma stutter_invariantD [dest]: assumes "stutter_invariant \" and "\ \ \" shows "(\ \ \) = (\ \ \)" using assms by (auto simp: stutter_invariant_def) text {* We first show that next-free PLTL formulas are indeed stuttering invariant. The proof proceeds by straightforward induction on the syntax of PLTL formulas. *} theorem next_free_stutter_invariant: "next_free \ \ stutter_invariant (\::'a pltl)" proof (induct "\") show "stutter_invariant false" by auto next fix p :: "'a \ bool" show "stutter_invariant (atom p)" proof fix \ \ assume "\ \ \" "\ \ atom p" thus "\ \ atom p" by (simp add: stutter_equiv_0) qed next fix \ \ :: "'a pltl" assume ih: "next_free \ \ stutter_invariant \" "next_free \ \ stutter_invariant \" assume "next_free (implies \ \)" with ih show "stutter_invariant (implies \ \)" by auto next fix \ :: "'a pltl" assume "next_free (next \)" -- {* hence contradiction *} thus "stutter_invariant (next \)" by simp next fix \ \ :: "'a pltl" assume ih: "next_free \ \ stutter_invariant \" "next_free \ \ stutter_invariant \" assume "next_free (until \ \)" with ih have stinv: "stutter_invariant \" "stutter_invariant \" by auto show "stutter_invariant (until \ \)" proof fix \ \ assume st: "\ \ \" and unt: "\ \ until \ \" from unt obtain m where 1: "\[m..] \ \" and 2: "\j[j..] \ \" by auto from st obtain n where 3: "\[m..] \ \[n..]" and 4: "\ij[j..] \ \[i..]" by (rule stutter_equiv_suffixes_right) from 1 3 stinv have "\[n..] \ \" by auto moreover from 2 4 stinv have "\i[i..] \ \" by force ultimately show "\ \ until \ \" by auto qed qed subsection {* Atoms, Canonical State Sequences, and Characteristic Formulas *} text {* We now address the converse implication: any stutter invariant PLTL formula @{text "\"} can be equivalently expressed by a next-free formula. The construction of that formula requires attention to the atomic formulas that appear in @{text "\"}. We will also prove that the next-free formula does not need any new atoms beyond those present in @{text "\"}. The following function collects the atoms (of type @{text "'a \ bool"}) of a PLTL formula. *} -fun atoms where - "atoms false = {}" -| "atoms (atom p) = {p}" -| "atoms (implies \ \) = atoms \ \ atoms \" -| "atoms (next \) = atoms \" -| "atoms (until \ \) = atoms \ \ atoms \" - -lemma atoms_finite [iff]: "finite (atoms \)" - by (induct \) auto - -lemma atoms_not [simp]: "atoms (not \) = atoms \" - by (simp add: not_def) - -lemma atoms_true [simp]: "atoms true = {}" - by (simp add: true_def) - -lemma atoms_or [simp]: "atoms (or \ \) = atoms \ \ atoms \" - by (simp add: or_def) - -lemma atoms_and [simp]: "atoms (and \ \) = atoms \ \ atoms \" - by (simp add: and_def) - -lemma atoms_eventually [simp]: "atoms (eventually \) = atoms \" - by (simp add: eventually_def) - -lemma atoms_always [simp]: "atoms (always \) = atoms \" - by (simp add: always_def) lemma atoms_OR [simp]: assumes fin: "finite (\::'a pltl set)" shows "atoms (OR \) = (\\\\. atoms \)" proof - { fix \::"'a pltl" - assume "fold_graph or false \ \" + assume "fold_graph LTLpOr LTLpFalse \ \" hence "atoms \ = (\\\\. atoms \)" by (rule fold_graph.induct) auto } with fold_graph_OR[OF fin] show ?thesis by simp qed lemma atoms_AND [simp]: assumes fin: "finite (\::'a pltl set)" shows "atoms (AND \) = (\\\\. atoms \)" proof - { fix \::"'a pltl" - assume "fold_graph and true \ \" + assume "fold_graph LTLpAnd LTLpTrue \ \" hence "atoms \ = (\\\\. atoms \)" by (rule fold_graph.induct) auto } with fold_graph_AND[OF fin] show ?thesis by simp qed text {* Given a set of atoms @{text A} as above, we say that two states are @{text A}-similar if they agree on all atoms in @{text A}. Two state sequences @{text "\"} and @{text "\"} are @{text A}-similar if corresponding states are @{text A}-equal. *} definition state_sim :: "['a, ('a \ bool) set, 'a] \ bool" ("_ ~_~ _" [70,100,70] 50) where "s ~A~ t = (\p\A. p s \ p t)" definition seq_sim :: "[nat \ 'a, ('a \ bool) set, nat \ 'a] \ bool" ("_ \_\ _" [70,100,70] 50) where "\ \A\ \ = (\n. (\ n) ~A~ (\ n))" text {* These relations are (indexed) equivalence relations. Moreover @{text "s ~A~ t"} implies @{text "s ~B~ t"} for @{text "B \ A"}, and similar for @{text "\ \A\ \"} and @{text "\ \B\ \"}. *} lemma state_sim_refl [simp]: "s ~A~ s" by (simp add: state_sim_def) lemma state_sim_sym: "s ~A~ t \ t ~A~ s" by (auto simp: state_sim_def) lemma state_sim_trans[trans]: "s ~A~ t \ t ~A~ u \ s ~A~ u" unfolding state_sim_def by blast lemma state_sim_mono: assumes "s ~A~ t" and "B \ A" shows "s ~B~ t" using assms unfolding state_sim_def by auto lemma seq_sim_refl [simp]: "\ \A\ \" by (simp add: seq_sim_def) lemma seq_sim_sym: "\ \A\ \ \ \ \A\ \" by (auto simp: seq_sim_def state_sim_sym) lemma seq_sim_trans[trans]: "\ \A\ \ \ \ \A\ \ \ \ \A\ \" unfolding seq_sim_def by (blast intro: state_sim_trans) lemma seq_sim_mono: assumes "\ \A\ \" and "B \ A" shows "\ \B\ \" using assms unfolding seq_sim_def by (blast intro: state_sim_mono) text {* State sequences that are similar w.r.t. the atoms of a PLTL formula evaluate that formula to the same value. *} lemma pltl_seq_sim: "\ \(atoms \)\ \ \ (\ \ \) = (\ \ \)" (is "?sim \ \ \ \ ?P \ \ \") proof (induct \ arbitrary: \ \) fix \ \ show "?P \ false \" by simp next fix p \ \ assume "?sim \ (atom p) \" thus "?P \ (atom p) \" by (auto simp: seq_sim_def state_sim_def) next fix \ \ \ \ assume ih: "\\ \. ?sim \ \ \ \ ?P \ \ \" "\\ \. ?sim \ \ \ \ ?P \ \ \" and sim: "?sim \ (implies \ \) \" from sim have "?sim \ \ \" "?sim \ \ \" by (auto elim: seq_sim_mono) with ih show "?P \ (implies \ \) \" by simp next fix \ \ \ assume ih: "\\ \. ?sim \ \ \ \ ?P \ \ \" and sim: "\ \atoms (next \)\ \" from sim have "\[1..] \atoms \\ \[1..]" by (auto simp: seq_sim_def) with ih show "?P \ (next \) \" by auto next fix \ \ \ \ assume ih: "\\ \. ?sim \ \ \ \ ?P \ \ \" "\\ \. ?sim \ \ \ \ ?P \ \ \" and sim: "?sim \ (until \ \) \" - from sim have "\i. \[i..] \atoms \\ \[i..]" "\j. \[j..] \atoms \\ \[j..]" + from sim have "\i. \[i..] \atoms \\ \[i..]" "\j. \[j..] \ atoms \ \ \[j..]" by (auto simp: seq_sim_def state_sim_def) with ih have "\i. ?P (\[i..]) \ (\[i..])" "\j. ?P (\[j..]) \ (\[j..])" - by (auto simp del: suffix_elt) - thus "?P \ (until \ \) \" by (auto simp del: suffix_elt) + by blast+ + thus "?P \ (until \ \) \" + by (meson pltl_semantics.simps(5)) qed text {* The following function picks an arbitrary representative among @{text A}-similar states. Because the choice is functional, any two @{text A}-similar states are mapped to the same state. *} definition canonize where "canonize A s \ SOME t. t ~A~ s" lemma canonize_state_sim: "canonize A s ~A~ s" unfolding canonize_def by (rule someI, rule state_sim_refl) lemma canonize_canonical: assumes st: "s ~A~ t" shows "canonize A s = canonize A t" proof - from st have "\u. (u ~A~s) = (u ~A~ t)" by (auto elim: state_sim_sym state_sim_trans) thus ?thesis unfolding canonize_def by simp qed lemma canonize_idempotent: "canonize A (canonize A s) = canonize A s" by (rule canonize_canonical[OF canonize_state_sim]) text {* In a canonical state sequence, any two @{text A}-similar states are in fact equal. *} definition canonical_sequence where "canonical_sequence A \ \ \m (n::nat). \ m ~A~ \ n \ \ m = \ n" text {* Every suffix of a canonical sequence is canonical, as is any (sampled) subsequence, in particular any stutter-sampling. *} lemma canonical_suffix: "canonical_sequence A \ \ canonical_sequence A (\[k..])" by (auto simp: canonical_sequence_def) lemma canonical_sampled: "canonical_sequence A \ \ canonical_sequence A (\ \ f)" by (auto simp: canonical_sequence_def) lemma canonical_reduced: "canonical_sequence A \ \ canonical_sequence A (\\)" unfolding stutter_reduced_def by (rule canonical_sampled) text {* For any sequence @{text "\"} there exists a canonical @{text A}-similar sequence @{text "\"}. Such a @{text "\"} can be obtained by canonizing all states of @{text "\"}. *} lemma canonical_exists: obtains \ where "\ \A\ \" "canonical_sequence A \" proof - have "(canonize A \ \) \A\ \" by (simp add: seq_sim_def canonize_state_sim) moreover have "canonical_sequence A (canonize A \ \)" by (auto simp: canonical_sequence_def canonize_idempotent dest: canonize_canonical) ultimately show ?thesis using that by blast qed text {* Given a state @{text s} and a set @{text A} of atoms, we define the characteristic formula of @{text s} as the conjunction of all atoms in @{text A} that hold of @{text s} and the negation of the atoms in @{text A} that do not hold of @{text s}. *} definition characteristic_formula where "characteristic_formula A s \ - (and (AND { atom p | p . p \ A \ p s }) - (AND { not (atom p) | p . p \ A \ \(p s) }))" + ((AND { atom p | p . p \ A \ p s }) and\<^sub>p (AND { not\<^sub>p (atom p) | p . p \ A \ \(p s) }))" lemma characteristic_holds: "finite A \ \ \ characteristic_formula A (\ 0)" by (auto simp: characteristic_formula_def) lemma characteristic_state_sim: assumes fin: "finite A" shows "(\ 0 ~A~ \ 0) = (\ \ characteristic_formula A (\ (0::nat)))" proof assume sim: "\ 0 ~A~ \ 0" { fix p assume "p \ A" with sim have "p (\ 0) = p (\ 0)" by (auto simp: state_sim_def) } with fin show "\ \ characteristic_formula A (\ 0)" by (auto simp: characteristic_formula_def) (blast+) next assume "\ \ characteristic_formula A (\ 0)" with fin show "\ 0 ~A~ \ 0" by (auto simp: characteristic_formula_def state_sim_def) qed subsection {* Stuttering Invariant PLTL Formulas Don't Need Next *} text {* The following is the main lemma used in the proof of the completeness theorem: for any PLTL formula @{text "\"} there exists a next-free formula @{text "\"} such that the two formulas evaluate to the same value over stutter-free and canonical sequences (w.r.t. some @{text "A \ atoms \"}). *} lemma ex_next_free_stutter_free_canonical: assumes A: "atoms \ \ A" and fin: "finite A" shows "\\. next_free \ \ atoms \ \ A \ (\\. stutter_free \ \ canonical_sequence A \ \ (\ \ \) = (\ \ \))" (is "\\. ?P \ \") using A proof (induct \) txt {* The cases of @{text "false"} and atomic formulas are trivial. *} have "?P false false" by auto thus "\\. ?P false \" .. next fix p assume "atoms (atom p) \ A" hence "?P (atom p) (atom p)" by auto thus "\\. ?P (atom p) \" .. next txt {* Implication is easy, using the induction hypothesis. *} fix \ \ assume "atoms \ \ A \ \\'. ?P \ \'" and "atoms \ \ A \ \\'. ?P \ \'" and "atoms (implies \ \) \ A" then obtain \' \' where "?P \ \'" "?P \ \'" by auto hence "?P (implies \ \) (implies \' \')" by auto thus "\\. ?P (implies \ \) \" .. next txt {* The case of @{text "until"} follows similarly. *} fix \ \ assume "atoms \ \ A \ \\'. ?P \ \'" and "atoms \ \ A \ \\'. ?P \ \'" and "atoms (until \ \) \ A" then obtain \' \' where 1: "?P \ \'" and 2: "?P \ \'" by auto { fix \ assume sigma: "stutter_free \" "canonical_sequence A \" hence "\k. stutter_free (\[k..])" "\k. canonical_sequence A (\[k..])" by (auto simp: stutter_free_suffix canonical_suffix) with 1 2 have "\k. (\[k..] \ \') = (\[k..] \ \)" and "\k. (\[k..] \ \') = (\[k..] \ \)" by (blast+) hence "(\ \ until \' \') = (\ \ until \ \)" by auto } with 1 2 have "?P (until \ \) (until \' \')" by auto thus "\\. ?P (until \ \) \" .. next txt {* The interesting case is the one of the @{text "next"}-operator. *} fix \ assume ih: "atoms \ \ A \ \\. ?P \ \" and at: "atoms (next \) \ A" then obtain \ where psi: "?P \ \" by auto txt {* A valuation (over @{text A}) is a set @{text "val \ A"} of atoms. We define some auxiliary notions: the valuation corresponding to a state and the characteristic formula for a valuation. Finally, we define the formula @{text "psi'"} that we will prove to be equivalent to @{text "next \"} over the stutter-free and canonical sequence @{text "\"}. *} def stval \ "\s. { p \ A . p s }" - def chi \ "\val. (and (AND {atom p | p . p \ val}) - (AND {not (atom p) | p . p \ A - val}))" - def psi' \ "(or (and \ (OR { always (chi val) | val . val \ A })) - (OR { and (chi val) (until (chi val) (and \ (chi val'))) | val val'. + def chi \ "\val. ((AND {atom p | p . p \ val}) and\<^sub>p + (AND {not\<^sub>p (atom p) | p . p \ A - val}))" + def psi' \ "(LTLpOr (\ and\<^sub>p (OR {G\<^sub>p (chi val) | val . val \ A })) + (OR {(chi val) and\<^sub>p (until (chi val) ( \ and\<^sub>p (chi val'))) | val val'. val \ A \ val' \ A \ val' \ val }))" - (is "(or (and _ (OR ?ALW)) (OR ?UNT))") + (is "(LTLpOr ( _ and\<^sub>p (OR ?ALW)) (OR ?UNT))") - have "\s. {not (atom p) | p . p \ A - stval s} - = {not (atom p) | p . p \ A \ \(p s)}" + have "\s. {not\<^sub>p (atom p) | p . p \ A - stval s} + = {not\<^sub>p (atom p) | p . p \ A \ \(p s)}" by (auto simp: stval_def) hence chi1: "\s. chi (stval s) = characteristic_formula A s" by (auto simp: chi_def stval_def characteristic_formula_def) { fix val \ assume val: "val \ A" and tau: "\ \ chi val" with fin have "val = stval (\ 0)" by (auto simp: stval_def chi_def finite_subset) } note chi2 = this - have "?UNT \ (\(val,val'). and (chi val) (until (chi val) (and \ (chi val')))) + have "?UNT \ (\(val,val'). (chi val) and\<^sub>p (until (chi val) (\ and\<^sub>p (chi val')))) ` (Pow A \ Pow A)" (is "_ \ ?S") by auto with fin have fin_UNT: "finite ?UNT" by (auto simp: finite_subset) have nf: "next_free psi'" proof - from fin have "\val. val \ A \ next_free (chi val)" by (auto simp: chi_def finite_subset) with fin fin_UNT psi show ?thesis by (force simp: psi'_def finite_subset) qed have atoms: "atoms psi' \ A" proof - from fin have at_chi: "\val. val \ A \ atoms (chi val) \ A" by (auto simp: chi_def finite_subset) - with fin psi have at_alw: "atoms (and \ (OR ?ALW)) \ A" + with fin psi have at_alw: "atoms (\ and\<^sub>p (OR ?ALW)) \ A" by auto blast? (** FRAGILE: auto leaves trivial goal about subset **) from fin fin_UNT psi at_chi have "atoms (OR ?UNT) \ A" by auto (blast+)? (** FRAGILE: even more left-over goals here **) with at_alw show ?thesis by (auto simp: psi'_def) qed { fix \ assume st: "stutter_free \" and can: "canonical_sequence A \" have "(\ \ next \) = (\ \ psi')" proof (cases "\ (Suc 0) = \ 0") case True txt {* In the case of a stuttering transition at the beginning, we must have infinite stuttering, and the first disjunct of @{text "psi'"} holds, whereas the second does not. *} { fix n have "\ n = \ 0" proof (cases n) case 0 thus ?thesis by simp next case Suc hence "n > 0" by simp with True st show ?thesis unfolding stutter_free_def by blast qed } note alleq = this have suffix: "\n. \[n..] = \" proof (rule ext) fix n i have "(\[n..]) i = \ 0" by (auto intro: alleq) moreover have "\ i = \ 0" by (rule alleq) ultimately show "(\[n..]) i = \ i" by simp qed with st can psi have 1: "(\ \ next \) = (\ \ \)" by simp from fin have "\ \ chi (stval (\ 0))" by (simp add: chi1 characteristic_holds) - with suffix have "\ \ always (chi (stval (\ 0)))" (is "_ \ ?alw") by simp + with suffix have "\ \ G\<^sub>p (chi (stval (\ 0)))" (is "_ \ ?alw") by simp moreover have "?alw \ ?ALW" by (auto simp: stval_def) ultimately have 2: "\ \ OR ?ALW" - using fin by (auto simp: finite_subset simp del: holds_of_always) + using fin by (auto simp: finite_subset simp del: pltl_semantics_sugar) have 3: "\(\ \ OR ?UNT)" proof assume unt: "\ \ OR ?UNT" with fin_UNT obtain val val' k where val: "val \ A" "val' \ A" "val' \ val" and now: "\ \ chi val" and k: "\[k..] \ chi val'" by auto (blast+)? (* FRAGILE: similar as above *) from `val \ A` now have "val = stval (\ 0)" by (rule chi2) moreover from `val' \ A` k suffix have "val' = stval (\ 0)" by (simp add: chi2) moreover note `val' \ val` ultimately show "False" by simp qed from 1 2 3 show ?thesis by (simp add: psi'_def) next case False txt {* Otherwise, @{text "\ \ next \"} is equivalent to @{text "\"} satisfying the second disjunct of @{text "psi'"}. We show both implications separately. *} let ?val = "stval (\ 0)" let ?val' = "stval (\ 1)" from False can have vals: "?val' \ ?val" by (auto simp: canonical_sequence_def state_sim_def stval_def) show ?thesis proof assume phi: "\ \ next \" from fin have 1: "\ \ chi ?val" by (simp add: chi1 characteristic_holds) from st can have "stutter_free (\[1..])" "canonical_sequence A (\[1..])" by (auto simp: stutter_free_suffix canonical_suffix) with phi psi have 2: "\[1..] \ \" by auto from fin have "\[1..] \ characteristic_formula A ((\[1..]) 0)" by (rule characteristic_holds) hence 3: "\[1..] \ chi ?val'" by (simp add: chi1) - from 1 2 3 have "\ \ and (chi ?val) (until (chi ?val) (and \ (chi ?val')))" + from 1 2 3 have "\ \ LTLpAnd (chi ?val) (until (chi ?val) (LTLpAnd \ (chi ?val')))" (is "_ \ ?unt") by auto moreover from vals have "?unt \ ?UNT" by (auto simp: stval_def) ultimately have "\ \ OR ?UNT" using fin_UNT[THEN holds_of_OR] by blast thus "\ \ psi'" by (simp add: psi'_def) next assume psi': "\ \ psi'" have "\(\ \ OR ?ALW)" proof assume "\ \ OR ?ALW" with fin obtain val where 1: "val \ A" and 2: "\n. \[n..] \ chi val" by (force simp: finite_subset) from 2 have "\[0..] \ chi val" .. with 1 have "val = ?val" by (simp add: chi2) moreover from 2 have "\[1..] \ chi val" .. with 1 have "val = ?val'" by (force dest: chi2) ultimately show "False" using vals by simp qed with psi' have "\ \ OR ?UNT" by (simp add: psi'_def) with fin_UNT obtain val val' k where val: "val \ A" "val' \ A" "val' \ val" and now: "\ \ chi val" and k: "\[k..] \ \" "\[k..] \ chi val'" and i: "\i[i..] \ chi val" by auto (blast+)? (* FRAGILE: similar as above *) from val now have 1: "val = ?val" by (simp add: chi2) have 2: "k \ 0" proof assume "k=0" with val k have "val' = ?val" by (simp add: chi2) with 1 `val' \ val` show "False" by simp qed have 3: "k \ 1" proof (rule ccontr) assume "\(k \ 1)" with i have "\[1..] \ chi val" by simp with 1 have "\[1..] \ characteristic_formula A (\ 0)" by (simp add: chi1) hence "(\ 0) ~A~ ((\[1..]) 0)" using characteristic_state_sim[OF fin] by blast with can have "\ 0 = \ 1" by (simp add: canonical_sequence_def) with `\ (Suc 0) \ \ 0` show "False" by simp qed from 2 3 have "k=1" by simp moreover from st can have "stutter_free (\[1..])" "canonical_sequence A (\[1..])" by (auto simp: stutter_free_suffix canonical_suffix) ultimately show "\ \ next \" using `\[k..] \ \` psi by auto qed qed } with nf atoms show "\\'. ?P (next \) \'" by blast qed text {* Comparing the definition of the next-free formula in the case of formulas @{text "next \"} with the one that appears in~\cite{peled:ltl-x}, there is a subtle difference. Peled and Wilke define the second disjunct as a disjunction of formulas % \begin{center}\( @{text "until (chi val) (and \ (chi val'))"} \)\end{center} % for subsets @{text "val, val' \ A"} whereas we conjoin the formula @{text "chi val"} to the ``until'' formula. This conjunct is indeed necessary in order to rule out the case of the ``until'' formula being true because of @{text "chi val'"} being true immediately. The subtle error in the definition of the formula was acknowledged by Peled and Wilke and apparently had not been noticed since the publication of~\cite{peled:ltl-x} in 1996 (which has been cited more than a hundred times according to Google Scholar). Although the error was corrected easily, the fact that authors, reviewers, and readers appear to have missed it for so long underscores the usefulness of formal proofs. We now show that any stuttering invariant PLTL formula can be expressed without the @{text "next"} operator. *} theorem stutter_invariant_next_free: assumes phi: "stutter_invariant \" obtains \ where "next_free \" "atoms \ \ atoms \" "\\. (\ \ \) = (\ \ \)" proof - have "atoms \ \ atoms \" "finite (atoms \)" by simp_all then obtain \ where psi: "next_free \" "atoms \ \ atoms \" and equiv: "\\. stutter_free \ \ canonical_sequence (atoms \) \ \ (\ \ \) = (\ \ \)" by (blast dest: ex_next_free_stutter_free_canonical) from `next_free \` have sinv: "stutter_invariant \" by (rule next_free_stutter_invariant) { fix \ obtain \ where 1: "\ \ atoms \ \ \" and 2: "canonical_sequence (atoms \) \" by (rule canonical_exists) from 1 `atoms \ \ atoms \` have 3: "\ \ atoms \ \ \" by (rule seq_sim_mono) from 1 have "(\ \ \) = (\ \ \)" by (simp add: pltl_seq_sim) also from phi stutter_reduced_equivalent have "... = (\\ \ \)" by auto also from 2[THEN canonical_reduced] equiv stutter_reduced_stutter_free have "... = (\\ \ \)" by auto also from sinv stutter_reduced_equivalent have "... = (\ \ \)" by auto also from 3 have "... = (\ \ \)" by (simp add: pltl_seq_sim) finally have "(\ \ \) = (\ \ \)" by (rule sym) } with psi that show ?thesis by blast qed text {* Combining theorems @{text "next_free_stutter_invariant"} and @{text "stutter_invariant_next_free"}, it follows that a PLTL formula is stuttering invariant iff it is equivalent to a next-free formula. *} theorem pltl_stutter_invariant: "stutter_invariant \ \ (\\. next_free \ \ atoms \ \ atoms \ \ (\\. \ \ \ \ \ \ \))" proof - { assume "stutter_invariant \" hence "\\. next_free \ \ atoms \ \ atoms \ \ (\\. \ \ \ \ \ \ \)" by (rule stutter_invariant_next_free) blast } moreover { fix \ assume 1: "next_free \" and 2: "\\. \ \ \ \ \ \ \" from 1 have "stutter_invariant \" by (rule next_free_stutter_invariant) with 2 have "stutter_invariant \" by blast } ultimately show ?thesis by blast qed end diff --git a/thys/Stuttering_Equivalence/Samplers.thy b/thys/Stuttering_Equivalence/Samplers.thy --- a/thys/Stuttering_Equivalence/Samplers.thy +++ b/thys/Stuttering_Equivalence/Samplers.thy @@ -1,587 +1,567 @@ theory Samplers -imports Main + imports Main "~~/src/HOL/Library/Omega_Words_Fun" begin section {* Utility Lemmas *} text {* The following lemmas about strictly monotonic functions could go to the standard library of Isabelle/HOL. *} text {* Strongly monotonic functions over the integers grow without bound. *} lemma strict_mono_exceeds: assumes f: "strict_mono (f::nat \ nat)" shows "\k. n < f k" proof (induct n) from f have "f 0 < f 1" by (rule strict_monoD) simp hence "0 < f 1" by simp thus "\k. 0 < f k" .. next fix n assume "\k. n < f k" then obtain k where "n < f k" .. hence "Suc n \ f k" by simp also from f have "f k < f (Suc k)" by (rule strict_monoD) simp finally show "\k. Suc n < f k" .. qed text {* More precisely, any natural number @{text "n \ f 0"} lies in the interval between @{text "f k"} and @{text "f (Suc k)"}, for some @{text k}. *} lemma strict_mono_interval: assumes f: "strict_mono (f::nat \ nat)" and n: "f 0 \ n" obtains k where "f k \ n" and "n < f (Suc k)" proof - from f[THEN strict_mono_exceeds] obtain m where m: "n < f m" .. have "m \ 0" proof assume "m = 0" with m n show "False" by simp qed with m obtain m' where m': "n < f (Suc m')" by (auto simp: gr0_conv_Suc) let ?k = "LEAST k. n < f (Suc k)" from m' have 1: "n < f (Suc ?k)" by (rule LeastI) have "f ?k \ n" proof (rule ccontr) assume "\ ?thesis" hence k: "n < f ?k" by simp show "False" proof (cases "?k") case 0 with k n show "False" by simp next case Suc with k show "False" by (auto dest: Least_le) qed qed with 1 that show ?thesis by simp qed lemma strict_mono_comp: assumes g: "strict_mono (g::'a::order \ 'b::order)" and f: "strict_mono (f::'b::order \ 'c::order)" shows "strict_mono (f \ g)" using assms by (auto simp: strict_mono_def) - -text {* - We represent @{text \}-words as functions of type @{text "nat \ 'a"}. - Suffixes of @{text \}-words are simply obtained by index shifting, - and we introduce a convenient notation. -*} - -definition suffix :: "(nat \ 'a) \ nat \ (nat \ 'a)" ("_[_ ..]" [80, 15] 80) where - "\[n..] \ \i. \ (n + i)" - -lemma suffix_elt [simp]: "(\[n..]) i = \ (n+i)" - by (simp add: suffix_def) - -lemma suffix_0 [simp]: "\[0..] = \" - by (simp add: suffix_def) - -lemma suffix_suffix [simp]: "\[n..][m..] = \[n+m ..]" - by (simp add: suffix_def add.assoc) - - section {* Stuttering Sampling Functions *} text {* Given an @{text \}-sequence @{text "\"}, a stuttering sampling function is a strictly monotonic function @{text "f::nat \ nat"} such that @{text "f 0 = 0"} and for all @{text i} and all @{text "f i \ k < f (i+1)"}, the elements @{text "\ k"} are the same. In other words, @{text f} skips some (but not necessarily all) stuttering steps, but never skips a non-stuttering step. Given such @{text "\"} and @{text f}, the (stuttering-)sampled reduction of @{text "\"} is the sequence of elements of @{text "\"} at the indices @{text "f i"}, which can simply be written as @{text "\ \ f"}. *} subsection {* Definition and elementary properties *} definition stutter_sampler where -- {* f is a stuttering sampling function for @{text "\"} *} "stutter_sampler (f::nat \ nat) \ \ f 0 = 0 \ strict_mono f \ (\k n. f k < n \ n < f (Suc k) \ \ n = \ (f k))" lemma stutter_sampler_0: "stutter_sampler f \ \ f 0 = 0" by (simp add: stutter_sampler_def) lemma stutter_sampler_mono: "stutter_sampler f \ \ strict_mono f" by (simp add: stutter_sampler_def) lemma stutter_sampler_between: assumes f: "stutter_sampler f \" and lo: "f k \ n" and hi: "n < f (Suc k)" shows "\ n = \ (f k)" using assms by (auto simp: stutter_sampler_def less_le) lemma stutter_sampler_interval: assumes f: "stutter_sampler f \" obtains k where "f k \ n" and "n < f (Suc k)" using f[THEN stutter_sampler_mono] proof (rule strict_mono_interval) from f show "f 0 \ n" by (simp add: stutter_sampler_0) qed text {* The identity function is a stuttering sampling function for any @{text "\"}. *} lemma id_stutter_sampler [iff]: "stutter_sampler id \" by (auto simp: stutter_sampler_def strict_mono_def) text {* Stuttering sampling functions compose, sort of. *} lemma stutter_sampler_comp: assumes f: "stutter_sampler f \" and g: "stutter_sampler g (\ \ f)" shows "stutter_sampler (f \ g) \" proof (auto simp: stutter_sampler_def) from f g show "f (g 0) = 0" by (simp add: stutter_sampler_0) next from g[THEN stutter_sampler_mono] f[THEN stutter_sampler_mono] show "strict_mono (f \ g)" by (rule strict_mono_comp) next fix i k assume lo: "f (g i) < k" and hi: "k < f (g (Suc i))" from f obtain m where 1: "f m \ k" and 2: "k < f (Suc m)" by (rule stutter_sampler_interval) with f have 3: "\ k = \ (f m)" by (rule stutter_sampler_between) from lo 2 have "f (g i) < f (Suc m)" by simp with f[THEN stutter_sampler_mono] have 4: "g i \ m" by (simp add: strict_mono_less) from 1 hi have "f m < f (g (Suc i))" by simp with f[THEN stutter_sampler_mono] have 5: "m < g (Suc i)"by (simp add: strict_mono_less) from g 4 5 have "(\ \ f) m = (\ \ f) (g i)" by (rule stutter_sampler_between) with 3 show "\ k = \ (f (g i))" by simp qed text {* Stuttering sampling functions can be extended to suffixes. *} lemma stutter_sampler_suffix: assumes f: "stutter_sampler f \" - shows "stutter_sampler (\k. f (n+k) - f n) (\[f n ..])" + shows "stutter_sampler (\k. f (n+k) - f n) (suffix (f n) \)" proof (auto simp: stutter_sampler_def strict_mono_def) fix i j assume ij: "(i::nat) < j" from f have mono: "strict_mono f" by (rule stutter_sampler_mono) from mono[THEN strict_mono_mono] have "f n \ f (n+i)" by (rule monoD) simp moreover from mono[THEN strict_mono_mono] have "f n \ f (n+j)" by (rule monoD) simp moreover from mono ij have "f (n+i) < f (n+j)" by (auto intro: strict_monoD) ultimately show "f (n+i) - f n < f (n+j) - f n" by simp next fix i k assume lo: "f (n+i) - f n < k" and hi: "k < f (Suc (n+i)) - f n" from lo have "f (n+i) \ f n + k" by simp moreover from hi have "f n + k < f (Suc (n + i))" by simp moreover from f[THEN stutter_sampler_mono, THEN strict_mono_mono] have "f n \ f (n+i)" by (rule monoD) simp ultimately show "\ (f n + k) = \ (f n + (f (n+i) - f n))" by (auto dest: stutter_sampler_between[OF f]) qed subsection {* Preservation of properties through stuttering sampling *} text {* Stuttering sampling preserves the initial element of the sequence, as well as the presence and relative ordering of different elements. *} lemma stutter_sampled_0: assumes "stutter_sampler f \" shows "\ (f 0) = \ 0" using assms[THEN stutter_sampler_0] by simp lemma stutter_sampled_in_range: assumes f: "stutter_sampler f \" and s: "s \ range \" shows "s \ range (\ \ f)" proof - from s obtain n where n: "\ n = s" by auto from f obtain k where "f k \ n" "n < f (Suc k)" by (rule stutter_sampler_interval) with f have "\ n = \ (f k)" by (rule stutter_sampler_between) with n show ?thesis by auto qed lemma stutter_sampled_range: assumes "stutter_sampler f \" shows "range (\ \ f) = range \" using assms by (auto intro: stutter_sampled_in_range) lemma stutter_sampled_precedence: assumes f: "stutter_sampler f \" and ij: "i \ j" obtains k l where "k \ l" "\ (f k) = \ i" "\ (f l) = \ j" proof - from f obtain k where k: "f k \ i" "i < f (Suc k)" by (rule stutter_sampler_interval) with f have 1: "\ i = \ (f k)" by (rule stutter_sampler_between) from f obtain l where l: "f l \ j" "j < f (Suc l)" by (rule stutter_sampler_interval) with f have 2: "\ j = \ (f l)" by (rule stutter_sampler_between) from k l ij have "f k < f (Suc l)" by simp with f[THEN stutter_sampler_mono] have "k \ l" by (simp add: strict_mono_less) with 1 2 that show ?thesis by simp qed subsection {* Maximal stuttering sampling *} text {* We define a particular sampling function that is maximal in the sense that it eliminates all finite stuttering. If a sequence ends with infinite stuttering then it behaves as the identity over the (maximal such) suffix. *} fun max_stutter_sampler where "max_stutter_sampler \ 0 = 0" | "max_stutter_sampler \ (Suc n) = (let prev = max_stutter_sampler \ n in if (\k > prev. \ k = \ prev) then Suc prev else (LEAST k. prev < k \ \ k \ \ prev))" text {* @{text max_stutter_sampler} is indeed a stuttering sampling function. *} lemma max_stutter_sampler: "stutter_sampler (max_stutter_sampler \) \" (is "stutter_sampler ?ms _") proof - have "?ms 0 = 0" by simp moreover have "\n. ?ms n < ?ms (Suc n)" proof fix n show "?ms n < ?ms (Suc n)" (is "?prev < ?next") proof (cases "\k > ?prev. \ k = \ ?prev") case True thus ?thesis by (simp add: Let_def) next case False hence "\k. ?prev < k \ \ k \ \ ?prev" by simp from this[THEN LeastI_ex] have "?prev < (LEAST k. ?prev < k \ \ k \ \ ?prev)" .. with False show ?thesis by (simp add: Let_def) qed qed hence "strict_mono ?ms" unfolding strict_mono_def by (blast intro: lift_Suc_mono_less) moreover have "\n k. ?ms n < k \ k < ?ms (Suc n) \ \ k = \ (?ms n)" proof (clarify) fix n k assume lo: "?ms n < k" (is "?prev < k") and hi: "k < ?ms (Suc n)" (is "k < ?next") show "\ k = \ ?prev" proof (cases "\k > ?prev. \ k = \ ?prev") case True hence "?next = Suc ?prev" by (simp add: Let_def) with lo hi show ?thesis by simp -- {* no room for intermediate index *} next case False hence "?next = (LEAST k. ?prev < k \ \ k \ \ ?prev)" by (auto simp add: Let_def) with lo hi show ?thesis by (auto dest: not_less_Least) qed qed ultimately show ?thesis unfolding stutter_sampler_def by blast qed text {* We write @{text "\\"} for the sequence @{text "\"} sampled by the maximal stuttering sampler. Also, a sequence is \emph{stutter free} if it contains no finite stuttering: whenever two subsequent elements are equal then all subsequent elements are the same. *} definition stutter_reduced ("\_" [100] 100) where "\\ = \ \ (max_stutter_sampler \)" definition stutter_free where "stutter_free \ \ \k. \ (Suc k) = \ k \ (\n>k. \ n = \ k)" lemma stutter_freeI: assumes "\k n. \\ (Suc k) = \ k; n>k\ \ \ n = \ k" shows "stutter_free \" using assms unfolding stutter_free_def by blast lemma stutter_freeD: assumes "stutter_free \" and "\ (Suc k) = \ k" and "n>k" shows "\ n = \ k" using assms unfolding stutter_free_def by blast text {* Any suffix of a stutter free sequence is itself stutter free. *} lemma stutter_free_suffix: assumes sigma: "stutter_free \" - shows "stutter_free (\[k..])" + shows "stutter_free (suffix k \)" proof (rule stutter_freeI) fix j n - assume j: "(\[k..]) (Suc j) = (\[k..]) j" and n: "j < n" + assume j: "(suffix k \) (Suc j) = (suffix k \) j" and n: "j < n" from j have "\ (Suc (k+j)) = \ (k+j)" by simp moreover from n have "k+n > k+j" by simp ultimately have "\ (k+n) = \ (k+j)" by (rule stutter_freeD[OF sigma]) - thus "(\[k..]) n = (\[k..]) j" by simp + thus "(suffix k \) n = (suffix k \) j" by simp qed lemma stutter_reduced_0: "(\\) 0 = \ 0" by (simp add: stutter_reduced_def stutter_sampled_0 max_stutter_sampler) lemma stutter_free_reduced: assumes sigma: "stutter_free \" shows "\\ = \" proof - { fix n have "max_stutter_sampler \ n = n" (is "?ms n = n") proof (induct n) show "?ms 0 = 0" by simp next fix n assume ih: "?ms n = n" show "?ms (Suc n) = Suc n" proof (cases "\ (Suc n) = \ (?ms n)") case True with ih have "\ (Suc n) = \ n" by simp with sigma have "\k > n. \ k = \ n" unfolding stutter_free_def by blast with ih show ?thesis by (simp add: Let_def) next case False with ih have "(LEAST k. k>n \ \ k \ \ (?ms n)) = Suc n" by (auto intro: Least_equality) with ih False show ?thesis by (simp add: Let_def) qed qed } thus ?thesis by (auto simp: stutter_reduced_def) qed text {* Whenever two sequence elements at two consecutive sampling points of the maximal stuttering sampler are equal then the sequence stutters infinitely from the first sampling point onwards. In particular, @{text "\\"} is stutter free. *} lemma max_stutter_sampler_nostuttering: assumes stut: "\ (max_stutter_sampler \ (Suc k)) = \ (max_stutter_sampler \ k)" and n: "n > max_stutter_sampler \ k" (is "_ > ?ms k") shows "\ n = \ (?ms k)" proof (rule ccontr) assume contr: "\ ?thesis" with n have "?ms k < n \ \ n \ \ (?ms k)" (is "?diff n") .. hence "?diff (LEAST n. ?diff n)" by (rule LeastI) with contr have "\ (?ms (Suc k)) \ \ (?ms k)" by (auto simp add: Let_def) from this stut show "False" .. qed lemma stutter_reduced_stutter_free: "stutter_free (\\)" proof (rule stutter_freeI) fix k n assume k: "(\\) (Suc k) = (\\) k" and n: "k < n" from n have "max_stutter_sampler \ k < max_stutter_sampler \ n" using max_stutter_sampler[THEN stutter_sampler_mono, THEN strict_monoD] by blast with k show "(\\) n = (\\) k" unfolding stutter_reduced_def by (auto elim: max_stutter_sampler_nostuttering simp del: max_stutter_sampler.simps) qed -lemma stutter_reduced_suffix: "\((\\)[k..]) = (\\)[k..]" +lemma stutter_reduced_suffix: "\ (suffix k (\\)) = suffix k (\\)" proof (rule stutter_free_reduced) have "stutter_free (\\)" by (rule stutter_reduced_stutter_free) - thus "stutter_free ((\\)[k..])" by (rule stutter_free_suffix) + thus "stutter_free (suffix k (\\))" by (rule stutter_free_suffix) qed lemma stutter_reduced_reduced: "\\\ = \\" - by (rule stutter_reduced_suffix[of "\" "0", simplified]) - + by (insert stutter_reduced_suffix[of 0 "\", simplified]) + text {* One can define a partial order on sampling functions for a given sequence @{text "\"} by saying that function @{text g} is better than function @{text f} if the reduced sequence induced by @{text f} can be further reduced to obtain the reduced sequence corresponding to @{text g}, i.e. if there exists a stuttering sampling function @{text h} for the reduced sequence @{text "\ \ f"} such that @{text "\ \ f \ h = \ \ g"}. (Note that @{text "f \ h"} is indeed a stuttering sampling function for @{text "\"}, by theorem @{text "stutter_sampler_comp"}.) We do not formalize this notion but prove that @{text "max_stutter_sampler \"} is the best sampling function according to this order. *} theorem sample_max_sample: assumes f: "stutter_sampler f \" shows "\(\ \ f) = \\" proof - let ?mss = "max_stutter_sampler \" let ?mssf = "max_stutter_sampler (\ \ f)" from f have mssf: "stutter_sampler (f \ ?mssf) \" by (blast intro: stutter_sampler_comp max_stutter_sampler) txt {* The following is the core invariant of the proof: the sampling functions @{text "max_stutter_sampler \"} and @{text "f \ (max_stutter_sampler (\ \ f))"} work in lock-step (i.e., sample the same points), except if @{text "\"} ends in infinite stuttering, at which point function @{text f} may make larger steps than the maximal sampling functions. *} { fix k have " ?mss k = f (?mssf k) \ ?mss k \ f (?mssf k) \ (\n \ ?mss k. \ (?mss k) = \ n)" (is "?P k" is "?A k \ ?B k") proof (induct k) from f mssf have "?mss 0 = f (?mssf 0)" by (simp add: max_stutter_sampler stutter_sampler_0) thus "?P 0" .. next fix k assume ih: "?P k" have b: "?B k \ ?B (Suc k)" proof assume 0: "?B k" hence 1: "?mss k \ f (?mssf k)" .. (* NB: For some reason "... hence 1: ... and 2: ..." cannot be proved *) from 0 have 2: "\n \ ?mss k. \ (?mss k) = \ n" .. hence "\n > ?mss k. \ (?mss k) = \ n" by auto hence "\n > ?mss k. \ n = \ (?mss k)" by auto hence 3: "?mss (Suc k) = Suc (?mss k)" by (simp add: Let_def) with 2 have "\ (?mss k) = \ (?mss (Suc k))" by (auto simp del: max_stutter_sampler.simps) from sym[OF this] 2 3 have "\n \ ?mss (Suc k). \ (?mss (Suc k)) = \ n" by (auto simp del: max_stutter_sampler.simps) moreover from mssf[THEN stutter_sampler_mono, THEN strict_monoD] have "f (?mssf k) < f (?mssf (Suc k))" by (simp del: max_stutter_sampler.simps) with 1 3 have "?mss (Suc k) \ f (?mssf (Suc k))" by (simp del: max_stutter_sampler.simps) ultimately show "?B (Suc k)" by blast qed from ih show "?P (Suc k)" proof assume a: "?A k" show ?thesis proof (cases "\n > ?mss k. \ n = \ (?mss k)") case True hence "\n \ ?mss k. \ (?mss k) = \ n" by (auto simp: le_less) with a have "?B k" by simp with b show ?thesis by (simp del: max_stutter_sampler.simps) next case False hence diff: "\ (?mss (Suc k)) \ \ (?mss k)" by (blast dest: max_stutter_sampler_nostuttering) have "?A (Suc k)" proof (rule antisym) show "f (?mssf (Suc k)) \ ?mss (Suc k)" proof (rule ccontr) assume "\ ?thesis" hence contr: "?mss (Suc k) < f (?mssf (Suc k))" by simp from mssf have "\ (?mss (Suc k)) = \ ((f \ ?mssf) k)" proof (rule stutter_sampler_between) from max_stutter_sampler[of "\", THEN stutter_sampler_mono] have "?mss k < ?mss (Suc k)" by (rule strict_monoD) simp with a show "(f \ ?mssf) k \ ?mss (Suc k)" by (simp add: o_def del: max_stutter_sampler.simps) next from contr show "?mss (Suc k) < (f \ ?mssf) (Suc k)" by simp qed with a have "\ (?mss (Suc k)) = \ (?mss k)" by (simp add: o_def del: max_stutter_sampler.simps) with diff show "False" .. qed next have "\m > ?mssf k. f m = ?mss (Suc k)" proof (rule ccontr) assume "\ ?thesis" hence contr: "\i. f ((?mssf k) + Suc i) \ ?mss (Suc k)" by simp { fix i have "f (?mssf k + i) < ?mss (Suc k)" (is "?F i") proof (induct i) from a have "f (?mssf k + 0) = ?mss k" by (simp add: o_def) also from max_stutter_sampler[of "\", THEN stutter_sampler_mono] have "... < ?mss (Suc k)" by (rule strict_monoD) simp finally show "?F 0" . next fix i assume ih: "?F i" show "?F (Suc i)" proof (rule ccontr) assume "\ ?thesis" then have "?mss (Suc k) \ f (?mssf k + Suc i)" by (simp add: o_def) moreover from contr have "f (?mssf k + Suc i) \ ?mss (Suc k)" by blast ultimately have i: "?mss (Suc k) < f (?mssf k + Suc i)" by (simp add: less_le) from f have "\ (?mss (Suc k)) = \ (f (?mssf k + i))" proof (rule stutter_sampler_between) from ih show "f (?mssf k + i) \ ?mss (Suc k)" by (simp add: o_def) next from i show "?mss (Suc k) < f (Suc (?mssf k + i))" by simp qed also from max_stutter_sampler have "... = \ (?mss k)" proof (rule stutter_sampler_between) from f[THEN stutter_sampler_mono, THEN strict_mono_mono] have "f (?mssf k) \ f (?mssf k + i)" by (rule monoD) simp with a show "?mss k \ f (?mssf k + i)" by (simp add: o_def) qed (rule ih) also note diff finally show "False" by simp qed qed } note bounded = this from f[THEN stutter_sampler_mono] have "strict_mono (\i. f (?mssf k + i))" by (auto simp: strict_mono_def) then obtain i where i: "?mss (Suc k) < f (?mssf k + i)" by (blast dest: strict_mono_exceeds) from bounded have "f (?mssf k + i) < ?mss (Suc k)" . with i show "False" by (simp del: max_stutter_sampler.simps) qed then obtain m where m: "m > ?mssf k" and m': "f m = ?mss (Suc k)" by blast show "?mss (Suc k) \ f (?mssf (Suc k))" proof (rule ccontr) assume "\ ?thesis" hence contr: "f (?mssf (Suc k)) < ?mss (Suc k)" by simp from mssf[THEN stutter_sampler_mono] have "(f \ ?mssf) k < (f \ ?mssf) (Suc k)" by (rule strict_monoD) simp with a have "?mss k \ f (?mssf (Suc k))" by (simp add: o_def) from this contr have "\ (f (?mssf (Suc k))) = \ (?mss k)" by (rule stutter_sampler_between[OF max_stutter_sampler]) with a have stut: "(\ \ f) (?mssf (Suc k)) = (\ \ f) (?mssf k)" by (simp add: o_def) from this m have "(\ \ f) m = (\ \ f) (?mssf k)" by (blast intro: max_stutter_sampler_nostuttering) with diff m' a show "False" by (simp add: o_def) qed qed thus ?thesis .. qed next assume "?B k" with b show ?thesis by (simp del: max_stutter_sampler.simps) qed qed } hence "\\ = \(\ \ f)" unfolding stutter_reduced_def by force thus ?thesis by (rule sym) qed end (* theory Samplers *) diff --git a/thys/Stuttering_Equivalence/StutterEquivalence.thy b/thys/Stuttering_Equivalence/StutterEquivalence.thy --- a/thys/Stuttering_Equivalence/StutterEquivalence.thy +++ b/thys/Stuttering_Equivalence/StutterEquivalence.thy @@ -1,251 +1,255 @@ theory StutterEquivalence imports Samplers begin section {* Stuttering Equivalence *} text {* Stuttering equivalence of two sequences is formally defined as the equality of their maximally reduced versions. *} definition stutter_equiv (infix "\" 50) where "\ \ \ \ \\ = \\" text {* Stuttering equivalence is an equivalence relation. *} lemma stutter_equiv_refl: "\ \ \" unfolding stutter_equiv_def .. lemma stutter_equiv_sym [sym]: "\ \ \ \ \ \ \" unfolding stutter_equiv_def by (rule sym) lemma stutter_equiv_trans [trans]: "\ \ \ \ \ \ \ \ \ \ \" unfolding stutter_equiv_def by simp text {* In particular, any sequence sampled by a stuttering sampler is stuttering equivalent to the original one. *} lemma sampled_stutter_equiv: assumes "stutter_sampler f \" shows "\ \ f \ \" using assms unfolding stutter_equiv_def by (rule sample_max_sample) lemma stutter_reduced_equivalent: "\\ \ \" unfolding stutter_equiv_def by (rule stutter_reduced_reduced) text {* For proving stuttering equivalence of two sequences, it is enough to exhibit two arbitrary sampling functions that equalize the reductions of the sequences. This can be more convenient than computing the maximal stutter-reduced version of the sequences. *} lemma stutter_equivI: assumes f: "stutter_sampler f \" and g: "stutter_sampler g \" and eq: "\ \ f = \ \ g" shows "\ \ \" proof - from f have "\\ = \(\ \ f)" by (rule sample_max_sample[THEN sym]) also from eq have "... = \(\ \ g)" by simp also from g have "... = \\" by (rule sample_max_sample) finally show ?thesis by (unfold stutter_equiv_def) qed text {* The corresponding elimination rule is easy to prove, given that the maximal stuttering sampling function is a stuttering sampling function. *} lemma stutter_equivE: assumes eq: "\ \ \" and p: "\f g. \ stutter_sampler f \; stutter_sampler g \; \ \ f = \ \ g \ \ P" shows "P" proof (rule p) from eq show "\ \ (max_stutter_sampler \) = \ \ (max_stutter_sampler \)" by (unfold stutter_equiv_def stutter_reduced_def) qed (rule max_stutter_sampler)+ text {* Therefore we get the following alternative characterization: two sequences are stuttering equivalent iff there are stuttering sampling functions that equalize the two sequences. *} lemma stutter_equiv_eq: "\ \ \ = (\f g. stutter_sampler f \ \ stutter_sampler g \ \ \ \ f = \ \ g)" by (blast intro: stutter_equivI elim: stutter_equivE) text {* The initial elements of stutter equivalent sequences are equal. *} lemma stutter_equiv_0: assumes "\ \ \" shows "\ 0 = \ 0" proof - have "\ 0 = (\\) 0" by (rule stutter_reduced_0[THEN sym]) with assms[unfolded stutter_equiv_def] show ?thesis by (simp add: stutter_reduced_0) qed +abbreviation suffix_notation ("_ [_..]") +where + "suffix_notation w k \ suffix k w" + text {* Given any stuttering sampling function @{text f} for sequence @{text "\"}, any suffix of @{text "\"} starting at index @{text "f n"} is stuttering equivalent to the suffix of the stutter-reduced version of @{text "\"} starting at @{text "n"}. *} lemma suffix_stutter_equiv: assumes f: "stutter_sampler f \" - shows "\[f n ..] \ (\ \ f)[n ..]" + shows "suffix (f n) \ \ suffix n (\ \ f)" proof - from f have "stutter_sampler (\k. f (n+k) - f n) (\[f n ..])" by (rule stutter_sampler_suffix) moreover have "stutter_sampler id ((\ \ f)[n ..])" by (rule id_stutter_sampler) moreover have "(\[f n ..]) \ (\k. f (n+k) - f n) = ((\ \ f)[n ..]) \ id" proof (rule ext, auto) fix i from f[THEN stutter_sampler_mono, THEN strict_mono_mono] have "f n \ f (n+i)" by (rule monoD) simp thus "\ (f n + (f (n+i) - f n)) = \ (f (n+i))" by simp qed ultimately show ?thesis by (rule stutter_equivI) qed text {* Given a stuttering sampling function @{text f} and a point @{text "n"} within the interval from @{text "f k"} to @{text "f (k+1)"}, the suffix starting at @{text n} is stuttering equivalent to the suffix starting at @{text "f k"}. *} lemma stutter_equiv_within_interval: assumes f: "stutter_sampler f \" and lo: "f k \ n" and hi: "n < f (Suc k)" shows "\[n ..] \ \[f k ..]" proof - have "stutter_sampler id (\[n ..])" by (rule id_stutter_sampler) moreover from lo have "stutter_sampler (\i. if i=0 then 0 else n + i - f k) (\[f k ..])" (is "stutter_sampler ?f _") proof (auto simp: stutter_sampler_def strict_mono_def) fix i assume i: "i < Suc n - f k" from f show "\ (f k + i) = \ (f k)" proof (rule stutter_sampler_between) from i hi show "f k + i < f (Suc k)" by simp qed simp qed moreover have "(\[n ..]) \ id = (\[f k ..]) \ ?f" proof (rule ext, auto) from f lo hi show "\ n = \ (f k)" by (rule stutter_sampler_between) next fix i from lo show "\ (n+i) = \ (f k + (n + i - f k))" by simp qed ultimately show ?thesis by (rule stutter_equivI) qed text {* Given two stuttering equivalent sequences @{text "\"} and @{text "\"}, we obtain a zig-zag relationship as follows: for any suffix @{text "\[n..]"} there is a suffix @{text "\[m..]"} such that: \begin{enumerate} \item @{text "\[m..] \ \[n..]"} and \item for every suffix @{text "\[j..]"} where @{text "j[k..]"} for some @{text "k \ \" obtains m where "\[m..] \ \[n..]" and "\jk[j..] \ \[k..]" using assms proof (rule stutter_equivE) fix f g assume f: "stutter_sampler f \" and g: "stutter_sampler g \" and eq: "\ \ f = \ \ g" from g obtain i where i: "g i \ n" "n < g (Suc i)" by (rule stutter_sampler_interval) with g have "\[n..] \ \[g i ..]" by (rule stutter_equiv_within_interval) also from g have "... \ (\ \ g)[i ..]" by (rule suffix_stutter_equiv) also from eq have "... = (\ \ f)[i ..]" by simp also from f have "... \ \[f i ..]" by (rule suffix_stutter_equiv[THEN stutter_equiv_sym]) finally have "\[f i ..] \ \[n ..]" by (rule stutter_equiv_sym) moreover { fix j assume j: "j < f i" from f obtain a where a: "f a \ j" "j < f (Suc a)" by (rule stutter_sampler_interval) from a j have "f a < f i" by simp with f[THEN stutter_sampler_mono] have "a < i" by (simp add: strict_mono_less) with g[THEN stutter_sampler_mono] have "g a < g i" by (simp add: strict_mono_less) with i have 1: "g a < n" by simp from f a have "\[j..] \ \[f a ..]" by (rule stutter_equiv_within_interval) also from f have "... \ (\ \ f)[a ..]" by (rule suffix_stutter_equiv) also from eq have "... = (\ \ g)[a ..]" by simp also from g have "... \ \[g a ..]" by (rule suffix_stutter_equiv[THEN stutter_equiv_sym]) finally have "\[j ..] \ \[g a ..]" . with 1 have "\k[j..] \ \[k ..]" by blast } moreover note that ultimately show ?thesis by blast qed theorem stutter_equiv_suffixes_right: assumes "\ \ \" obtains n where "\[m..] \ \[n..]" and "\jk[k..] \ \[j..]" proof - from assms have "\ \ \" by (rule stutter_equiv_sym) then obtain n where "\[n..] \ \[m..]" "\jk[j..] \ \[k..]" by (rule stutter_equiv_suffixes_left) with that show ?thesis by (blast dest: stutter_equiv_sym) qed text {* In particular, if @{text "\"} and @{text "\"} are stutter equivalent then every element that occurs in one sequence also occurs in the other. *} lemma stutter_equiv_element_left: assumes "\ \ \" obtains m where "\ m = \ n" and "\jk j = \ k" proof - from assms obtain m where "\[m..] \ \[n..]" "\jk[j..] \ \[k..]" by (rule stutter_equiv_suffixes_left) with that show ?thesis by (force dest: stutter_equiv_0) qed lemma stutter_equiv_element_right: assumes "\ \ \" obtains n where "\ m = \ n" and "\jk k = \ j" proof - from assms obtain n where "\[m..] \ \[n..]" "\jk[k..] \ \[j..]" by (rule stutter_equiv_suffixes_right) with that show ?thesis by (force dest: stutter_equiv_0) qed end (* theory StutterEquivalence *)