diff --git a/thys/Algebraic_VCs/Domain_Quantale.thy b/thys/Algebraic_VCs/Domain_Quantale.thy --- a/thys/Algebraic_VCs/Domain_Quantale.thy +++ b/thys/Algebraic_VCs/Domain_Quantale.thy @@ -1,524 +1,511 @@ (* Title: Domain Quantales Author: Victor Gomes, Georg Struth Maintainer: Victor Gomes Georg Struth *) section \Component for Recursive Programs\ theory Domain_Quantale imports KAD.Modal_Kleene_Algebra begin text \This component supports the verification and step-wise refinement of recursive programs in a partial correctness setting.\ notation times (infixl "\" 70) and bot ("\") and top ("\") and inf (infixl "\" 65) and sup (infixl "\" 65) -(* - Inf ("\_" [900] 900) and - Sup ("\_" [900] 900) -*) - -(* -syntax - "_INF1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) - "_INF" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) - "_SUP1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) - "_SUP" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) - -*) subsection \Lattice-Ordered Monoids with Domain\ class bd_lattice_ordered_monoid = bounded_lattice + distrib_lattice + monoid_mult + assumes left_distrib: "x \ (y \ z) = x \ y \ x \ z" and right_distrib: "(x \ y) \ z = x \ z \ y \ z" and bot_annil [simp]: "\ \ x = \" and bot_annir [simp]: "x \ \ = \" begin sublocale semiring_one_zero "(\)" "(\)" "1" "bot" by (standard, auto simp: sup.assoc sup.commute sup_left_commute left_distrib right_distrib sup_absorb1) sublocale dioid_one_zero "(\)" "(\)" "1" bot "(\)" "(<)" by (standard, simp add: le_iff_sup, auto) end no_notation ads_d ("d") and ars_r ("r") and antirange_op ("ar _" [999] 1000) class domain_bdlo_monoid = bd_lattice_ordered_monoid + assumes rdv: "(z \ x \ top) \ y = z \ y \ x \ top" begin definition "d x = 1 \ x \ \" sublocale ds: domain_semiring "(\)" "(\)" "1" "\" "d" "(\)" "(<)" proof standard fix x y show "x \ d x \ x = d x \ x" by (metis d_def inf_sup_absorb left_distrib mult_1_left mult_1_right rdv sup.absorb_iff1 sup.idem sup.left_commute top_greatest) show "d (x \ d y) = d (x \ y)" by (simp add: d_def inf_absorb2 rdv mult_assoc) show "d x \ 1 = 1" by (simp add: d_def sup.commute) show "d bot = bot" by (simp add: d_def inf.absorb1 inf.commute) show "d (x \ y) = d x \ d y" by (simp add: d_def inf_sup_distrib1) qed end subsection\Boolean Monoids with Domain\ class boolean_monoid = boolean_algebra + monoid_mult + assumes left_distrib': "x \ (y \ z) = x \ y \ x \ z" and right_distrib': "(x \ y) \ z = x \ z \ y \ z" and bot_annil' [simp]: "\ \ x = \" and bot_annir' [simp]: "x \ \ = \" begin subclass bd_lattice_ordered_monoid by (standard, simp_all add: left_distrib' right_distrib') lemma inf_bot_iff_le: "x \ y = \ \ x \ -y" by (metis le_iff_inf inf_sup_distrib1 inf_top_right sup_bot.left_neutral sup_compl_top compl_inf_bot inf.assoc inf_bot_right) end class domain_boolean_monoid = boolean_monoid + assumes rdv': "(z \ x \ \) \ y = z \ y \ x \ \" begin sublocale dblo: domain_bdlo_monoid "1" "(\)" "(\)" "(\)" "(<)" "(\)" "\" "\" by (standard, simp add: rdv') definition "a x = 1 \ -(dblo.d x)" lemma a_d_iff: "a x = 1 \ -(x \ \)" by (clarsimp simp: a_def dblo.d_def inf_sup_distrib1) lemma topr: "-(x \ \) \ \ = -(x \ \)" proof (rule order.antisym) show "-(x \ \) \ -(x \ \) \ \" by (metis mult_isol_var mult_oner order_refl top_greatest) have "-(x \ \) \ (x \ \) = \" by simp hence "(-(x \ \) \ (x \ \)) \ \ = \" by simp hence "-(x \ \) \ \ \ (x \ \) = \" by (metis rdv') thus "-(x \ \) \ \ \ -(x \ \)" by (simp add: inf_bot_iff_le) qed lemma dd_a: "dblo.d x = a (a x)" by (metis a_d_iff dblo.d_def double_compl inf_top.left_neutral mult_1_left rdv' topr) lemma ad_a [simp]: "a (dblo.d x) = a x" by (simp add: a_def) lemma da_a [simp]: "dblo.d (a x) = a x" using ad_a dd_a by auto lemma a1 [simp]: "a x \ x = \" proof - have "a x \ x \ \ = \" by (metis a_d_iff inf_compl_bot mult_1_left rdv' topr) then show ?thesis by (metis (no_types) dblo.d_def dblo.ds.domain_very_strict inf_bot_right) qed lemma a2 [simp]: "a (x \ y) \ a (x \ a (a y)) = a (x \ a (a y))" by (metis a_def dblo.ds.dsr2 dd_a sup.idem) lemma a3 [simp]: "a (a x) \ a x = 1" by (metis a_def da_a inf.commute sup.commute sup_compl_top sup_inf_absorb sup_inf_distrib1) subclass domain_bdlo_monoid .. text \The next statement shows that every boolean monoid with domain is an antidomain semiring. In this setting the domain operation has been defined explicitly.\ sublocale ad: antidomain_semiring a "(\)" "(\)" "1" "\" "(\)" "(<)" rewrites ad_eq: "ad.ads_d x = d x" proof - show "class.antidomain_semiring a (\) (\) 1 \ (\) (<)" by (standard, simp_all) then interpret ad: antidomain_semiring a "(\)" "(\)" "1" "\" "(\)" "(<)" . show "ad.ads_d x = d x" by (simp add: ad.ads_d_def dd_a) qed end subsection\Boolean Monoids with Range\ class range_boolean_monoid = boolean_monoid + assumes ldv': "y \ (z \ \ \ x) = y \ z \ \ \ x" begin definition "r x = 1 \ \ \ x" definition "ar x = 1 \ -(r x)" lemma ar_r_iff: "ar x = 1 \ -(\ \ x)" by (simp add: ar_def inf_sup_distrib1 r_def) lemma topl: "\\(-(\ \ x)) = -(\ \ x)" proof (rule order.antisym) show "\ \ - (\ \ x) \ - (\ \ x)" by (metis bot_annir' compl_inf_bot inf_bot_iff_le ldv') show "- (\ \ x) \ \ \ - (\ \ x)" by (metis inf_le2 inf_top.right_neutral mult_1_left mult_isor) qed lemma r_ar: "r x = ar (ar x)" by (metis ar_r_iff double_compl inf.commute inf_top.right_neutral ldv' mult_1_right r_def topl) lemma ar_ar [simp]: "ar (r x) = ar x" by (simp add: ar_def ldv' r_def) lemma rar_ar [simp]: "r (ar x) = ar x" using r_ar ar_ar by force lemma ar1 [simp]: "x \ ar x = \" proof - have "\ \ x \ ar x = \" by (metis ar_r_iff inf_compl_bot ldv' mult_oner topl) then show ?thesis by (metis inf_bot_iff_le inf_le2 inf_top.right_neutral mult_1_left mult_isor mult_oner topl) qed lemma ars: "r (r x \ y) = r (x \ y)" by (metis inf.commute inf_top.right_neutral ldv' mult_oner mult_assoc r_def) lemma ar2 [simp]: "ar (x \ y) \ ar (ar (ar x) \ y) = ar (ar (ar x) \ y)" by (metis ar_def ars r_ar sup.idem) lemma ar3 [simp]: "ar (ar x) \ ar x = 1 " by (metis ar_def rar_ar inf.commute sup.commute sup_compl_top sup_inf_absorb sup_inf_distrib1) sublocale ar: antirange_semiring "(\)" "(\)" "1" "\" ar "(\)" "(<)" rewrites ar_eq: "ar.ars_r x = r x" proof - show "class.antirange_semiring (\) (\) 1 \ ar (\) (<)" by (standard, simp_all) then interpret ar: antirange_semiring "(\)" "(\)" "1" "\" ar "(\)" "(<)" . show "ar.ars_r x = r x" by (simp add: ar.ars_r_def r_ar) qed end subsection \Quantales\ text \This part will eventually move into an AFP quantale entry.\ class quantale = complete_lattice + monoid_mult + assumes Sup_distr: "Sup X \ y = Sup {z. \x \ X. z = x \ y}" and Sup_distl: "x \ Sup Y = Sup {z. \y \ Y. z = x \ y}" begin lemma bot_annil'' [simp]: "\ \ x = \" using Sup_distr[where X="{}"] by auto lemma bot_annirr'' [simp]: "x \ \ = \" using Sup_distl[where Y="{}"] by auto lemma sup_distl: "x \ (y \ z) = x \ y \ x \ z" using Sup_distl[where Y="{y, z}"] by (fastforce intro!: Sup_eqI) lemma sup_distr: "(x \ y) \ z = x \ z \ y \ z" using Sup_distr[where X="{x, y}"] by (fastforce intro!: Sup_eqI) sublocale semiring_one_zero "(\)" "(\)" "1" "\" by (standard, auto simp: sup.assoc sup.commute sup_left_commute sup_distl sup_distr) sublocale dioid_one_zero "(\)" "(\)" "1" "\" "(\)" "(<)" by (standard, simp add: le_iff_sup, auto) lemma Sup_sup_pred: "x \ Sup{y. P y} = Sup{y. y = x \ P y}" apply (rule order.antisym) apply (simp add: Collect_mono Sup_subset_mono Sup_upper) using Sup_least Sup_upper le_supI2 by fastforce definition star :: "'a \ 'a" where "star x = (SUP i. x ^ i)" lemma star_def_var1: "star x = Sup{y. \i. y = x ^ i}" by (simp add: star_def full_SetCompr_eq) lemma star_def_var2: "star x = Sup{x ^ i |i. True}" by (simp add: star_def full_SetCompr_eq) lemma star_unfoldl' [simp]: "1 \ x \ (star x) = star x" proof - have "1 \ x \ (star x) = x ^ 0 \ x \ Sup{y. \i. y = x ^ i}" by (simp add: star_def_var1) also have "... = x ^ 0 \ Sup{y. \i. y = x ^ (i + 1)}" by (simp add: Sup_distl, metis) also have "... = Sup{y. y = x ^ 0 \ (\i. y = x ^ (i + 1))}" using Sup_sup_pred by simp also have "... = Sup{y. \i. y = x ^ i}" by (metis Suc_eq_plus1 power.power.power_Suc power.power_eq_if) finally show ?thesis by (simp add: star_def_var1) qed lemma star_unfoldr' [simp]: "1 \ (star x) \ x = star x" proof - have "1 \ (star x) \ x = x ^ 0 \ Sup{y. \i. y = x ^ i} \ x" by (simp add: star_def_var1) also have "... = x ^ 0 \ Sup{y. \i. y = x ^ i \ x}" by (simp add: Sup_distr, metis) also have "... = x ^ 0 \ Sup{y. \i. y = x ^ (i + 1)}" using power_Suc2 by simp also have "... = Sup{y. y = x ^ 0 \ (\i. y = x ^ (i + 1))}" using Sup_sup_pred by simp also have "... = Sup{y. \i. y = x ^ i}" by (metis Suc_eq_plus1 power.power.power_Suc power.power_eq_if) finally show ?thesis by (simp add: star_def_var1) qed lemma (in dioid_one_zero) power_inductl: "z + x \ y \ y \ (x ^ n) \ z \ y" proof (induct n) case 0 show ?case using "0.prems" by simp case Suc thus ?case by (simp, metis mult.assoc mult_isol order_trans) qed lemma (in dioid_one_zero) power_inductr: "z + y \ x \ y \ z \ (x ^ n) \ y" proof (induct n) case 0 show ?case using "0.prems" by auto case Suc { fix n assume "z + y \ x \ y \ z \ x ^ n \ y" and "z + y \ x \ y" hence "z \ x ^ n \ y" by simp also have "z \ x ^ Suc n = z \ x \ x ^ n" by (metis mult.assoc power_Suc) moreover have "... = (z \ x ^ n) \ x" by (metis mult.assoc power_commutes) moreover have "... \ y \ x" by (metis calculation(1) mult_isor) moreover have "... \ y" using \z + y \ x \ y\ by simp ultimately have "z \ x ^ Suc n \ y" by simp } thus ?case by (metis Suc) qed lemma star_inductl': "z \ x \ y \ y \ (star x) \ z \ y" proof - assume "z \ x \ y \ y" hence "\i. x ^ i \ z \ y" by (simp add: power_inductl) hence "Sup{w. \i. w = x ^ i \ z} \ y" by (intro Sup_least, fast) hence "Sup{w. \i. w = x ^ i} \ z \ y" using Sup_distr Sup_le_iff by auto thus "(star x) \ z \ y" by (simp add: star_def_var1) qed lemma star_inductr': "z \ y \ x \ y \ z \ (star x) \ y" proof - assume "z \ y \ x \ y" hence "\i. z \ x ^ i \ y" by (simp add: power_inductr) hence "Sup{w. \i. w = z \ x ^ i} \ y" by (intro Sup_least, fast) hence "z \ Sup{w. \i. w = x ^ i} \ y" using Sup_distl Sup_le_iff by auto thus "z \ (star x) \ y" by (simp add: star_def_var1) qed sublocale ka: kleene_algebra "(\)" "(\)" "1" "\" "(\)" "(<)" star by standard (simp_all add: star_inductl' star_inductr') end text \Distributive quantales are often assumed to satisfy infinite distributivity laws between joins and meets, but finite ones suffice for our purposes.\ class distributive_quantale = quantale + distrib_lattice begin subclass bd_lattice_ordered_monoid by (standard, simp_all add: distrib_left) lemma "(1 \ x \ \) \ x = x" (* nitpick [expect=genuine]*) oops end subsection \Domain Quantales\ class domain_quantale = distributive_quantale + assumes rdv'': "(z \ x \ \) \ y = z \ y \ x \ \" begin subclass domain_bdlo_monoid by (standard, simp add: rdv'') end class range_quantale = distributive_quantale + assumes ldv'': "y \ (z \ \ \ x) = y \ z \ \ \ x" class boolean_quantale = quantale + complete_boolean_algebra begin subclass boolean_monoid by (standard, simp_all add: sup_distl) lemma "(1 \ x \ \) \ x = x" (*nitpick[expect=genuine]*) oops lemma "(1 \ -(x \ \)) \ x = \" (*nitpick[expect=genuine]*) oops end subsection\Boolean Domain Quantales\ class domain_boolean_quantale = domain_quantale + boolean_quantale begin subclass domain_boolean_monoid by (standard, simp add: rdv'') lemma fbox_eq: "ad.fbox x q = Sup{d p |p. d p \ x \ x \ d q}" apply (rule Sup_eqI[symmetric]) apply clarsimp using ad.fbox_demodalisation3 ad.fbox_simp apply auto[1] apply clarsimp by (metis ad.fbox_def ad.fbox_demodalisation3 ad.fbox_simp da_a eq_refl) lemma fdia_eq: "ad.fdia x p = Inf{d q |q. x \ d p \ d q \ x}" apply (rule Inf_eqI[symmetric]) apply clarsimp using ds.fdemodalisation2 apply auto[1] apply clarsimp by (metis ad.fd_eq_fdia ad.fdia_def da_a double_compl ds.fdemodalisation2 inf_bot_iff_le inf_compl_bot) text \The specification statement can be defined explicitly.\ definition R :: "'a \ 'a \ 'a" where "R p q \ Sup{x. (d p) \ x \ x \ d q}" lemma "x \ R p q \ d p \ ad.fbox x (d q)" proof (simp add: R_def ad.kat_1_equiv ad.kat_2_equiv) assume "x \ Sup{x. d p \ x \ a q = \}" hence "d p \ x \ a q \ d p \ Sup{x. d p \ x \ a q = \} \ a q " using mult_double_iso by blast also have "... = Sup{d p \ x \ a q |x. d p \ x \ a q = \}" apply (subst Sup_distl) apply (subst Sup_distr) apply clarsimp by metis also have "... = \" by (auto simp: Sup_eqI) finally show ?thesis using ad.fbox_demodalisation3 ad.kat_3 ad.kat_4 le_bot by blast qed lemma "d p \ ad.fbox x (d q) \ x \ R p q" apply (simp add: R_def) apply (rule Sup_upper) apply simp using ad.fbox_demodalisation3 ad.fbox_simp apply auto[1] done end subsection\Relational Model of Boolean Domain Quantales\ interpretation rel_dbq: domain_boolean_quantale \(-)\ uminus \(\)\ \(\)\ \(\)\ \(\)\ \{}\ UNIV \\\ \\\ Id \(O)\ by standard auto subsection\Modal Boolean Quantales\ class range_boolean_quantale = range_quantale + boolean_quantale begin subclass range_boolean_monoid by (standard, simp add: ldv'') lemma fbox_eq: "ar.bbox x (r q) = Sup{r p |p. x \ r p \ (r q) \ x}" apply (rule Sup_eqI[symmetric]) apply clarsimp using ar.ardual.fbox_demodalisation3 ar.ardual.fbox_simp apply auto[1] apply clarsimp by (metis ar.ardual.fbox_def ar.ardual.fbox_demodalisation3 eq_refl rar_ar) lemma fdia_eq: "ar.bdia x (r p) = Inf{r q |q. (r p) \ x \ x \ r q}" apply (rule Inf_eqI[symmetric]) apply clarsimp using ar.ars_r_def ar.ardual.fdemodalisation22 ar.ardual.kat_3_equiv_opp ar.ardual.kat_4_equiv_opp apply auto[1] apply clarsimp using ar.bdia_def ar.ardual.ds.fdemodalisation2 r_ar by fastforce end class modal_boolean_quantale = domain_boolean_quantale + range_boolean_quantale + assumes domrange' [simp]: "d (r x) = r x" and rangedom' [simp]: "r (d x) = d x" begin sublocale mka: modal_kleene_algebra "(\)" "(\)" 1 \ "(\)" "(<)" star a ar by standard (simp_all add: ar_eq ad_eq) end no_notation fbox ("( |_] _)" [61,81] 82) and antidomain_semiringl_class.fbox ("( |_] _)" [61,81] 82) notation ad.fbox ("( |_] _)" [61,81] 82) subsection \Recursion Rule\ lemma recursion: "mono (f :: 'a \ 'a :: domain_boolean_quantale) \ (\x. d p \ |x] d q \ d p \ |f x] d q) \ d p \ |lfp f] d q" apply (erule lfp_ordinal_induct [where f=f], simp) by (auto simp: ad.addual.ardual.fbox_demodalisation3 Sup_distr Sup_distl intro: Sup_mono) text \We have already tested this rule in the context of test quantales~\cite{ArmstrongGS15}, which is based on a formalisation of quantales that is currently not in the AFP. The two theories will be merged as soon as the quantale is available in the AFP.\ end diff --git a/thys/Automatic_Refinement/Lib/Misc.thy b/thys/Automatic_Refinement/Lib/Misc.thy --- a/thys/Automatic_Refinement/Lib/Misc.thy +++ b/thys/Automatic_Refinement/Lib/Misc.thy @@ -1,4523 +1,4507 @@ (* Title: Miscellaneous Definitions and Lemmas Author: Peter Lammich Maintainer: Peter Lammich Thomas Tuerk *) (* CHANGELOG: 2010-05-09: Removed AC, AI locales, they are superseeded by concepts from OrderedGroups 2010-09-22: Merges with ext/Aux 2017-06-12: Added a bunch of lemmas from various other projects *) section \Miscellaneous Definitions and Lemmas\ theory Misc imports Main "HOL-Library.Multiset" "HOL-ex.Quicksort" "HOL-Library.Option_ord" "HOL-Library.Infinite_Set" "HOL-Eisbach.Eisbach" begin text_raw \\label{thy:Misc}\ subsection \Isabelle Distribution Move Proposals\ subsubsection \Pure\ lemma TERMI: "TERM x" unfolding Pure.term_def . subsubsection \HOL\ (* Stronger disjunction elimination rules. *) lemma disjE1: "\ P \ Q; P \ R; \\P;Q\ \ R \ \ R" by metis lemma disjE2: "\ P \ Q; \P; \Q\ \ R; Q \ R \ \ R" by blast lemma imp_mp_iff[simp]: "a \ (a \ b) \ a \ b" "(a \ b) \ a \ a \ b" (* is Inductive.imp_conj_iff, but not in simpset by default *) by blast+ lemma atomize_Trueprop_eq[atomize]: "(Trueprop x \ Trueprop y) \ Trueprop (x=y)" apply rule apply (rule) apply (erule equal_elim_rule1) apply assumption apply (erule equal_elim_rule2) apply assumption apply simp done subsubsection \Set\ lemma inter_compl_diff_conv[simp]: "A \ -B = A - B" by auto lemma pair_set_inverse[simp]: "{(a,b). P a b}\ = {(b,a). P a b}" by auto lemma card_doubleton_eq_2_iff[simp]: "card {a,b} = 2 \ a\b" by auto subsubsection \List\ (* TODO: Move, analogous to List.length_greater_0_conv *) thm List.length_greater_0_conv lemma length_ge_1_conv[iff]: "Suc 0 \ length l \ l\[]" by (cases l) auto \ \Obtains a list from the pointwise characterization of its elements\ lemma obtain_list_from_elements: assumes A: "\ili. P li i)" obtains l where "length l = n" "\il. length l=n \ (\iii l!j" by (simp add: sorted_iff_nth_mono) also from nth_eq_iff_index_eq[OF D] B have "l!i \ l!j" by auto finally show ?thesis . qed lemma distinct_sorted_strict_mono_iff: assumes "distinct l" "sorted l" assumes "i il!j \ i\j" by (metis assms distinct_sorted_strict_mono_iff leD le_less_linear) (* List.thy has: declare map_eq_Cons_D [dest!] Cons_eq_map_D [dest!] We could, analogously, declare rules for "map _ _ = _@_" as dest!, or use elim!, or declare the _conv-rule as simp *) lemma map_eq_appendE: assumes "map f ls = fl@fl'" obtains l l' where "ls=l@l'" and "map f l=fl" and "map f l' = fl'" using assms proof (induction fl arbitrary: ls thesis) case (Cons x xs) then obtain l ls' where [simp]: "ls = l#ls'" "f l = x" by force with Cons.prems(2) have "map f ls' = xs @ fl'" by simp from Cons.IH[OF _ this] guess ll ll' . with Cons.prems(1)[of "l#ll" ll'] show thesis by simp qed simp lemma map_eq_append_conv: "map f ls = fl@fl' \ (\l l'. ls = l@l' \ map f l = fl \ map f l' = fl')" by (auto elim!: map_eq_appendE) lemmas append_eq_mapE = map_eq_appendE[OF sym] lemma append_eq_map_conv: "fl@fl' = map f ls \ (\l l'. ls = l@l' \ map f l = fl \ map f l' = fl')" by (auto elim!: append_eq_mapE) lemma distinct_mapI: "distinct (map f l) \ distinct l" by (induct l) auto lemma map_distinct_upd_conv: "\i \ (map f l)[i := x] = map (f(l!i := x)) l" \ \Updating a mapped distinct list is equal to updating the mapping function\ by (auto simp: nth_eq_iff_index_eq intro: nth_equalityI) lemma zip_inj: "\length a = length b; length a' = length b'; zip a b = zip a' b'\ \ a=a' \ b=b'" proof (induct a b arbitrary: a' b' rule: list_induct2) case Nil then show ?case by (cases a'; cases b'; auto) next case (Cons x xs y ys) then show ?case by (cases a'; cases b'; auto) qed lemma zip_eq_zip_same_len[simp]: "\ length a = length b; length a' = length b' \ \ zip a b = zip a' b' \ a=a' \ b=b'" by (auto dest: zip_inj) lemma upt_merge[simp]: "i\j \ j\k \ [i.. (ys \ [] \ butlast ys = xs \ last ys = x)" by auto (* Case distinction how two elements of a list can be related to each other *) lemma list_match_lel_lel: assumes "c1 @ qs # c2 = c1' @ qs' # c2'" obtains (left) c21' where "c1 = c1' @ qs' # c21'" "c2' = c21' @ qs # c2" | (center) "c1' = c1" "qs' = qs" "c2' = c2" | (right) c21 where "c1' = c1 @ qs # c21" "c2 = c21 @ qs' # c2'" using assms apply (clarsimp simp: append_eq_append_conv2) subgoal for us by (cases us) auto done lemma xy_in_set_cases[consumes 2, case_names EQ XY YX]: assumes A: "x\set l" "y\set l" and C: "!!l1 l2. \ x=y; l=l1@y#l2 \ \ P" "!!l1 l2 l3. \ x\y; l=l1@x#l2@y#l3 \ \ P" "!!l1 l2 l3. \ x\y; l=l1@y#l2@x#l3 \ \ P" shows P proof (cases "x=y") case True with A(1) obtain l1 l2 where "l=l1@y#l2" by (blast dest: split_list) with C(1) True show ?thesis by blast next case False from A(1) obtain l1 l2 where S1: "l=l1@x#l2" by (blast dest: split_list) from A(2) obtain l1' l2' where S2: "l=l1'@y#l2'" by (blast dest: split_list) from S1 S2 have M: "l1@x#l2 = l1'@y#l2'" by simp thus P proof (cases rule: list_match_lel_lel[consumes 1, case_names 1 2 3]) case (1 c) with S1 have "l=l1'@y#c@x#l2" by simp with C(3) False show ?thesis by blast next case 2 with False have False by blast thus ?thesis .. next case (3 c) with S1 have "l=l1@x#c@y#l2'" by simp with C(2) False show ?thesis by blast qed qed lemma list_e_eq_lel[simp]: "[e] = l1@e'#l2 \ l1=[] \ e'=e \ l2=[]" "l1@e'#l2 = [e] \ l1=[] \ e'=e \ l2=[]" apply (cases l1, auto) [] apply (cases l1, auto) [] done lemma list_ee_eq_leel[simp]: "([e1,e2] = l1@e1'#e2'#l2) \ (l1=[] \ e1=e1' \ e2=e2' \ l2=[])" "(l1@e1'#e2'#l2 = [e1,e2]) \ (l1=[] \ e1=e1' \ e2=e2' \ l2=[])" apply (cases l1, auto) [] apply (cases l1, auto) [] done subsubsection \Transitive Closure\ text \A point-free induction rule for elements reachable from an initial set\ lemma rtrancl_reachable_induct[consumes 0, case_names base step]: assumes I0: "I \ INV" assumes IS: "E``INV \ INV" shows "E\<^sup>*``I \ INV" by (metis I0 IS Image_closed_trancl Image_mono subset_refl) lemma acyclic_empty[simp, intro!]: "acyclic {}" by (unfold acyclic_def) auto lemma acyclic_union: "acyclic (A\B) \ acyclic A" "acyclic (A\B) \ acyclic B" by (metis Un_upper1 Un_upper2 acyclic_subset)+ - - -subsubsection \Lattice Syntax\ -(* Providing the syntax in a locale makes it more usable, without polluting the global namespace*) -locale Lattice_Syntax begin - notation - bot ("\") and - top ("\") and - inf (infixl "\" 70) and - sup (infixl "\" 65) and - Inf ("\_" [900] 900) and - Sup ("\_" [900] 900) - -end - - - + text \Here we provide a collection of miscellaneous definitions and helper lemmas\ subsection "Miscellaneous (1)" text \This stuff is used in this theory itself, and thus occurs in first place or is simply not sorted into any other section of this theory.\ lemma IdD: "(a,b)\Id \ a=b" by simp text \Conversion Tag\ definition [simp]: "CNV x y \ x=y" lemma CNV_I: "CNV x x" by simp lemma CNV_eqD: "CNV x y \ x=y" by simp lemma CNV_meqD: "CNV x y \ x\y" by simp (* TODO: Move. Shouldn't this be detected by simproc? *) lemma ex_b_b_and_simp[simp]: "(\b. b \ Q b) \ Q True" by auto lemma ex_b_not_b_and_simp[simp]: "(\b. \b \ Q b) \ Q False" by auto method repeat_all_new methods m = m;(repeat_all_new \m\)? subsubsection "AC-operators" text \Locale to declare AC-laws as simplification rules\ locale Assoc = fixes f assumes assoc[simp]: "f (f x y) z = f x (f y z)" locale AC = Assoc + assumes commute[simp]: "f x y = f y x" lemma (in AC) left_commute[simp]: "f x (f y z) = f y (f x z)" by (simp only: assoc[symmetric]) simp lemmas (in AC) AC_simps = commute assoc left_commute text \Locale to define functions from surjective, unique relations\ locale su_rel_fun = fixes F and f assumes unique: "\(A,B)\F; (A,B')\F\ \ B=B'" assumes surjective: "\!!B. (A,B)\F \ P\ \ P" assumes f_def: "f A == THE B. (A,B)\F" lemma (in su_rel_fun) repr1: "(A,f A)\F" proof (unfold f_def) obtain B where "(A,B)\F" by (rule surjective) with theI[where P="\B. (A,B)\F", OF this] show "(A, THE x. (A, x) \ F) \ F" by (blast intro: unique) qed lemma (in su_rel_fun) repr2: "(A,B)\F \ B=f A" using repr1 by (blast intro: unique) lemma (in su_rel_fun) repr: "(f A = B) = ((A,B)\F)" using repr1 repr2 by (blast) \ \Contract quantification over two variables to pair\ lemma Ex_prod_contract: "(\a b. P a b) \ (\z. P (fst z) (snd z))" by auto lemma All_prod_contract: "(\a b. P a b) \ (\z. P (fst z) (snd z))" by auto lemma nat_geq_1_eq_neqz: "x\1 \ x\(0::nat)" by auto lemma nat_in_between_eq: "(a b\Suc a) \ b = Suc a" "(a\b \ b b = a" by auto lemma Suc_n_minus_m_eq: "\ n\m; m>1 \ \ Suc (n - m) = n - (m - 1)" by simp lemma Suc_to_right: "Suc n = m \ n = m - Suc 0" by simp lemma Suc_diff[simp]: "\n m. n\m \ m\1 \ Suc (n - m) = n - (m - 1)" by simp lemma if_not_swap[simp]: "(if \c then a else b) = (if c then b else a)" by auto lemma all_to_meta: "Trueprop (\a. P a) \ (\a. P a)" apply rule by auto lemma imp_to_meta: "Trueprop (P\Q) \ (P\Q)" apply rule by auto (* for some reason, there is no such rule in HOL *) lemma iffI2: "\P \ Q; \ P \ \ Q\ \ P \ Q" by metis lemma iffExI: "\ \x. P x \ Q x; \x. Q x \ P x \ \ (\x. P x) \ (\x. Q x)" by metis lemma bex2I[intro?]: "\ (a,b)\S; (a,b)\S \ P a b \ \ \a b. (a,b)\S \ P a b" by blast (* TODO: Move lemma to HOL! *) lemma cnv_conj_to_meta: "(P \ Q \ PROP X) \ (\P;Q\ \ PROP X)" by (rule BNF_Fixpoint_Base.conj_imp_eq_imp_imp) subsection \Sets\ lemma remove_subset: "x\S \ S-{x} \ S" by auto lemma subset_minus_empty: "A\B \ A-B = {}" by auto lemma insert_minus_eq: "x\y \ insert x A - {y} = insert x (A - {y})" by auto lemma set_notEmptyE: "\S\{}; !!x. x\S \ P\ \ P" by (metis equals0I) lemma subset_Collect_conv: "S \ Collect P \ (\x\S. P x)" by auto lemma memb_imp_not_empty: "x\S \ S\{}" by auto lemma disjoint_mono: "\ a\a'; b\b'; a'\b'={} \ \ a\b={}" by auto lemma disjoint_alt_simp1: "A-B = A \ A\B = {}" by auto lemma disjoint_alt_simp2: "A-B \ A \ A\B \ {}" by auto lemma disjoint_alt_simp3: "A-B \ A \ A\B \ {}" by auto lemma disjointI[intro?]: "\ \x. \x\a; x\b\ \ False \ \ a\b={}" by auto lemmas set_simps = subset_minus_empty disjoint_alt_simp1 disjoint_alt_simp2 disjoint_alt_simp3 Un_absorb1 Un_absorb2 lemma set_minus_singleton_eq: "x\X \ X-{x} = X" by auto lemma set_diff_diff_left: "A-B-C = A-(B\C)" by auto lemma image_update[simp]: "x\A \ f(x:=n)`A = f`A" by auto lemma eq_or_mem_image_simp[simp]: "{f l |l. l = a \ l\B} = insert (f a) {f l|l. l\B}" by blast lemma set_union_code [code_unfold]: "set xs \ set ys = set (xs @ ys)" by auto lemma in_fst_imageE: assumes "x \ fst`S" obtains y where "(x,y)\S" using assms by auto lemma in_snd_imageE: assumes "y \ snd`S" obtains x where "(x,y)\S" using assms by auto lemma fst_image_mp: "\fst`A \ B; (x,y)\A \ \ x\B" by (metis Domain.DomainI fst_eq_Domain in_mono) lemma snd_image_mp: "\snd`A \ B; (x,y)\A \ \ y\B" by (metis Range.intros rev_subsetD snd_eq_Range) lemma inter_eq_subsetI: "\ S\S'; A\S' = B\S' \ \ A\S = B\S" by auto text \ Decompose general union over sum types. \ lemma Union_plus: "(\ x \ A <+> B. f x) = (\ a \ A. f (Inl a)) \ (\b \ B. f (Inr b))" by auto lemma Union_sum: "(\x. f (x::'a+'b)) = (\l. f (Inl l)) \ (\r. f (Inr r))" (is "?lhs = ?rhs") proof - have "?lhs = (\x \ UNIV <+> UNIV. f x)" by simp thus ?thesis by (simp only: Union_plus) qed subsubsection \Finite Sets\ lemma card_1_singletonI: "\finite S; card S = 1; x\S\ \ S={x}" proof (safe, rule ccontr, goal_cases) case prems: (1 x') hence "finite (S-{x})" "S-{x} \ {}" by auto hence "card (S-{x}) \ 0" by auto moreover from prems(1-3) have "card (S-{x}) = 0" by auto ultimately have False by simp thus ?case .. qed lemma card_insert_disjoint': "\finite A; x \ A\ \ card (insert x A) - Suc 0 = card A" by (drule (1) card_insert_disjoint) auto lemma card_eq_UNIV[simp]: "card (S::'a::finite set) = card (UNIV::'a set) \ S=UNIV" proof (auto) fix x assume A: "card S = card (UNIV::'a set)" show "x\S" proof (rule ccontr) assume "x\S" hence "S\UNIV" by auto with psubset_card_mono[of UNIV S] have "card S < card (UNIV::'a set)" by auto with A show False by simp qed qed lemma card_eq_UNIV2[simp]: "card (UNIV::'a set) = card (S::'a::finite set) \ S=UNIV" using card_eq_UNIV[of S] by metis lemma card_ge_UNIV[simp]: "card (UNIV::'a::finite set) \ card (S::'a set) \ S=UNIV" using card_mono[of "UNIV::'a::finite set" S, simplified] by auto lemmas length_remdups_card = length_remdups_concat[of "[l]", simplified] for l lemma fs_contract: "fst ` { p | p. f (fst p) (snd p) \ S } = { a . \b. f a b \ S }" by (simp add: image_Collect) lemma finite_Collect: "finite S \ inj f \ finite {a. f a : S}" by(simp add: finite_vimageI vimage_def[symmetric]) \ \Finite sets have an injective mapping to an initial segments of the natural numbers\ (* This lemma is also in the standard library (from Isabelle2009-1 on) as @{thm [source] Finite_Set.finite_imp_inj_to_nat_seg}. However, it is formulated with HOL's \ there rather then with the meta-logic obtain *) lemma finite_imp_inj_to_nat_seg': fixes A :: "'a set" assumes A: "finite A" obtains f::"'a \ nat" and n::"nat" where "f`A = {i. i finite (lists P \ { l. length l = n })" proof (induct n) case 0 thus ?case by auto next case (Suc n) have "lists P \ { l. length l = Suc n } = (\(a,l). a#l) ` (P \ (lists P \ {l. length l = n}))" apply auto apply (case_tac x) apply auto done moreover from Suc have "finite \" by auto ultimately show ?case by simp qed lemma lists_of_len_fin2: "finite P \ finite (lists P \ { l. n = length l })" proof - assume A: "finite P" have S: "{ l. n = length l } = { l. length l = n }" by auto have "finite (lists P \ { l. n = length l }) \ finite (lists P \ { l. length l = n })" by (subst S) simp thus ?thesis using lists_of_len_fin1[OF A] by auto qed lemmas lists_of_len_fin = lists_of_len_fin1 lists_of_len_fin2 (* Try (simp only: cset_fin_simps, fastforce intro: cset_fin_intros) when reasoning about finiteness of collected sets *) lemmas cset_fin_simps = Ex_prod_contract fs_contract[symmetric] image_Collect[symmetric] lemmas cset_fin_intros = finite_imageI finite_Collect inj_onI lemma Un_interval: fixes b1 :: "'a::linorder" assumes "b1\b2" and "b2\b3" shows "{ f i | i. b1\i \ i { f i | i. b2\i \ ii \ i The standard library proves that a generalized union is finite if the index set is finite and if for every index the component set is itself finite. Conversely, we show that every component set must be finite when the union is finite. \ lemma finite_UNION_then_finite: "finite (\(B ` A)) \ a \ A \ finite (B a)" by (metis Set.set_insert UN_insert Un_infinite) lemma finite_if_eq_beyond_finite: "finite S \ finite {s. s - S = s' - S}" proof (rule finite_subset[where B="(\s. s \ (s' - S)) ` Pow S"], clarsimp) fix s have "s = (s \ S) \ (s - S)" by auto also assume "s - S = s' - S" finally show "s \ (\s. s \ (s' - S)) ` Pow S" by blast qed blast lemma distinct_finite_subset: assumes "finite x" shows "finite {ys. set ys \ x \ distinct ys}" (is "finite ?S") proof (rule finite_subset) from assms show "?S \ {ys. set ys \ x \ length ys \ card x}" by clarsimp (metis distinct_card card_mono) from assms show "finite ..." by (rule finite_lists_length_le) qed lemma distinct_finite_set: shows "finite {ys. set ys = x \ distinct ys}" (is "finite ?S") proof (cases "finite x") case False hence "{ys. set ys = x} = {}" by auto thus ?thesis by simp next case True show ?thesis proof (rule finite_subset) show "?S \ {ys. set ys \ x \ length ys \ card x}" using distinct_card by force from True show "finite ..." by (rule finite_lists_length_le) qed qed lemma finite_set_image: assumes f: "finite (set ` A)" and dist: "\xs. xs \ A \ distinct xs" shows "finite A" proof (rule finite_subset) from f show "finite (set -` (set ` A) \ {xs. distinct xs})" proof (induct rule: finite_induct) case (insert x F) have "finite (set -` {x} \ {xs. distinct xs})" apply (simp add: vimage_def) by (metis Collect_conj_eq distinct_finite_set) with insert show ?case apply (subst vimage_insert) apply (subst Int_Un_distrib2) apply (rule finite_UnI) apply simp_all done qed simp moreover from dist show "A \ ..." by (auto simp add: vimage_image_eq) qed subsubsection \Infinite Set\ lemma INFM_nat_inductI: assumes P0: "P (0::nat)" assumes PS: "\i. P i \ \j>i. P j \ Q j" shows "\\<^sub>\i. Q i" proof - have "\i. \j>i. P j \ Q j" proof fix i show "\j>i. P j \ Q j" apply (induction i) using PS[OF P0] apply auto [] by (metis PS Suc_lessI) qed thus ?thesis unfolding INFM_nat by blast qed subsection \Functions\ lemma fun_neq_ext_iff: "m\m' \ (\x. m x \ m' x)" by auto definition "inv_on f A x == SOME y. y\A \ f y = x" lemma inv_on_f_f[simp]: "\inj_on f A; x\A\ \ inv_on f A (f x) = x" by (auto simp add: inv_on_def inj_on_def) lemma f_inv_on_f: "\ y\f`A \ \ f (inv_on f A y) = y" by (auto simp add: inv_on_def intro: someI2) lemma inv_on_f_range: "\ y \ f`A \ \ inv_on f A y \ A" by (auto simp add: inv_on_def intro: someI2) lemma inj_on_map_inv_f [simp]: "\set l \ A; inj_on f A\ \ map (inv_on f A) (map f l) = l" apply (simp) apply (induct l) apply auto done lemma comp_cong_right: "x = y \ f o x = f o y" by (simp) lemma comp_cong_left: "x = y \ x o f = y o f" by (simp) lemma fun_comp_eq_conv: "f o g = fg \ (\x. f (g x) = fg x)" by auto abbreviation comp2 (infixl "oo" 55) where "f oo g \ \x. f o (g x)" abbreviation comp3 (infixl "ooo" 55) where "f ooo g \ \x. f oo (g x)" notation comp2 (infixl "\\" 55) and comp3 (infixl "\\\" 55) definition [code_unfold, simp]: "swap_args2 f x y \ f y x" subsection \Multisets\ (* The following is a syntax extension for multisets. Unfortunately, it depends on a change in the Library/Multiset.thy, so it is commented out here, until it will be incorporated into Library/Multiset.thy by its maintainers. The required change in Library/Multiset.thy is removing the syntax for single: - single :: "'a => 'a multiset" ("{#_#}") + single :: "'a => 'a multiset" And adding the following translations instead: + syntax + "_multiset" :: "args \ 'a multiset" ("{#(_)#}") + translations + "{#x, xs#}" == "{#x#} + {#xs#}" + "{# x #}" == "single x" This translates "{# \ #}" into a sum of singletons, that is parenthesized to the right. ?? Can we also achieve left-parenthesizing ?? *) (* Let's try what happens if declaring AC-rules for multiset union as simp-rules *) (*declare union_ac[simp] -- don't do it !*) lemma count_mset_set_finite_iff: "finite S \ count (mset_set S) a = (if a \ S then 1 else 0)" by simp lemma ex_Melem_conv: "(\x. x \# A) = (A \ {#})" by (simp add: ex_in_conv) subsubsection \Count\ lemma count_ne_remove: "\ x ~= t\ \ count S x = count (S-{#t#}) x" by (auto) lemma mset_empty_count[simp]: "(\p. count M p = 0) = (M={#})" by (auto simp add: multiset_eq_iff) subsubsection \Union, difference and intersection\ lemma size_diff_se: "t \# S \ size S = size (S - {#t#}) + 1" proof (unfold size_multiset_overloaded_eq) let ?SIZE = "sum (count S) (set_mset S)" assume A: "t \# S" from A have SPLITPRE: "finite (set_mset S) & {t}\(set_mset S)" by auto hence "?SIZE = sum (count S) (set_mset S - {t}) + sum (count S) {t}" by (blast dest: sum.subset_diff) hence "?SIZE = sum (count S) (set_mset S - {t}) + count (S) t" by auto moreover with A have "count S t = count (S-{#t#}) t + 1" by auto ultimately have D: "?SIZE = sum (count S) (set_mset S - {t}) + count (S-{#t#}) t + 1" by (arith) moreover have "sum (count S) (set_mset S - {t}) = sum (count (S-{#t#})) (set_mset S - {t})" proof - have "\x\(set_mset S - {t}). count S x = count (S-{#t#}) x" by (auto iff add: count_ne_remove) thus ?thesis by simp qed ultimately have D: "?SIZE = sum (count (S-{#t#})) (set_mset S - {t}) + count (S-{#t#}) t + 1" by (simp) moreover { assume CASE: "t \# S - {#t#}" from CASE have "set_mset S - {t} = set_mset (S - {#t#})" by (auto simp add: in_diff_count split: if_splits) with CASE D have "?SIZE = sum (count (S-{#t#})) (set_mset (S - {#t#})) + 1" by (simp add: not_in_iff) } moreover { assume CASE: "t \# S - {#t#}" from CASE have "t \# S" by (rule in_diffD) with CASE have 1: "set_mset S = set_mset (S-{#t#})" by (auto simp add: in_diff_count split: if_splits) moreover from D have "?SIZE = sum (count (S-{#t#})) (set_mset S - {t}) + sum (count (S-{#t#})) {t} + 1" by simp moreover from SPLITPRE sum.subset_diff have "sum (count (S-{#t#})) (set_mset S) = sum (count (S-{#t#})) (set_mset S - {t}) + sum (count (S-{#t#})) {t}" by (blast) ultimately have "?SIZE = sum (count (S-{#t#})) (set_mset (S-{#t#})) + 1" by simp } ultimately show "?SIZE = sum (count (S-{#t#})) (set_mset (S - {#t#})) + 1" by blast qed (* TODO: Check whether this proof can be done simpler *) lemma mset_union_diff_comm: "t \# S \ T + (S - {#t#}) = (T + S) - {#t#}" proof - assume "t \# S" then obtain S' where S: "S = add_mset t S'" by (metis insert_DiffM) then show ?thesis by auto qed (* lemma mset_diff_diff_left: "A-B-C = A-((B::'a multiset)+C)" proof - have "\e . count (A-B-C) e = count (A-(B+C)) e" by auto thus ?thesis by (simp add: multiset_eq_conv_count_eq) qed lemma mset_diff_commute: "A-B-C = A-C-(B::'a multiset)" proof - have "A-B-C = A-(B+C)" by (simp add: mset_diff_diff_left) also have "\ = A-(C+B)" by (simp add: union_commute) thus ?thesis by (simp add: mset_diff_diff_left) qed lemma mset_diff_same_empty[simp]: "(S::'a multiset) - S = {#}" proof - have "\e . count (S-S) e = 0" by auto hence "\e . ~ (e : set_mset (S-S))" by auto hence "set_mset (S-S) = {}" by blast thus ?thesis by (auto) qed *) lemma mset_right_cancel_union: "\a \# A+B; ~(a \# B)\ \ a\#A" by (simp) lemma mset_left_cancel_union: "\a \# A+B; ~(a \# A)\ \ a\#B" by (simp) lemmas mset_cancel_union = mset_right_cancel_union mset_left_cancel_union lemma mset_right_cancel_elem: "\a \# A+{#b#}; a~=b\ \ a\#A" by simp lemma mset_left_cancel_elem: "\a \# {#b#}+A; a~=b\ \ a\#A" by simp lemmas mset_cancel_elem = mset_right_cancel_elem mset_left_cancel_elem lemma mset_diff_cancel1elem[simp]: "~(a \# B) \ {#a#}-B = {#a#}" by (auto simp add: not_in_iff intro!: multiset_eqI) (* lemma diff_union_inverse[simp]: "A + B - B = (A::'a multiset)" by (auto iff add: multiset_eq_conv_count_eq) lemma diff_union_inverse2[simp]: "B + A - B = (A::'a multiset)" by (auto iff add: multiset_eq_conv_count_eq) *) (*lemma union_diff_assoc_se2: "t \# A \ (A+B)-{#t#} = (A-{#t#}) + B" by (auto iff add: multiset_eq_conv_count_eq) lemmas union_diff_assoc_se = union_diff_assoc_se1 union_diff_assoc_se2*) lemma union_diff_assoc: "C-B={#} \ (A+B)-C = A + (B-C)" by (simp add: multiset_eq_iff) lemmas mset_neutral_cancel1 = union_left_cancel[where N="{#}", simplified] union_right_cancel[where N="{#}", simplified] declare mset_neutral_cancel1[simp] lemma mset_union_2_elem: "{#a, b#} = add_mset c M \ {#a#}=M & b=c | a=c & {#b#}=M" by (auto simp: add_eq_conv_diff) lemma mset_un_cases[cases set, case_names left right]: "\a \# A + B; a \# A \ P; a \# B \ P\ \ P" by (auto) lemma mset_unplusm_dist_cases[cases set, case_names left right]: assumes A: "{#s#}+A = B+C" assumes L: "\B={#s#}+(B-{#s#}); A=(B-{#s#})+C\ \ P" assumes R: "\C={#s#}+(C-{#s#}); A=B+(C-{#s#})\ \ P" shows P proof - from A[symmetric] have "s \# B+C" by simp thus ?thesis proof (cases rule: mset_un_cases) case left hence 1: "B={#s#}+(B-{#s#})" by simp with A have "{#s#}+A = {#s#}+((B-{#s#})+C)" by (simp add: union_ac) hence 2: "A = (B-{#s#})+C" by (simp) from L[OF 1 2] show ?thesis . next case right hence 1: "C={#s#}+(C-{#s#})" by simp with A have "{#s#}+A = {#s#}+(B+(C-{#s#}))" by (simp add: union_ac) hence 2: "A = B+(C-{#s#})" by (simp) from R[OF 1 2] show ?thesis . qed qed lemma mset_unplusm_dist_cases2[cases set, case_names left right]: assumes A: "B+C = {#s#}+A" assumes L: "\B={#s#}+(B-{#s#}); A=(B-{#s#})+C\ \ P" assumes R: "\C={#s#}+(C-{#s#}); A=B+(C-{#s#})\ \ P" shows P using mset_unplusm_dist_cases[OF A[symmetric]] L R by blast lemma mset_single_cases[cases set, case_names loc env]: assumes A: "add_mset s c = add_mset r' c'" assumes CASES: "\s=r'; c=c'\ \ P" "\c'={#s#}+(c'-{#s#}); c={#r'#}+(c-{#r'#}); c-{#r'#} = c'-{#s#} \ \ P" shows "P" proof - { assume CASE: "s=r'" with A have "c=c'" by simp with CASE CASES have ?thesis by auto } moreover { assume CASE: "s\r'" have "s \# {#s#}+c" by simp with A have "s \# {#r'#}+c'" by simp with CASE have "s \# c'" by simp hence 1: "c' = {#s#} + (c' - {#s#})" by simp with A have "{#s#}+c = {#s#}+({#r'#}+(c' - {#s#}))" by (auto simp add: union_ac) hence 2: "c={#r'#}+(c' - {#s#})" by (auto) hence 3: "c-{#r'#} = (c' - {#s#})" by auto from 1 2 3 CASES have ?thesis by auto } ultimately show ?thesis by blast qed lemma mset_single_cases'[cases set, case_names loc env]: assumes A: "add_mset s c = add_mset r' c'" assumes CASES: "\s=r'; c=c'\ \ P" "!!cc. \c'={#s#}+cc; c={#r'#}+cc; c'-{#s#}=cc; c-{#r'#}=cc\ \ P" shows "P" using A CASES by (auto elim!: mset_single_cases) lemma mset_single_cases2[cases set, case_names loc env]: assumes A: "add_mset s c = add_mset r' c'" assumes CASES: "\s=r'; c=c'\ \ P" "\c'=(c'-{#s#})+{#s#}; c=(c-{#r'#})+{#r'#}; c-{#r'#} = c'-{#s#} \ \ P" shows "P" proof - from A have "add_mset s c = add_mset r' c'" by (simp add: union_ac) thus ?thesis using CASES by (cases rule: mset_single_cases) simp_all qed lemma mset_single_cases2'[cases set, case_names loc env]: assumes A: "add_mset s c = add_mset r' c'" assumes CASES: "\s=r'; c=c'\ \ P" "!!cc. \c'=cc+{#s#}; c=cc+{#r'#}; c'-{#s#}=cc; c-{#r'#}=cc\ \ P" shows "P" using A CASES by (auto elim!: mset_single_cases2) lemma mset_un_single_un_cases [consumes 1, case_names left right]: assumes A: "add_mset a A = B + C" and CASES: "a \# B \ A = (B - {#a#}) + C \ P" "a \# C \ A = B + (C - {#a#}) \ P" shows "P" proof - have "a \# A+{#a#}" by simp with A have "a \# B+C" by auto thus ?thesis proof (cases rule: mset_un_cases) case left hence "B=B-{#a#}+{#a#}" by auto with A have "A+{#a#} = (B-{#a#})+C+{#a#}" by (auto simp add: union_ac) hence "A=(B-{#a#})+C" by simp with CASES(1)[OF left] show ?thesis by blast next case right hence "C=C-{#a#}+{#a#}" by auto with A have "A+{#a#} = B+(C-{#a#})+{#a#}" by (auto simp add: union_ac) hence "A=B+(C-{#a#})" by simp with CASES(2)[OF right] show ?thesis by blast qed qed (* TODO: Can this proof be done more automatically ? *) lemma mset_distrib[consumes 1, case_names dist]: assumes A: "(A::'a multiset)+B = M+N" "!!Am An Bm Bn. \A=Am+An; B=Bm+Bn; M=Am+Bm; N=An+Bn\ \ P" shows "P" proof - have BN_MA: "B - N = M - A" by (metis (no_types) add_diff_cancel_right assms(1) union_commute) have H: "A = A\# C + (A - C) \# D" if "A + B = C + D" for A B C D :: "'a multiset" by (metis add.commute diff_intersect_left_idem mset_subset_eq_add_left subset_eq_diff_conv subset_mset.add_diff_inverse subset_mset.inf_absorb1 subset_mset.inf_le1 that) have A': "A = A\# M + (A - M) \# N" using A(1) H by blast moreover have B': "B = (B - N) \# M + B\# N" using A(1) H[of B A N M] by (auto simp: ac_simps) moreover have "M = A \# M + (B - N) \# M" using H[of M N A B] BN_MA[symmetric] A(1) by (metis (no_types) diff_intersect_left_idem diff_union_cancelR multiset_inter_commute subset_mset.diff_add subset_mset.inf.cobounded1 union_commute) moreover have "N = (A - M) \# N + B \# N" by (metis A' assms(1) diff_union_cancelL inter_union_distrib_left inter_union_distrib_right mset_subset_eq_multiset_union_diff_commute subset_mset.inf.cobounded1 subset_mset.inf.commute) ultimately show P using A(2) by blast qed subsubsection \Singleton multisets\ lemma mset_size_le1_cases[case_names empty singleton,consumes 1]: "\ size M \ Suc 0; M={#} \ P; !!m. M={#m#} \ P \ \ P" by (cases M) auto lemma diff_union_single_conv2: "a \# J \ J + I - {#a#} = (J - {#a#}) + I" by simp lemmas diff_union_single_convs = diff_union_single_conv diff_union_single_conv2 lemma mset_contains_eq: "(m \# M) = ({#m#}+(M-{#m#})=M)" using diff_single_trivial by fastforce subsubsection \Pointwise ordering\ (*declare mset_le_trans[trans] Seems to be in there now. Why is this not done in Multiset.thy or order-class ? *) lemma mset_le_incr_right1: "a\#(b::'a multiset) \ a\#b+c" using mset_subset_eq_mono_add[of a b "{#}" c, simplified] . lemma mset_le_incr_right2: "a\#(b::'a multiset) \ a\#c+b" using mset_le_incr_right1 by (auto simp add: union_commute) lemmas mset_le_incr_right = mset_le_incr_right1 mset_le_incr_right2 lemma mset_le_decr_left1: "a+c\#(b::'a multiset) \ a\#b" using mset_le_incr_right1 mset_subset_eq_mono_add_right_cancel by blast lemma mset_le_decr_left2: "c+a\#(b::'a multiset) \ a\#b" using mset_le_decr_left1 by (auto simp add: union_ac) lemma mset_le_add_mset_decr_left1: "add_mset c a\#(b::'a multiset) \ a\#b" by (simp add: mset_subset_eq_insertD subset_mset.dual_order.strict_implies_order) lemma mset_le_add_mset_decr_left2: "add_mset c a\#(b::'a multiset) \ {#c#}\#b" by (simp add: mset_subset_eq_insertD subset_mset.dual_order.strict_implies_order) lemmas mset_le_decr_left = mset_le_decr_left1 mset_le_decr_left2 mset_le_add_mset_decr_left1 mset_le_add_mset_decr_left2 lemma mset_le_subtract: "A\#B \ A-C \# B-(C::'a multiset)" apply (unfold subseteq_mset_def count_diff) apply clarify apply (subgoal_tac "count A a \ count B a") apply arith apply simp done lemma mset_union_subset: "A+B \# C \ A\#C \ B\#(C::'a multiset)" by (auto dest: mset_le_decr_left) lemma mset_le_add_mset: "add_mset x B \# C \ {#x#}\#C \ B\#(C::'a multiset)" by (auto dest: mset_le_decr_left) lemma mset_le_subtract_left: "A+B \# (X::'a multiset) \ B \# X-A \ A\#X" by (auto dest: mset_le_subtract[of "A+B" "X" "A"] mset_union_subset) lemma mset_le_subtract_right: "A+B \# (X::'a multiset) \ A \# X-B \ B\#X" by (auto dest: mset_le_subtract[of "A+B" "X" "B"] mset_union_subset) lemma mset_le_subtract_add_mset_left: "add_mset x B \# (X::'a multiset) \ B \# X-{#x#} \ {#x#}\#X" by (auto dest: mset_le_subtract[of "add_mset x B" "X" "{#x#}"] mset_le_add_mset) lemma mset_le_subtract_add_mset_right: "add_mset x B \# (X::'a multiset) \ {#x#} \# X-B \ B\#X" by (auto dest: mset_le_subtract[of "add_mset x B" "X" "B"] mset_le_add_mset) lemma mset_le_addE: "\ xs \# (ys::'a multiset); !!zs. ys=xs+zs \ P \ \ P" using mset_subset_eq_exists_conv by blast declare subset_mset.diff_add[simp, intro] lemma mset_2dist2_cases: assumes A: "{#a#}+{#b#} \# A+B" assumes CASES: "{#a#}+{#b#} \# A \ P" "{#a#}+{#b#} \# B \ P" "\a \# A; b \# B\ \ P" "\a \# B; b \# A\ \ P" shows "P" proof - { assume C: "a \# A" "b \# A-{#a#}" with mset_subset_eq_mono_add[of "{#a#}" "{#a#}" "{#b#}" "A-{#a#}"] have "{#a#}+{#b#} \# A" by auto } moreover { assume C: "a \# A" "\ (b \# A-{#a#})" with A have "b \# B" by (metis diff_union_single_conv2 mset_le_subtract_left mset_subset_eq_insertD mset_un_cases) } moreover { assume C: "\ (a \# A)" "b \# B-{#a#}" with A have "a \# B" by (auto dest: mset_subset_eqD) with C mset_subset_eq_mono_add[of "{#a#}" "{#a#}" "{#b#}" "B-{#a#}"] have "{#a#}+{#b#} \# B" by auto } moreover { assume C: "\ (a \# A)" "\ (b \# B-{#a#})" with A have "a \# B \ b \# A" apply (intro conjI) apply (auto dest!: mset_subset_eq_insertD simp: insert_union_subset_iff; fail)[] by (metis mset_diff_cancel1elem mset_le_subtract_left multiset_diff_union_assoc single_subset_iff subset_eq_diff_conv) } ultimately show P using CASES by blast qed lemma mset_union_subset_s: "{#a#}+B \# C \ a \# C \ B \# C" by (drule mset_union_subset) simp (* TODO: Check which of these lemmas are already introduced by order-classes ! *) lemma mset_le_single_cases[consumes 1, case_names empty singleton]: "\M\#{#a#}; M={#} \ P; M={#a#} \ P\ \ P" by (induct M) auto lemma mset_le_distrib[consumes 1, case_names dist]: "\(X::'a multiset)\#A+B; !!Xa Xb. \X=Xa+Xb; Xa\#A; Xb\#B\ \ P \ \ P" by (auto elim!: mset_le_addE mset_distrib) lemma mset_le_mono_add_single: "\a \# ys; b \# ws\ \ {#a#} + {#b#} \# ys + ws" by (meson mset_subset_eq_mono_add single_subset_iff) lemma mset_size1elem: "\size P \ 1; q \# P\ \ P={#q#}" by (auto elim: mset_size_le1_cases) lemma mset_size2elem: "\size P \ 2; {#q#}+{#q'#} \# P\ \ P={#q#}+{#q'#}" by (auto elim: mset_le_addE) subsubsection \Image under function\ notation image_mset (infixr "`#" 90) lemma mset_map_split_orig: "!!M1 M2. \f `# P = M1+M2; !!P1 P2. \P=P1+P2; f `# P1 = M1; f `# P2 = M2\ \ Q \ \ Q" by (induct P) (force elim!: mset_un_single_un_cases)+ lemma mset_map_id: "\!!x. f (g x) = x\ \ f `# g `# X = X" by (induct X) auto text \The following is a very specialized lemma. Intuitively, it splits the original multiset by a splitting of some pointwise supermultiset of its image. Application: This lemma came in handy when proving the correctness of a constraint system that collects at most k sized submultisets of the sets of spawned threads. \ lemma mset_map_split_orig_le: assumes A: "f `# P \# M1+M2" and EX: "!!P1 P2. \P=P1+P2; f `# P1 \# M1; f `# P2 \# M2\ \ Q" shows "Q" using A EX by (auto elim: mset_le_distrib mset_map_split_orig) subsection \Lists\ lemma len_greater_imp_nonempty[simp]: "length l > x \ l\[]" by auto lemma list_take_induct_tl2: "\length xs = length ys; \n \ \n < length (tl xs). P ((tl ys) ! n) ((tl xs) ! n)" by (induct xs ys rule: list_induct2) auto lemma not_distinct_split_distinct: assumes "\ distinct xs" obtains y ys zs where "distinct ys" "y \ set ys" "xs = ys@[y]@zs" using assms proof (induct xs rule: rev_induct) case Nil thus ?case by simp next case (snoc x xs) thus ?case by (cases "distinct xs") auto qed lemma distinct_length_le: assumes d: "distinct ys" and eq: "set ys = set xs" shows "length ys \ length xs" proof - from d have "length ys = card (set ys)" by (simp add: distinct_card) also from eq List.card_set have "card (set ys) = length (remdups xs)" by simp also have "... \ length xs" by simp finally show ?thesis . qed lemma find_SomeD: "List.find P xs = Some x \ P x" "List.find P xs = Some x \ x\set xs" by (auto simp add: find_Some_iff) lemma in_hd_or_tl_conv[simp]: "l\[] \ x=hd l \ x\set (tl l) \ x\set l" by (cases l) auto lemma length_dropWhile_takeWhile: assumes "x < length (dropWhile P xs)" shows "x + length (takeWhile P xs) < length xs" using assms by (induct xs) auto text \Elim-version of @{thm neq_Nil_conv}.\ lemma neq_NilE: assumes "l\[]" obtains x xs where "l=x#xs" using assms by (metis list.exhaust) lemma length_Suc_rev_conv: "length xs = Suc n \ (\ys y. xs=ys@[y] \ length ys = n)" by (cases xs rule: rev_cases) auto subsubsection \List Destructors\ lemma not_hd_in_tl: "x \ hd xs \ x \ set xs \ x \ set (tl xs)" by (induct xs) simp_all lemma distinct_hd_tl: "distinct xs \ x = hd xs \ x \ set (tl (xs))" by (induct xs) simp_all lemma in_set_tlD: "x \ set (tl xs) \ x \ set xs" by (induct xs) simp_all lemma nth_tl: "xs \ [] \ tl xs ! n = xs ! Suc n" by (induct xs) simp_all lemma tl_subset: "xs \ [] \ set xs \ A \ set (tl xs) \ A" by (metis in_set_tlD rev_subsetD subsetI) lemma tl_last: "tl xs \ [] \ last xs = last (tl xs)" by (induct xs) simp_all lemma tl_obtain_elem: assumes "xs \ []" "tl xs = []" obtains e where "xs = [e]" using assms by (induct xs rule: list_nonempty_induct) simp_all lemma butlast_subset: "xs \ [] \ set xs \ A \ set (butlast xs) \ A" by (metis in_set_butlastD rev_subsetD subsetI) lemma butlast_rev_tl: "xs \ [] \ butlast (rev xs) = rev (tl xs)" by (induct xs rule: rev_induct) simp_all lemma hd_butlast: "length xs > 1 \ hd (butlast xs) = hd xs" by (induct xs) simp_all lemma butlast_upd_last_eq[simp]: "length l \ 2 \ (butlast l) [ length l - 2 := x ] = take (length l - 2) l @ [x]" apply (case_tac l rule: rev_cases) apply simp apply simp apply (case_tac ys rule: rev_cases) apply simp apply simp done lemma distinct_butlast_swap[simp]: "distinct pq \ distinct (butlast (pq[i := last pq]))" apply (cases pq rule: rev_cases) apply (auto simp: list_update_append distinct_list_update split: nat.split) done subsubsection \Splitting list according to structure of other list\ context begin private definition "SPLIT_ACCORDING m l \ length l = length m" private lemma SPLIT_ACCORDINGE: assumes "length m = length l" obtains "SPLIT_ACCORDING m l" unfolding SPLIT_ACCORDING_def using assms by auto private lemma SPLIT_ACCORDING_simp: "SPLIT_ACCORDING m (l1@l2) \ (\m1 m2. m=m1@m2 \ SPLIT_ACCORDING m1 l1 \ SPLIT_ACCORDING m2 l2)" "SPLIT_ACCORDING m (x#l') \ (\y m'. m=y#m' \ SPLIT_ACCORDING m' l')" apply (fastforce simp: SPLIT_ACCORDING_def intro: exI[where x = "take (length l1) m"] exI[where x = "drop (length l1) m"]) apply (cases m;auto simp: SPLIT_ACCORDING_def) done text \Split structure of list @{term m} according to structure of list @{term l}.\ method split_list_according for m :: "'a list" and l :: "'b list" = (rule SPLIT_ACCORDINGE[of m l], (simp; fail), ( simp only: SPLIT_ACCORDING_simp, elim exE conjE, simp only: SPLIT_ACCORDING_def ) ) end subsubsection \\list_all2\\ lemma list_all2_induct[consumes 1, case_names Nil Cons]: assumes "list_all2 P l l'" assumes "Q [] []" assumes "\x x' ls ls'. \ P x x'; list_all2 P ls ls'; Q ls ls' \ \ Q (x#ls) (x'#ls')" shows "Q l l'" using list_all2_lengthD[OF assms(1)] assms apply (induct rule: list_induct2) apply auto done subsubsection \Indexing\ lemma ran_nth_set_encoding_conv[simp]: "ran (\i. if ii l[j:=x]!i = l!i" apply (induction l arbitrary: i j) apply (auto split: nat.splits) done lemma nth_list_update': "l[i:=x]!j = (if i=j \ i length l \ n\0 \ last (take n l) = l!(n - 1)" apply (induction l arbitrary: n) apply (auto simp: take_Cons split: nat.split) done lemma nth_append_first[simp]: "i (l@l')!i = l!i" by (simp add: nth_append) lemma in_set_image_conv_nth: "f x \ f`set l \ (\ii. i f (l!i) = f (l'!i)" shows "f`set l = f`set l'" using assms by (fastforce simp: in_set_conv_nth in_set_image_conv_nth) lemma insert_swap_set_eq: "i insert (l!i) (set (l[i:=x])) = insert x (set l)" by (auto simp: in_set_conv_nth nth_list_update split: if_split_asm) subsubsection \Reverse lists\ lemma neq_Nil_revE: assumes "l\[]" obtains ll e where "l = ll@[e]" using assms by (cases l rule: rev_cases) auto lemma neq_Nil_rev_conv: "l\[] \ (\xs x. l = xs@[x])" by (cases l rule: rev_cases) auto text \Caution: Same order of case variables in snoc-case as @{thm [source] rev_exhaust}, the other way round than @{thm [source] rev_induct} !\ lemma length_compl_rev_induct[case_names Nil snoc]: "\P []; !! l e . \!! ll . length ll <= length l \ P ll\ \ P (l@[e])\ \ P l" apply(induct_tac l rule: length_induct) apply(case_tac "xs" rule: rev_cases) apply(auto) done lemma list_append_eq_Cons_cases[consumes 1]: "\ys@zs = x#xs; \ys=[]; zs=x#xs\ \ P; !!ys'. \ ys=x#ys'; ys'@zs=xs \ \ P \ \ P" by (auto iff add: append_eq_Cons_conv) lemma list_Cons_eq_append_cases[consumes 1]: "\x#xs = ys@zs; \ys=[]; zs=x#xs\ \ P; !!ys'. \ ys=x#ys'; ys'@zs=xs \ \ P \ \ P" by (auto iff add: Cons_eq_append_conv) lemma map_of_rev_distinct[simp]: "distinct (map fst m) \ map_of (rev m) = map_of m" apply (induct m) apply simp apply simp apply (subst map_add_comm) apply force apply simp done \ \Tail-recursive, generalized @{const rev}. May also be used for tail-recursively getting a list with all elements of the two operands, if the order does not matter, e.g. when implementing sets by lists.\ fun revg where "revg [] b = b" | "revg (a#as) b = revg as (a#b)" lemma revg_fun[simp]: "revg a b = rev a @ b" by (induct a arbitrary: b) auto lemma rev_split_conv[simp]: "l \ [] \ rev (tl l) @ [hd l] = rev l" by (induct l) simp_all lemma rev_butlast_is_tl_rev: "rev (butlast l) = tl (rev l)" by (induct l) auto lemma hd_last_singletonI: "\xs \ []; hd xs = last xs; distinct xs\ \ xs = [hd xs]" by (induct xs rule: list_nonempty_induct) auto lemma last_filter: "\xs \ []; P (last xs)\ \ last (filter P xs) = last xs" by (induct xs rule: rev_nonempty_induct) simp_all (* As the following two rules are similar in nature to list_induct2', they are named accordingly. *) lemma rev_induct2' [case_names empty snocl snocr snoclr]: assumes empty: "P [] []" and snocl: "\x xs. P (xs@[x]) []" and snocr: "\y ys. P [] (ys@[y])" and snoclr: "\x xs y ys. P xs ys \ P (xs@[x]) (ys@[y])" shows "P xs ys" proof (induct xs arbitrary: ys rule: rev_induct) case Nil thus ?case using empty snocr by (cases ys rule: rev_exhaust) simp_all next case snoc thus ?case using snocl snoclr by (cases ys rule: rev_exhaust) simp_all qed lemma rev_nonempty_induct2' [case_names single snocl snocr snoclr, consumes 2]: assumes "xs \ []" "ys \ []" assumes single': "\x y. P [x] [y]" and snocl: "\x xs y. xs \ [] \ P (xs@[x]) [y]" and snocr: "\x y ys. ys \ [] \ P [x] (ys@[y])" and snoclr: "\x xs y ys. \P xs ys; xs \ []; ys\[]\ \ P (xs@[x]) (ys@[y])" shows "P xs ys" using assms(1,2) proof (induct xs arbitrary: ys rule: rev_nonempty_induct) case single then obtain z zs where "ys = zs@[z]" by (metis rev_exhaust) thus ?case using single' snocr by (cases "zs = []") simp_all next case (snoc x xs) then obtain z zs where zs: "ys = zs@[z]" by (metis rev_exhaust) thus ?case using snocl snoclr snoc by (cases "zs = []") simp_all qed subsubsection "Folding" text "Ugly lemma about foldl over associative operator with left and right neutral element" lemma foldl_A1_eq: "!!i. \ !! e. f n e = e; !! e. f e n = e; !!a b c. f a (f b c) = f (f a b) c \ \ foldl f i ww = f i (foldl f n ww)" proof (induct ww) case Nil thus ?case by simp next case (Cons a ww i) note IHP[simplified]=this have "foldl f i (a # ww) = foldl f (f i a) ww" by simp also from IHP have "\ = f (f i a) (foldl f n ww)" by blast also from IHP(4) have "\ = f i (f a (foldl f n ww))" by simp also from IHP(1)[OF IHP(2,3,4), where i=a] have "\ = f i (foldl f a ww)" by simp also from IHP(2)[of a] have "\ = f i (foldl f (f n a) ww)" by simp also have "\ = f i (foldl f n (a#ww))" by simp finally show ?case . qed lemmas foldl_conc_empty_eq = foldl_A1_eq[of "(@)" "[]", simplified] lemmas foldl_un_empty_eq = foldl_A1_eq[of "(\)" "{}", simplified, OF Un_assoc[symmetric]] lemma foldl_set: "foldl (\) {} l = \{x. x\set l}" apply (induct l) apply simp_all apply (subst foldl_un_empty_eq) apply auto done lemma (in monoid_mult) foldl_absorb1: "x*foldl (*) 1 zs = foldl (*) x zs" apply (rule sym) apply (rule foldl_A1_eq) apply (auto simp add: mult.assoc) done text \Towards an invariant rule for foldl\ lemma foldl_rule_aux: fixes I :: "'\ \ 'a list \ bool" assumes initial: "I \0 l0" assumes step: "!!l1 l2 x \. \ l0=l1@x#l2; I \ (x#l2) \ \ I (f \ x) l2" shows "I (foldl f \0 l0) []" using initial step apply (induct l0 arbitrary: \0) apply auto done lemma foldl_rule_aux_P: fixes I :: "'\ \ 'a list \ bool" assumes initial: "I \0 l0" assumes step: "!!l1 l2 x \. \ l0=l1@x#l2; I \ (x#l2) \ \ I (f \ x) l2" assumes final: "!!\. I \ [] \ P \" shows "P (foldl f \0 l0)" using foldl_rule_aux[of I \0 l0, OF initial, OF step] final by simp lemma foldl_rule: fixes I :: "'\ \ 'a list \ 'a list \ bool" assumes initial: "I \0 [] l0" assumes step: "!!l1 l2 x \. \ l0=l1@x#l2; I \ l1 (x#l2) \ \ I (f \ x) (l1@[x]) l2" shows "I (foldl f \0 l0) l0 []" using initial step apply (rule_tac I="\\ lr. \ll. l0=ll@lr \ I \ ll lr" in foldl_rule_aux_P) apply auto done text \ Invariant rule for foldl. The invariant is parameterized with the state, the list of items that have already been processed and the list of items that still have to be processed. \ lemma foldl_rule_P: fixes I :: "'\ \ 'a list \ 'a list \ bool" \ \The invariant holds for the initial state, no items processed yet and all items to be processed:\ assumes initial: "I \0 [] l0" \ \The invariant remains valid if one item from the list is processed\ assumes step: "!!l1 l2 x \. \ l0=l1@x#l2; I \ l1 (x#l2) \ \ I (f \ x) (l1@[x]) l2" \ \The proposition follows from the invariant in the final state, i.e. all items processed and nothing to be processed\ assumes final: "!!\. I \ l0 [] \ P \" shows "P (foldl f \0 l0)" using foldl_rule[of I, OF initial step] by (simp add: final) text \Invariant reasoning over @{const foldl} for distinct lists. Invariant rule makes no assumptions about ordering.\ lemma distinct_foldl_invar: "\ distinct S; I (set S) \0; \x it \. \x \ it; it \ set S; I it \\ \ I (it - {x}) (f \ x) \ \ I {} (foldl f \0 S)" proof (induct S arbitrary: \0) case Nil thus ?case by auto next case (Cons x S) note [simp] = Cons.prems(1)[simplified] show ?case apply simp apply (rule Cons.hyps) proof - from Cons.prems(1) show "distinct S" by simp from Cons.prems(3)[of x "set (x#S)", simplified, OF Cons.prems(2)[simplified]] show "I (set S) (f \0 x)" . fix xx it \ assume A: "xx\it" "it \ set S" "I it \" show "I (it - {xx}) (f \ xx)" using A(2) apply (rule_tac Cons.prems(3)) apply (simp_all add: A(1,3)) apply blast done qed qed lemma foldl_length_aux: "foldl (\i x. Suc i) a l = a + length l" by (induct l arbitrary: a) auto lemmas foldl_length[simp] = foldl_length_aux[where a=0, simplified] lemma foldr_length_aux: "foldr (\x i. Suc i) l a = a + length l" by (induct l arbitrary: a rule: rev_induct) auto lemmas foldr_length[simp] = foldr_length_aux[where a=0, simplified] context comp_fun_commute begin lemma foldl_f_commute: "f a (foldl (\a b. f b a) b xs) = foldl (\a b. f b a) (f a b) xs" by(induct xs arbitrary: b)(simp_all add: fun_left_comm) lemma foldr_conv_foldl: "foldr f xs a = foldl (\a b. f b a) a xs" by(induct xs arbitrary: a)(simp_all add: foldl_f_commute) end lemma filter_conv_foldr: "filter P xs = foldr (\x xs. if P x then x # xs else xs) xs []" by(induct xs) simp_all lemma foldr_Cons: "foldr Cons xs [] = xs" by(induct xs) simp_all lemma foldr_snd_zip: "length xs \ length ys \ foldr (\(x, y). f y) (zip xs ys) b = foldr f ys b" proof(induct ys arbitrary: xs) case (Cons y ys) thus ?case by(cases xs) simp_all qed simp lemma foldl_snd_zip: "length xs \ length ys \ foldl (\b (x, y). f b y) b (zip xs ys) = foldl f b ys" proof(induct ys arbitrary: xs b) case (Cons y ys) thus ?case by(cases xs) simp_all qed simp lemma fst_foldl: "fst (foldl (\(a, b) x. (f a x, g a b x)) (a, b) xs) = foldl f a xs" by(induct xs arbitrary: a b) simp_all lemma foldl_foldl_conv_concat: "foldl (foldl f) a xs = foldl f a (concat xs)" by(induct xs arbitrary: a) simp_all lemma foldl_list_update: "n < length xs \ foldl f a (xs[n := x]) = foldl f (f (foldl f a (take n xs)) x) (drop (Suc n) xs)" by(simp add: upd_conv_take_nth_drop) lemma map_by_foldl: fixes l :: "'a list" and f :: "'a \ 'b" shows "foldl (\l x. l@[f x]) [] l = map f l" proof - { fix l' have "foldl (\l x. l@[f x]) l' l = l'@map f l" by (induct l arbitrary: l') auto } thus ?thesis by simp qed subsubsection \Sorting\ lemma sorted_in_between: assumes A: "0\i" "i x" "xk" and "kx" and "xk. i\k \ k l!k\x \ x x") case True from True Suc.hyps have "d = j - (i + 1)" by simp moreover from True have "i+1 < j" by (metis Suc.prems Suc_eq_plus1 Suc_lessI not_less) moreover from True have "0\i+1" by simp ultimately obtain k where "i+1\k" "k x" "xk" "k x" "xsorted l; l\[]\ \ hd l \ last l" by (metis eq_iff hd_Cons_tl last_in_set not_hd_in_tl sorted_wrt.simps(2)) lemma (in linorder) sorted_hd_min: "\xs \ []; sorted xs\ \ \x \ set xs. hd xs \ x" by (induct xs, auto) lemma sorted_append_bigger: "\sorted xs; \x \ set xs. x \ y\ \ sorted (xs @ [y])" proof (induct xs) case Nil then show ?case by simp next case (Cons x xs) then have s: "sorted xs" by (cases xs) simp_all from Cons have a: "\x\set xs. x \ y" by simp from Cons(1)[OF s a] Cons(2-) show ?case by (cases xs) simp_all qed lemma sorted_filter': "sorted l \ sorted (filter P l)" using sorted_filter[where f=id, simplified] . subsubsection \Map\ (* List.thy has: declare map_eq_Cons_D [dest!] Cons_eq_map_D [dest!] *) lemma map_eq_consE: "\map f ls = fa#fl; !!a l. \ ls=a#l; f a=fa; map f l = fl \ \ P\ \ P" by auto lemma map_fst_mk_snd[simp]: "map fst (map (\x. (x,k)) l) = l" by (induct l) auto lemma map_snd_mk_fst[simp]: "map snd (map (\x. (k,x)) l) = l" by (induct l) auto lemma map_fst_mk_fst[simp]: "map fst (map (\x. (k,x)) l) = replicate (length l) k" by (induct l) auto lemma map_snd_mk_snd[simp]: "map snd (map (\x. (x,k)) l) = replicate (length l) k" by (induct l) auto lemma map_zip1: "map (\x. (x,k)) l = zip l (replicate (length l) k)" by (induct l) auto lemma map_zip2: "map (\x. (k,x)) l = zip (replicate (length l) k) l" by (induct l) auto lemmas map_zip=map_zip1 map_zip2 (* TODO/FIXME: hope nobody changes nth to be underdefined! *) lemma map_eq_nth_eq: assumes A: "map f l = map f l'" shows "f (l!i) = f (l'!i)" proof - from A have "length l = length l'" by (metis length_map) thus ?thesis using A apply (induct arbitrary: i rule: list_induct2) apply simp apply (simp add: nth_def split: nat.split) done qed lemma map_upd_eq: "\i f (l!i) = f x\ \ map f (l[i:=x]) = map f l" by (metis list_update_beyond list_update_id map_update not_le_imp_less) lemma inj_map_inv_f [simp]: "inj f \ map (inv f) (map f l) = l" by (simp) lemma inj_on_map_the: "\D \ dom m; inj_on m D\ \ inj_on (the\m) D" apply (rule inj_onI) apply simp apply (case_tac "m x") apply (case_tac "m y") apply (auto intro: inj_onD) [1] apply (auto intro: inj_onD) [1] apply (case_tac "m y") apply (auto intro: inj_onD) [1] apply simp apply (rule inj_onD) apply assumption apply auto done lemma map_consI: "w=map f ww \ f a#w = map f (a#ww)" "w@l=map f ww@l \ f a#w@l = map f (a#ww)@l" by auto lemma restrict_map_subset_eq: fixes R shows "\m |` R = m'; R'\R\ \ m|` R' = m' |` R'" by (auto simp add: Int_absorb1) lemma restrict_map_self[simp]: "m |` dom m = m" apply (rule ext) apply (case_tac "m x") apply (auto simp add: restrict_map_def) done lemma restrict_map_UNIV[simp]: "f |` UNIV = f" by (auto simp add: restrict_map_def) lemma restrict_map_inv[simp]: "f |` (- dom f) = Map.empty" by (auto simp add: restrict_map_def intro: ext) lemma restrict_map_upd: "(f |` S)(k \ v) = f(k\v) |` (insert k S)" by (auto simp add: restrict_map_def intro: ext) lemma map_upd_eq_restrict: "m (x:=None) = m |` (-{x})" by (auto intro: ext) declare Map.finite_dom_map_of [simp, intro!] lemma dom_const'[simp]: "dom (\x. Some (f x)) = UNIV" by auto lemma restrict_map_eq : "((m |` A) k = None) \ (k \ dom m \ A)" "((m |` A) k = Some v) \ (m k = Some v \ k \ A)" unfolding restrict_map_def by (simp_all add: dom_def) definition "rel_of m P == {(k,v). m k = Some v \ P (k, v)}" lemma rel_of_empty[simp]: "rel_of Map.empty P = {}" by (auto simp add: rel_of_def) lemma remove1_tl: "xs \ [] \ remove1 (hd xs) xs = tl xs" by (cases xs) auto lemma set_oo_map_alt: "(set \\ map) f = (\l. f ` set l)" by auto subsubsection "Filter and Revert" primrec filter_rev_aux where "filter_rev_aux a P [] = a" | "filter_rev_aux a P (x#xs) = ( if P x then filter_rev_aux (x#a) P xs else filter_rev_aux a P xs)" lemma filter_rev_aux_alt: "filter_rev_aux a P l = filter P (rev l) @ a" by (induct l arbitrary: a) auto definition "filter_rev == filter_rev_aux []" lemma filter_rev_alt: "filter_rev P l = filter P (rev l)" unfolding filter_rev_def by (simp add: filter_rev_aux_alt) definition "remove_rev x == filter_rev (Not o (=) x)" lemma remove_rev_alt_def : "remove_rev x xs = (filter (\y. y \ x) (rev xs))" unfolding remove_rev_def apply (simp add: filter_rev_alt comp_def) by metis subsubsection "zip" declare zip_map_fst_snd[simp] lemma pair_list_split: "\ !!l1 l2. \ l = zip l1 l2; length l1=length l2; length l=length l2 \ \ P \ \ P" proof (induct l arbitrary: P) case Nil thus ?case by auto next case (Cons a l) from Cons.hyps obtain l1 l2 where IHAPP: "l=zip l1 l2" "length l1 = length l2" "length l=length l2" . obtain a1 a2 where [simp]: "a=(a1,a2)" by (cases a) auto from IHAPP have "a#l = zip (a1#l1) (a2#l2)" "length (a1#l1) = length (a2#l2)" "length (a#l) = length (a2#l2)" by (simp_all only:) (simp_all (no_asm_use)) with Cons.prems show ?case by blast qed lemma set_zip_cart: "x\set (zip l l') \ x\set l \ set l'" by (auto simp add: set_zip) lemma map_prod_fun_zip: "map (\(x, y). (f x, g y)) (zip xs ys) = zip (map f xs) (map g ys)" proof(induct xs arbitrary: ys) case Nil thus ?case by simp next case (Cons x xs) thus ?case by(cases ys) simp_all qed subsubsection \Generalized Zip\ text \Zip two lists element-wise, where the combination of two elements is specified by a function. Note that this function is underdefined for lists of different length.\ fun zipf :: "('a\'b\'c) \ 'a list \ 'b list \ 'c list" where "zipf f [] [] = []" | "zipf f (a#as) (b#bs) = f a b # zipf f as bs" lemma zipf_zip: "\length l1 = length l2\ \ zipf Pair l1 l2 = zip l1 l2" apply (induct l1 arbitrary: l2) apply auto apply (case_tac l2) apply auto done \ \All quantification over zipped lists\ fun list_all_zip where "list_all_zip P [] [] \ True" | "list_all_zip P (a#as) (b#bs) \ P a b \ list_all_zip P as bs" | "list_all_zip P _ _ \ False" lemma list_all_zip_alt: "list_all_zip P as bs \ length as = length bs \ (\iP as bs rule: list_all_zip.induct) apply auto apply (case_tac i) apply auto done lemma list_all_zip_map1: "list_all_zip P (List.map f as) bs \ list_all_zip (\a b. P (f a) b) as bs" apply (induct as arbitrary: bs) apply (case_tac bs) apply auto [2] apply (case_tac bs) apply auto [2] done lemma list_all_zip_map2: "list_all_zip P as (List.map f bs) \ list_all_zip (\a b. P a (f b)) as bs" apply (induct as arbitrary: bs) apply (case_tac bs) apply auto [2] apply (case_tac bs) apply auto [2] done declare list_all_zip_alt[mono] lemma lazI[intro?]: "\ length a = length b; !!i. i P (a!i) (b!i) \ \ list_all_zip P a b" by (auto simp add: list_all_zip_alt) lemma laz_conj[simp]: "list_all_zip (\x y. P x y \ Q x y) a b \ list_all_zip P a b \ list_all_zip Q a b" by (auto simp add: list_all_zip_alt) lemma laz_len: "list_all_zip P a b \ length a = length b" by (simp add: list_all_zip_alt) lemma laz_eq: "list_all_zip (=) a b \ a=b" apply (induct a arbitrary: b) apply (case_tac b) apply simp apply simp apply (case_tac b) apply simp apply simp done lemma laz_swap_ex: assumes A: "list_all_zip (\a b. \c. P a b c) A B" obtains C where "list_all_zip (\a c. \b. P a b c) A C" "list_all_zip (\b c. \a. P a b c) B C" proof - from A have [simp]: "length A = length B" and IC: "\ici. P (A!i) (B!i) ci" by (auto simp add: list_all_zip_alt) from obtain_list_from_elements[OF IC] obtain C where "length C = length B" "\ia b. P a) A B \ (length A = length B) \ (\a\set A. P a)" by (auto simp add: list_all_zip_alt set_conv_nth) lemma laz_weak_Pb[simp]: "list_all_zip (\a b. P b) A B \ (length A = length B) \ (\b\set B. P b)" by (force simp add: list_all_zip_alt set_conv_nth) subsubsection "Collecting Sets over Lists" definition "list_collect_set f l == \{ f a | a. a\set l }" lemma list_collect_set_simps[simp]: "list_collect_set f [] = {}" "list_collect_set f [a] = f a" "list_collect_set f (a#l) = f a \ list_collect_set f l" "list_collect_set f (l@l') = list_collect_set f l \ list_collect_set f l'" by (unfold list_collect_set_def) auto lemma list_collect_set_map_simps[simp]: "list_collect_set f (map x []) = {}" "list_collect_set f (map x [a]) = f (x a)" "list_collect_set f (map x (a#l)) = f (x a) \ list_collect_set f (map x l)" "list_collect_set f (map x (l@l')) = list_collect_set f (map x l) \ list_collect_set f (map x l')" by simp_all lemma list_collect_set_alt: "list_collect_set f l = \{ f (l!i) | i. i(set (map f l))" by (unfold list_collect_set_def) auto subsubsection \Sorted List with arbitrary Relations\ lemma (in linorder) sorted_wrt_rev_linord [simp] : "sorted_wrt (\) l \ sorted (rev l)" by (simp add: sorted_wrt_rev) lemma (in linorder) sorted_wrt_map_linord [simp] : "sorted_wrt (\(x::'a \ 'b) y. fst x \ fst y) l \ sorted (map fst l)" by (simp add: sorted_wrt_map) lemma (in linorder) sorted_wrt_map_rev_linord [simp] : "sorted_wrt (\(x::'a \ 'b) y. fst x \ fst y) l \ sorted (rev (map fst l))" by (induct l) (auto simp add: sorted_append) subsubsection \Take and Drop\ lemma take_update[simp]: "take n (l[i:=x]) = (take n l)[i:=x]" apply (induct l arbitrary: n i) apply (auto split: nat.split) apply (case_tac n) apply simp_all apply (case_tac n) apply simp_all done lemma take_update_last: "length list>n \ (take (Suc n) list) [n:=x] = take n list @ [x]" by (induct list arbitrary: n) (auto split: nat.split) lemma drop_upd_irrelevant: "m < n \ drop n (l[m:=x]) = drop n l" apply (induct n arbitrary: l m) apply simp apply (case_tac l) apply (auto split: nat.split) done lemma set_drop_conv: "set (drop n l) = { l!i | i. n\i \ i < length l }" (is "?L=?R") proof (intro equalityI subsetI) fix x assume "x\?L" then obtain i where L: "i = l!(n+i)" using L by simp finally show "x\?R" using L by auto next fix x assume "x\?R" then obtain i where L: "n\i" "i?L" by (auto simp add: in_set_conv_nth) qed lemma filter_upt_take_conv: "[i\[n..[n..set (drop n l) \ (\i. n\i \ i x = l!i)" apply (clarsimp simp: in_set_conv_nth) apply safe apply simp apply (metis le_add2 less_diff_conv add.commute) apply (rule_tac x="i-n" in exI) apply auto [] done lemma Union_take_drop_id: "\(set (drop n l)) \ \(set (take n l)) = \(set l)" by (metis Union_Un_distrib append_take_drop_id set_union_code sup_commute) lemma Un_set_drop_extend: "\j\Suc 0; j < length l\ \ l ! (j - Suc 0) \ \(set (drop j l)) = \(set (drop (j - Suc 0) l))" apply safe apply simp_all apply (metis diff_Suc_Suc diff_zero gr0_implies_Suc in_set_drop_conv_nth le_refl less_eq_Suc_le order.strict_iff_order) apply (metis Nat.diff_le_self set_drop_subset_set_drop subset_code(1)) by (metis diff_Suc_Suc gr0_implies_Suc in_set_drop_conv_nth less_eq_Suc_le order.strict_iff_order minus_nat.diff_0) lemma drop_take_drop_unsplit: "i\j \ drop i (take j l) @ drop j l = drop i l" proof - assume "i \ j" then obtain skf where "i + skf = j" by (metis le_iff_add) thus "drop i (take j l) @ drop j l = drop i l" by (metis append_take_drop_id diff_add_inverse drop_drop drop_take add.commute) qed lemma drop_last_conv[simp]: "l\[] \ drop (length l - Suc 0) l = [last l]" by (cases l rule: rev_cases) auto lemma take_butlast_conv[simp]: "take (length l - Suc 0) l = butlast l" by (cases l rule: rev_cases) auto lemma drop_takeWhile: assumes "i\length (takeWhile P l)" shows "drop i (takeWhile P l) = takeWhile P (drop i l)" using assms proof (induction l arbitrary: i) case Nil thus ?case by auto next case (Cons x l) thus ?case by (cases i) auto qed lemma less_length_takeWhile_conv: "i < length (takeWhile P l) \ (i (\j\i. P (l!j)))" apply safe subgoal using length_takeWhile_le less_le_trans by blast subgoal by (metis dual_order.strict_trans2 nth_mem set_takeWhileD takeWhile_nth) subgoal by (meson less_le_trans not_le_imp_less nth_length_takeWhile) done lemma eq_len_takeWhile_conv: "i=length (takeWhile P l) \ i\length l \ (\j (i \P (l!i))" apply safe subgoal using length_takeWhile_le less_le_trans by blast subgoal by (auto simp: less_length_takeWhile_conv) subgoal using nth_length_takeWhile by blast subgoal by (metis length_takeWhile_le nth_length_takeWhile order.order_iff_strict) subgoal by (metis dual_order.strict_trans2 leI less_length_takeWhile_conv linorder_neqE_nat nth_length_takeWhile) done subsubsection \Up-to\ lemma upt_eq_append_conv: "i\j \ [i.. (\k. i\k \ k\j \ [i.. [k..j" thus "\k\i. k \ j \ [i.. [k.. length l \ map (nth l) [M.. is1 = [l.. is2 = [Suc i.. l\i \ ii. i + ofs) [a..Last and butlast\ lemma butlast_upt: "butlast [m.. butlast [n..length l \ take (n - Suc 0) l = butlast (take n l)" by (simp add: butlast_take) lemma butlast_eq_cons_conv: "butlast l = x#xs \ (\xl. l=x#xs@[xl])" by (metis Cons_eq_appendI append_butlast_last_id butlast.simps butlast_snoc eq_Nil_appendI) lemma butlast_eq_consE: assumes "butlast l = x#xs" obtains xl where "l=x#xs@[xl]" using assms by (auto simp: butlast_eq_cons_conv) lemma drop_eq_ConsD: "drop n xs = x # xs' \ drop (Suc n) xs = xs'" by(induct xs arbitrary: n)(simp_all add: drop_Cons split: nat.split_asm) subsubsection \List Slices\ text \Based on Lars Hupel's code.\ definition slice :: "nat \ nat \ 'a list \ 'a list" where "slice from to list = take (to - from) (drop from list)" lemma slice_len[simp]: "\ from \ to; to \ length xs \ \ length (slice from to xs) = to - from" unfolding slice_def by simp lemma slice_head: "\ from < to; to \ length xs \ \ hd (slice from to xs) = xs ! from" unfolding slice_def proof - assume a1: "from < to" assume "to \ length xs" then have "\n. to - (to - n) \ length (take to xs)" by (metis (no_types) slice_def diff_le_self drop_take length_drop slice_len) then show "hd (take (to - from) (drop from xs)) = xs ! from" using a1 by (metis diff_diff_cancel drop_take hd_drop_conv_nth leI le_antisym less_or_eq_imp_le nth_take) qed lemma slice_eq_bounds_empty[simp]: "slice i i xs = []" unfolding slice_def by auto lemma slice_nth: "\ from < to; to \ length xs; i < to - from \ \ slice from to xs ! i = xs ! (from + i)" unfolding slice_def by (induction "to - from" arbitrary: "from" to i) simp+ lemma slice_prepend: "\ i \ k; k \ length xs \ \ let p = length ys in slice i k xs = slice (i + p) (k + p) (ys @ xs)" unfolding slice_def Let_def by force lemma slice_Nil[simp]: "slice begin end [] = []" unfolding slice_def by auto lemma slice_Cons: "slice begin end (x#xs) = (if begin=0 \ end>0 then x#slice begin (end-1) xs else slice (begin - 1) (end - 1) xs)" unfolding slice_def by (auto simp: take_Cons' drop_Cons') lemma slice_complete[simp]: "slice 0 (length xs) xs = xs" unfolding slice_def by simp subsubsection \Miscellaneous\ lemma length_compl_induct[case_names Nil Cons]: "\P []; !! e l . \!! ll . length ll <= length l \ P ll\ \ P (e#l)\ \ P l" apply(induct_tac l rule: length_induct) apply(case_tac "xs") apply(auto) done lemma in_set_list_format: "\ e\set l; !!l1 l2. l=l1@e#l2 \ P \ \ P" proof (induct l arbitrary: P) case Nil thus ?case by auto next case (Cons a l) show ?case proof (cases "a=e") case True with Cons show ?thesis by force next case False with Cons.prems(1) have "e\set l" by auto with Cons.hyps obtain l1 l2 where "l=l1@e#l2" by blast hence "a#l = (a#l1)@e#l2" by simp with Cons.prems(2) show P by blast qed qed lemma in_set_upd_cases: assumes "x\set (l[i:=y])" obtains "iset l" by (metis assms in_set_conv_nth length_list_update nth_list_update_eq nth_list_update_neq) lemma in_set_upd_eq_aux: assumes "iset (l[i:=y]) \ x=y \ (\y. x\set (l[i:=y]))" by (metis in_set_upd_cases assms list_update_overwrite set_update_memI) lemma in_set_upd_eq: assumes "iset (l[i:=y]) \ x=y \ (x\set l \ (\y. x\set (l[i:=y])))" by (metis in_set_upd_cases in_set_upd_eq_aux assms) text \Simultaneous induction over two lists, prepending an element to one of the lists in each step\ lemma list_2pre_induct[case_names base left right]: assumes BASE: "P [] []" and LEFT: "!!e w1' w2. P w1' w2 \ P (e#w1') w2" and RIGHT: "!!e w1 w2'. P w1 w2' \ P w1 (e#w2')" shows "P w1 w2" proof - { \ \The proof is done by induction over the sum of the lengths of the lists\ fix n have "!!w1 w2. \length w1 + length w2 = n; P [] []; !!e w1' w2. P w1' w2 \ P (e#w1') w2; !!e w1 w2'. P w1 w2' \ P w1 (e#w2') \ \ P w1 w2 " apply (induct n) apply simp apply (case_tac w1) apply auto apply (case_tac w2) apply auto done } from this[OF _ BASE LEFT RIGHT] show ?thesis by blast qed lemma list_decomp_1: "length l=1 \ \a. l=[a]" by (case_tac l, auto) lemma list_decomp_2: "length l=2 \ \a b. l=[a,b]" by (case_tac l, auto simp add: list_decomp_1) lemma list_rest_coinc: "\length s2 \ length s1; s1@r1 = s2@r2\ \ \r1p. r2=r1p@r1" by (metis append_eq_append_conv_if) lemma list_tail_coinc: "n1#r1 = n2#r2 \ n1=n2 & r1=r2" by (auto) lemma last_in_set[intro]: "\l\[]\ \ last l \ set l" by (induct l) auto lemma empty_append_eq_id[simp]: "(@) [] = (\x. x)" by auto lemma op_conc_empty_img_id[simp]: "((@) [] ` L) = L" by auto lemma distinct_match: "\ distinct (al@e#bl) \ \ (al@e#bl = al'@e#bl') \ (al=al' \ bl=bl')" proof (rule iffI, induct al arbitrary: al') case Nil thus ?case by (cases al') auto next case (Cons a al) note Cprems=Cons.prems note Chyps=Cons.hyps show ?case proof (cases al') case Nil with Cprems have False by auto thus ?thesis .. next case [simp]: (Cons a' all') with Cprems have [simp]: "a=a'" and P: "al@e#bl = all'@e#bl'" by auto from Cprems(1) have D: "distinct (al@e#bl)" by auto from Chyps[OF D P] have [simp]: "al=all'" "bl=bl'" by auto show ?thesis by simp qed qed simp lemma prop_match: "\ list_all P al; \P e; \P e'; list_all P bl \ \ (al@e#bl = al'@e'#bl') \ (al=al' \ e=e' \ bl=bl')" apply (rule iffI, induct al arbitrary: al') apply (case_tac al', fastforce, fastforce)+ done lemmas prop_matchD = rev_iffD1[OF _ prop_match[where P=P]] for P declare distinct_tl[simp] lemma list_se_match[simp]: "l1 \ [] \ l1@l2 = [a] \ l1 = [a] \ l2 = []" "l2 \ [] \ l1@l2 = [a] \ l1 = [] \ l2 = [a]" "l1 \ [] \ [a] = l1@l2 \ l1 = [a] \ l2 = []" "l2 \ [] \ [a] = l1@l2 \ l1 = [] \ l2 = [a]" apply (cases l1, simp_all) apply (cases l1, simp_all) apply (cases l1, auto) [] apply (cases l1, auto) [] done (* Placed here because it depends on xy_in_set_cases *) lemma distinct_map_eq: "\ distinct (List.map f l); f x = f y; x\set l; y\set l \ \ x=y" by (erule (2) xy_in_set_cases) auto lemma upt_append: assumes "iu'" assumes NP: "\i. u\i \ i \P i" shows "[i\[0..[0..[0..[0..[u..[u..[0.. P (l!i)" proof assume A: "P (l!i)" have "[0..i by (simp add: upt_append) also have "[i..i by (auto simp: upt_conv_Cons) finally have "[k\[0..[Suc i..P (l!i)\ by simp hence "j = last (i#[k\[Suc i.. \ i" proof - have "sorted (i#[k\[Suc i.. last ?l" by (rule sorted_hd_last) simp thus ?thesis by simp qed finally have "i\j" . thus False using \j by simp qed lemma all_set_conv_nth: "(\x\set l. P x) \ (\i *) lemma upt_0_eq_Nil_conv[simp]: "[0.. j=0" by auto lemma filter_eq_snocD: "filter P l = l'@[x] \ x\set l \ P x" proof - assume A: "filter P l = l'@[x]" hence "x\set (filter P l)" by simp thus ?thesis by simp qed lemma lists_image_witness: assumes A: "x\lists (f`Q)" obtains xo where "xo\lists Q" "x=map f xo" proof - have "\ x\lists (f`Q) \ \ \xo\lists Q. x=map f xo" proof (induct x) case Nil thus ?case by auto next case (Cons x xs) then obtain xos where "xos\lists Q" "xs=map f xos" by force moreover from Cons.prems have "x\f`Q" by auto then obtain xo where "xo\Q" "x=f xo" by auto ultimately show ?case by (rule_tac x="xo#xos" in bexI) auto qed thus ?thesis apply (simp_all add: A) apply (erule_tac bexE) apply (rule_tac that) apply assumption+ done qed lemma map_of_None_filterD: "map_of xs x = None \ map_of (filter P xs) x = None" by(induct xs) auto lemma map_of_concat: "map_of (concat xss) = foldr (\xs f. f ++ map_of xs) xss Map.empty" by(induct xss) simp_all lemma map_of_Some_split: "map_of xs k = Some v \ \ys zs. xs = ys @ (k, v) # zs \ map_of ys k = None" proof(induct xs) case (Cons x xs) obtain k' v' where x: "x = (k', v')" by(cases x) show ?case proof(cases "k' = k") case True with \map_of (x # xs) k = Some v\ x have "x # xs = [] @ (k, v) # xs" "map_of [] k = None" by simp_all thus ?thesis by blast next case False with \map_of (x # xs) k = Some v\ x have "map_of xs k = Some v" by simp from \map_of xs k = Some v \ \ys zs. xs = ys @ (k, v) # zs \ map_of ys k = None\[OF this] obtain ys zs where "xs = ys @ (k, v) # zs" "map_of ys k = None" by blast with False x have "x # xs = (x # ys) @ (k, v) # zs" "map_of (x # ys) k = None" by simp_all thus ?thesis by blast qed qed simp lemma map_add_find_left: "g k = None \ (f ++ g) k = f k" by(simp add: map_add_def) lemma map_add_left_None: "f k = None \ (f ++ g) k = g k" by(simp add: map_add_def split: option.split) lemma map_of_Some_filter_not_in: "\ map_of xs k = Some v; \ P (k, v); distinct (map fst xs) \ \ map_of (filter P xs) k = None" apply(induct xs) apply(auto) apply(auto simp add: map_of_eq_None_iff) done lemma distinct_map_fst_filterI: "distinct (map fst xs) \ distinct (map fst (filter P xs))" by(induct xs) auto lemma distinct_map_fstD: "\ distinct (map fst xs); (x, y) \ set xs; (x, z) \ set xs \ \ y = z" by(induct xs)(fastforce elim: notE rev_image_eqI)+ lemma concat_filter_neq_Nil: "concat [ys\xs. ys \ Nil] = concat xs" by(induct xs) simp_all lemma distinct_concat': "\distinct [ys\xs. ys \ Nil]; \ys. ys \ set xs \ distinct ys; \ys zs. \ys \ set xs; zs \ set xs; ys \ zs\ \ set ys \ set zs = {}\ \ distinct (concat xs)" by(erule distinct_concat[of "[ys\xs. ys \ Nil]", unfolded concat_filter_neq_Nil]) auto lemma distinct_idx: assumes "distinct (map f l)" assumes "im. n \ m \ m < length xs \ filter P xs ! n = xs ! m \ filter P (take m xs) = take n (filter P xs)" using assms proof(induct xs rule: rev_induct) case Nil thus ?case by simp next case (snoc x xs) show ?case proof(cases "P x") case [simp]: False from \n < length (filter P (xs @ [x]))\ have "n < length (filter P xs)" by simp hence "\m\n. m < length xs \ filter P xs ! n = xs ! m \ filter P (take m xs) = take n (filter P xs)" by(rule snoc) thus ?thesis by(auto simp add: nth_append) next case [simp]: True show ?thesis proof(cases "n = length (filter P xs)") case False with \n < length (filter P (xs @ [x]))\ have "n < length (filter P xs)" by simp moreover hence "\m\n. m < length xs \ filter P xs ! n = xs ! m \ filter P (take m xs) = take n (filter P xs)" by(rule snoc) ultimately show ?thesis by(auto simp add: nth_append) next case [simp]: True hence "filter P (xs @ [x]) ! n = (xs @ [x]) ! length xs" by simp moreover have "length xs < length (xs @ [x])" by simp moreover have "length xs \ n" by simp moreover have "filter P (take (length xs) (xs @ [x])) = take n (filter P (xs @ [x]))" by simp ultimately show ?thesis by blast qed qed qed lemma set_map_filter: "set (List.map_filter g xs) = {y. \x. x \ set xs \ g x = Some y}" by (induct xs) (auto simp add: List.map_filter_def set_eq_iff) subsection \Quicksort by Relation\ text \A functional implementation of quicksort on lists. It it similar to the one in Isabelle/HOL's example directory. However, it uses tail-recursion for append and arbitrary relations.\ fun partition_rev :: "('a \ bool) \ ('a list \ 'a list) \ 'a list \ ('a list \ 'a list)" where "partition_rev P (yes, no) [] = (yes, no)" | "partition_rev P (yes, no) (x # xs) = partition_rev P (if P x then (x # yes, no) else (yes, x # no)) xs" lemma partition_rev_filter_conv : "partition_rev P (yes, no) xs = (rev (filter P xs) @ yes, rev (filter (Not \ P) xs) @ no)" by (induct xs arbitrary: yes no) (simp_all) function quicksort_by_rel :: "('a \ 'a \ bool) \ 'a list \ 'a list \ 'a list" where "quicksort_by_rel R sl [] = sl" | "quicksort_by_rel R sl (x#xs) = (let (xs_s, xs_b) = partition_rev (\y. R y x) ([],[]) xs in quicksort_by_rel R (x # (quicksort_by_rel R sl xs_b)) xs_s)" by pat_completeness simp_all termination by (relation "measure (\(_, _, xs). length xs)") (simp_all add: partition_rev_filter_conv less_Suc_eq_le) lemma quicksort_by_rel_remove_acc : "quicksort_by_rel R sl xs = (quicksort_by_rel R [] xs @ sl)" proof (induct xs arbitrary: sl rule: measure_induct_rule[of "length"]) case (less xs) note ind_hyp = this show ?case proof (cases xs) case Nil thus ?thesis by simp next case (Cons x xs') note xs_eq[simp] = this obtain xs1 xs2 where part_rev_eq[simp]: "partition_rev (\y. R y x) ([], []) xs' = (xs1, xs2)" by (rule prod.exhaust) from part_rev_eq[symmetric] have length_le: "length xs1 < length xs" "length xs2 < length xs" unfolding partition_rev_filter_conv by (simp_all add: less_Suc_eq_le) note ind_hyp1a = ind_hyp[OF length_le(1), of "x # quicksort_by_rel R [] xs2"] note ind_hyp1b = ind_hyp[OF length_le(1), of "x # quicksort_by_rel R [] xs2 @ sl"] note ind_hyp2 = ind_hyp[OF length_le(2), of sl] show ?thesis by (simp add: ind_hyp1a ind_hyp1b ind_hyp2) qed qed lemma quicksort_by_rel_remove_acc_guared : "sl \ [] \ quicksort_by_rel R sl xs = (quicksort_by_rel R [] xs @ sl)" by (metis quicksort_by_rel_remove_acc) lemma quicksort_by_rel_permutes [simp]: "mset (quicksort_by_rel R sl xs) = mset (xs @ sl)" proof (induct xs arbitrary: sl rule: measure_induct_rule[of "length"]) case (less xs) note ind_hyp = this show ?case proof (cases xs) case Nil thus ?thesis by simp next case (Cons x xs') note xs_eq[simp] = this obtain xs1 xs2 where part_rev_eq[simp]: "partition_rev (\y. R y x) ([], []) xs' = (xs1, xs2)" by (rule prod.exhaust) from part_rev_eq[symmetric] have xs'_multi_eq : "mset xs' = mset xs1 + mset xs2" unfolding partition_rev_filter_conv by (simp add: mset_filter) from part_rev_eq[symmetric] have length_le: "length xs1 < length xs" "length xs2 < length xs" unfolding partition_rev_filter_conv by (simp_all add: less_Suc_eq_le) note ind_hyp[OF length_le(1)] ind_hyp[OF length_le(2)] thus ?thesis by (simp add: xs'_multi_eq union_assoc) qed qed lemma set_quicksort_by_rel [simp]: "set (quicksort_by_rel R sl xs) = set (xs @ sl)" unfolding set_mset_comp_mset [symmetric] o_apply by simp lemma sorted_wrt_quicksort_by_rel: fixes R:: "'x \ 'x \ bool" assumes lin : "\x y. (R x y) \ (R y x)" and trans_R: "\x y z. R x y \ R y z \ R x z" shows "sorted_wrt R (quicksort_by_rel R [] xs)" proof (induct xs rule: measure_induct_rule[of "length"]) case (less xs) note ind_hyp = this show ?case proof (cases xs) case Nil thus ?thesis by simp next case (Cons x xs') note xs_eq[simp] = this obtain xs1 xs2 where part_rev_eq[simp]: "partition_rev (\y. R y x) ([], []) xs' = (xs1, xs2)" by (rule prod.exhaust) from part_rev_eq[symmetric] have xs1_props: "\y. y \ set xs1 \ (R y x)" and xs2_props: "\y. y \ set xs2 \ \(R y x)" unfolding partition_rev_filter_conv by simp_all from xs2_props lin have xs2_props': "\y. y \ set xs2 \ (R x y)" by blast from xs2_props' xs1_props trans_R have xs1_props': "\y1 y2. y1 \ set xs1 \ y2 \ set xs2 \ (R y1 y2)" by metis from part_rev_eq[symmetric] have length_le: "length xs1 < length xs" "length xs2 < length xs" unfolding partition_rev_filter_conv by (simp_all add: less_Suc_eq_le) note ind_hyps = ind_hyp[OF length_le(1)] ind_hyp[OF length_le(2)] thus ?thesis by (simp add: quicksort_by_rel_remove_acc_guared sorted_wrt_append Ball_def xs1_props xs2_props' xs1_props') qed qed lemma sorted_quicksort_by_rel: "sorted (quicksort_by_rel (\) [] xs)" by (rule sorted_wrt_quicksort_by_rel) auto lemma sort_quicksort_by_rel: "sort = quicksort_by_rel (\) []" apply (rule ext, rule properties_for_sort) apply(simp_all add: sorted_quicksort_by_rel) done lemma [code]: "quicksort = quicksort_by_rel (\) []" apply (subst sort_quicksort[symmetric]) by (rule sort_quicksort_by_rel) subsection \Mergesort by Relation\ text \A functional implementation of mergesort on lists. It it similar to the one in Isabelle/HOL's example directory. However, it uses tail-recursion for append and arbitrary relations.\ fun mergesort_by_rel_split :: "('a list \ 'a list) \ 'a list \ ('a list \ 'a list)" where "mergesort_by_rel_split (xs1, xs2) [] = (xs1, xs2)" | "mergesort_by_rel_split (xs1, xs2) [x] = (x # xs1, xs2)" | "mergesort_by_rel_split (xs1, xs2) (x1 # x2 # xs) = mergesort_by_rel_split (x1 # xs1, x2 # xs2) xs" lemma list_induct_first2 [consumes 0, case_names Nil Sing Cons2]: assumes "P []" "\x. P [x]" "\x1 x2 xs. P xs \ P (x1 # x2 #xs)" shows "P xs" proof (induct xs rule: length_induct) case (1 xs) note ind_hyp = this show ?case proof (cases xs) case Nil thus ?thesis using assms(1) by simp next case (Cons x1 xs') note xs_eq[simp] = this thus ?thesis proof (cases xs') case Nil thus ?thesis using assms(2) by simp next case (Cons x2 xs'') note xs'_eq[simp] = this show ?thesis by (simp add: ind_hyp assms(3)) qed qed qed lemma mergesort_by_rel_split_length : "length (fst (mergesort_by_rel_split (xs1, xs2) xs)) = length xs1 + (length xs div 2) + (length xs mod 2) \ length (snd (mergesort_by_rel_split (xs1, xs2) xs)) = length xs2 + (length xs div 2)" by (induct xs arbitrary: xs1 xs2 rule: list_induct_first2) (simp_all) lemma mset_mergesort_by_rel_split [simp]: "mset (fst (mergesort_by_rel_split (xs1, xs2) xs)) + mset (snd (mergesort_by_rel_split (xs1, xs2) xs)) = mset xs + mset xs1 + mset xs2" apply (induct xs arbitrary: xs1 xs2 rule: list_induct_first2) apply (simp_all add: ac_simps) done fun mergesort_by_rel_merge :: "('a \ 'a \ bool) \ 'a list \ 'a list \ 'a list" where "mergesort_by_rel_merge R (x#xs) (y#ys) = (if R x y then x # mergesort_by_rel_merge R xs (y#ys) else y # mergesort_by_rel_merge R (x#xs) ys)" | "mergesort_by_rel_merge R xs [] = xs" | "mergesort_by_rel_merge R [] ys = ys" declare mergesort_by_rel_merge.simps [simp del] lemma mergesort_by_rel_merge_simps[simp] : "mergesort_by_rel_merge R (x#xs) (y#ys) = (if R x y then x # mergesort_by_rel_merge R xs (y#ys) else y # mergesort_by_rel_merge R (x#xs) ys)" "mergesort_by_rel_merge R xs [] = xs" "mergesort_by_rel_merge R [] ys = ys" apply (simp_all add: mergesort_by_rel_merge.simps) apply (cases ys) apply (simp_all add: mergesort_by_rel_merge.simps) done lemma mergesort_by_rel_merge_induct [consumes 0, case_names Nil1 Nil2 Cons1 Cons2]: assumes "\xs::'a list. P xs []" "\ys::'b list. P [] ys" "\x xs y ys. R x y \ P xs (y # ys) \ P (x # xs) (y # ys)" "\x xs y ys. \(R x y) \ P (x # xs) ys \ P (x # xs) (y # ys)" shows "P xs ys" proof (induct xs arbitrary: ys) case Nil thus ?case using assms(2) by simp next case (Cons x xs) note P_xs = this show ?case proof (induct ys) case Nil thus ?case using assms(1) by simp next case (Cons y ys) note P_x_xs_ys = this show ?case using assms(3,4)[of x y xs ys] P_x_xs_ys P_xs by metis qed qed lemma mset_mergesort_by_rel_merge [simp]: "mset (mergesort_by_rel_merge R xs ys) = mset xs + mset ys" by (induct R xs ys rule: mergesort_by_rel_merge.induct) (simp_all add: ac_simps) lemma set_mergesort_by_rel_merge [simp]: "set (mergesort_by_rel_merge R xs ys) = set xs \ set ys" by (induct R xs ys rule: mergesort_by_rel_merge.induct) auto lemma sorted_wrt_mergesort_by_rel_merge [simp]: assumes lin : "\x y. (R x y) \ (R y x)" and trans_R: "\x y z. R x y \ R y z \ R x z" shows "sorted_wrt R (mergesort_by_rel_merge R xs ys) \ sorted_wrt R xs \ sorted_wrt R ys" proof (induct xs ys rule: mergesort_by_rel_merge_induct[where R = R]) case Nil1 thus ?case by simp next case Nil2 thus ?case by simp next case (Cons1 x xs y ys) thus ?case by (simp add: Ball_def) (metis trans_R) next case (Cons2 x xs y ys) thus ?case apply (auto simp add: Ball_def) apply (metis lin) apply (metis lin trans_R) done qed function mergesort_by_rel :: "('a \ 'a \ bool) \ 'a list \ 'a list" where "mergesort_by_rel R xs = (if length xs < 2 then xs else (mergesort_by_rel_merge R (mergesort_by_rel R (fst (mergesort_by_rel_split ([], []) xs))) (mergesort_by_rel R (snd (mergesort_by_rel_split ([], []) xs)))))" by auto termination by (relation "measure (\(_, xs). length xs)") (simp_all add: mergesort_by_rel_split_length not_less minus_div_mult_eq_mod [symmetric]) declare mergesort_by_rel.simps [simp del] lemma mergesort_by_rel_simps [simp, code] : "mergesort_by_rel R [] = []" "mergesort_by_rel R [x] = [x]" "mergesort_by_rel R (x1 # x2 # xs) = (let (xs1, xs2) = (mergesort_by_rel_split ([x1], [x2]) xs) in mergesort_by_rel_merge R (mergesort_by_rel R xs1) (mergesort_by_rel R xs2))" apply (simp add: mergesort_by_rel.simps) apply (simp add: mergesort_by_rel.simps) apply (simp add: mergesort_by_rel.simps[of _ "x1 # x2 # xs"] split: prod.splits) done lemma mergesort_by_rel_permutes [simp]: "mset (mergesort_by_rel R xs) = mset xs" proof (induct xs rule: length_induct) case (1 xs) note ind_hyp = this show ?case proof (cases xs) case Nil thus ?thesis by simp next case (Cons x1 xs') note xs_eq[simp] = this show ?thesis proof (cases xs') case Nil thus ?thesis by simp next case (Cons x2 xs'') note xs'_eq[simp] = this have "length (fst (mergesort_by_rel_split ([], []) xs)) < length xs" "length (snd (mergesort_by_rel_split ([], []) xs)) < length xs" by (simp_all add: mergesort_by_rel_split_length) with ind_hyp show ?thesis unfolding mergesort_by_rel.simps[of _ xs] by (simp add: ac_simps) qed qed qed lemma set_mergesort_by_rel [simp]: "set (mergesort_by_rel R xs) = set xs" unfolding set_mset_comp_mset [symmetric] o_apply by simp lemma sorted_wrt_mergesort_by_rel: fixes R:: "'x \ 'x \ bool" assumes lin : "\x y. (R x y) \ (R y x)" and trans_R: "\x y z. R x y \ R y z \ R x z" shows "sorted_wrt R (mergesort_by_rel R xs)" proof (induct xs rule: measure_induct_rule[of "length"]) case (less xs) note ind_hyp = this show ?case proof (cases xs) case Nil thus ?thesis by simp next case (Cons x xs') note xs_eq[simp] = this thus ?thesis proof (cases xs') case Nil thus ?thesis by simp next case (Cons x2 xs'') note xs'_eq[simp] = this have "length (fst (mergesort_by_rel_split ([], []) xs)) < length xs" "length (snd (mergesort_by_rel_split ([], []) xs)) < length xs" by (simp_all add: mergesort_by_rel_split_length) with ind_hyp show ?thesis unfolding mergesort_by_rel.simps[of _ xs] by (simp add: sorted_wrt_mergesort_by_rel_merge[OF lin trans_R]) qed qed qed lemma sorted_mergesort_by_rel: "sorted (mergesort_by_rel (\) xs)" by (rule sorted_wrt_mergesort_by_rel) auto lemma sort_mergesort_by_rel: "sort = mergesort_by_rel (\)" apply (rule ext, rule properties_for_sort) apply(simp_all add: sorted_mergesort_by_rel) done definition "mergesort = mergesort_by_rel (\)" lemma sort_mergesort: "sort = mergesort" unfolding mergesort_def by (rule sort_mergesort_by_rel) subsubsection \Mergesort with Remdup\ term merge fun merge :: "'a::{linorder} list \ 'a list \ 'a list" where "merge [] l2 = l2" | "merge l1 [] = l1" | "merge (x1 # l1) (x2 # l2) = (if (x1 < x2) then x1 # (merge l1 (x2 # l2)) else (if (x1 = x2) then x1 # (merge l1 l2) else x2 # (merge (x1 # l1) l2)))" lemma merge_correct : assumes l1_OK: "distinct l1 \ sorted l1" assumes l2_OK: "distinct l2 \ sorted l2" shows "distinct (merge l1 l2) \ sorted (merge l1 l2) \ set (merge l1 l2) = set l1 \ set l2" using assms proof (induct l1 arbitrary: l2) case Nil thus ?case by simp next case (Cons x1 l1 l2) note x1_l1_props = Cons(2) note l2_props = Cons(3) from x1_l1_props have l1_props: "distinct l1 \ sorted l1" and x1_nin_l1: "x1 \ set l1" and x1_le: "\x. x \ set l1 \ x1 \ x" by (simp_all add: Ball_def) note ind_hyp_l1 = Cons(1)[OF l1_props] show ?case using l2_props proof (induct l2) case Nil with x1_l1_props show ?case by simp next case (Cons x2 l2) note x2_l2_props = Cons(2) from x2_l2_props have l2_props: "distinct l2 \ sorted l2" and x2_nin_l2: "x2 \ set l2" and x2_le: "\x. x \ set l2 \ x2 \ x" by (simp_all add: Ball_def) note ind_hyp_l2 = Cons(1)[OF l2_props] show ?case proof (cases "x1 < x2") case True note x1_less_x2 = this from ind_hyp_l1[OF x2_l2_props] x1_less_x2 x1_nin_l1 x1_le x2_le show ?thesis apply (auto simp add: Ball_def) apply (metis linorder_not_le) apply (metis linorder_not_less xt1(6) xt1(9)) done next case False note x2_le_x1 = this show ?thesis proof (cases "x1 = x2") case True note x1_eq_x2 = this from ind_hyp_l1[OF l2_props] x1_le x2_le x2_nin_l2 x1_eq_x2 x1_nin_l1 show ?thesis by (simp add: x1_eq_x2 Ball_def) next case False note x1_neq_x2 = this with x2_le_x1 have x2_less_x1 : "x2 < x1" by auto from ind_hyp_l2 x2_le_x1 x1_neq_x2 x2_le x2_nin_l2 x1_le show ?thesis apply (simp add: x2_less_x1 Ball_def) apply (metis linorder_not_le x2_less_x1 xt1(7)) done qed qed qed qed function merge_list :: "'a::{linorder} list list \ 'a list list \ 'a list" where "merge_list [] [] = []" | "merge_list [] [l] = l" | "merge_list (la # acc2) [] = merge_list [] (la # acc2)" | "merge_list (la # acc2) [l] = merge_list [] (l # la # acc2)" | "merge_list acc2 (l1 # l2 # ls) = merge_list ((merge l1 l2) # acc2) ls" by pat_completeness simp_all termination by (relation "measure (\(acc, ls). 3 * length acc + 2 * length ls)") (simp_all) lemma merge_list_correct : assumes ls_OK: "\l. l \ set ls \ distinct l \ sorted l" assumes as_OK: "\l. l \ set as \ distinct l \ sorted l" shows "distinct (merge_list as ls) \ sorted (merge_list as ls) \ set (merge_list as ls) = set (concat (as @ ls))" using assms proof (induct as ls rule: merge_list.induct) case 1 thus ?case by simp next case 2 thus ?case by simp next case 3 thus ?case by simp next case 4 thus ?case by auto next case (5 acc l1 l2 ls) note ind_hyp = 5(1) note l12_l_OK = 5(2) note acc_OK = 5(3) from l12_l_OK acc_OK merge_correct[of l1 l2] have set_merge_eq: "set (merge l1 l2) = set l1 \ set l2" by auto from l12_l_OK acc_OK merge_correct[of l1 l2] have "distinct (merge_list (merge l1 l2 # acc) ls) \ sorted (merge_list (merge l1 l2 # acc) ls) \ set (merge_list (merge l1 l2 # acc) ls) = set (concat ((merge l1 l2 # acc) @ ls))" by (rule_tac ind_hyp) auto with set_merge_eq show ?case by auto qed definition mergesort_remdups where "mergesort_remdups xs = merge_list [] (map (\x. [x]) xs)" lemma mergesort_remdups_correct : "distinct (mergesort_remdups l) \ sorted (mergesort_remdups l) \ (set (mergesort_remdups l) = set l)" proof - let ?l' = "map (\x. [x]) l" { fix xs assume "xs \ set ?l'" then obtain x where xs_eq: "xs = [x]" by auto hence "distinct xs \ sorted xs" by simp } note l'_OK = this from merge_list_correct[of "?l'" "[]", OF l'_OK] show ?thesis unfolding mergesort_remdups_def by simp qed (* TODO: Move *) lemma ex1_eqI: "\\!x. P x; P a; P b\ \ a=b" by blast lemma remdup_sort_mergesort_remdups: "remdups o sort = mergesort_remdups" (is "?lhs=?rhs") proof fix l have "set (?lhs l) = set l" and "sorted (?lhs l)" and "distinct (?lhs l)" by simp_all moreover note mergesort_remdups_correct ultimately show "?lhs l = ?rhs l" by (auto intro!: ex1_eqI[OF finite_sorted_distinct_unique[OF finite_set]]) qed subsection \Native Integers\ lemma int_of_integer_less_iff: "int_of_integer x < int_of_integer y \ x0 \ y\0 \ nat_of_integer x < nat_of_integer y \ xn' < n. \ P n') \ P (n::nat)" shows "\n' \ n. P n'" proof (rule classical) assume contra: "\ (\n'\n. P n')" hence "\n' < n. \ P n'" by auto hence "P n" by (rule hyp) thus "\n'\n. P n'" by auto qed subsubsection \Induction on nat\ lemma nat_compl_induct[case_names 0 Suc]: "\P 0; \n . \nn. nn \ n \ P nn \ P (Suc n)\ \ P n" apply(induct_tac n rule: nat_less_induct) apply(case_tac n) apply(auto) done lemma nat_compl_induct'[case_names 0 Suc]: "\P 0; !! n . \!! nn . nn \ n \ P nn\ \ P (Suc n)\ \ P n" apply(induct_tac n rule: nat_less_induct) apply(case_tac n) apply(auto) done lemma nz_le_conv_less: "0 k \ m \ k - Suc 0 < m" by auto lemma min_Suc_gt[simp]: "a min (Suc a) b = Suc a" "a min b (Suc a) = Suc a" by auto subsection \Integer\ text \Some setup from \int\ transferred to \integer\\ lemma atLeastLessThanPlusOne_atLeastAtMost_integer: "{l.. u \ {(0::integer).. u") apply (subst image_atLeastZeroLessThan_integer, assumption) apply (rule finite_imageI) apply auto done lemma finite_atLeastLessThan_integer [iff]: "finite {l..Functions of type @{typ "bool\bool"}\ lemma boolfun_cases_helper: "g=(\x. False) | g=(\x. x) | g=(\x. True) | g= (\x. \x)" proof - { assume "g False" "g True" hence "g = (\x. True)" by (rule_tac ext, case_tac x, auto) } moreover { assume "g False" "\g True" hence "g = (\x. \x)" by (rule_tac ext, case_tac x, auto) } moreover { assume "\g False" "g True" hence "g = (\x. x)" by (rule_tac ext, case_tac x, auto) } moreover { assume "\g False" "\g True" hence "g = (\x. False)" by (rule_tac ext, case_tac x, auto) } ultimately show ?thesis by fast qed lemma boolfun_cases[case_names False Id True Neg]: "\g=(\x. False) \ P; g=(\x. x) \ P; g=(\x. True) \ P; g=(\x. \x) \ P\ \ P" proof - note boolfun_cases_helper[of g] moreover assume "g=(\x. False) \ P" "g=(\x. x) \ P" "g=(\x. True) \ P" "g=(\x. \x) \ P" ultimately show ?thesis by fast qed subsection \Definite and indefinite description\ text "Combined definite and indefinite description for binary predicate" lemma some_theI: assumes EX: "\a b . P a b" and BUN: "!! b1 b2 . \\a . P a b1; \a . P a b2\ \ b1=b2" shows "P (SOME a . \b . P a b) (THE b . \a . P a b)" proof - from EX have "\b. P (SOME a. \b. P a b) b" by (rule someI_ex) moreover from EX have "\b. \a. P a b" by blast with BUN theI'[of "\b. \a. P a b"] have "\a. P a (THE b. \a. P a b)" by (unfold Ex1_def, blast) moreover note BUN ultimately show ?thesis by (fast) qed lemma some_insert_self[simp]: "S\{} \ insert (SOME x. x\S) S = S" by (auto intro: someI) lemma some_elem[simp]: "S\{} \ (SOME x. x\S) \ S" by (auto intro: someI) subsubsection\Hilbert Choice with option\ definition Eps_Opt where "Eps_Opt P = (if (\x. P x) then Some (SOME x. P x) else None)" lemma some_opt_eq_trivial[simp] : "Eps_Opt (\y. y = x) = Some x" unfolding Eps_Opt_def by simp lemma some_opt_sym_eq_trivial[simp] : "Eps_Opt ((=) x) = Some x" unfolding Eps_Opt_def by simp lemma some_opt_false_trivial[simp] : "Eps_Opt (\_. False) = None" unfolding Eps_Opt_def by simp lemma Eps_Opt_eq_None[simp] : "Eps_Opt P = None \ \(Ex P)" unfolding Eps_Opt_def by simp lemma Eps_Opt_eq_Some_implies : "Eps_Opt P = Some x \ P x" unfolding Eps_Opt_def by (metis option.inject option.simps(2) someI_ex) lemma Eps_Opt_eq_Some : assumes P_prop: "\x'. P x \ P x' \ x' = x" shows "Eps_Opt P = Some x \ P x" using P_prop unfolding Eps_Opt_def by (metis option.inject option.simps(2) someI_ex) subsection \Product Type\ lemma nested_case_prod_simp: "(\(a,b) c. f a b c) x y = (case_prod (\a b. f a b y) x)" by (auto split: prod.split) lemma fn_fst_conv: "(\x. (f (fst x))) = (\(a,_). f a)" by auto lemma fn_snd_conv: "(\x. (f (snd x))) = (\(_,b). f b)" by auto fun pairself where "pairself f (a,b) = (f a, f b)" lemma pairself_image_eq[simp]: "pairself f ` {(a,b). P a b} = {(f a, f b)| a b. P a b}" by force lemma pairself_image_cart[simp]: "pairself f ` (A\B) = f`A \ f`B" by (auto simp: image_def) lemma in_prod_fst_sndI: "fst x \ A \ snd x \ B \ x\A\B" by (cases x) auto lemma inj_Pair[simp]: "inj_on (\x. (x,c x)) S" "inj_on (\x. (c x,x)) S" by (auto intro!: inj_onI) declare Product_Type.swap_inj_on[simp] lemma img_fst [intro]: assumes "(a,b) \ S" shows "a \ fst ` S" by (rule image_eqI[OF _ assms]) simp lemma img_snd [intro]: assumes "(a,b) \ S" shows "b \ snd ` S" by (rule image_eqI[OF _ assms]) simp lemma range_prod: "range f \ (range (fst \ f)) \ (range (snd \ f))" proof fix y assume "y \ range f" then obtain x where y: "y = f x" by auto hence "y = (fst(f x), snd(f x))" by simp thus "y \ (range (fst \ f)) \ (range (snd \ f))" by (fastforce simp add: image_def) qed lemma finite_range_prod: assumes fst: "finite (range (fst \ f))" and snd: "finite (range (snd \ f))" shows "finite (range f)" proof - from fst snd have "finite (range (fst \ f) \ range (snd \ f))" by (rule finite_SigmaI) thus ?thesis by (rule finite_subset[OF range_prod]) qed lemma fstE: "x = (a,b) \ P (fst x) \ P a" by (metis fst_conv) lemma sndE: "x = (a,b) \ P (snd x) \ P b" by (metis snd_conv) subsubsection \Uncurrying\ (* TODO: Move to HOL/Product_Type? Lars H: "It's equal to case_prod, should use an abbreviation"*) definition uncurry :: "('a \ 'b \ 'c) \ 'a \ 'b \ 'c" where "uncurry f \ \(a,b). f a b" lemma uncurry_apply[simp]: "uncurry f (a,b) = f a b" unfolding uncurry_def by simp lemma curry_uncurry_id[simp]: "curry (uncurry f) = f" unfolding uncurry_def by simp lemma uncurry_curry_id[simp]: "uncurry (curry f) = f" unfolding uncurry_def by simp lemma do_curry: "f (a,b) = curry f a b" by simp lemma do_uncurry: "f a b = uncurry f (a,b)" by simp subsection \Sum Type\ lemma map_sum_Inr_conv: "map_sum fl fr s = Inr y \ (\x. s=Inr x \ y = fr x)" by (cases s) auto lemma map_sum_Inl_conv: "map_sum fl fr s = Inl y \ (\x. s=Inl x \ y = fl x)" by (cases s) auto subsection \Directed Graphs and Relations\ subsubsection "Reflexive-Transitive Closure" lemma r_le_rtrancl[simp]: "S\S\<^sup>*" by auto lemma rtrancl_mono_rightI: "S\S' \ S\S'\<^sup>*" by auto lemma trancl_sub: "R \ R\<^sup>+" by auto lemma trancl_single[simp]: "{(a,b)}\<^sup>+ = {(a,b)}" by (auto simp: trancl_insert) text \Pick first non-reflexive step\ lemma converse_rtranclE'[consumes 1, case_names base step]: assumes "(u,v)\R\<^sup>*" obtains "u=v" | vh where "u\vh" and "(u,vh)\R" and "(vh,v)\R\<^sup>*" using assms apply (induct rule: converse_rtrancl_induct) apply auto [] apply (case_tac "y=z") apply auto done lemma in_rtrancl_insert: "x\R\<^sup>* \ x\(insert r R)\<^sup>*" by (metis in_mono rtrancl_mono subset_insertI) lemma rtrancl_apply_insert: "R\<^sup>*``(insert x S) = insert x (R\<^sup>*``(S\R``{x}))" apply (auto) apply (erule converse_rtranclE) apply auto [2] apply (erule converse_rtranclE) apply (auto intro: converse_rtrancl_into_rtrancl) [2] done text \A path in a graph either does not use nodes from S at all, or it has a prefix leading to a node in S and a suffix that does not use nodes in S\ lemma rtrancl_last_visit[cases set, case_names no_visit last_visit_point]: shows "\ (q,q')\R\<^sup>*; (q,q')\(R-UNIV\S)\<^sup>* \ P; !!qt. \ qt\S; (q,qt)\R\<^sup>+; (qt,q')\(R-UNIV\S)\<^sup>* \ \ P \ \ P" proof (induct rule: converse_rtrancl_induct[case_names refl step]) case refl thus ?case by auto next case (step q qh) show P proof (rule step.hyps(3)) assume A: "(qh,q')\(R-UNIV\S)\<^sup>*" show P proof (cases "qh\S") case False with step.hyps(1) A have "(q,q')\(R-UNIV\S)\<^sup>*" by (auto intro: converse_rtrancl_into_rtrancl) with step.prems(1) show P . next case True from step.hyps(1) have "(q,qh)\R\<^sup>+" by auto with step.prems(2) True A show P by blast qed next fix qt assume A: "qt\S" "(qh,qt)\R\<^sup>+" "(qt,q')\(R-UNIV\S)\<^sup>*" with step.hyps(1) have "(q,qt)\R\<^sup>+" by auto with step.prems(2) A(1,3) show P by blast qed qed text \Less general version of \rtrancl_last_visit\, but there's a short automatic proof\ lemma rtrancl_last_visit': "\ (q,q')\R\<^sup>*; (q,q')\(R-UNIV\S)\<^sup>* \ P; !!qt. \ qt\S; (q,qt)\R\<^sup>*; (qt,q')\(R-UNIV\S)\<^sup>* \ \ P \ \ P" by (induct rule: converse_rtrancl_induct) (auto intro: converse_rtrancl_into_rtrancl) lemma rtrancl_last_visit_node: assumes "(s,s')\R\<^sup>*" shows "s\sh \ (s,s')\(R \ (UNIV \ (-{sh})))\<^sup>* \ (s,sh)\R\<^sup>* \ (sh,s')\(R \ (UNIV \ (-{sh})))\<^sup>*" using assms proof (induct rule: converse_rtrancl_induct) case base thus ?case by auto next case (step s st) moreover { assume P: "(st,s')\ (R \ UNIV \ - {sh})\<^sup>*" { assume "st=sh" with step have ?case by auto } moreover { assume "st\sh" with \(s,st)\R\ have "(s,st)\(R \ UNIV \ - {sh})\<^sup>*" by auto also note P finally have ?case by blast } ultimately have ?case by blast } moreover { assume P: "(st, sh) \ R\<^sup>* \ (sh, s') \ (R \ UNIV \ - {sh})\<^sup>*" with step(1) have ?case by (auto dest: converse_rtrancl_into_rtrancl) } ultimately show ?case by blast qed text \Find last point where a path touches a set\ lemma rtrancl_last_touch: "\ (q,q')\R\<^sup>*; q\S; !!qt. \ qt\S; (q,qt)\R\<^sup>*; (qt,q')\(R-UNIV\S)\<^sup>* \ \ P \ \ P" by (erule rtrancl_last_visit') auto text \A path either goes over edge once, or not at all\ lemma trancl_over_edgeE: assumes "(u,w)\(insert (v1,v2) E)\<^sup>+" obtains "(u,w)\E\<^sup>+" | "(u,v1)\E\<^sup>*" and "(v2,w)\E\<^sup>*" using assms proof induct case (base z) thus ?thesis by (metis insertE prod.inject r_into_trancl' rtrancl_eq_or_trancl) next case (step y z) thus ?thesis by (metis (opaque_lifting, no_types) Pair_inject insertE rtrancl.simps trancl.simps trancl_into_rtrancl) qed lemma rtrancl_image_advance: "\q\R\<^sup>* `` Q0; (q,x)\R\ \ x\R\<^sup>* `` Q0" by (auto intro: rtrancl_into_rtrancl) lemma trancl_image_by_rtrancl: "(E\<^sup>+)``Vi \ Vi = (E\<^sup>*)``Vi" by (metis Image_Id Un_Image rtrancl_trancl_reflcl) lemma reachable_mono: "\R\R'; X\X'\ \ R\<^sup>*``X \ R'\<^sup>*``X'" by (metis Image_mono rtrancl_mono) lemma finite_reachable_advance: "\ finite (E\<^sup>*``{v0}); (v0,v)\E\<^sup>* \ \ finite (E\<^sup>*``{v})" by (erule finite_subset[rotated]) auto lemma rtrancl_mono_mp: "U\V \ x\U\<^sup>* \ x\V\<^sup>*" by (metis in_mono rtrancl_mono) lemma trancl_mono_mp: "U\V \ x\U\<^sup>+ \ x\V\<^sup>+" by (metis trancl_mono) lemma rtrancl_mapI: "(a,b)\E\<^sup>* \ (f a, f b)\(pairself f `E)\<^sup>*" apply (induction rule: rtrancl_induct) apply (force intro: rtrancl.intros)+ done lemma rtrancl_image_advance_rtrancl: assumes "q \ R\<^sup>*``Q0" assumes "(q,x) \ R\<^sup>*" shows "x \ R\<^sup>*``Q0" using assms by (metis rtrancl_idemp rtrancl_image_advance) lemma nth_step_trancl: "\n m. \ \ n. n < length xs - 1 \ (xs ! Suc n, xs ! n) \ R \ \ n < length xs \ m < n \ (xs ! n, xs ! m) \ R\<^sup>+" proof (induction xs) case (Cons x xs) hence "\n. n < length xs - 1 \ (xs ! Suc n, xs ! n) \ R" apply clarsimp by (metis One_nat_def diff_Suc_eq_diff_pred nth_Cons_Suc zero_less_diff) note IH = this[THEN Cons.IH] from Cons obtain n' where n': "Suc n' = n" by (cases n) blast+ show ?case proof (cases m) case "0" with Cons have "xs \ []" by auto with "0" Cons.prems(1)[of m] have "(xs ! 0, x) \ R" by simp moreover from IH[where m = 0] have "\n. n < length xs \ n > 0 \ (xs ! n, xs ! 0) \ R\<^sup>+" by simp ultimately have "\n. n < length xs \ (xs ! n, x) \ R\<^sup>+" by (metis trancl_into_trancl gr0I r_into_trancl') with Cons "0" show ?thesis by auto next case (Suc m') with Cons.prems n' have "n' < length xs" "m' < n'" by auto with IH have "(xs ! n', xs ! m') \ R\<^sup>+" by simp with Suc n' show ?thesis by auto qed qed simp lemma Image_empty_trancl_Image_empty: "R `` {v} = {} \ R\<^sup>+ `` {v} = {}" unfolding Image_def by (auto elim: converse_tranclE) lemma Image_empty_rtrancl_Image_id: "R `` {v} = {} \ R\<^sup>* `` {v} = {v}" unfolding Image_def by (auto elim: converse_rtranclE) lemma trans_rtrancl_eq_reflcl: "trans A \ A^* = A^=" by (simp add: rtrancl_trancl_reflcl) lemma refl_on_reflcl_Image: "refl_on B A \ C \ B \ A^= `` C = A `` C" by (auto simp add: Image_def dest: refl_onD) lemma Image_absorb_rtrancl: "\ trans A; refl_on B A; C \ B \ \ A^* `` C = A `` C" by (simp add: trans_rtrancl_eq_reflcl refl_on_reflcl_Image) lemma trancl_Image_unfold_left: "E\<^sup>+``S = E\<^sup>*``E``S" by (auto simp: trancl_unfold_left) lemma trancl_Image_unfold_right: "E\<^sup>+``S = E``E\<^sup>*``S" by (auto simp: trancl_unfold_right) lemma trancl_Image_advance_ss: "(u,v)\E \ E\<^sup>+``{v} \ E\<^sup>+``{u}" by auto lemma rtrancl_Image_advance_ss: "(u,v)\E \ E\<^sup>*``{v} \ E\<^sup>*``{u}" by auto (* FIXME: nicer name *) lemma trancl_union_outside: assumes "(v,w) \ (E\U)\<^sup>+" and "(v,w) \ E\<^sup>+" shows "\x y. (v,x) \ (E\U)\<^sup>* \ (x,y) \ U \ (y,w) \ (E\U)\<^sup>*" using assms proof (induction) case base thus ?case by auto next case (step w x) show ?case proof (cases "(v,w)\E\<^sup>+") case True from step have "(v,w)\(E\U)\<^sup>*" by simp moreover from True step have "(w,x) \ U" by (metis Un_iff trancl.simps) moreover have "(x,x) \ (E\U)\<^sup>*" by simp ultimately show ?thesis by blast next case False with step.IH obtain a b where "(v,a) \ (E\U)\<^sup>*" "(a,b) \ U" "(b,w) \ (E\U)\<^sup>*" by blast moreover with step have "(b,x) \ (E\U)\<^sup>*" by (metis rtrancl_into_rtrancl) ultimately show ?thesis by blast qed qed lemma trancl_restrict_reachable: assumes "(u,v) \ E\<^sup>+" assumes "E``S \ S" assumes "u\S" shows "(u,v) \ (E\S\S)\<^sup>+" using assms by (induction rule: converse_trancl_induct) (auto intro: trancl_into_trancl2) lemma rtrancl_image_unfold_right: "E``E\<^sup>*``V \ E\<^sup>*``V" by (auto intro: rtrancl_into_rtrancl) lemma trancl_Image_in_Range: "R\<^sup>+ `` V \ Range R" by (auto elim: trancl.induct) lemma rtrancl_Image_in_Field: "R\<^sup>* `` V \ Field R \ V" proof - from trancl_Image_in_Range have "R\<^sup>+ `` V \ Field R" unfolding Field_def by fast hence "R\<^sup>+ `` V \ V \ Field R \ V" by blast with trancl_image_by_rtrancl show ?thesis by metis qed lemma rtrancl_sub_insert_rtrancl: "R\<^sup>* \ (insert x R)\<^sup>*" by (auto elim: rtrancl.induct rtrancl_into_rtrancl) lemma trancl_sub_insert_trancl: "R\<^sup>+ \ (insert x R)\<^sup>+" by (auto elim: trancl.induct trancl_into_trancl) lemma Restr_rtrancl_mono: "(v,w) \ (Restr E U)\<^sup>* \ (v,w) \ E\<^sup>*" by (metis inf_le1 rtrancl_mono subsetCE) lemma Restr_trancl_mono: "(v,w) \ (Restr E U)\<^sup>+ \ (v,w) \ E\<^sup>+" by (metis inf_le1 trancl_mono) subsubsection "Converse Relation" lemmas converse_add_simps = converse_Times trancl_converse[symmetric] converse_Un converse_Int lemma dom_ran_disj_comp[simp]: "Domain R \ Range R = {} \ R O R = {}" by auto lemma below_Id_inv[simp]: "R\\Id \ R\Id" by (auto) subsubsection "Cyclicity" lemma cyclicE: "\\acyclic g; !!x. (x,x)\g\<^sup>+ \ P\ \ P" by (unfold acyclic_def) blast lemma acyclic_insert_cyclic: "\acyclic g; \acyclic (insert (x,y) g)\ \ (y,x)\g\<^sup>*" by (unfold acyclic_def) (auto simp add: trancl_insert) text \ This lemma makes a case distinction about a path in a graph where a couple of edges with the same endpoint have been inserted: If there is a path from a to b, then there's such a path in the original graph, or there's a path that uses an inserted edge only once. Originally, this lemma was used to reason about the graph of an updated acquisition history. Any path in this graph is either already contained in the original graph, or passes via an inserted edge. Because all the inserted edges point to the same target node, in the second case, the path can be short-circuited to use exactly one inserted edge. \ lemma trancl_multi_insert[cases set, case_names orig via]: "\ (a,b)\(r\X\{m})\<^sup>+; (a,b)\r\<^sup>+ \ P; !!x. \ x\X; (a,x)\r\<^sup>*; (m,b)\r\<^sup>* \ \ P \ \ P" proof (induct arbitrary: P rule: trancl_induct) case (base b) thus ?case by auto next case (step b c) show ?case proof (rule step.hyps(3)) assume A: "(a,b)\r\<^sup>+" note step.hyps(2) moreover { assume "(b,c)\r" with A have "(a,c)\r\<^sup>+" by auto with step.prems have P by blast } moreover { assume "b\X" "c=m" with A have P by (rule_tac step.prems(2)) simp+ } ultimately show P by auto next fix x assume A: "x \ X" "(a, x) \ r\<^sup>*" "(m, b) \ r\<^sup>*" note step.hyps(2) moreover { assume "(b,c)\r" with A(3) have "(m,c)\r\<^sup>*" by auto with step.prems(2)[OF A(1,2)] have P by blast } moreover { assume "b\X" "c=m" with A have P by (rule_tac step.prems(2)) simp+ } ultimately show P by auto qed qed text \ Version of @{thm [source] trancl_multi_insert} for inserted edges with the same startpoint. \ lemma trancl_multi_insert2[cases set, case_names orig via]: "\(a,b)\(r\{m}\X)\<^sup>+; (a,b)\r\<^sup>+ \ P; !!x. \ x\X; (a,m)\r\<^sup>*; (x,b)\r\<^sup>* \ \ P \ \ P" proof goal_cases case prems: 1 from prems(1) have "(b,a)\((r\{m}\X)\<^sup>+)\" by simp also have "((r\{m}\X)\<^sup>+)\ = (r\\X\{m})\<^sup>+" by (simp add: converse_add_simps) finally have "(b, a) \ (r\ \ X \ {m})\<^sup>+" . thus ?case by (auto elim!: trancl_multi_insert intro: prems(2,3) simp add: trancl_converse rtrancl_converse ) qed lemma cyclic_subset: "\ \ acyclic R; R \ S \ \ \ acyclic S" unfolding acyclic_def by (blast intro: trancl_mono) subsubsection \Wellfoundedness\ lemma wf_min: assumes A: "wf R" "R\{}" "!!m. m\Domain R - Range R \ P" shows P proof - have H: "!!x. wf R \ \y. (x,y)\R \ x\Domain R - Range R \ (\m. m\Domain R - Range R)" by (erule_tac wf_induct_rule[where P="\x. \y. (x,y)\R \ x\Domain R - Range R \ (\m. m\Domain R - Range R)"]) auto from A(2) obtain x y where "(x,y)\R" by auto with A(1,3) H show ?thesis by blast qed lemma finite_wf_eq_wf_converse: "finite R \ wf (R\) \ wf R" by (metis acyclic_converse finite_acyclic_wf finite_acyclic_wf_converse wf_acyclic) lemma wf_max: assumes A: "wf (R\)" "R\{}" and C: "!!m. m\Range R - Domain R \ P" shows "P" proof - from A(2) have NE: "R\\{}" by auto from wf_min[OF A(1) NE] obtain m where "m\Range R - Domain R" by auto thus P by (blast intro: C) qed \ \Useful lemma to show well-foundedness of some process approaching a finite upper bound\ lemma wf_bounded_supset: "finite S \ wf {(Q',Q). Q'\Q \ Q'\ S}" proof - assume [simp]: "finite S" hence [simp]: "!!x. finite (S-x)" by auto have "{(Q',Q). Q\Q' \ Q'\ S} \ inv_image ({(s'::nat,s). s'Q. card (S-Q))" proof (intro subsetI, case_tac x, simp) fix a b assume A: "b\a \ a\S" hence "S-a \ S-b" by blast thus "card (S-a) < card (S-b)" by (auto simp add: psubset_card_mono) qed moreover have "wf ({(s'::nat,s). s' Range R = {} \ wf R" apply (rule wf_no_loop) by simp text \Extend a wf-relation by a break-condition\ definition "brk_rel R \ {((False,x),(False,y)) | x y. (x,y)\R} \ {((True,x),(False,y)) | x y. True}" lemma brk_rel_wf[simp,intro!]: assumes WF[simp]: "wf R" shows "wf (brk_rel R)" proof - have "wf {((False,x),(False,y)) | x y. (x,y)\R}" proof - have "{((False,x),(False,y)) | x y. (x,y)\R} \ inv_image R snd" by auto from wf_subset[OF wf_inv_image[OF WF] this] show ?thesis . qed moreover have "wf {((True,x),(False,y)) | x y. True}" by (rule wf_no_path) auto ultimately show ?thesis unfolding brk_rel_def apply (subst Un_commute) by (blast intro: wf_Un) qed subsubsection \Restrict Relation\ definition rel_restrict :: "('a \ 'a) set \ 'a set \ ('a \ 'a) set" where "rel_restrict R A \ {(v,w). (v,w) \ R \ v \ A \ w \ A}" lemma rel_restrict_alt_def: "rel_restrict R A = R \ (-A) \ (-A)" unfolding rel_restrict_def by auto lemma rel_restrict_empty[simp]: "rel_restrict R {} = R" by (simp add: rel_restrict_def) lemma rel_restrict_notR: assumes "(x,y) \ rel_restrict A R" shows "x \ R" and "y \ R" using assms unfolding rel_restrict_def by auto lemma rel_restrict_sub: "rel_restrict R A \ R" unfolding rel_restrict_def by auto lemma rel_restrict_Int_empty: "A \ Field R = {} \ rel_restrict R A = R" unfolding rel_restrict_def Field_def by auto lemma Domain_rel_restrict: "Domain (rel_restrict R A) \ Domain R - A" unfolding rel_restrict_def by auto lemma Range_rel_restrict: "Range (rel_restrict R A) \ Range R - A" unfolding rel_restrict_def by auto lemma Field_rel_restrict: "Field (rel_restrict R A) \ Field R - A" unfolding rel_restrict_def Field_def by auto lemma rel_restrict_compl: "rel_restrict R A \ rel_restrict R (-A) = {}" unfolding rel_restrict_def by auto lemma finite_rel_restrict: "finite R \ finite (rel_restrict R A)" by (metis finite_subset rel_restrict_sub) lemma R_subset_Field: "R \ Field R \ Field R" unfolding Field_def by auto lemma homo_rel_restrict_mono: "R \ B \ B \ rel_restrict R A \ (B - A) \ (B - A)" proof - assume A: "R \ B \ B" hence "Field R \ B" unfolding Field_def by auto with Field_rel_restrict have "Field (rel_restrict R A) \ B - A" by (metis Diff_mono order_refl order_trans) with R_subset_Field show ?thesis by blast qed lemma rel_restrict_union: "rel_restrict R (A \ B) = rel_restrict (rel_restrict R A) B" unfolding rel_restrict_def by auto lemma rel_restrictI: "x \ R \ y \ R \ (x,y) \ E \ (x,y) \ rel_restrict E R" unfolding rel_restrict_def by auto lemma rel_restrict_lift: "(x,y) \ rel_restrict E R \ (x,y) \ E" unfolding rel_restrict_def by simp lemma rel_restrict_trancl_mem: "(a,b) \ (rel_restrict A R)\<^sup>+ \ (a,b) \ rel_restrict (A\<^sup>+) R" by (induction rule: trancl_induct) (auto simp add: rel_restrict_def) lemma rel_restrict_trancl_sub: "(rel_restrict A R)\<^sup>+ \ rel_restrict (A\<^sup>+) R" by (metis subrelI rel_restrict_trancl_mem) lemma rel_restrict_mono: "A \ B \ rel_restrict A R \ rel_restrict B R" unfolding rel_restrict_def by auto lemma rel_restrict_mono2: "R \ S \ rel_restrict A S \ rel_restrict A R" unfolding rel_restrict_def by auto lemma rel_restrict_Sigma_sub: "rel_restrict ((A\A)\<^sup>+) R \ ((A - R) \ (A - R))\<^sup>+" unfolding rel_restrict_def by auto (metis DiffI converse_tranclE mem_Sigma_iff r_into_trancl tranclE) lemma finite_reachable_restrictedI: assumes F: "finite Q" assumes I: "I\Q" assumes R: "Range E \ Q" shows "finite (E\<^sup>*``I)" proof - from I R have "E\<^sup>*``I \ Q" by (force elim: rtrancl.cases) also note F finally (finite_subset) show ?thesis . qed context begin private lemma rtrancl_restrictI_aux: assumes "(u,v)\(E-UNIV\R)\<^sup>*" assumes "u\R" shows "(u,v)\(rel_restrict E R)\<^sup>* \ v\R" using assms by (induction) (auto simp: rel_restrict_def intro: rtrancl.intros) corollary rtrancl_restrictI: assumes "(u,v)\(E-UNIV\R)\<^sup>*" assumes "u\R" shows "(u,v)\(rel_restrict E R)\<^sup>*" using rtrancl_restrictI_aux[OF assms] .. end lemma E_closed_restr_reach_cases: assumes P: "(u,v)\E\<^sup>*" assumes CL: "E``R \ R" obtains "v\R" | "u\R" "(u,v)\(rel_restrict E R)\<^sup>*" using P proof (cases rule: rtrancl_last_visit[where S=R]) case no_visit show ?thesis proof (cases "u\R") case True with P have "v\R" using rtrancl_reachable_induct[OF _ CL, where I="{u}"] by auto thus ?thesis .. next case False with no_visit have "(u,v)\(rel_restrict E R)\<^sup>*" by (rule rtrancl_restrictI) with False show ?thesis .. qed next case (last_visit_point x) from \(x, v) \ (E - UNIV \ R)\<^sup>*\ have "(x,v)\E\<^sup>*" by (rule rtrancl_mono_mp[rotated]) auto with \x\R\ have "v\R" using rtrancl_reachable_induct[OF _ CL, where I="{x}"] by auto thus ?thesis .. qed lemma rel_restrict_trancl_notR: assumes "(v,w) \ (rel_restrict E R)\<^sup>+" shows "v \ R" and "w \ R" using assms by (metis rel_restrict_trancl_mem rel_restrict_notR)+ lemma rel_restrict_tranclI: assumes "(x,y) \ E\<^sup>+" and "x \ R" "y \ R" and "E `` R \ R" shows "(x,y) \ (rel_restrict E R)\<^sup>+" using assms proof (induct) case base thus ?case by (metis r_into_trancl rel_restrictI) next case (step y z) hence "y \ R" by auto with step show ?case by (metis trancl_into_trancl rel_restrictI) qed subsubsection \Single-Valued Relations\ lemma single_valued_inter1: "single_valued R \ single_valued (R\S)" by (auto intro: single_valuedI dest: single_valuedD) lemma single_valued_inter2: "single_valued R \ single_valued (S\R)" by (auto intro: single_valuedI dest: single_valuedD) lemma single_valued_below_Id: "R\Id \ single_valued R" by (auto intro: single_valuedI) subsubsection \Bijective Relations\ definition "bijective R \ (\x y z. (x,y)\R \ (x,z)\R \ y=z) \ (\x y z. (x,z)\R \ (y,z)\R \ x=y)" lemma bijective_alt: "bijective R \ single_valued R \ single_valued (R\)" unfolding bijective_def single_valued_def by blast lemma bijective_Id[simp, intro!]: "bijective Id" by (auto simp: bijective_def) lemma bijective_Empty[simp, intro!]: "bijective {}" by (auto simp: bijective_def) subsubsection \Miscellaneous\ lemma pair_vimage_is_Image[simp]: "(Pair u -` E) = E``{u}" by auto lemma fst_in_Field: "fst ` R \ Field R" by (simp add: Field_def fst_eq_Domain) lemma snd_in_Field: "snd ` R \ Field R" by (simp add: Field_def snd_eq_Range) lemma ran_map_of: "ran (map_of xs) \ snd ` set (xs)" by (induct xs) (auto simp add: ran_def) lemma Image_subset_snd_image: "A `` B \ snd ` A" unfolding Image_def image_def by force lemma finite_Image_subset: "finite (A `` B) \ C \ A \ finite (C `` B)" by (metis Image_mono order_refl rev_finite_subset) lemma finite_Field_eq_finite[simp]: "finite (Field R) \ finite R" by (metis finite_cartesian_product finite_subset R_subset_Field finite_Field) definition "fun_of_rel R x \ SOME y. (x,y)\R" lemma for_in_RI(*[intro]*): "x\Domain R \ (x,fun_of_rel R x)\R" unfolding fun_of_rel_def by (auto intro: someI) lemma Field_not_elem: "v \ Field R \ \(x,y) \ R. x \ v \ y \ v" unfolding Field_def by auto lemma Sigma_UNIV_cancel[simp]: "(A \ X - A \ UNIV) = {}" by auto lemma same_fst_trancl[simp]: "(same_fst P R)\<^sup>+ = same_fst P (\x. (R x)\<^sup>+)" proof - { fix x y assume "(x,y)\(same_fst P R)\<^sup>+" hence "(x,y)\same_fst P (\x. (R x)\<^sup>+)" by induction (auto simp: same_fst_def) } moreover { fix f f' x y assume "((f,x),(f',y))\same_fst P (\x. (R x)\<^sup>+)" hence [simp]: "f'=f" "P f" and 1: "(x,y)\(R f)\<^sup>+" by (auto simp: same_fst_def) from 1 have "((f,x),(f',y))\(same_fst P R)\<^sup>+" apply induction subgoal by (rule r_into_trancl) auto subgoal by (erule trancl_into_trancl) auto done } ultimately show ?thesis by auto qed subsection \\option\ Datatype\ lemma le_some_optE: "\Some m\x; !!m'. \x=Some m'; m\m'\ \ P\ \ P" by (cases x) auto lemma not_Some_eq2[simp]: "(\x y. v \ Some (x,y)) = (v = None)" by (cases v) auto subsection "Maps" primrec the_default where "the_default _ (Some x) = x" | "the_default x None = x" declare map_add_dom_app_simps[simp] declare map_add_upd_left[simp] lemma ran_add[simp]: "dom f \ dom g = {} \ ran (f++g) = ran f \ ran g" by (fastforce simp add: ran_def map_add_def split: option.split_asm option.split) lemma nempty_dom: "\e\Map.empty; !!m. m\dom e \ P \ \ P" by (subgoal_tac "dom e \ {}") (blast, auto) lemma le_map_dom_mono: "m\m' \ dom m \ dom m'" apply (safe) apply (drule_tac x=x in le_funD) apply simp apply (erule le_some_optE) apply simp done lemma map_add_first_le: fixes m::"'a\('b::order)" shows "\ m\m' \ \ m++n \ m'++n" apply (rule le_funI) apply (auto simp add: map_add_def split: option.split elim: le_funE) done lemma map_add_distinct_le: shows "\ m\m'; n\n'; dom m' \ dom n' = {} \ \ m++n \ m'++n'" apply (rule le_funI) apply (auto simp add: map_add_def split: option.split) apply (fastforce elim: le_funE) apply (drule le_map_dom_mono) apply (drule le_map_dom_mono) apply (case_tac "m x") apply simp apply (force) apply (fastforce dest!: le_map_dom_mono) apply (erule le_funE) apply (erule_tac x=x in le_funE) apply simp done lemma map_add_left_comm: assumes A: "dom A \ dom B = {}" shows "A ++ (B ++ C) = B ++ (A ++ C)" proof - have "A ++ (B ++ C) = (A++B)++C" by simp also have "\ = (B++A)++C" by (simp add: map_add_comm[OF A]) also have "\ = B++(A++C)" by simp finally show ?thesis . qed lemmas map_add_ac = map_add_assoc map_add_comm map_add_left_comm lemma le_map_restrict[simp]: fixes m :: "'a \ ('b::order)" shows "m |` X \ m" by (rule le_funI) (simp add: restrict_map_def) lemma map_of_distinct_upd: "x \ set (map fst xs) \ [x \ y] ++ map_of xs = (map_of xs) (x \ y)" by (induct xs) (auto simp add: fun_upd_twist) lemma map_of_distinct_upd2: assumes "x \ set(map fst xs)" "x \ set (map fst ys)" shows "map_of (xs @ (x,y) # ys) = (map_of (xs @ ys))(x \ y)" apply(insert assms) apply(induct xs) apply (auto intro: ext) done lemma map_of_distinct_upd3: assumes "x \ set(map fst xs)" "x \ set (map fst ys)" shows "map_of (xs @ (x,y) # ys) = (map_of (xs @ (x,y') # ys))(x \ y)" apply(insert assms) apply(induct xs) apply (auto intro: ext) done lemma map_of_distinct_upd4: assumes "x \ set(map fst xs)" "x \ set (map fst ys)" shows "map_of (xs @ ys) = (map_of (xs @ (x,y) # ys))(x := None)" using assms by (induct xs) (auto simp: map_of_eq_None_iff) lemma map_of_distinct_lookup: assumes "x \ set(map fst xs)" "x \ set (map fst ys)" shows "map_of (xs @ (x,y) # ys) x = Some y" proof - have "map_of (xs @ (x,y) # ys) = (map_of (xs @ ys)) (x \ y)" using assms map_of_distinct_upd2 by simp thus ?thesis by simp qed lemma ran_distinct: assumes dist: "distinct (map fst al)" shows "ran (map_of al) = snd ` set al" using assms proof (induct al) case Nil then show ?case by simp next case (Cons kv al) then have "ran (map_of al) = snd ` set al" by simp moreover from Cons.prems have "map_of al (fst kv) = None" by (simp add: map_of_eq_None_iff) ultimately show ?case by (simp only: map_of.simps ran_map_upd) simp qed lemma ran_is_image: "ran M = (the \ M) ` (dom M)" unfolding ran_def dom_def image_def by auto lemma map_card_eq_iff: assumes finite: "finite (dom M)" and card_eq: "card (dom M) = card (ran M)" and indom: "x \ dom M" shows "(M x = M y) \ (x = y)" proof - from ran_is_image finite card_eq have *: "inj_on (the \ M) (dom M)" using eq_card_imp_inj_on by metis thus ?thesis proof (cases "y \ dom M") case False with indom show ?thesis by auto next case True with indom have "the (M x) = the (M y) \ (x = y)" using inj_on_eq_iff[OF *] by auto thus ?thesis by auto qed qed lemma map_dom_ran_finite: "finite (dom M) \ finite (ran M)" by (simp add: ran_is_image) lemma map_update_eta_repair[simp]: (* An update operation may get simplified, if it happens to be eta-expanded. This lemma tries to repair some common expressions *) "dom (\x. if x=k then Some v else m x) = insert k (dom m)" "m k = None \ ran (\x. if x=k then Some v else m x) = insert v (ran m)" apply auto [] apply (force simp: ran_def) done lemma map_leI[intro?]: "\\x v. m1 x = Some v \ m2 x = Some v\ \ m1\\<^sub>mm2" unfolding map_le_def by force lemma map_leD: "m1\\<^sub>mm2 \ m1 k = Some v \ m2 k = Some v" unfolding map_le_def by force lemma map_restrict_insert_none_simp: "m x = None \ m|`(-insert x s) = m|`(-s)" by (auto intro!: ext simp:restrict_map_def) (* TODO: Move *) lemma eq_f_restr_conv: "s\dom (f A) \ A = f A |` (-s) \ A \\<^sub>m f A \ s = dom (f A) - dom A" apply auto subgoal by (metis map_leI restrict_map_eq(2)) subgoal by (metis ComplD restrict_map_eq(2)) subgoal by (metis Compl_iff restrict_in) subgoal by (force simp: map_le_def restrict_map_def) done corollary eq_f_restr_ss_eq: "\ s\dom (f A) \ \ A = f A |` (-s) \ A \\<^sub>m f A \ s = dom (f A) - dom A" using eq_f_restr_conv by blast subsubsection \Simultaneous Map Update\ definition "map_mmupd m K v k \ if k\K then Some v else m k" lemma map_mmupd_empty[simp]: "map_mmupd m {} v = m" by (auto simp: map_mmupd_def) lemma mmupd_in_upd[simp]: "k\K \ map_mmupd m K v k = Some v" by (auto simp: map_mmupd_def) lemma mmupd_notin_upd[simp]: "k\K \ map_mmupd m K v k = m k" by (auto simp: map_mmupd_def) lemma map_mmupdE: assumes "map_mmupd m K v k = Some x" obtains "k\K" "m k = Some x" | "k\K" "x=v" using assms by (auto simp: map_mmupd_def split: if_split_asm) lemma dom_mmupd[simp]: "dom (map_mmupd m K v) = dom m \ K" by (auto simp: map_mmupd_def split: if_split_asm) lemma le_map_mmupd_not_dom[simp, intro!]: "m \\<^sub>m map_mmupd m (K-dom m) v" by (auto simp: map_le_def) lemma map_mmupd_update_less: "K\K' \ map_mmupd m (K - dom m) v \\<^sub>m map_mmupd m (K'-dom m) v" by (auto simp: map_le_def map_mmupd_def) subsection\Connection between Maps and Sets of Key-Value Pairs\ definition map_to_set where "map_to_set m = {(k, v) . m k = Some v}" definition set_to_map where "set_to_map S k = Eps_Opt (\v. (k, v) \ S)" lemma set_to_map_simp : assumes inj_on_fst: "inj_on fst S" shows "(set_to_map S k = Some v) \ (k, v) \ S" proof (cases "\v. (k, v) \ S") case True note kv_ex = this then obtain v' where kv'_in: "(k, v') \ S" by blast with inj_on_fst have kv''_in: "\v''. (k, v'') \ S \ v' = v''" unfolding inj_on_def Ball_def by auto show ?thesis unfolding set_to_map_def by (simp add: kv_ex kv''_in) next case False hence kv''_nin: "\v''. (k, v'') \ S" by simp thus ?thesis by (simp add: set_to_map_def) qed lemma inj_on_fst_map_to_set : "inj_on fst (map_to_set m)" unfolding map_to_set_def inj_on_def by simp lemma map_to_set_inverse : "set_to_map (map_to_set m) = m" proof fix k show "set_to_map (map_to_set m) k = m k" proof (cases "m k") case None note mk_eq = this hence "\v. (k, v) \ map_to_set m" unfolding map_to_set_def by simp with set_to_map_simp [OF inj_on_fst_map_to_set, of m k] show ?thesis unfolding mk_eq by auto next case (Some v) note mk_eq = this hence "(k, v) \ map_to_set m" unfolding map_to_set_def by simp with set_to_map_simp [OF inj_on_fst_map_to_set, of m k v] show ?thesis unfolding mk_eq by auto qed qed lemma set_to_map_inverse : assumes inj_on_fst_S: "inj_on fst S" shows "map_to_set (set_to_map S) = S" proof (rule set_eqI) fix kv from set_to_map_simp [OF inj_on_fst_S, of "fst kv" "snd kv"] show "(kv \ map_to_set (set_to_map S)) = (kv \ S)" unfolding map_to_set_def by auto qed lemma map_to_set_empty[simp]: "map_to_set Map.empty = {}" unfolding map_to_set_def by simp lemma set_to_map_empty[simp]: "set_to_map {} = Map.empty" unfolding set_to_map_def[abs_def] by simp lemma map_to_set_empty_iff: "map_to_set m = {} \ m = Map.empty" "{} = map_to_set m \ m = Map.empty" unfolding map_to_set_def by auto lemma set_to_map_empty_iff: "set_to_map S = Map.empty \ S = {}" (is ?T1) "Map.empty = set_to_map S \ S = {}" (is ?T2) proof - show T1: ?T1 apply (simp only: set_eq_iff) apply (simp only: fun_eq_iff) apply (simp add: set_to_map_def) apply auto done from T1 show ?T2 by auto qed lemma map_to_set_upd[simp]: "map_to_set (m (k \ v)) = insert (k, v) (map_to_set m - {(k, v') |v'. True})" unfolding map_to_set_def apply (simp add: set_eq_iff) apply metis done lemma set_to_map_insert: assumes k_nin: "fst kv \ fst ` S" shows "set_to_map (insert kv S) = (set_to_map S) (fst kv \ snd kv)" proof fix k' obtain k v where kv_eq[simp]: "kv = (k, v)" by (rule prod.exhaust) from k_nin have k_nin': "\v'. (k, v') \ S" by (auto simp add: image_iff Ball_def) show "set_to_map (insert kv S) k' = (set_to_map S(fst kv \ snd kv)) k'" by (simp add: set_to_map_def k_nin') qed lemma map_to_set_dom : "dom m = fst ` (map_to_set m)" unfolding dom_def map_to_set_def by (auto simp add: image_iff) lemma map_to_set_ran : "ran m = snd ` (map_to_set m)" unfolding ran_def map_to_set_def by (auto simp add: image_iff) lemma set_to_map_dom : "dom (set_to_map S) = fst ` S" unfolding set_to_map_def[abs_def] dom_def by (auto simp add: image_iff Bex_def) lemma set_to_map_ran : "ran (set_to_map S) \ snd ` S" unfolding set_to_map_def[abs_def] ran_def subset_iff by (auto simp add: image_iff Bex_def) (metis Eps_Opt_eq_Some) lemma finite_map_to_set: "finite (map_to_set m) = finite (dom m)" unfolding map_to_set_def map_to_set_dom apply (intro iffI finite_imageI) apply assumption apply (rule finite_imageD[of fst]) apply assumption apply (simp add: inj_on_def) done lemma card_map_to_set : "card (map_to_set m) = card (dom m)" unfolding map_to_set_def map_to_set_dom apply (rule card_image[symmetric]) apply (simp add: inj_on_def) done lemma map_of_map_to_set : "distinct (map fst l) \ map_of l = m \ set l = map_to_set m" proof (induct l arbitrary: m) case Nil thus ?case by (simp add: map_to_set_empty_iff) blast next case (Cons kv l m) obtain k v where kv_eq[simp]: "kv = (k, v)" by (rule prod.exhaust) from Cons(2) have dist_l: "distinct (map fst l)" and kv'_nin: "\v'. (k, v') \ set l" by (auto simp add: image_iff) note ind_hyp = Cons(1)[OF dist_l] from kv'_nin have l_eq: "set (kv # l) = map_to_set m \ (set l = map_to_set (m (k := None))) \ m k = Some v" apply (simp add: map_to_set_def restrict_map_def set_eq_iff) apply (auto) apply (metis) apply (metis option.inject) done from kv'_nin have m_eq: "map_of (kv # l) = m \ map_of l = (m (k := None)) \ m k = Some v" apply (simp add: fun_eq_iff restrict_map_def map_of_eq_None_iff image_iff Ball_def) apply metis done show ?case unfolding m_eq l_eq using ind_hyp[of "m (k := None)"] by metis qed lemma map_to_set_map_of : "distinct (map fst l) \ map_to_set (map_of l) = set l" by (metis map_of_map_to_set) subsubsection \Mapping empty set to None\ definition "dflt_None_set S \ if S={} then None else Some S" lemma the_dflt_None_empty[simp]: "dflt_None_set {} = None" unfolding dflt_None_set_def by simp lemma the_dflt_None_nonempty[simp]: "S\{} \ dflt_None_set S = Some S" unfolding dflt_None_set_def by simp lemma the_dflt_None_set[simp]: "the_default {} (dflt_None_set x) = x" unfolding dflt_None_set_def by auto subsection \Orderings\ lemma (in order) min_arg_le[simp]: "n \ min m n \ min m n = n" "m \ min m n \ min m n = m" by (auto simp: min_def) lemma (in linorder) min_arg_not_ge[simp]: "\ min m n < m \ min m n = m" "\ min m n < n \ min m n = n" by (auto simp: min_def) lemma (in linorder) min_eq_arg[simp]: "min m n = m \ m\n" "min m n = n \ n\m" by (auto simp: min_def) lemma min_simps[simp]: "a<(b::'a::order) \ min a b = a" "b<(a::'a::order) \ min a b = b" by (auto simp add: min_def dest: less_imp_le) lemma (in -) min_less_self_conv[simp]: "min a b < a \ b < (a::_::linorder)" "min a b < b \ a < (b::_::linorder)" by (auto simp: min_def) lemma ord_eq_le_eq_trans: "\ a=b; b\c; c=d \ \ a\d" by auto lemma zero_comp_diff_simps[simp]: "(0::'a::linordered_idom) \ a - b \ b \ a" "(0::'a::linordered_idom) < a - b \ b < a" by auto subsubsection \Termination Measures\ text \Lexicographic measure, assuming upper bound for second component\ lemma mlex_fst_decrI: fixes a a' b b' N :: nat assumes "a a*N + N" using \b by linarith also have "\ \ a'*N" using \a by (metis Suc_leI ab_semigroup_add_class.add.commute ab_semigroup_mult_class.mult.commute mult_Suc_right mult_le_mono2) also have "\ \ a'*N + b'" by auto finally show ?thesis by auto qed lemma mlex_leI: fixes a a' b b' N :: nat assumes "a\a'" assumes "b\b'" shows "a*N + b \ a'*N + b'" using assms by (auto intro!: add_mono) lemma mlex_snd_decrI: fixes a a' b b' N :: nat assumes "a=a'" assumes "bCCPOs\ context ccpo begin lemma ccpo_Sup_mono: assumes C: "Complete_Partial_Order.chain (\) A" "Complete_Partial_Order.chain (\) B" assumes B: "\x\A. \y\B. x\y" shows "Sup A \ Sup B" proof (rule ccpo_Sup_least) fix x assume "x\A" with B obtain y where I: "y\B" and L: "x\y" by blast note L also from I ccpo_Sup_upper have "y\Sup B" by (blast intro: C) finally show "x\Sup B" . qed (rule C) lemma fixp_mono: assumes M: "monotone (\) (\) f" "monotone (\) (\) g" assumes LE: "\Z. f Z \ g Z" shows "ccpo_class.fixp f \ ccpo_class.fixp g" unfolding fixp_def[abs_def] apply (rule ccpo_Sup_mono) apply (rule chain_iterates M)+ proof rule fix x assume "x\ccpo_class.iterates f" thus "\y\ccpo_class.iterates g. x\y" proof (induct) case (step x) then obtain y where I: "y\ccpo_class.iterates g" and L: "x\y" by blast hence "g y \ ccpo_class.iterates g" and "f x \ g y" apply - apply (erule iterates.step) apply (rule order_trans) apply (erule monotoneD[OF M(1)]) apply (rule LE) done thus "\y\ccpo_class.iterates g. f x \ y" .. next case (Sup M) define N where "N = {SOME y. y\ccpo_class.iterates g \ x\y | x. x\M}" have N1: "\y\N. y\ccpo_class.iterates g \ (\x\M. x\y)" unfolding N_def apply auto apply (metis (lifting) Sup.hyps(2) tfl_some) by (metis (lifting) Sup.hyps(2) tfl_some) have N2: "\x\M. \y\N. x\y" unfolding N_def apply auto by (metis (lifting) Sup.hyps(2) tfl_some) have "N \ ccpo_class.iterates g" using Sup using N1 by auto hence C_chain: "Complete_Partial_Order.chain (\) N" using chain_iterates[OF M(2)] unfolding chain_def by auto have "Sup N \ ccpo_class.iterates g" and "Sup M \ Sup N" apply - apply (rule iterates.Sup[OF C_chain]) using N1 apply blast apply (rule ccpo_Sup_mono) apply (rule Sup.hyps) apply (rule C_chain) apply (rule N2) done thus ?case by blast qed qed end subsection \Code\ text \Constant for code-abort. If that gets executed, an abort-exception is raised.\ definition [simp]: "CODE_ABORT f = f ()" declare [[code abort: CODE_ABORT]] end diff --git a/thys/Buchi_Complementation/Complementation_Implement.thy b/thys/Buchi_Complementation/Complementation_Implement.thy --- a/thys/Buchi_Complementation/Complementation_Implement.thy +++ b/thys/Buchi_Complementation/Complementation_Implement.thy @@ -1,974 +1,975 @@ section \Complementation Implementation\ theory Complementation_Implement imports - "HOL-Library.Lattice_Syntax" "Transition_Systems_and_Automata.NBA_Implement" "Complementation" begin + unbundle lattice_syntax + type_synonym item = "nat \ bool" type_synonym 'state items = "'state \ item" type_synonym state = "(nat \ item) list" abbreviation "item_rel \ nat_rel \\<^sub>r bool_rel" abbreviation "state_rel \ \nat_rel, item_rel\ list_map_rel" abbreviation "pred A a q \ {p. q \ transition A a p}" subsection \Phase 1\ definition cs_lr :: "'state items \ 'state lr" where "cs_lr f \ map_option fst \ f" definition cs_st :: "'state items \ 'state st" where "cs_st f \ f -` Some ` snd -` {True}" abbreviation cs_abs :: "'state items \ 'state cs" where "cs_abs f \ (cs_lr f, cs_st f)" definition cs_rep :: "'state cs \ 'state items" where "cs_rep \ \ (g, P) p. map_option (\ k. (k, p \ P)) (g p)" lemma cs_abs_rep[simp]: "cs_rep (cs_abs f) = f" proof show "cs_rep (cs_abs f) x = f x" for x unfolding cs_lr_def cs_st_def cs_rep_def by (cases "f x") (force+) qed lemma cs_rep_lr[simp]: "cs_lr (cs_rep (g, P)) = g" proof show "cs_lr (cs_rep (g, P)) x = g x" for x unfolding cs_rep_def cs_lr_def by (cases "g x") (auto) qed lemma cs_rep_st[simp]: "cs_st (cs_rep (g, P)) = P \ dom g" unfolding cs_rep_def cs_st_def by force lemma cs_lr_dom[simp]: "dom (cs_lr f) = dom f" unfolding cs_lr_def by simp lemma cs_lr_apply[simp]: assumes "p \ dom f" shows "the (cs_lr f p) = fst (the (f p))" using assms unfolding cs_lr_def by auto lemma cs_rep_dom[simp]: "dom (cs_rep (g, P)) = dom g" unfolding cs_rep_def by auto lemma cs_rep_apply[simp]: assumes "p \ dom f" shows "fst (the (cs_rep (f, P) p)) = the (f p)" using assms unfolding cs_rep_def by auto abbreviation cs_rel :: "('state items \ 'state cs) set" where "cs_rel \ br cs_abs top" lemma cs_rel_inv_single_valued: "single_valued (cs_rel\)" by (auto intro!: inj_onI) (metis cs_abs_rep) definition refresh_1 :: "'state items \ 'state items" where "refresh_1 f \ if True \ snd ` ran f then f else map_option (apsnd top) \ f" definition ranks_1 :: "('label, 'state) nba \ 'label \ 'state items \ 'state items set" where "ranks_1 A a f \ {g. dom g = \((transition A a) ` (dom f)) \ (\ p \ dom f. \ q \ transition A a p. fst (the (g q)) \ fst (the (f p))) \ (\ q \ dom g. accepting A q \ even (fst (the (g q)))) \ cs_st g = {q \ \((transition A a) ` (cs_st f)). even (fst (the (g q)))}}" definition complement_succ_1 :: "('label, 'state) nba \ 'label \ 'state items \ 'state items set" where "complement_succ_1 A a = ranks_1 A a \ refresh_1" definition complement_1 :: "('label, 'state) nba \ ('label, 'state items) nba" where "complement_1 A \ nba (alphabet A) ({const (Some (2 * card (nodes A), False)) |` initial A}) (complement_succ_1 A) (\ f. cs_st f = {})" lemma refresh_1_dom[simp]: "dom (refresh_1 f) = dom f" unfolding refresh_1_def by simp lemma refresh_1_apply[simp]: "fst (the (refresh_1 f p)) = fst (the (f p))" unfolding refresh_1_def by (cases "f p") (auto) lemma refresh_1_cs_st[simp]: "cs_st (refresh_1 f) = (if cs_st f = {} then dom f else cs_st f)" unfolding refresh_1_def cs_st_def ran_def image_def vimage_def by auto lemma complement_succ_1_abs: assumes "g \ complement_succ_1 A a f" shows "cs_abs g \ complement_succ A a (cs_abs f)" unfolding complement_succ_def proof (simp, rule) have 1: "dom g = \((transition A a) ` (dom f))" "\ p \ dom f. \ q \ transition A a p. fst (the (g q)) \ fst (the (f p))" "\ p \ dom g. accepting A p \ even (fst (the (g p)))" using assms unfolding complement_succ_1_def ranks_1_def by simp_all show "cs_lr g \ lr_succ A a (cs_lr f)" unfolding lr_succ_def proof (intro CollectI conjI ballI impI) show "dom (cs_lr g) = \ (transition A a ` dom (cs_lr f))" using 1 by simp next fix p q assume 2: "p \ dom (cs_lr f)" "q \ transition A a p" have 3: "q \ dom (cs_lr g)" using 1 2 by auto show "the (cs_lr g q) \ the (cs_lr f p)" using 1 2 3 by simp next fix p assume 2: "p \ dom (cs_lr g)" "accepting A p" show "even (the (cs_lr g p))" using 1 2 by auto qed have 2: "cs_st g = {q \ \ (transition A a ` cs_st (refresh_1 f)). even (fst (the (g q)))}" using assms unfolding complement_succ_1_def ranks_1_def by simp show "cs_st g = st_succ A a (cs_lr g) (cs_st f)" proof (cases "cs_st f = {}") case True have 3: "the (cs_lr g q) = fst (the (g q))" if "q \ \((transition A a) ` (dom f))" for q using that 1(1) by simp show ?thesis using 2 3 unfolding st_succ_def refresh_1_cs_st True cs_lr_dom 1(1) by force next case False have 3: "the (cs_lr g q) = fst (the (g q))" if "q \ \((transition A a) ` (cs_st f))" for q using that 1(1) by (auto intro!: cs_lr_apply) (metis IntE UN_iff cs_abs_rep cs_lr_dom cs_rep_st domD prod.collapse) have "cs_st g = {q \ \ (transition A a ` cs_st (refresh_1 f)). even (fst (the (g q)))}" using 2 by this also have "cs_st (refresh_1 f) = cs_st f" using False by simp also have "{q \ \((transition A a) ` (cs_st f)). even (fst (the (g q)))} = {q \ \((transition A a) ` (cs_st f)). even (the (cs_lr g q))}" using 3 by metis also have "\ = st_succ A a (cs_lr g) (cs_st f)" unfolding st_succ_def using False by simp finally show ?thesis by this qed qed lemma complement_succ_1_rep: assumes "P \ dom f" "(g, Q) \ complement_succ A a (f, P)" shows "cs_rep (g, Q) \ complement_succ_1 A a (cs_rep (f, P))" unfolding complement_succ_1_def ranks_1_def comp_apply proof (intro CollectI conjI ballI impI) have 1: "dom g = \((transition A a) ` (dom f))" "\ p \ dom f. \ q \ transition A a p. the (g q) \ the (f p)" "\ p \ dom g. accepting A p \ even (the (g p))" using assms(2) unfolding complement_succ_def lr_succ_def by simp_all have 2: "Q = {q \ if P = {} then dom g else \((transition A a) ` P). even (the (g q))}" using assms(2) unfolding complement_succ_def st_succ_def by simp have 3: "Q \ dom g" unfolding 2 1(1) using assms(1) by auto show "dom (cs_rep (g, Q)) = \ (transition A a ` dom (refresh_1 (cs_rep (f, P))))" using 1 by simp show "\ p q. p \ dom (refresh_1 (cs_rep (f, P))) \ q \ transition A a p \ fst (the (cs_rep (g, Q) q)) \ fst (the (refresh_1 (cs_rep (f, P)) p))" using 1(1, 2) by (auto) (metis UN_I cs_rep_apply domI option.sel) show "\ p. p \ dom (cs_rep (g, Q)) \ accepting A p \ even (fst (the (cs_rep (g, Q) p)))" using 1(1, 3) by auto show "cs_st (cs_rep (g, Q)) = {q \ \ (transition A a ` cs_st (refresh_1 (cs_rep (f, P)))). even (fst (the (cs_rep (g, Q) q)))}" proof (cases "P = {}") case True have "cs_st (cs_rep (g, Q)) = Q" using 3 by auto also have "\ = {q \ dom g. even (the (g q))}" unfolding 2 using True by auto also have "\ = {q \ dom g. even (fst (the (cs_rep (g, Q) q)))}" using cs_rep_apply by metis also have "dom g = \((transition A a) ` (dom f))" using 1(1) by this also have "dom f = cs_st (refresh_1 (cs_rep (f, P)))" using True by simp finally show ?thesis by this next case False have 4: "fst (the (cs_rep (g, Q) q)) = the (g q)" if "q \ \((transition A a) ` P)" for q using 1(1) that assms(1) by (fast intro: cs_rep_apply) have "cs_st (cs_rep (g, Q)) = Q" using 3 by auto also have "\ = {q \ \((transition A a) ` P). even (the (g q))}" unfolding 2 using False by auto also have "\ = {q \ \((transition A a) ` P). even (fst (the (cs_rep (g, Q) q)))}" using 4 by force also have "P = (cs_st (refresh_1 (cs_rep (f, P))))" using assms(1) False by auto finally show ?thesis by simp qed qed lemma complement_succ_1_refine: "(complement_succ_1, complement_succ) \ Id \ Id \ cs_rel \ \cs_rel\ set_rel" proof (clarsimp simp: br_set_rel_alt in_br_conv) fix A :: "('a, 'b) nba" fix a f show "complement_succ A a (cs_abs f) = cs_abs ` complement_succ_1 A a f" proof safe fix g Q assume 1: "(g, Q) \ complement_succ A a (cs_abs f)" have 2: "Q \ dom g" using 1 unfolding complement_succ_def lr_succ_def st_succ_def by (auto) (metis IntE cs_abs_rep cs_lr_dom cs_rep_st) have 3: "cs_st f \ dom (cs_lr f)" unfolding cs_st_def by auto show "(g, Q) \ cs_abs ` complement_succ_1 A a f" proof show "(g, Q) = cs_abs (cs_rep (g, Q))" using 2 by auto have "cs_rep (g, Q) \ complement_succ_1 A a (cs_rep (cs_abs f))" using complement_succ_1_rep 3 1 by this also have "cs_rep (cs_abs f) = f" by simp finally show "cs_rep (g, Q) \ complement_succ_1 A a f" by this qed next fix g assume 1: "g \ complement_succ_1 A a f" show "cs_abs g \ complement_succ A a (cs_abs f)" using complement_succ_1_abs 1 by this qed qed lemma complement_1_refine: "(complement_1, complement) \ \Id, Id\ nba_rel \ \Id, cs_rel\ nba_rel" unfolding complement_1_def complement_def proof parametricity fix A B :: "('a, 'b) nba" assume 1: "(A, B) \ \Id, Id\ nba_rel" have 2: "(const (Some (2 * card (nodes B), False)) |` initial B, const (Some (2 * card (nodes B))) |` initial B, {}) \ cs_rel" unfolding cs_lr_def cs_st_def in_br_conv by (force simp: restrict_map_def) show "(complement_succ_1 A, complement_succ B) \ Id \ cs_rel \ \cs_rel\ set_rel" using complement_succ_1_refine 1 by parametricity auto show "({const (Some (2 * card (nodes A), False)) |` initial A}, {const (Some (2 * card (nodes B))) |` initial B} \ {{}}) \ \cs_rel\ set_rel" using 1 2 by simp parametricity show "(\ f. cs_st f = {}, \ (f, P). P = {}) \ cs_rel \ bool_rel" by (auto simp: in_br_conv) qed subsection \Phase 2\ definition ranks_2 :: "('label, 'state) nba \ 'label \ 'state items \ 'state items set" where "ranks_2 A a f \ {g. dom g = \((transition A a) ` (dom f)) \ (\ q l d. g q = Some (l, d) \ l \ \ (fst ` Some -` f ` pred A a q) \ (d \ \ (snd ` Some -` f ` pred A a q) \ even l) \ (accepting A q \ even l))}" definition complement_succ_2 :: "('label, 'state) nba \ 'label \ 'state items \ 'state items set" where "complement_succ_2 A a \ ranks_2 A a \ refresh_1" definition complement_2 :: "('label, 'state) nba \ ('label, 'state items) nba" where "complement_2 A \ nba (alphabet A) ({const (Some (2 * card (nodes A), False)) |` initial A}) (complement_succ_2 A) (\ f. True \ snd ` ran f)" lemma ranks_2_refine: "ranks_2 = ranks_1" proof (intro ext) fix A :: "('a, 'b) nba" and a f show "ranks_2 A a f = ranks_1 A a f" proof safe fix g assume 1: "g \ ranks_2 A a f" have 2: "dom g = \((transition A a) ` (dom f))" using 1 unfolding ranks_2_def by auto have 3: "g q = Some (l, d) \ l \ \ (fst ` Some -` f ` pred A a q)" for q l d using 1 unfolding ranks_2_def by auto have 4: "g q = Some (l, d) \ d \ \ (snd ` Some -` f ` pred A a q) \ even l" for q l d using 1 unfolding ranks_2_def by auto have 5: "g q = Some (l, d) \ accepting A q \ even l" for q l d using 1 unfolding ranks_2_def by auto show "g \ ranks_1 A a f" unfolding ranks_1_def proof (intro CollectI conjI ballI impI) show "dom g = \((transition A a) ` (dom f))" using 2 by this next fix p q assume 10: "p \ dom f" "q \ transition A a p" obtain k c where 11: "f p = Some (k, c)" using 10(1) by auto have 12: "q \ dom g" using 10 2 by auto obtain l d where 13: "g q = Some (l, d)" using 12 by auto have "fst (the (g q)) = l" unfolding 13 by simp also have "\ \ \ (fst ` Some -` f ` pred A a q)" using 3 13 by this also have "\ \ k" proof (rule cInf_lower) show "k \ fst ` Some -` f ` pred A a q" using 11 10(2) by force show "bdd_below (fst ` Some -` f ` pred A a q)" by simp qed also have "\ = fst (the (f p))" unfolding 11 by simp finally show "fst (the (g q)) \ fst (the (f p))" by this next fix q assume 10: "q \ dom g" "accepting A q" show "even (fst (the (g q)))" using 10 5 by auto next show "cs_st g = {q \ \((transition A a) ` (cs_st f)). even (fst (the (g q)))}" proof show "cs_st g \ {q \ \((transition A a) ` (cs_st f)). even (fst (the (g q)))}" using 4 unfolding cs_st_def image_def vimage_def by auto metis+ show "{q \ \((transition A a) ` (cs_st f)). even (fst (the (g q)))} \ cs_st g" proof safe fix p q assume 10: "even (fst (the (g q)))" "p \ cs_st f" "q \ transition A a p" have 12: "q \ dom g" using 10 2 unfolding cs_st_def by auto show "q \ cs_st g" using 10 4 12 unfolding cs_st_def image_def by force qed qed qed next fix g assume 1: "g \ ranks_1 A a f" have 2: "dom g = \((transition A a) ` (dom f))" using 1 unfolding ranks_1_def by auto have 3: "\ p q. p \ dom f \ q \ transition A a p \ fst (the (g q)) \ fst (the (f p))" using 1 unfolding ranks_1_def by auto have 4: "\ q. q \ dom g \ accepting A q \ even (fst (the (g q)))" using 1 unfolding ranks_1_def by auto have 5: "cs_st g = {q \ \((transition A a) ` (cs_st f)). even (fst (the (g q)))}" using 1 unfolding ranks_1_def by auto show "g \ ranks_2 A a f" unfolding ranks_2_def proof (intro CollectI conjI allI impI) show "dom g = \((transition A a) ` (dom f))" using 2 by this next fix q l d assume 10: "g q = Some (l, d)" have 11: "q \ dom g" using 10 by auto show "l \ \ (fst ` Some -` f ` pred A a q)" proof (rule cInf_greatest) show "fst ` Some -` f ` pred A a q \ {}" using 11 unfolding 2 image_def vimage_def by force show "\ x. x \ fst ` Some -` f ` pred A a q \ l \ x" using 3 10 by (auto) (metis domI fst_conv option.sel) qed have "d \ q \ cs_st g" unfolding cs_st_def by (force simp: 10) also have "cs_st g = {q \ \((transition A a) ` (cs_st f)). even (fst (the (g q)))}" using 5 by this also have "q \ \ \ (\ x \ cs_st f. q \ transition A a x) \ even l" unfolding mem_Collect_eq 10 by simp also have "\ \ \ (snd ` Some -` f ` pred A a q) \ even l" unfolding cs_st_def image_def vimage_def by auto metis+ finally show "d \ \ (snd ` Some -` f ` pred A a q) \ even l" by this show "accepting A q \ even l" using 4 10 11 by force qed qed qed lemma complement_2_refine: "(complement_2, complement_1) \ \Id, Id\ nba_rel \ \Id, Id\ nba_rel" unfolding complement_2_def complement_1_def complement_succ_2_def complement_succ_1_def unfolding ranks_2_refine cs_st_def image_def vimage_def ran_def by auto subsection \Phase 3\ definition bounds_3 :: "('label, 'state) nba \ 'label \ 'state items \ 'state items" where "bounds_3 A a f \ \ q. let S = Some -` f ` pred A a q in if S = {} then None else Some (\(fst ` S), \(snd ` S))" definition items_3 :: "('label, 'state) nba \ 'state \ item \ item set" where "items_3 A p \ \ (k, c). {(l, c \ even l) |l. l \ k \ (accepting A p \ even l)}" definition get_3 :: "('label, 'state) nba \ 'state items \ ('state \ item set)" where "get_3 A f \ \ p. map_option (items_3 A p) (f p)" definition complement_succ_3 :: "('label, 'state) nba \ 'label \ 'state items \ 'state items set" where "complement_succ_3 A a \ expand_map \ get_3 A \ bounds_3 A a \ refresh_1" definition complement_3 :: "('label, 'state) nba \ ('label, 'state items) nba" where "complement_3 A \ nba (alphabet A) ({(Some \ (const (2 * card (nodes A), False))) |` initial A}) (complement_succ_3 A) (\ f. \ (p, k, c) \ map_to_set f. \ c)" lemma bounds_3_dom[simp]: "dom (bounds_3 A a f) = \((transition A a) ` (dom f))" unfolding bounds_3_def Let_def dom_def by (force split: if_splits) lemma items_3_nonempty[intro!, simp]: "items_3 A p s \ {}" unfolding items_3_def by auto lemma items_3_finite[intro!, simp]: "finite (items_3 A p s)" unfolding items_3_def by (auto split: prod.splits) lemma get_3_dom[simp]: "dom (get_3 A f) = dom f" unfolding get_3_def by (auto split: bind_splits) lemma get_3_finite[intro, simp]: "S \ ran (get_3 A f) \ finite S" unfolding get_3_def ran_def by auto lemma get_3_update[simp]: "get_3 A (f (p \ s)) = (get_3 A f) (p \ items_3 A p s)" unfolding get_3_def by auto lemma expand_map_get_bounds_3: "expand_map \ get_3 A \ bounds_3 A a = ranks_2 A a" proof (intro ext set_eqI, unfold comp_apply) fix f g have 1: "(\ x S y. get_3 A (bounds_3 A a f) x = Some S \ g x = Some y \ y \ S) \ (\ q S l d. get_3 A (bounds_3 A a f) q = Some S \ g q = Some (l, d) \ (l, d) \ S)" by auto have 2: "(\ S. get_3 A (bounds_3 A a f) q = Some S \ g q = Some (l, d) \ (l, d) \ S) \ (g q = Some (l, d) \ l \ \(fst ` (Some -` f ` pred A a q)) \ (d \ \(snd ` (Some -` f ` pred A a q)) \ even l) \ (accepting A q \ even l))" if 3: "dom g = \((transition A a) ` (dom f))" for q l d proof - have 4: "q \ dom g" if "Some -` f ` pred A a q = {}" unfolding 3 using that by force show ?thesis unfolding get_3_def items_3_def bounds_3_def Let_def using 4 by auto qed show "g \ expand_map (get_3 A (bounds_3 A a f)) \ g \ ranks_2 A a f" unfolding expand_map_alt_def ranks_2_def mem_Collect_eq unfolding get_3_dom bounds_3_dom 1 using 2 by blast qed lemma complement_succ_3_refine: "complement_succ_3 = complement_succ_2" unfolding complement_succ_3_def complement_succ_2_def expand_map_get_bounds_3 by rule lemma complement_initial_3_refine: "{const (Some (2 * card (nodes A), False)) |` initial A} = {(Some \ (const (2 * card (nodes A), False))) |` initial A}" unfolding comp_apply by rule lemma complement_accepting_3_refine: "True \ snd ` ran f \ (\ (p, k, c) \ map_to_set f. \ c)" unfolding map_to_set_def ran_def by auto lemma complement_3_refine: "(complement_3, complement_2) \ \Id, Id\ nba_rel \ \Id, Id\ nba_rel" unfolding complement_3_def complement_2_def unfolding complement_succ_3_refine complement_initial_3_refine complement_accepting_3_refine by auto subsection \Phase 4\ definition items_4 :: "('label, 'state) nba \ 'state \ item \ item set" where "items_4 A p \ \ (k, c). {(l, c \ even l) |l. k \ Suc l \ l \ k \ (accepting A p \ even l)}" definition get_4 :: "('label, 'state) nba \ 'state items \ ('state \ item set)" where "get_4 A f \ \ p. map_option (items_4 A p) (f p)" definition complement_succ_4 :: "('label, 'state) nba \ 'label \ 'state items \ 'state items set" where "complement_succ_4 A a \ expand_map \ get_4 A \ bounds_3 A a \ refresh_1" definition complement_4 :: "('label, 'state) nba \ ('label, 'state items) nba" where "complement_4 A \ nba (alphabet A) ({(Some \ (const (2 * card (nodes A), False))) |` initial A}) (complement_succ_4 A) (\ f. \ (p, k, c) \ map_to_set f. \ c)" lemma get_4_dom[simp]: "dom (get_4 A f) = dom f" unfolding get_4_def by (auto split: bind_splits) definition R :: "'state items rel" where "R \ {(f, g). dom f = dom g \ (\ p \ dom f. fst (the (f p)) \ fst (the (g p))) \ (\ p \ dom f. snd (the (f p)) \ snd (the (g p)))}" lemma bounds_R: assumes "(f, g) \ R" assumes "bounds_3 A a (refresh_1 f) p = Some (n, e)" assumes "bounds_3 A a (refresh_1 g) p = Some (k, c)" shows "n \ k" "e \ c" proof - have 1: "dom f = dom g" "\ p \ dom f. fst (the (f p)) \ fst (the (g p))" "\ p \ dom f. snd (the (f p)) \ snd (the (g p))" using assms(1) unfolding R_def by auto have "n = \(fst ` (Some -` refresh_1 f ` pred A a p))" using assms(2) unfolding bounds_3_def by (auto simp: Let_def split: if_splits) also have "fst ` Some -` refresh_1 f ` pred A a p = fst ` Some -` f ` pred A a p" proof show " fst ` Some -` refresh_1 f ` pred A a p \ fst ` Some -` f ` pred A a p" unfolding refresh_1_def image_def by (auto simp: map_option_case split: option.split) (force) show "fst ` Some -` f ` pred A a p \ fst ` Some -` refresh_1 f ` pred A a p" unfolding refresh_1_def image_def by (auto simp: map_option_case split: option.split) (metis fst_conv option.sel) qed also have "\ = fst ` Some -` f ` (pred A a p \ dom f)" unfolding dom_def image_def Int_def by auto metis also have "\ = fst ` the ` f ` (pred A a p \ dom f)" unfolding dom_def by force also have "\ = (fst \ the \ f) ` (pred A a p \ dom f)" by force also have "\((fst \ the \ f) ` (pred A a p \ dom f)) \ \((fst \ the \ g) ` (pred A a p \ dom g))" proof (rule cINF_mono) show "pred A a p \ dom g \ {}" using assms(2) 1(1) unfolding bounds_3_def refresh_1_def by (auto simp: Let_def split: if_splits) (force+) show "bdd_below ((fst \ the \ f) ` (pred A a p \ dom f))" by rule show "\ n \ pred A a p \ dom f. (fst \ the \ f) n \ (fst \ the \ g) m" if "m \ pred A a p \ dom g" for m using 1 that by auto qed also have "(fst \ the \ g) ` (pred A a p \ dom g) = fst ` the ` g ` (pred A a p \ dom g)" by force also have "\ = fst ` Some -` g ` (pred A a p \ dom g)" unfolding dom_def by force also have "\ = fst ` Some -` g ` pred A a p" unfolding dom_def image_def Int_def by auto metis also have "\ = fst ` Some -` refresh_1 g ` pred A a p" proof show "fst ` Some -` g ` pred A a p \ fst ` Some -` refresh_1 g ` pred A a p" unfolding refresh_1_def image_def by (auto simp: map_option_case split: option.split) (metis fst_conv option.sel) show "fst ` Some -` refresh_1 g ` pred A a p \ fst ` Some -` g ` pred A a p" unfolding refresh_1_def image_def by (auto simp: map_option_case split: option.split) (force) qed also have "\(fst ` (Some -` refresh_1 g ` pred A a p)) = k" using assms(3) unfolding bounds_3_def by (auto simp: Let_def split: if_splits) finally show "n \ k" by this have "e \ \(snd ` (Some -` refresh_1 f ` pred A a p))" using assms(2) unfolding bounds_3_def by (auto simp: Let_def split: if_splits) also have "snd ` Some -` refresh_1 f ` pred A a p = snd ` Some -` refresh_1 f ` (pred A a p \ dom (refresh_1 f))" unfolding dom_def image_def Int_def by auto metis also have "\ = snd ` the ` refresh_1 f ` (pred A a p \ dom (refresh_1 f))" unfolding dom_def by force also have "\ = (snd \ the \ refresh_1 f) ` (pred A a p \ dom (refresh_1 f))" by force also have "\ = (snd \ the \ refresh_1 g) ` (pred A a p \ dom (refresh_1 g))" proof (rule image_cong) show "pred A a p \ dom (refresh_1 f) = pred A a p \ dom (refresh_1 g)" unfolding refresh_1_dom 1(1) by rule show "(snd \ the \ refresh_1 f) q \ (snd \ the \ refresh_1 g) q" if 2: "q \ pred A a p \ dom (refresh_1 g)" for q proof have 3: "\ x \ ran f. \ snd x \ (n, True) \ ran g \ g q = Some (k, c) \ c" for n k c using 1(1, 3) unfolding dom_def ran_def by (auto dest!: Collect_inj) (metis option.sel snd_conv) have 4: "g q = Some (n, True) \ f q = Some (k, c) \ c" for n k c using 1(3) unfolding dom_def by force have 5: "\ x \ ran g. \ snd x \ (k, True) \ ran f \ False" for k using 1(1, 3) unfolding dom_def ran_def by (auto dest!: Collect_inj) (metis option.sel snd_conv) show "(snd \ the \ refresh_1 f) q \ (snd \ the \ refresh_1 g) q" using 1(1, 3) 2 3 unfolding refresh_1_def by (force split: if_splits) show "(snd \ the \ refresh_1 g) q \ (snd \ the \ refresh_1 f) q" using 1(1, 3) 2 4 5 unfolding refresh_1_def by (auto simp: map_option_case split: option.splits if_splits) (force+) qed qed also have "\ = snd ` the ` refresh_1 g ` (pred A a p \ dom (refresh_1 g))" by force also have "\ = snd ` Some -` refresh_1 g ` (pred A a p \ dom (refresh_1 g))" unfolding dom_def by force also have "\ = snd ` Some -` refresh_1 g ` pred A a p" unfolding dom_def image_def Int_def by auto metis also have "\(snd ` (Some -` refresh_1 g ` pred A a p)) \ c" using assms(3) unfolding bounds_3_def by (auto simp: Let_def split: if_splits) finally show "e \ c" by this qed lemma complement_4_language_1: "language (complement_3 A) \ language (complement_4 A)" proof (rule simulation_language) show "alphabet (complement_3 A) \ alphabet (complement_4 A)" unfolding complement_3_def complement_4_def by simp show "\ q \ initial (complement_4 A). (p, q) \ R" if "p \ initial (complement_3 A)" for p using that unfolding complement_3_def complement_4_def R_def by simp show "\ g' \ transition (complement_4 A) a g. (f', g') \ R" if "f' \ transition (complement_3 A) a f" "(f, g) \ R" for a f f' g proof - have 1: "f' \ expand_map (get_3 A (bounds_3 A a (refresh_1 f)))" using that(1) unfolding complement_3_def complement_succ_3_def by auto have 2: "dom f = dom g" "\ p \ dom f. fst (the (f p)) \ fst (the (g p))" "\ p \ dom f. snd (the (f p)) \ snd (the (g p))" using that(2) unfolding R_def by auto have "dom f' = dom (get_3 A (bounds_3 A a (refresh_1 f)))" using expand_map_dom 1 by this also have "\ = dom (bounds_3 A a (refresh_1 f))" by simp finally have 3: "dom f' = dom (bounds_3 A a (refresh_1 f))" by this define g' where "g' p \ do { (k, c) \ bounds_3 A a (refresh_1 g) p; (l, d) \ f' p; Some (if even k = even l then k else k - 1, d) }" for p have 4: "g' p = do { kc \ bounds_3 A a (refresh_1 g) p; ld \ f' p; Some (if even (fst kc) = even (fst ld) then fst kc else fst kc - 1, snd ld) }" for p unfolding g'_def case_prod_beta by rule have "dom g' = dom (bounds_3 A a (refresh_1 g)) \ dom f'" using 4 bind_eq_Some_conv by fastforce also have "\ = dom f'" using 2 3 by simp finally have 5: "dom g' = dom f'" by this have 6: "(l, d) \ items_3 A p (k, c)" if "bounds_3 A a (refresh_1 f) p = Some (k, c)" "f' p = Some (l, d)" for p k c l d using 1 that unfolding expand_map_alt_def get_3_def by blast show ?thesis unfolding complement_4_def nba.sel complement_succ_4_def comp_apply proof show "(f', g') \ R" unfolding R_def mem_Collect_eq prod.case proof (intro conjI ballI) show "dom f' = dom g'" using 5 by rule next fix p assume 10: "p \ dom f'" have 11: "p \ dom (bounds_3 A a (refresh_1 g))" using 2(1) 3 10 by simp obtain k c where 12: "bounds_3 A a (refresh_1 g) p = Some (k, c)" using 11 by fast obtain l d where 13: "f' p = Some (l, d)" using 10 by auto obtain n e where 14: "bounds_3 A a (refresh_1 f) p = Some (n, e)" using 10 3 by fast have 15: "(l, d) \ items_3 A p (n, e)" using 6 14 13 by this have 16: "n \ k" using bounds_R(1) that(2) 14 12 by this have 17: "l \ k" using 15 16 unfolding items_3_def by simp have 18: "even k \ odd l \ l \ k \ l \ k - 1" by presburger have 19: "e \ c" using bounds_R(2) that(2) 14 12 by this show "fst (the (f' p)) \ fst (the (g' p))" using 17 18 unfolding 4 12 13 by simp show "snd (the (f' p)) \ snd (the (g' p))" using 19 unfolding 4 12 13 by simp qed show "g' \ expand_map (get_4 A (bounds_3 A a (refresh_1 g)))" unfolding expand_map_alt_def mem_Collect_eq proof (intro conjI allI impI) show "dom g' = dom (get_4 A (bounds_3 A a (refresh_1 g)))" using 2(1) 3 5 by simp fix p S xy assume 10: "get_4 A (bounds_3 A a (refresh_1 g)) p = Some S" assume 11: "g' p = Some xy" obtain k c where 12: "bounds_3 A a (refresh_1 g) p = Some (k, c)" "S = items_4 A p (k, c)" using 10 unfolding get_4_def by auto obtain l d where 13: "f' p = Some (l, d)" "xy = (if even k \ even l then k else k - 1, d)" using 11 12 unfolding g'_def by (auto split: bind_splits) obtain n e where 14: "bounds_3 A a (refresh_1 f) p = Some (n, e)" using 13(1) 3 by fast have 15: "(l, d) \ items_3 A p (n, e)" using 6 14 13(1) by this have 16: "n \ k" using bounds_R(1) that(2) 14 12(1) by this have 17: "e \ c" using bounds_R(2) that(2) 14 12(1) by this show "xy \ S" using 15 16 17 unfolding 12(2) 13(2) items_3_def items_4_def by auto qed qed qed show "\ p q. (p, q) \ R \ accepting (complement_3 A) p \ accepting (complement_4 A) q" unfolding complement_3_def complement_4_def R_def map_to_set_def by (auto) (metis domIff eq_snd_iff option.exhaust_sel option.sel) qed lemma complement_4_less: "complement_4 A \ complement_3 A" unfolding less_eq_nba_def unfolding complement_4_def complement_3_def nba.sel unfolding complement_succ_4_def complement_succ_3_def proof (safe intro!: le_funI, unfold comp_apply) fix a f g assume "g \ expand_map (get_4 A (bounds_3 A a (refresh_1 f)))" then show "g \ expand_map (get_3 A (bounds_3 A a (refresh_1 f)))" unfolding get_4_def get_3_def items_4_def items_3_def expand_map_alt_def by blast qed lemma complement_4_language_2: "language (complement_4 A) \ language (complement_3 A)" using language_mono complement_4_less by (auto dest: monoD) lemma complement_4_language: "language (complement_3 A) = language (complement_4 A)" using complement_4_language_1 complement_4_language_2 by blast lemma complement_4_finite[simp]: assumes "finite (nodes A)" shows "finite (nodes (complement_4 A))" proof - have "(nodes (complement_3 A), nodes (complement_2 A)) \ \Id\ set_rel" using complement_3_refine by parametricity auto also have "(nodes (complement_2 A), nodes (complement_1 A)) \ \Id\ set_rel" using complement_2_refine by parametricity auto also have "(nodes (complement_1 A), nodes (complement A)) \ \cs_rel\ set_rel" using complement_1_refine by parametricity auto finally have 1: "(nodes (complement_3 A), nodes (complement A)) \ \cs_rel\ set_rel" by simp have 2: "finite (nodes (complement A))" using complement_finite assms(1) by this have 3: "finite (nodes (complement_3 A))" using finite_set_rel_transfer_back 1 cs_rel_inv_single_valued 2 by this have 4: "nodes (complement_4 A) \ nodes (complement_3 A)" using nodes_mono complement_4_less by (auto dest: monoD) show "finite (nodes (complement_4 A))" using finite_subset 4 3 by this qed lemma complement_4_correct: assumes "finite (nodes A)" shows "language (complement_4 A) = streams (alphabet A) - language A" proof - have "language (complement_4 A) = language (complement_3 A)" using complement_4_language by rule also have "(language (complement_3 A), language (complement_2 A)) \ \\Id\ stream_rel\ set_rel" using complement_3_refine by parametricity auto also have "(language (complement_2 A), language (complement_1 A)) \ \\Id\ stream_rel\ set_rel" using complement_2_refine by parametricity auto also have "(language (complement_1 A), language (complement A)) \ \\Id\ stream_rel\ set_rel" using complement_1_refine by parametricity auto also have "language (complement A) = streams (alphabet A) - language A" using complement_language assms(1) by this finally show "language (complement_4 A) = streams (alphabet A) - language A" by simp qed subsection \Phase 5\ definition refresh_5 :: "'state items \ 'state items nres" where "refresh_5 f \ if \ (p, k, c) \ map_to_set f. c then RETURN f else do { ASSUME (finite (dom f)); FOREACH (map_to_set f) (\ (p, k, c) m. do { ASSERT (p \ dom m); RETURN (m (p \ (k, True))) } ) Map.empty }" definition merge_5 :: "item \ item option \ item" where "merge_5 \ \ (k, c). \ None \ (k, c) | Some (l, d) \ (k \ l, c \ d)" definition bounds_5 :: "('label, 'state) nba \ 'label \ 'state items \ 'state items nres" where "bounds_5 A a f \ do { ASSUME (finite (dom f)); ASSUME (\ p. finite (transition A a p)); FOREACH (map_to_set f) (\ (p, s) m. FOREACH (transition A a p) (\ q f. RETURN (f (q \ merge_5 s (f q)))) m) Map.empty }" definition items_5 :: "('label, 'state) nba \ 'state \ item \ item set" where "items_5 A p \ \ (k, c). do { let values = if accepting A p then Set.filter even {k - 1 .. k} else {k - 1 .. k}; let item = \ l. (l, c \ even l); item ` values }" definition get_5 :: "('label, 'state) nba \ 'state items \ ('state \ item set)" where "get_5 A f \ \ p. map_option (items_5 A p) (f p)" definition expand_5 :: "('a \ 'b set) \ ('a \ 'b) set nres" where "expand_5 f \ FOREACH (map_to_set f) (\ (x, S) X. do { ASSERT (\ g \ X. x \ dom g); ASSERT (\ a \ S. \ b \ S. a \ b \ (\ y. (\ g. g (x \ y)) ` X) a \ (\ y. (\ g. g (x \ y)) ` X) b = {}); RETURN (\ y \ S. (\ g. g (x \ y)) ` X) }) {Map.empty}" definition complement_succ_5 :: "('label, 'state) nba \ 'label \ 'state items \ 'state items set nres" where "complement_succ_5 A a f \ do { f \ refresh_5 f; f \ bounds_5 A a f; ASSUME (finite (dom f)); expand_5 (get_5 A f) }" lemma bounds_3_empty: "bounds_3 A a Map.empty = Map.empty" unfolding bounds_3_def Let_def by auto lemma bounds_3_update: "bounds_3 A a (f (p \ s)) = override_on (bounds_3 A a f) (Some \ merge_5 s \ bounds_3 A a (f (p := None))) (transition A a p)" proof note fun_upd_image[simp] fix q show "bounds_3 A a (f (p \ s)) q = override_on (bounds_3 A a f) (Some \ merge_5 s \ bounds_3 A a (f (p := None))) (transition A a p) q" proof (cases "q \ transition A a p") case True define S where "S \ Some -` f ` (pred A a q - {p})" have 1: "Some -` f (p := Some s) ` pred A a q = insert s S" using True unfolding S_def by auto have 2: "Some -` f (p := None) ` pred A a q = S" unfolding S_def by auto have "bounds_3 A a (f (p \ s)) q = Some (\(fst ` (insert s S)), \(snd ` (insert s S)))" unfolding bounds_3_def 1 by simp also have "\ = Some (merge_5 s (bounds_3 A a (f (p := None)) q))" unfolding 2 bounds_3_def merge_5_def by (cases s) (simp_all add: cInf_insert) also have "\ = override_on (bounds_3 A a f) (Some \ merge_5 s \ bounds_3 A a (f (p := None))) (transition A a p) q" using True by simp finally show ?thesis by this next case False then have "pred A a q \ {x. x \ p} = pred A a q" by auto with False show ?thesis by (simp add: bounds_3_def) qed qed lemma refresh_5_refine: "(refresh_5, \ f. RETURN (refresh_1 f)) \ Id \ \Id\ nres_rel" proof safe fix f :: "'a \ item option" have 1: "(\ (p, k, c) \ map_to_set f. c) \ True \ snd ` ran f" unfolding image_def map_to_set_def ran_def by force show "(refresh_5 f, RETURN (refresh_1 f)) \ \Id\ nres_rel" unfolding refresh_5_def refresh_1_def 1 by (refine_vcg FOREACH_rule_map_eq[where X = "\ m. map_option (apsnd \) \ m"]) (auto) qed lemma bounds_5_refine: "(bounds_5 A a, \ f. RETURN (bounds_3 A a f)) \ Id \ \Id\ nres_rel" unfolding bounds_5_def by (refine_vcg FOREACH_rule_map_eq[where X = "bounds_3 A a"] FOREACH_rule_insert_eq) (auto simp: override_on_insert bounds_3_empty bounds_3_update) lemma items_5_refine: "items_5 = items_4" unfolding items_5_def items_4_def by (intro ext) (auto split: if_splits) lemma get_5_refine: "get_5 = get_4" unfolding get_5_def get_4_def items_5_refine by rule lemma expand_5_refine: "(expand_5 f, ASSERT (finite (dom f)) \ RETURN (expand_map f)) \ \Id\ nres_rel" unfolding expand_5_def by (refine_vcg FOREACH_rule_map_eq[where X = expand_map]) (auto dest!: expand_map_dom map_upd_eqD1) lemma complement_succ_5_refine: "(complement_succ_5, RETURN \\\ complement_succ_4) \ Id \ Id \ Id \ \Id\ nres_rel" unfolding complement_succ_5_def complement_succ_4_def get_5_refine comp_apply by (refine_vcg vcg1[OF refresh_5_refine] vcg1[OF bounds_5_refine] vcg0[OF expand_5_refine]) (auto) subsection \Phase 6\ definition expand_map_get_6 :: "('label, 'state) nba \ 'state items \ 'state items set nres" where "expand_map_get_6 A f \ FOREACH (map_to_set f) (\ (k, v) X. do { ASSERT (\ g \ X. k \ dom g); ASSERT (\ a \ (items_5 A k v). \ b \ (items_5 A k v). a \ b \ (\ y. (\ g. g (k \ y)) ` X) a \ (\ y. (\ g. g (k \ y)) ` X) b = {}); RETURN (\ y \ items_5 A k v. (\ g. g (k \ y)) ` X) }) {Map.empty}" lemma expand_map_get_6_refine: "(expand_map_get_6, expand_5 \\ get_5) \ Id \ Id \ \Id\ nres_rel" unfolding expand_map_get_6_def expand_5_def get_5_def by (auto intro: FOREACH_rule_map_map[param_fo]) definition complement_succ_6 :: "('label, 'state) nba \ 'label \ 'state items \ 'state items set nres" where "complement_succ_6 A a f \ do { f \ refresh_5 f; f \ bounds_5 A a f; ASSUME (finite (dom f)); expand_map_get_6 A f }" lemma complement_succ_6_refine: "(complement_succ_6, complement_succ_5) \ Id \ Id \ Id \ \Id\ nres_rel" unfolding complement_succ_6_def complement_succ_5_def by (refine_vcg vcg2[OF expand_map_get_6_refine]) (auto intro: refine_IdI) subsection \Phase 7\ interpretation autoref_syn by this context fixes fi f assumes fi[autoref_rules]: "(fi, f) \ state_rel" begin private lemma [simp]: "finite (dom f)" using list_map_rel_finite fi unfolding finite_map_rel_def by force schematic_goal refresh_7: "(?f :: ?'a, refresh_5 f) \ ?R" unfolding refresh_5_def by (autoref_monadic (plain)) end concrete_definition refresh_7 uses refresh_7 lemma refresh_7_refine: "(\ f. RETURN (refresh_7 f), refresh_5) \ state_rel \ \state_rel\ nres_rel" using refresh_7.refine by fast context fixes A :: "('label, nat) nba" fixes succi a fi f assumes succi[autoref_rules]: "(succi, transition A a) \ nat_rel \ \nat_rel\ list_set_rel" assumes fi[autoref_rules]: "(fi, f) \ state_rel" begin private lemma [simp]: "finite (transition A a p)" using list_set_rel_finite succi[param_fo] unfolding finite_set_rel_def by blast private lemma [simp]: "finite (dom f)" using fi by force private lemma [autoref_op_pat]: "transition A a \ OP (transition A a)" by simp private lemma [autoref_rules]: "(min, min) \ nat_rel \ nat_rel \ nat_rel" by simp schematic_goal bounds_7: notes ty_REL[where R = "\nat_rel, item_rel\ dflt_ahm_rel", autoref_tyrel] shows "(?f :: ?'a, bounds_5 A a f) \ ?R" unfolding bounds_5_def merge_5_def sup_bool_def inf_nat_def by (autoref_monadic (plain)) end concrete_definition bounds_7 uses bounds_7 lemma bounds_7_refine: "(si, transition A a) \ nat_rel \ \nat_rel\ list_set_rel \ (\ p. RETURN (bounds_7 si p), bounds_5 A a) \ state_rel \ \\nat_rel, item_rel\ dflt_ahm_rel\ nres_rel" using bounds_7.refine by auto context fixes A :: "('label, nat) nba" fixes acci assumes [autoref_rules]: "(acci, accepting A) \ nat_rel \ bool_rel" begin private lemma [autoref_op_pat]: "accepting A \ OP (accepting A)" by simp private lemma [autoref_rules]: "((dvd), (dvd)) \ nat_rel \ nat_rel \ bool_rel" by simp private lemma [autoref_rules]: "(\ k l. upt k (Suc l), atLeastAtMost) \ nat_rel \ nat_rel \ \nat_rel\ list_set_rel" by (auto simp: list_set_rel_def in_br_conv) schematic_goal items_7: "(?f :: ?'a, items_5 A) \ ?R" unfolding items_5_def Let_def Set.filter_def by autoref end concrete_definition items_7 uses items_7 (* TODO: use generic expand_map implementation *) context fixes A :: "('label, nat) nba" fixes ai fixes fi f assumes ai: "(ai, accepting A) \ nat_rel \ bool_rel" assumes fi[autoref_rules]: "(fi, f) \ \nat_rel, item_rel\ dflt_ahm_rel" begin private lemma [simp]: "finite (dom f)" using dflt_ahm_rel_finite_nat fi unfolding finite_map_rel_def by force private lemma [simp]: assumes "\ m. m \ S \ x \ dom m" shows "inj_on (\ m. m (x \ y)) S" using assms unfolding dom_def inj_on_def by (auto) (metis fun_upd_triv fun_upd_upd) private lemmas [simp] = op_map_update_def[abs_def] private lemma [autoref_op_pat]: "items_5 A \ OP (items_5 A)" by simp private lemmas [autoref_rules] = items_7.refine[OF ai] schematic_goal expand_map_get_7: "(?f, expand_map_get_6 A f) \ \\state_rel\ list_set_rel\ nres_rel" unfolding expand_map_get_6_def by (autoref_monadic (plain)) end concrete_definition expand_map_get_7 uses expand_map_get_7 lemma expand_map_get_7_refine: assumes "(ai, accepting A) \ nat_rel \ bool_rel" shows "(\ fi. RETURN (expand_map_get_7 ai fi), \ f. ASSUME (finite (dom f)) \ expand_map_get_6 A f) \ \nat_rel, item_rel\ dflt_ahm_rel \ \\state_rel\ list_set_rel\ nres_rel" using expand_map_get_7.refine[OF assms] by auto context fixes A :: "('label, nat) nba" fixes a :: "'label" fixes p :: "nat items" fixes Ai fixes ai fixes pi assumes Ai: "(Ai, A) \ \Id, Id\ nbai_nba_rel" assumes ai: "(ai, a) \ Id" assumes pi[autoref_rules]: "(pi, p) \ state_rel" begin private lemmas succi = nbai_nba_param(4)[THEN fun_relD, OF Ai, THEN fun_relD, OF ai] private lemmas acceptingi = nbai_nba_param(5)[THEN fun_relD, OF Ai] private lemma [autoref_op_pat]: "(\ g. ASSUME (finite (dom g)) \ expand_map_get_6 A g) \ OP (\ g. ASSUME (finite (dom g)) \ expand_map_get_6 A g)" by simp private lemma [autoref_op_pat]: "bounds_5 A a \ OP (bounds_5 A a)" by simp private lemmas [autoref_rules] = refresh_7_refine bounds_7_refine[OF succi] expand_map_get_7_refine[OF acceptingi] schematic_goal complement_succ_7: "(?f :: ?'a, complement_succ_6 A a p) \ ?R" unfolding complement_succ_6_def by (autoref_monadic (plain)) end concrete_definition complement_succ_7 uses complement_succ_7 lemma complement_succ_7_refine: "(RETURN \\\ complement_succ_7, complement_succ_6) \ \Id, Id\ nbai_nba_rel \ Id \ state_rel \ \\state_rel\ list_set_rel\ nres_rel" using complement_succ_7.refine unfolding comp_apply by parametricity context fixes A :: "('label, nat) nba" fixes Ai fixes n ni :: nat assumes Ai: "(Ai, A) \ \Id, Id\ nbai_nba_rel" assumes ni[autoref_rules]: "(ni, n) \ Id" begin private lemma [autoref_op_pat]: "initial A \ OP (initial A)" by simp private lemmas [autoref_rules] = nbai_nba_param(3)[THEN fun_relD, OF Ai] schematic_goal complement_initial_7: "(?f, {(Some \ (const (2 * n, False))) |` initial A}) \ \state_rel\ list_set_rel" by autoref end concrete_definition complement_initial_7 uses complement_initial_7 schematic_goal complement_accepting_7: "(?f, \ f. \ (p, k, c) \ map_to_set f. \ c) \ state_rel \ bool_rel" by autoref concrete_definition complement_accepting_7 uses complement_accepting_7 definition complement_7 :: "('label, nat) nbai \ nat \ ('label, state) nbai" where "complement_7 Ai ni \ nbai (alphabeti Ai) (complement_initial_7 Ai ni) (complement_succ_7 Ai) (complement_accepting_7)" lemma complement_7_refine[autoref_rules]: assumes "(Ai, A) \ \Id, Id\ nbai_nba_rel" assumes "(ni, (OP card ::: \Id\ ahs_rel bhc \ nat_rel) $ ((OP nodes ::: \Id, Id\ nbai_nba_rel \ \Id\ ahs_rel bhc) $ A)) \ nat_rel" shows "(complement_7 Ai ni, (OP complement_4 ::: \Id, Id\ nbai_nba_rel \ \Id, state_rel\ nbai_nba_rel) $ A) \ \Id, state_rel\ nbai_nba_rel" proof - note complement_succ_7_refine also note complement_succ_6_refine also note complement_succ_5_refine finally have 1: "(complement_succ_7, complement_succ_4) \ \Id, Id\ nbai_nba_rel \ Id \ state_rel \ \state_rel\ list_set_rel" unfolding nres_rel_comp unfolding nres_rel_def unfolding fun_rel_def by auto show ?thesis unfolding complement_7_def complement_4_def using 1 complement_initial_7.refine complement_accepting_7.refine assms unfolding autoref_tag_defs by parametricity qed end \ No newline at end of file diff --git a/thys/Concurrent_Ref_Alg/Refinement_Lattice.thy b/thys/Concurrent_Ref_Alg/Refinement_Lattice.thy --- a/thys/Concurrent_Ref_Alg/Refinement_Lattice.thy +++ b/thys/Concurrent_Ref_Alg/Refinement_Lattice.thy @@ -1,115 +1,116 @@ section \Refinement Lattice \label{S:lattice}\ theory Refinement_Lattice imports Main - "HOL-Library.Lattice_Syntax" begin +unbundle lattice_syntax + text \ The underlying lattice of commands is complete and distributive. We follow the refinement calculus tradition so that $\nondet$ is non-deterministic choice and $c \refsto d$ means $c$ is refined (or implemented) by $d$. \ declare [[show_sorts]] text \Remove existing notation for quotient as it interferes with the rely quotient\ no_notation Equiv_Relations.quotient (infixl "'/'/" 90) class refinement_lattice = complete_distrib_lattice begin text \The refinement lattice infimum corresponds to non-deterministic choice for commands.\ abbreviation refine :: "'a \ 'a \ bool" (infix "\" 50) where "c \ d \ less_eq c d" abbreviation refine_strict :: "'a \ 'a \ bool" (infix "\" 50) where "c \ d \ less c d" text \Non-deterministic choice is monotonic in both arguments\ lemma inf_mono_left: "a \ b \ a \ c \ b \ c" using inf_mono by auto lemma inf_mono_right: "c \ d \ a \ c \ a \ d" using inf_mono by auto text \Binary choice is a special case of choice over a set.\ lemma Inf2_inf: "\{ f x | x. x \ {c, d}} = f c \ f d" proof - have "{ f x | x. x \ {c, d}} = {f c, f d}" by blast then have "\{ f x | x. x \ {c, d}} = \{f c, f d}" by simp also have "... = f c \ f d" by simp finally show ?thesis . qed text \Helper lemma for choice over indexed set.\ lemma INF_Inf: "(\x\X. f x) = (\{f x |x. x \ X})" by (simp add: Setcompr_eq_image) lemma (in -) INF_absorb_args: "(\i j. (f::nat \ 'c::complete_lattice) (i + j)) = (\k. f k)" proof (rule order_class.order.antisym) show "(\k. f k) \ (\i j. f (i + j))" by (simp add: complete_lattice_class.INF_lower complete_lattice_class.le_INF_iff) next have "\k. \i j. f (i + j) \ f k" by (metis Nat.add_0_right order_refl) then have "\k. \i. (\j. f (i + j)) \ f k" by (meson UNIV_I complete_lattice_class.INF_lower2) then show "(\i j. f (i + j)) \ (\k. f k)" by (simp add: complete_lattice_class.INF_mono) qed lemma (in -) nested_Collect: "{f y |y. y \ {g x |x. x \ X}} = {f (g x) |x. x \ X}" by blast text \A transition lemma for INF distributivity properties, going from Inf to INF, qualified version followed by a straightforward one.\ lemma Inf_distrib_INF_qual: fixes f :: "'a \ 'a \ 'a" assumes qual: "P {d x |x. x \ X}" assumes f_Inf_distrib: "\c D. P D \ f c (\ D) = \ {f c d | d . d \ D }" shows "f c (\x\X. d x) = (\x\X. f c (d x))" proof - have "f c (\x\X. d x) = f c (\{d x |x. x \ X})" by (simp add: INF_Inf) also have "... = (\{f c dx |dx. dx \ {d x | x. x \ X}})" by (simp add: qual f_Inf_distrib) also have "... = (\{f c (d x) |x. x \ X})" by (simp only: nested_Collect) also have "... = (\x\X. f c (d x))" by (simp add: INF_Inf) finally show ?thesis . qed lemma Inf_distrib_INF: fixes f :: "'a \ 'a \ 'a" assumes f_Inf_distrib: "\c D. f c (\ D) = \ {f c d | d . d \ D }" shows "f c (\x\X. d x) = (\x\X. f c (d x))" by (simp add: Setcompr_eq_image f_Inf_distrib image_comp) end lemmas refine_trans = order.trans text \More transitivity rules to make calculational reasoning smoother\ declare ord_eq_le_trans[trans] declare ord_le_eq_trans[trans] declare dual_order.trans[trans] abbreviation dist_over_sup :: "('a::refinement_lattice \ 'a) \ bool" where "dist_over_sup F \ (\ X . F (\ X) = (\x\X. F (x)))" abbreviation dist_over_inf :: "('a::refinement_lattice \ 'a) \ bool" where "dist_over_inf F \ (\ X . F (\ X) = (\x\X. F (x)))" end diff --git a/thys/Dependent_SIFUM_Type_Systems/Preliminaries.thy b/thys/Dependent_SIFUM_Type_Systems/Preliminaries.thy --- a/thys/Dependent_SIFUM_Type_Systems/Preliminaries.thy +++ b/thys/Dependent_SIFUM_Type_Systems/Preliminaries.thy @@ -1,124 +1,124 @@ (* Title: Value-Dependent SIFUM-Type-Systems Authors: Toby Murray, Robert Sison, Edward Pierzchalski, Christine Rizkallah (Based on the SIFUM-Type-Systems AFP entry, whose authors are: Sylvia Grewe, Heiko Mantel, Daniel Schoepe) *) section \Preliminaries\ theory Preliminaries -imports Main (* "HOL-Library.Lattice_Syntax" *) +imports Main begin text \Possible modes for variables:\ datatype Mode = AsmNoReadOrWrite | AsmNoWrite | GuarNoReadOrWrite | GuarNoWrite text \We consider a two-element security lattice:\ datatype Sec = High | Low text \@{term Sec} forms a (complete) lattice:\ instantiation Sec :: complete_lattice begin definition top_Sec_def: "top = High" definition sup_Sec_def: "sup d1 d2 = (if (d1 = High \ d2 = High) then High else Low)" definition inf_Sec_def: "inf d1 d2 = (if (d1 = Low \ d2 = Low) then Low else High)" definition bot_Sec_def: "bot = Low" definition less_eq_Sec_def: "d1 \ d2 = (d1 = d2 \ d1 = Low)" definition less_Sec_def: "d1 < d2 = (d1 = Low \ d2 = High)" definition Sup_Sec_def: "Sup S = (if (High \ S) then High else Low)" definition Inf_Sec_def: "Inf S = (if (Low \ S) then Low else High)" instance apply (intro_classes) using Sec.exhaust less_Sec_def less_eq_Sec_def inf_Sec_def sup_Sec_def apply auto[10] apply (metis Inf_Sec_def Sec.exhaust less_eq_Sec_def) apply (metis Inf_Sec_def Sec.exhaust less_eq_Sec_def) using Sec.exhaust less_Sec_def less_eq_Sec_def inf_Sec_def sup_Sec_def Inf_Sec_def Sup_Sec_def top_Sec_def bot_Sec_def by auto end text \Memories are mappings from variables to values\ type_synonym ('var, 'val) Mem = "'var \ 'val" text \A mode state maps modes to the set of variables for which the given mode is set.\ type_synonym 'var Mds = "Mode \ 'var set" text \Local configurations:\ type_synonym ('com, 'var, 'val) LocalConf = "('com \ 'var Mds) \ ('var, 'val) Mem" text \Global configurations:\ type_synonym ('com, 'var, 'val) GlobalConf = "('com \ 'var Mds) list \ ('var, 'val) Mem" text \A locale to fix various parametric components in Mantel et. al, and assumptions about them:\ locale sifum_security_init = fixes dma :: "('Var,'Val) Mem \ 'Var \ Sec" fixes \_vars :: "'Var \ 'Var set" fixes \ :: "'Var set" (* classification control variables *) fixes eval :: "('Com, 'Var, 'Val) LocalConf rel" fixes some_val :: "'Val" fixes INIT :: "('Var,'Val) Mem \ bool" assumes deterministic: "\ (lc, lc') \ eval; (lc, lc'') \ eval \ \ lc' = lc''" assumes finite_memory: "finite {(x::'Var). True}" defines "\ \ \x. \_vars x" assumes \_vars_\: "x \ \ \ \_vars x = {}" assumes dma_\_vars: "\x\\_vars y. mem\<^sub>1 x = mem\<^sub>2 x \ dma mem\<^sub>1 y = dma mem\<^sub>2 y" assumes \_Low: "\x\\. dma mem x = Low" locale sifum_security = sifum_security_init dma \_vars \ eval some_val "\_.True" for dma :: "('Var,'Val) Mem \ 'Var \ Sec" and \_vars :: "'Var \ 'Var set" and \ :: "'Var set" (* classification control variables *) and eval :: "('Com, 'Var, 'Val) LocalConf rel" and some_val :: "'Val" context sifum_security_init begin lemma \_vars_subset_\: "\_vars x \ \" by(force simp: \_def) lemma dma_\: "\x\\. mem\<^sub>1 x = mem\<^sub>2 x \ dma mem\<^sub>1 = dma mem\<^sub>2" proof fix y assume "\x\\. mem\<^sub>1 x = mem\<^sub>2 x" hence "\x\\_vars y. mem\<^sub>1 x = mem\<^sub>2 x" using \_vars_subset_\ by blast thus "dma mem\<^sub>1 y = dma mem\<^sub>2 y" by(rule dma_\_vars) qed end (* induction tools, moved up as far as possible *) lemma my_trancl_induct [consumes 1, case_names base step]: "\(a, b) \ r\<^sup>+; P a ; \x y. \(x, y) \ r; P x\ \ P y\ \ P b" by (induct rule: trancl.induct, blast+) lemma my_trancl_step_induct [consumes 1, case_names base step]: "\(a, b) \ r\<^sup>+; \x y. (x, y) \ r \ P x y; \x y z. P x y \ (y, z) \ r \ P x z\ \ P a b" by (induct rule: trancl_induct, blast+) lemma my_trancl_big_step_induct [consumes 1, case_names base step]: "\(a, b) \ r\<^sup>+; \x y. (x, y) \ r \ P x y; \x y z. (x, y) \ r\<^sup>+ \ P x y \ (y, z) \ r \ P y z \ P x z\ \ P a b" by (induct rule: trancl.induct, blast+) lemmas my_trancl_step_induct3 = my_trancl_step_induct[of "((ax,ay), az)" "((bx,by), bz)", split_format (complete), consumes 1, case_names step] lemmas my_trancl_big_step_induct3 = my_trancl_big_step_induct[of "((ax,ay), az)" "((bx,by), bz)", split_format (complete), consumes 1, case_names base step] end diff --git a/thys/Free-Boolean-Algebra/Free_Boolean_Algebra.thy b/thys/Free-Boolean-Algebra/Free_Boolean_Algebra.thy --- a/thys/Free-Boolean-Algebra/Free_Boolean_Algebra.thy +++ b/thys/Free-Boolean-Algebra/Free_Boolean_Algebra.thy @@ -1,753 +1,748 @@ (* Title: Free_Boolean_Algebra.thy Author: Brian Huffman, Portland State University *) section \Free Boolean algebras\ theory Free_Boolean_Algebra imports Main begin (*<*) notation bot ("\") and top ("\") and inf (infixl "\" 70) and sup (infixl "\" 65) lemma sup_conv_inf: fixes x y :: "'a::boolean_algebra" shows "x \ y = - (- x \ - y)" by simp (*>*) subsection \Free boolean algebra as a set\ text \ We start by defining the free boolean algebra over type @{typ 'a} as an inductive set. Here \i :: 'a\ represents a variable; \A :: 'a set\ represents a valuation, assigning a truth value to each variable; and \S :: 'a set set\ represents a formula, as the set of valuations that make the formula true. The set \fba\ contains representatives of formulas built from finite combinations of variables with negation and conjunction. \ inductive_set fba :: "'a set set set" where var: "{A. i \ A} \ fba" | Compl: "S \ fba \ - S \ fba" | inter: "S \ fba \ T \ fba \ S \ T \ fba" lemma fba_Diff: "S \ fba \ T \ fba \ S - T \ fba" unfolding Diff_eq by (intro fba.inter fba.Compl) lemma fba_union: "S \ fba \ T \ fba \ S \ T \ fba" proof - assume "S \ fba" and "T \ fba" hence "- (- S \ - T) \ fba" by (intro fba.intros) thus "S \ T \ fba" by simp qed lemma fba_empty: "({} :: 'a set set) \ fba" proof - obtain S :: "'a set set" where "S \ fba" by (fast intro: fba.var) hence "S \ - S \ fba" by (intro fba.intros) thus ?thesis by simp qed lemma fba_UNIV: "(UNIV :: 'a set set) \ fba" proof - have "- {} \ fba" using fba_empty by (rule fba.Compl) thus "UNIV \ fba" by simp qed subsection \Free boolean algebra as a type\ text \ The next step is to use \typedef\ to define a type isomorphic to the set @{const fba}. We also define a constructor \var\ that corresponds with the similarly-named introduction rule for @{const fba}. \ typedef 'a formula = "fba :: 'a set set set" by (auto intro: fba_empty) definition var :: "'a \ 'a formula" where "var i = Abs_formula {A. i \ A}" lemma Rep_formula_var: "Rep_formula (var i) = {A. i \ A}" unfolding var_def using fba.var by (rule Abs_formula_inverse) text \ \medskip Now we make type @{typ "'a formula"} into a Boolean algebra. This involves defining the various operations (ordering relations, binary infimum and supremum, complement, difference, top and bottom elements) and proving that they satisfy the appropriate laws. \ instantiation formula :: (type) boolean_algebra begin definition "x \ y = Abs_formula (Rep_formula x \ Rep_formula y)" definition "x \ y = Abs_formula (Rep_formula x \ Rep_formula y)" definition "\ = Abs_formula UNIV" definition "\ = Abs_formula {}" definition "x \ y \ Rep_formula x \ Rep_formula y" definition "x < y \ Rep_formula x \ Rep_formula y" definition "- x = Abs_formula (- Rep_formula x)" definition "x - y = Abs_formula (Rep_formula x - Rep_formula y)" lemma Rep_formula_inf: "Rep_formula (x \ y) = Rep_formula x \ Rep_formula y" unfolding inf_formula_def by (intro Abs_formula_inverse fba.inter Rep_formula) lemma Rep_formula_sup: "Rep_formula (x \ y) = Rep_formula x \ Rep_formula y" unfolding sup_formula_def by (intro Abs_formula_inverse fba_union Rep_formula) lemma Rep_formula_top: "Rep_formula \ = UNIV" unfolding top_formula_def by (intro Abs_formula_inverse fba_UNIV) lemma Rep_formula_bot: "Rep_formula \ = {}" unfolding bot_formula_def by (intro Abs_formula_inverse fba_empty) lemma Rep_formula_compl: "Rep_formula (- x) = - Rep_formula x" unfolding uminus_formula_def by (intro Abs_formula_inverse fba.Compl Rep_formula) lemma Rep_formula_diff: "Rep_formula (x - y) = Rep_formula x - Rep_formula y" unfolding minus_formula_def by (intro Abs_formula_inverse fba_Diff Rep_formula) lemmas eq_formula_iff = Rep_formula_inject [symmetric] lemmas Rep_formula_simps = less_eq_formula_def less_formula_def eq_formula_iff Rep_formula_sup Rep_formula_inf Rep_formula_top Rep_formula_bot Rep_formula_compl Rep_formula_diff Rep_formula_var instance proof qed (unfold Rep_formula_simps, auto) end text \ \medskip The laws of a Boolean algebra do not require the top and bottom elements to be distinct, so the following rules must be proved separately: \ lemma bot_neq_top_formula [simp]: "(\ :: 'a formula) \ \" unfolding Rep_formula_simps by auto lemma top_neq_bot_formula [simp]: "(\ :: 'a formula) \ \" unfolding Rep_formula_simps by auto text \ \medskip Here we prove an essential property of a free Boolean algebra: all generators are independent. \ lemma var_le_var_simps [simp]: "var i \ var j \ i = j" "\ var i \ - var j" "\ - var i \ var j" unfolding Rep_formula_simps by fast+ lemma var_eq_var_simps [simp]: "var i = var j \ i = j" "var i \ - var j" "- var i \ var j" unfolding Rep_formula_simps set_eq_subset by fast+ text \ \medskip We conclude this section by proving an induction principle for formulas. It mirrors the definition of the inductive set \fba\, with cases for variables, complements, and conjunction. \ lemma formula_induct [case_names var compl inf, induct type: formula]: fixes P :: "'a formula \ bool" assumes 1: "\i. P (var i)" assumes 2: "\x. P x \ P (- x)" assumes 3: "\x y. P x \ P y \ P (x \ y)" shows "P x" proof (induct x rule: Abs_formula_induct) fix y :: "'a set set" assume "y \ fba" thus "P (Abs_formula y)" proof (induct rule: fba.induct) case (var i) have "P (var i)" by (rule 1) thus ?case unfolding var_def . next case (Compl S) from \P (Abs_formula S)\ have "P (- Abs_formula S)" by (rule 2) with \S \ fba\ show ?case unfolding uminus_formula_def by (simp add: Abs_formula_inverse) next case (inter S T) from \P (Abs_formula S)\ and \P (Abs_formula T)\ have "P (Abs_formula S \ Abs_formula T)" by (rule 3) with \S \ fba\ and \T \ fba\ show ?case unfolding inf_formula_def by (simp add: Abs_formula_inverse) qed qed subsection \If-then-else for Boolean algebras\ text \ This is a generic if-then-else operator for arbitrary Boolean algebras. \ definition ifte :: "'a::boolean_algebra \ 'a \ 'a \ 'a" where "ifte a x y = (a \ x) \ (- a \ y)" lemma ifte_top [simp]: "ifte \ x y = x" unfolding ifte_def by simp lemma ifte_bot [simp]: "ifte \ x y = y" unfolding ifte_def by simp lemma ifte_same: "ifte a x x = x" unfolding ifte_def by (simp add: inf_sup_distrib2 [symmetric] sup_compl_top) lemma compl_ifte: "- ifte a x y = ifte a (- x) (- y)" unfolding ifte_def apply (rule order_antisym) apply (simp add: inf_sup_distrib1 inf_sup_distrib2 compl_inf_bot) apply (simp add: sup_inf_distrib1 sup_inf_distrib2 sup_compl_top) apply (simp add: le_infI1 le_infI2 le_supI1 le_supI2) apply (simp add: le_infI1 le_infI2 le_supI1 le_supI2) done lemma inf_ifte_distrib: "ifte x a b \ ifte x c d = ifte x (a \ c) (b \ d)" unfolding ifte_def apply (simp add: inf_sup_distrib1 inf_sup_distrib2) apply (simp add: inf_sup_aci inf_compl_bot) done lemma ifte_ifte_distrib: "ifte x (ifte y a b) (ifte y c d) = ifte y (ifte x a c) (ifte x b d)" unfolding ifte_def [of x] sup_conv_inf by (simp only: compl_ifte [symmetric] inf_ifte_distrib [symmetric] ifte_same) subsection \Formulas over a set of generators\ text \ The set \formulas S\ consists of those formulas that only depend on variables in the set \S\. It is analogous to the @{const lists} operator for the list datatype. \ definition formulas :: "'a set \ 'a formula set" where "formulas S = {x. \A B. (\i\S. i \ A \ i \ B) \ A \ Rep_formula x \ B \ Rep_formula x}" lemma formulasI: assumes "\A B. \i\S. i \ A \ i \ B \ A \ Rep_formula x \ B \ Rep_formula x" shows "x \ formulas S" using assms unfolding formulas_def by simp lemma formulasD: assumes "x \ formulas S" assumes "\i\S. i \ A \ i \ B" shows "A \ Rep_formula x \ B \ Rep_formula x" using assms unfolding formulas_def by simp lemma formulas_mono: "S \ T \ formulas S \ formulas T" by (fast intro!: formulasI elim!: formulasD) lemma formulas_insert: "x \ formulas S \ x \ formulas (insert a S)" unfolding formulas_def by simp lemma formulas_var: "i \ S \ var i \ formulas S" unfolding formulas_def by (simp add: Rep_formula_simps) lemma formulas_var_iff: "var i \ formulas S \ i \ S" unfolding formulas_def by (simp add: Rep_formula_simps, fast) lemma formulas_bot: "\ \ formulas S" unfolding formulas_def by (simp add: Rep_formula_simps) lemma formulas_top: "\ \ formulas S" unfolding formulas_def by (simp add: Rep_formula_simps) lemma formulas_compl: "x \ formulas S \ - x \ formulas S" unfolding formulas_def by (simp add: Rep_formula_simps) lemma formulas_inf: "x \ formulas S \ y \ formulas S \ x \ y \ formulas S" unfolding formulas_def by (auto simp add: Rep_formula_simps) lemma formulas_sup: "x \ formulas S \ y \ formulas S \ x \ y \ formulas S" unfolding formulas_def by (auto simp add: Rep_formula_simps) lemma formulas_diff: "x \ formulas S \ y \ formulas S \ x - y \ formulas S" unfolding formulas_def by (auto simp add: Rep_formula_simps) lemma formulas_ifte: "a \ formulas S \ x \ formulas S \ y \ formulas S \ ifte a x y \ formulas S" unfolding ifte_def by (intro formulas_sup formulas_inf formulas_compl) lemmas formulas_intros = formulas_var formulas_bot formulas_top formulas_compl formulas_inf formulas_sup formulas_diff formulas_ifte subsection \Injectivity of if-then-else\ text \ The if-then-else operator is injective in some limited circumstances: when the scrutinee is a variable that is not mentioned in either branch. \ lemma ifte_inject: assumes "ifte (var i) x y = ifte (var i) x' y'" assumes "i \ S" assumes "x \ formulas S" and "x' \ formulas S" assumes "y \ formulas S" and "y' \ formulas S" shows "x = x' \ y = y'" proof have 1: "\A. i \ A \ A \ Rep_formula x \ A \ Rep_formula x'" using assms(1) by (simp add: Rep_formula_simps ifte_def set_eq_iff, fast) have 2: "\A. i \ A \ A \ Rep_formula y \ A \ Rep_formula y'" using assms(1) by (simp add: Rep_formula_simps ifte_def set_eq_iff, fast) show "x = x'" unfolding Rep_formula_simps proof (rule set_eqI) fix A have "A \ Rep_formula x \ insert i A \ Rep_formula x" using \x \ formulas S\ by (rule formulasD, force simp add: \i \ S\) also have "\ \ insert i A \ Rep_formula x'" by (rule 1, simp) also have "\ \ A \ Rep_formula x'" using \x' \ formulas S\ by (rule formulasD, force simp add: \i \ S\) finally show "A \ Rep_formula x \ A \ Rep_formula x'" . qed show "y = y'" unfolding Rep_formula_simps proof (rule set_eqI) fix A have "A \ Rep_formula y \ A - {i} \ Rep_formula y" using \y \ formulas S\ by (rule formulasD, force simp add: \i \ S\) also have "\ \ A - {i} \ Rep_formula y'" by (rule 2, simp) also have "\ \ A \ Rep_formula y'" using \y' \ formulas S\ by (rule formulasD, force simp add: \i \ S\) finally show "A \ Rep_formula y \ A \ Rep_formula y'" . qed qed subsection \Specification of homomorphism operator\ text \ Our goal is to define a homomorphism operator \hom\ such that for any function \f\, \hom f\ is the unique Boolean algebra homomorphism satisfying \hom f (var i) = f i\ for all \i\. Instead of defining \hom\ directly, we will follow the approach used to define Isabelle's \fold\ operator for finite sets. First we define the graph of the \hom\ function as a relation; later we will define the \hom\ function itself using definite choice. The \hom_graph\ relation is defined inductively, with introduction rules based on the if-then-else normal form of Boolean formulas. The relation is also indexed by an extra set parameter \S\, to ensure that branches of each if-then-else do not use the same variable again. \ inductive hom_graph :: "('a \ 'b::boolean_algebra) \ 'a set \ 'a formula \ 'b \ bool" for f :: "'a \ 'b::boolean_algebra" where bot: "hom_graph f {} bot bot" | top: "hom_graph f {} top top" | ifte: "i \ S \ hom_graph f S x a \ hom_graph f S y b \ hom_graph f (insert i S) (ifte (var i) x y) (ifte (f i) a b)" text \ \medskip The next two lemmas establish a stronger elimination rule for assumptions of the form @{term "hom_graph f (insert i S) x a"}. Essentially, they say that we can arrange the top-level if-then-else to use the variable of our choice. The proof makes use of the distributive properties of if-then-else. \ lemma hom_graph_dest: "hom_graph f S x a \ k \ S \ \y z b c. x = ifte (var k) y z \ a = ifte (f k) b c \ hom_graph f (S - {k}) y b \ hom_graph f (S - {k}) z c" proof (induct set: hom_graph) case (ifte i S x a y b) show ?case proof (cases "i = k") assume "i = k" with ifte(1,2,4) show ?case by auto next assume "i \ k" with \k \ insert i S\ have k: "k \ S" by simp have *: "insert i S - {k} = insert i (S - {k})" using \i \ k\ by (simp add: insert_Diff_if) have **: "i \ S - {k}" using \i \ S\ by simp from ifte(1) ifte(3) [OF k] ifte(5) [OF k] show ?case unfolding * apply clarify apply (simp only: ifte_ifte_distrib [of "var i"]) apply (simp only: ifte_ifte_distrib [of "f i"]) apply (fast intro: hom_graph.ifte [OF **]) done qed qed simp_all lemma hom_graph_insert_elim: assumes "hom_graph f (insert i S) x a" and "i \ S" obtains y z b c where "x = ifte (var i) y z" and "a = ifte (f i) b c" and "hom_graph f S y b" and "hom_graph f S z c" using hom_graph_dest [OF assms(1) insertI1] by (clarify, simp add: assms(2)) text \ \medskip Now we prove the first uniqueness property of the @{const hom_graph} relation. This version of uniqueness says that for any particular value of \S\, the relation @{term "hom_graph f S"} maps each \x\ to at most one \a\. The proof uses the injectiveness of if-then-else, which we proved earlier. \ lemma hom_graph_imp_formulas: "hom_graph f S x a \ x \ formulas S" by (induct set: hom_graph, simp_all add: formulas_intros formulas_insert) lemma hom_graph_unique: "hom_graph f S x a \ hom_graph f S x a' \ a = a'" proof (induct arbitrary: a' set: hom_graph) case (ifte i S y b z c a') from ifte(6,1) obtain y' z' b' c' where 1: "ifte (var i) y z = ifte (var i) y' z'" and 2: "a' = ifte (f i) b' c'" and 3: "hom_graph f S y' b'" and 4: "hom_graph f S z' c'" by (rule hom_graph_insert_elim) from 1 3 4 ifte(1,2,4) have "y = y' \ z = z'" by (intro ifte_inject hom_graph_imp_formulas) with 2 3 4 ifte(3,5) show "ifte (f i) b c = a'" by simp qed (erule hom_graph.cases, simp_all)+ text \ \medskip The next few lemmas will help to establish a stronger version of the uniqueness property of @{const hom_graph}. They show that the @{const hom_graph} relation is preserved if we replace \S\ with a larger finite set. \ lemma hom_graph_insert: assumes "hom_graph f S x a" shows "hom_graph f (insert i S) x a" proof (cases "i \ S") assume "i \ S" with assms show ?thesis by (simp add: insert_absorb) next assume "i \ S" hence "hom_graph f (insert i S) (ifte (var i) x x) (ifte (f i) a a)" by (intro hom_graph.ifte assms) thus "hom_graph f (insert i S) x a" by (simp only: ifte_same) qed lemma hom_graph_finite_superset: assumes "hom_graph f S x a" and "finite T" and "S \ T" shows "hom_graph f T x a" proof - from \finite T\ have "hom_graph f (S \ T) x a" by (induct set: finite, simp add: assms, simp add: hom_graph_insert) with \S \ T\ show "hom_graph f T x a" by (simp only: subset_Un_eq) qed lemma hom_graph_imp_finite: "hom_graph f S x a \ finite S" by (induct set: hom_graph) simp_all text \ \medskip This stronger uniqueness property says that @{term "hom_graph f"} maps each \x\ to at most one \a\, even for \emph{different} values of the set parameter. \ lemma hom_graph_unique': assumes "hom_graph f S x a" and "hom_graph f T x a'" shows "a = a'" proof (rule hom_graph_unique) have fin: "finite (S \ T)" using assms by (intro finite_UnI hom_graph_imp_finite) show "hom_graph f (S \ T) x a" using assms(1) fin Un_upper1 by (rule hom_graph_finite_superset) show "hom_graph f (S \ T) x a'" using assms(2) fin Un_upper2 by (rule hom_graph_finite_superset) qed text \ \medskip Finally, these last few lemmas establish that the @{term "hom_graph f"} relation is total: every \x\ is mapped to some \a\. \ lemma hom_graph_var: "hom_graph f {i} (var i) (f i)" proof - have "hom_graph f {i} (ifte (var i) top bot) (ifte (f i) top bot)" by (simp add: hom_graph.intros) thus "hom_graph f {i} (var i) (f i)" unfolding ifte_def by simp qed lemma hom_graph_compl: "hom_graph f S x a \ hom_graph f S (- x) (- a)" by (induct set: hom_graph, simp_all add: hom_graph.intros compl_ifte) lemma hom_graph_inf: "hom_graph f S x a \ hom_graph f S y b \ hom_graph f S (x \ y) (a \ b)" apply (induct arbitrary: y b set: hom_graph) apply (simp add: hom_graph.bot) apply simp apply (erule (1) hom_graph_insert_elim) apply (auto simp add: inf_ifte_distrib hom_graph.ifte) done lemma hom_graph_union_inf: assumes "hom_graph f S x a" and "hom_graph f T y b" shows "hom_graph f (S \ T) (x \ y) (a \ b)" proof (rule hom_graph_inf) have fin: "finite (S \ T)" using assms by (intro finite_UnI hom_graph_imp_finite) show "hom_graph f (S \ T) x a" using assms(1) fin Un_upper1 by (rule hom_graph_finite_superset) show "hom_graph f (S \ T) y b" using assms(2) fin Un_upper2 by (rule hom_graph_finite_superset) qed lemma hom_graph_exists: "\a S. hom_graph f S x a" by (induct x) (auto intro: hom_graph_var hom_graph_compl hom_graph_union_inf) subsection \Homomorphisms into other boolean algebras\ text \ Now that we have proved the necessary existence and uniqueness properties of @{const hom_graph}, we can define the function \hom\ using definite choice. \ definition hom :: "('a \ 'b::boolean_algebra) \ 'a formula \ 'b" where "hom f x = (THE a. \S. hom_graph f S x a)" lemma hom_graph_hom: "\S. hom_graph f S x (hom f x)" unfolding hom_def apply (rule theI') apply (rule ex_ex1I) apply (rule hom_graph_exists) apply (fast elim: hom_graph_unique') done lemma hom_equality: "hom_graph f S x a \ hom f x = a" unfolding hom_def apply (rule the_equality) apply (erule exI) apply (erule exE) apply (erule (1) hom_graph_unique') done text \ \medskip The @{const hom} function correctly implements its specification: \ lemma hom_var [simp]: "hom f (var i) = f i" by (rule hom_equality, rule hom_graph_var) lemma hom_bot [simp]: "hom f \ = \" by (rule hom_equality, rule hom_graph.bot) lemma hom_top [simp]: "hom f \ = \" by (rule hom_equality, rule hom_graph.top) lemma hom_compl [simp]: "hom f (- x) = - hom f x" proof - obtain S where "hom_graph f S x (hom f x)" using hom_graph_hom .. hence "hom_graph f S (- x) (- hom f x)" by (rule hom_graph_compl) thus "hom f (- x) = - hom f x" by (rule hom_equality) qed lemma hom_inf [simp]: "hom f (x \ y) = hom f x \ hom f y" proof - obtain S where S: "hom_graph f S x (hom f x)" using hom_graph_hom .. obtain T where T: "hom_graph f T y (hom f y)" using hom_graph_hom .. have "hom_graph f (S \ T) (x \ y) (hom f x \ hom f y)" using S T by (rule hom_graph_union_inf) thus ?thesis by (rule hom_equality) qed lemma hom_sup [simp]: "hom f (x \ y) = hom f x \ hom f y" unfolding sup_conv_inf by (simp only: hom_compl hom_inf) lemma hom_diff [simp]: "hom f (x - y) = hom f x - hom f y" unfolding diff_eq by (simp only: hom_compl hom_inf) lemma hom_ifte [simp]: "hom f (ifte x y z) = ifte (hom f x) (hom f y) (hom f z)" unfolding ifte_def by (simp only: hom_compl hom_inf hom_sup) lemmas hom_simps = hom_var hom_bot hom_top hom_compl hom_inf hom_sup hom_diff hom_ifte text \ \medskip The type @{typ "'a formula"} can be viewed as a monad, with @{const var} as the unit, and @{const hom} as the bind operator. We can prove the standard monad laws with simple proofs by induction. \ lemma hom_var_eq_id: "hom var x = x" by (induct x) simp_all lemma hom_hom: "hom f (hom g x) = hom (\i. hom f (g i)) x" by (induct x) simp_all subsection \Map operation on Boolean formulas\ text \ We can define a map functional in terms of @{const hom} and @{const var}. The properties of \fmap\ follow directly from the lemmas we have already proved about @{const hom}. \ definition fmap :: "('a \ 'b) \ 'a formula \ 'b formula" where "fmap f = hom (\i. var (f i))" lemma fmap_var [simp]: "fmap f (var i) = var (f i)" unfolding fmap_def by simp lemma fmap_bot [simp]: "fmap f \ = \" unfolding fmap_def by simp lemma fmap_top [simp]: "fmap f \ = \" unfolding fmap_def by simp lemma fmap_compl [simp]: "fmap f (- x) = - fmap f x" unfolding fmap_def by simp lemma fmap_inf [simp]: "fmap f (x \ y) = fmap f x \ fmap f y" unfolding fmap_def by simp lemma fmap_sup [simp]: "fmap f (x \ y) = fmap f x \ fmap f y" unfolding fmap_def by simp lemma fmap_diff [simp]: "fmap f (x - y) = fmap f x - fmap f y" unfolding fmap_def by simp lemma fmap_ifte [simp]: "fmap f (ifte x y z) = ifte (fmap f x) (fmap f y) (fmap f z)" unfolding fmap_def by simp lemmas fmap_simps = fmap_var fmap_bot fmap_top fmap_compl fmap_inf fmap_sup fmap_diff fmap_ifte text \ \medskip The map functional satisfies the functor laws: it preserves identity and function composition. \ lemma fmap_ident: "fmap (\i. i) x = x" by (induct x) simp_all lemma fmap_fmap: "fmap f (fmap g x) = fmap (f \ g) x" by (induct x) simp_all subsection \Hiding lattice syntax\ text \ The following command hides the lattice syntax, to avoid potential conflicts with other theories that import this one. To re-enable - the syntax, users should import theory \Lattice_Syntax\ from - the Isabelle library. + the syntax, users should unbundle \lattice_syntax\. \ -no_notation - top ("\") and - bot ("\") and - inf (infixl "\" 70) and - sup (infixl "\" 65) +unbundle no_lattice_syntax end diff --git a/thys/MFODL_Monitor_Optimized/Regex.thy b/thys/MFODL_Monitor_Optimized/Regex.thy --- a/thys/MFODL_Monitor_Optimized/Regex.thy +++ b/thys/MFODL_Monitor_Optimized/Regex.thy @@ -1,372 +1,373 @@ (*<*) theory Regex imports "MFOTL_Monitor.Trace" - "HOL-Library.Lattice_Syntax" "HOL-Library.Extended_Nat" begin + +unbundle lattice_syntax (*>*) section \Regular expressions\ context begin qualified datatype (atms: 'a) regex = Skip nat | Test 'a | Plus "'a regex" "'a regex" | Times "'a regex" "'a regex" | Star "'a regex" lemma finite_atms[simp]: "finite (atms r)" by (induct r) auto definition "Wild = Skip 1" lemma size_regex_estimation[termination_simp]: "x \ atms r \ y < f x \ y < size_regex f r" by (induct r) auto lemma size_regex_estimation'[termination_simp]: "x \ atms r \ y \ f x \ y \ size_regex f r" by (induct r) auto qualified definition "TimesL r S = Times r ` S" qualified definition "TimesR R s = (\r. Times r s) ` R" qualified primrec fv_regex where "fv_regex fv (Skip n) = {}" | "fv_regex fv (Test \) = fv \" | "fv_regex fv (Plus r s) = fv_regex fv r \ fv_regex fv s" | "fv_regex fv (Times r s) = fv_regex fv r \ fv_regex fv s" | "fv_regex fv (Star r) = fv_regex fv r" lemma fv_regex_cong[fundef_cong]: "r = r' \ (\z. z \ atms r \ fv z = fv' z) \ fv_regex fv r = fv_regex fv' r'" by (induct r arbitrary: r') auto lemma finite_fv_regex[simp]: "(\z. z \ atms r \ finite (fv z)) \ finite (fv_regex fv r)" by (induct r) auto lemma fv_regex_commute: "(\z. z \ atms r \ x \ fv z \ g x \ fv' z) \ x \ fv_regex fv r \ g x \ fv_regex fv' r" by (induct r) auto lemma fv_regex_alt: "fv_regex fv r = (\z \ atms r. fv z)" by (induct r) auto qualified definition nfv_regex where "nfv_regex fv r = Max (insert 0 (Suc ` fv_regex fv r))" lemma insert_Un: "insert x (A \ B) = insert x A \ insert x B" by auto lemma nfv_regex_simps[simp]: assumes [simp]: "(\z. z \ atms r \ finite (fv z))" "(\z. z \ atms s \ finite (fv z))" shows "nfv_regex fv (Skip n) = 0" "nfv_regex fv (Test \) = Max (insert 0 (Suc ` fv \))" "nfv_regex fv (Plus r s) = max (nfv_regex fv r) (nfv_regex fv s)" "nfv_regex fv (Times r s) = max (nfv_regex fv r) (nfv_regex fv s)" "nfv_regex fv (Star r) = nfv_regex fv r" unfolding nfv_regex_def by (auto simp add: image_Un Max_Un insert_Un simp del: Un_insert_right Un_insert_left) abbreviation "min_regex_default f r j \ (if atms r = {} then j else Min ((\z. f z j) ` atms r))" qualified primrec match :: "(nat \ 'a \ bool) \ 'a regex \ nat \ nat \ bool" where "match test (Skip n) = (\i j. j = i + n)" | "match test (Test \) = (\i j. i = j \ test i \)" | "match test (Plus r s) = match test r \ match test s" | "match test (Times r s) = match test r OO match test s" | "match test (Star r) = (match test r)\<^sup>*\<^sup>*" lemma match_cong[fundef_cong]: "r = r' \ (\i z. z \ atms r \ t i z = t' i z) \ match t r = match t' r'" by (induct r arbitrary: r') auto qualified primrec eps where "eps test i (Skip n) = (n = 0)" | "eps test i (Test \) = test i \" | "eps test i (Plus r s) = (eps test i r \ eps test i s)" | "eps test i (Times r s) = (eps test i r \ eps test i s)" | "eps test i (Star r) = True" qualified primrec lpd where "lpd test i (Skip n) = (case n of 0 \ {} | Suc m \ {Skip m})" | "lpd test i (Test \) = {}" | "lpd test i (Plus r s) = (lpd test i r \ lpd test i s)" | "lpd test i (Times r s) = TimesR (lpd test i r) s \ (if eps test i r then lpd test i s else {})" | "lpd test i (Star r) = TimesR (lpd test i r) (Star r)" qualified primrec lpd\ where "lpd\ \ test i (Skip n) = (case n of 0 \ {} | Suc m \ {\ (Skip m)})" | "lpd\ \ test i (Test \) = {}" | "lpd\ \ test i (Plus r s) = lpd\ \ test i r \ lpd\ \ test i s" | "lpd\ \ test i (Times r s) = lpd\ (\t. \ (Times t s)) test i r \ (if eps test i r then lpd\ \ test i s else {})" | "lpd\ \ test i (Star r) = lpd\ (\t. \ (Times t (Star r))) test i r" qualified primrec rpd where "rpd test i (Skip n) = (case n of 0 \ {} | Suc m \ {Skip m})" | "rpd test i (Test \) = {}" | "rpd test i (Plus r s) = (rpd test i r \ rpd test i s)" | "rpd test i (Times r s) = TimesL r (rpd test i s) \ (if eps test i s then rpd test i r else {})" | "rpd test i (Star r) = TimesL (Star r) (rpd test i r)" qualified primrec rpd\ where "rpd\ \ test i (Skip n) = (case n of 0 \ {} | Suc m \ {\ (Skip m)})" | "rpd\ \ test i (Test \) = {}" | "rpd\ \ test i (Plus r s) = rpd\ \ test i r \ rpd\ \ test i s" | "rpd\ \ test i (Times r s) = rpd\ (\t. \ (Times r t)) test i s \ (if eps test i s then rpd\ \ test i r else {})" | "rpd\ \ test i (Star r) = rpd\ (\t. \ (Times (Star r) t)) test i r" lemma lpd\_lpd: "lpd\ \ test i r = \ ` lpd test i r" by (induct r arbitrary: \) (auto simp: TimesR_def split: nat.splits) lemma rpd\_rpd: "rpd\ \ test i r = \ ` rpd test i r" by (induct r arbitrary: \) (auto simp: TimesL_def split: nat.splits) lemma match_le: "match test r i j \ i \ j" proof (induction r arbitrary: i j) case (Times r s) then show ?case using order.trans by fastforce next case (Star r) from Star.prems show ?case unfolding match.simps by (induct i j rule: rtranclp.induct) (force dest: Star.IH)+ qed auto lemma match_rtranclp_le: "(match test r)\<^sup>*\<^sup>* i j \ i \ j" by (metis match.simps(5) match_le) lemma eps_match: "eps test i r \ match test r i i" by (induction r) (auto dest: antisym[OF match_le match_le]) lemma lpd_match: "i < j \ match test r i j \ (\s \ lpd test i r. match test s) (i + 1) j" proof (induction r arbitrary: i j) case (Times r s) have "match test (Times r s) i j \ (\k. match test r i k \ match test s k j)" by auto also have "\ \ match test r i i \ match test s i j \ (\k>i. match test r i k \ match test s k j)" using match_le[of test r i] nat_less_le by auto also have "\ \ match test r i i \ (\t \ lpd test i s. match test t) (i + 1) j \ (\k>i. (\t \ lpd test i r. match test t) (i + 1) k \ match test s k j)" using Times.IH(1) Times.IH(2)[OF Times.prems] by metis also have "\ \ match test r i i \ (\t \ lpd test i s. match test t) (i + 1) j \ (\k. (\t \ lpd test i r. match test t) (i + 1) k \ match test s k j)" using Times.prems by (intro disj_cong[OF refl] iff_exI) (auto dest: match_le) also have "\ \ (\ (match test ` lpd test i (Times r s))) (i + 1) j" by (force simp: TimesL_def TimesR_def eps_match) finally show ?case . next case (Star r) have "\s\lpd test i r. (match test s OO (match test r)\<^sup>*\<^sup>*) (i + 1) j" if "(match test r)\<^sup>*\<^sup>* i j" using that Star.prems match_le[of test _ "i + 1"] proof (induct rule: converse_rtranclp_induct) case (step i k) then show ?case by (cases "i < k") (auto simp: not_less Star.IH dest: match_le) qed simp with Star.prems show ?case using match_le[of test _ "i + 1"] by (auto simp: TimesL_def TimesR_def Suc_le_eq Star.IH[of i] elim!: converse_rtranclp_into_rtranclp[rotated]) qed (auto split: nat.splits) lemma rpd_match: "i < j \ match test r i j \ (\s \ rpd test j r. match test s) i (j - 1)" proof (induction r arbitrary: i j) case (Times r s) have "match test (Times r s) i j \ (\k. match test r i k \ match test s k j)" by auto also have "\ \ match test r i j \ match test s j j \ (\k match test s k j)" using match_le[of test s _ j] nat_less_le by auto also have "\ \ (\t \ rpd test j r. match test t) i (j - 1) \ match test s j j \ (\k (\t \ rpd test j s. match test t) k (j - 1))" using Times.IH(1)[OF Times.prems] Times.IH(2) by metis also have "\ \ (\t \ rpd test j r. match test t) i (j - 1) \ match test s j j \ (\k. match test r i k \ (\t \ rpd test j s. match test t) k (j - 1))" using Times.prems by (intro disj_cong[OF refl] iff_exI) (auto dest: match_le) also have "\ \ (\ (match test ` rpd test j (Times r s))) i (j - 1)" by (force simp: TimesL_def TimesR_def eps_match) finally show ?case . next case (Star r) have "\s\rpd test j r. ((match test r)\<^sup>*\<^sup>* OO match test s) i (j - 1)" if "(match test r)\<^sup>*\<^sup>* i j" using that Star.prems match_le[of test _ _ "j - 1"] proof (induct rule: rtranclp_induct) case (step k j) then show ?case by (cases "k < j") (auto simp: not_less Star.IH dest: match_le) qed simp with Star.prems show ?case by (auto 0 3 simp: TimesL_def TimesR_def intro: Star.IH[of _ j, THEN iffD2] elim!: rtranclp.rtrancl_into_rtrancl dest: match_le[of test _ _ "j - 1", unfolded One_nat_def]) qed (auto split: nat.splits) lemma lpd_fv_regex: "s \ lpd test i r \ fv_regex fv s \ fv_regex fv r" by (induct r arbitrary: s) (auto simp: TimesR_def TimesL_def split: if_splits nat.splits)+ lemma rpd_fv_regex: "s \ rpd test i r \ fv_regex fv s \ fv_regex fv r" by (induct r arbitrary: s) (auto simp: TimesR_def TimesL_def split: if_splits nat.splits)+ lemma match_fv_cong: "(\i x. x \ atms r \ test i x = test' i x) \ match test r = match test' r" by (induct r) auto lemma eps_fv_cong: "(\i x. x \ atms r \ test i x = test' i x) \ eps test i r = eps test' i r" by (induct r) auto datatype modality = Past | Futu datatype safety = Strict | Lax context fixes fv :: "'a \ 'b set" and safe :: "safety \ 'a \ bool" begin qualified fun safe_regex :: "modality \ safety \ 'a regex \ bool" where "safe_regex m _ (Skip n) = True" | "safe_regex m g (Test \) = safe g \" | "safe_regex m g (Plus r s) = ((g = Lax \ fv_regex fv r = fv_regex fv s) \ safe_regex m g r \ safe_regex m g s)" | "safe_regex Futu g (Times r s) = ((g = Lax \ fv_regex fv r \ fv_regex fv s) \ safe_regex Futu g s \ safe_regex Futu Lax r)" | "safe_regex Past g (Times r s) = ((g = Lax \ fv_regex fv s \ fv_regex fv r) \ safe_regex Past g r \ safe_regex Past Lax s)" | "safe_regex m g (Star r) = ((g = Lax \ fv_regex fv r = {}) \ safe_regex m g r)" lemmas safe_regex_induct = safe_regex.induct[case_names Skip Test Plus TimesF TimesP Star] lemma safe_cosafe: "(\x. x \ atms r \ safe Strict x \ safe Lax x) \ safe_regex m Strict r \ safe_regex m Lax r" by (induct r; cases m) auto lemma safe_lpd_fv_regex_le: "safe_regex Futu Strict r \ s \ lpd test i r \ fv_regex fv r \ fv_regex fv s" by (induct r) (auto simp: TimesR_def split: if_splits) lemma safe_lpd_fv_regex: "safe_regex Futu Strict r \ s \ lpd test i r \ fv_regex fv s = fv_regex fv r" by (simp add: eq_iff lpd_fv_regex safe_lpd_fv_regex_le) lemma cosafe_lpd: "safe_regex Futu Lax r \ s \ lpd test i r \ safe_regex Futu Lax s" proof (induct r arbitrary: s) case (Plus r1 r2) from Plus(3,4) show ?case by (auto elim: Plus(1,2)) next case (Times r1 r2) from Times(3,4) show ?case by (auto simp: TimesR_def elim: Times(1,2) split: if_splits) qed (auto simp: TimesR_def split: nat.splits) lemma safe_lpd: "(\x \ atms r. safe Strict x \ safe Lax x) \ safe_regex Futu Strict r \ s \ lpd test i r \ safe_regex Futu Strict s" proof (induct r arbitrary: s) case (Plus r1 r2) from Plus(3,4,5) show ?case by (auto elim: Plus(1,2) simp: ball_Un) next case (Times r1 r2) from Times(3,4,5) show ?case by (force simp: TimesR_def ball_Un elim: Times(1,2) cosafe_lpd dest: lpd_fv_regex split: if_splits) next case (Star r) from Star(2,3,4) show ?case by (force simp: TimesR_def elim: Star(1) cosafe_lpd dest: safe_cosafe[rotated] lpd_fv_regex[where fv=fv] split: if_splits) qed (auto split: nat.splits) lemma safe_rpd_fv_regex_le: "safe_regex Past Strict r \ s \ rpd test i r \ fv_regex fv r \ fv_regex fv s" by (induct r) (auto simp: TimesL_def split: if_splits) lemma safe_rpd_fv_regex: "safe_regex Past Strict r \ s \ rpd test i r \ fv_regex fv s = fv_regex fv r" by (simp add: eq_iff rpd_fv_regex safe_rpd_fv_regex_le) lemma cosafe_rpd: "safe_regex Past Lax r \ s \ rpd test i r \ safe_regex Past Lax s" proof (induct r arbitrary: s) case (Plus r1 r2) from Plus(3,4) show ?case by (auto elim: Plus(1,2)) next case (Times r1 r2) from Times(3,4) show ?case by (auto simp: TimesL_def elim: Times(1,2) split: if_splits) qed (auto simp: TimesL_def split: nat.splits) lemma safe_rpd: "(\x \ atms r. safe Strict x \ safe Lax x) \ safe_regex Past Strict r \ s \ rpd test i r \ safe_regex Past Strict s" proof (induct r arbitrary: s) case (Plus r1 r2) from Plus(3,4,5) show ?case by (auto elim: Plus(1,2) simp: ball_Un) next case (Times r1 r2) from Times(3,4,5) show ?case by (force simp: TimesL_def ball_Un elim: Times(1,2) cosafe_rpd dest: rpd_fv_regex split: if_splits) next case (Star r) from Star(2,3,4) show ?case by (force simp: TimesL_def elim: Star(1) cosafe_rpd dest: safe_cosafe[rotated] rpd_fv_regex[where fv=fv] split: if_splits) qed (auto split: nat.splits) lemma safe_regex_safe: "(\g r. safe g r \ safe Lax r) \ safe_regex m g r \ x \ atms r \ safe Lax x" by (induct m g r rule: safe_regex.induct) auto lemma safe_regex_map_regex: "(\g x. x \ atms r \ safe g x \ safe g (f x)) \ (\x. x \ atms r \ fv (f x) = fv x) \ safe_regex m g r \ safe_regex m g (map_regex f r)" by (induct m g r rule: safe_regex.induct) (auto simp: fv_regex_alt regex.set_map) end lemma safe_regex_cong[fundef_cong]: "(\g x. x \ atms r \ safe g x = safe' g x) \ Regex.safe_regex fv safe m g r = Regex.safe_regex fv safe' m g r" by (induct m g r rule: safe_regex.induct) auto lemma safe_regex_mono: "(\g x. x \ atms r \ safe g x \ safe' g x) \ Regex.safe_regex fv safe m g r \ Regex.safe_regex fv safe' m g r" by (induct m g r rule: safe_regex.induct) auto lemma match_map_regex: "match t (map_regex f r) = match (\k z. t k (f z)) r" by (induct r) auto lemma match_cong_strong: "(\k z. k \ {i ..< j + 1} \ z \ atms r \ t k z = t' k z) \ match t r i j = match t' r i j" proof (induction r arbitrary: i j) case (Times r s) from Times.prems show ?case by (auto 0 4 simp: relcompp_apply intro: le_less_trans match_le less_Suc_eq_le dest: Times.IH[THEN iffD1, rotated -1] Times.IH[THEN iffD2, rotated -1] match_le) next case (Star r) show ?case unfolding match.simps proof (rule iffI) assume *: "(match t r)\<^sup>*\<^sup>* i j" then have "i \ j" unfolding match.simps(5)[symmetric] by (rule match_le) with * show "(match t' r)\<^sup>*\<^sup>* i j" using Star.prems proof (induction i j rule: rtranclp.induct) case (rtrancl_into_rtrancl a b c) from rtrancl_into_rtrancl(1,2,4,5) show ?case by (intro rtranclp.rtrancl_into_rtrancl[OF rtrancl_into_rtrancl.IH]) (auto dest!: Star.IH[THEN iffD1, rotated -1] dest: match_le match_rtranclp_le simp: less_Suc_eq_le) qed simp next assume *: "(match t' r)\<^sup>*\<^sup>* i j" then have "i \ j" unfolding match.simps(5)[symmetric] by (rule match_le) with * show "(match t r)\<^sup>*\<^sup>* i j" using Star.prems proof (induction i j rule: rtranclp.induct) case (rtrancl_into_rtrancl a b c) from rtrancl_into_rtrancl(1,2,4,5) show ?case by (intro rtranclp.rtrancl_into_rtrancl[OF rtrancl_into_rtrancl.IH]) (auto dest!: Star.IH[THEN iffD2, rotated -1] dest: match_le match_rtranclp_le simp: less_Suc_eq_le) qed simp qed qed auto end (*<*) end (*>*) diff --git a/thys/MonoBoolTranAlgebra/Mono_Bool_Tran.thy b/thys/MonoBoolTranAlgebra/Mono_Bool_Tran.thy --- a/thys/MonoBoolTranAlgebra/Mono_Bool_Tran.thy +++ b/thys/MonoBoolTranAlgebra/Mono_Bool_Tran.thy @@ -1,500 +1,488 @@ section \Monotonic Boolean Transformers\ theory Mono_Bool_Tran imports LatticeProperties.Complete_Lattice_Prop LatticeProperties.Conj_Disj begin text\ The type of monotonic transformers is the type associated to the set of monotonic functions from a partially ordered set (poset) to itself. The type of monotonic transformers with the pointwise extended order is also a poset. The monotonic transformers with composition and identity form a monoid, and the monoid operation is compatible with the order. Gradually we extend the algebraic structure of monotonic transformers to lattices, and complete lattices. We also introduce a dual operator ($(\mathsf{dual}\;f) p = - f (-p)$) on monotonic transformers over a boolean algebra. However the monotonic transformers over a boolean algebra are not closed to the pointwise extended negation operator. Finally we introduce an iteration operator on monotonic transformers over a complete lattice. \ -notation - bot ("\") and - top ("\") and - inf (infixl "\" 70) and - sup (infixl "\" 65) and - Inf ("\_" [900] 900) and - Sup ("\_" [900] 900) - -syntax - "_INF1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) - "_INF" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) - "_SUP1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) - "_SUP" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) +unbundle lattice_syntax lemma Inf_comp_fun: "\M \ f = (\m\M. m \ f)" by (simp add: fun_eq_iff image_comp) lemma INF_comp_fun: "(\a\A. g a) \ f = (\a\A. g a \ f)" by (simp add: fun_eq_iff image_comp) lemma Sup_comp_fun: "\M \ f = (\m\M. m \ f)" by (simp add: fun_eq_iff image_comp) lemma SUP_comp_fun: "(\a\A. g a) \ f = (\a\A. g a \ f)" by (simp add: fun_eq_iff image_comp) lemma (in order) mono_const [simp]: "mono (\_. c)" by (auto intro: monoI) lemma (in order) mono_id [simp]: "mono id" by (auto intro: order_class.monoI) lemma (in order) mono_comp [simp]: "mono f \ mono g \ mono (f \ g)" by (auto intro!: monoI elim!: monoE order_class.monoE) lemma (in bot) mono_bot [simp]: "mono \" by (auto intro: monoI) lemma (in top) mono_top [simp]: "mono \" by (auto intro: monoI) lemma (in semilattice_inf) mono_inf [simp]: assumes "mono f" and "mono g" shows "mono (f \ g)" proof fix a b assume "a \ b" have "f a \ g a \ f a" by simp also from \mono f\ \a \ b\ have "\ \ f b" by (auto elim: monoE) finally have *: "f a \ g a \ f b" . have "f a \ g a \ g a" by simp also from \mono g\ \a \ b\ have "\ \ g b" by (auto elim: monoE) finally have **: "f a \ g a \ g b" . from * ** show "(f \ g) a \ (f \ g) b" by auto qed lemma (in semilattice_sup) mono_sup [simp]: assumes "mono f" and "mono g" shows "mono (f \ g)" proof fix a b assume "a \ b" from \mono f\ \a \ b\ have "f a \ f b" by (auto elim: monoE) also have "f b \ f b \ g b" by simp finally have *: "f a \ f b \ g b" . from \mono g\ \a \ b\ have "g a \ g b" by (auto elim: monoE) also have "g b \ f b \ g b" by simp finally have **: "g a \ f b \ g b" . from * ** show "(f \ g) a \ (f \ g) b" by auto qed lemma (in complete_lattice) mono_Inf [simp]: assumes "A \ {f :: 'a \ 'b:: complete_lattice. mono f}" shows "mono (\A)" proof fix a b assume "a \ b" { fix f assume "f \ A" with assms have "mono f" by auto with \a \ b\ have "f a \ f b" by (auto elim: monoE) } then have "(\f\A. f a) \ (\f\A. f b)" by (auto intro: complete_lattice_class.INF_greatest complete_lattice_class.INF_lower2) then show "(\A) a \ (\A) b" by simp qed lemma (in complete_lattice) mono_Sup [simp]: assumes "A \ {f :: 'a \ 'b:: complete_lattice. mono f}" shows "mono (\A)" proof fix a b assume "a \ b" { fix f assume "f \ A" with assms have "mono f" by auto with \a \ b\ have "f a \ f b" by (auto elim: monoE) } then have "(\f\A. f a) \ (\f\A. f b)" by (auto intro: complete_lattice_class.SUP_least complete_lattice_class.SUP_upper2) then show "(\A) a \ (\A) b" by simp qed typedef (overloaded) 'a MonoTran = "{f::'a::order \ 'a . mono f}" proof show "id \ ?MonoTran" by simp qed lemma [simp]: "mono (Rep_MonoTran f)" using Rep_MonoTran [of f] by simp setup_lifting type_definition_MonoTran instantiation MonoTran :: (order) order begin lift_definition less_eq_MonoTran :: "'a MonoTran \ 'a MonoTran \ bool" is less_eq . lift_definition less_MonoTran :: "'a MonoTran \ 'a MonoTran \ bool" is less . instance by intro_classes (transfer, auto intro: order_antisym)+ end instantiation MonoTran :: (order) monoid_mult begin lift_definition one_MonoTran :: "'a MonoTran" is id by (fact mono_id) lift_definition times_MonoTran :: "'a MonoTran \ 'a MonoTran \ 'a MonoTran" is comp by (fact mono_comp) instance by intro_classes (transfer, auto)+ end instantiation MonoTran :: (order_bot) order_bot begin lift_definition bot_MonoTran :: "'a MonoTran" is \ by (fact mono_bot) instance by intro_classes (transfer, simp) end instantiation MonoTran :: (order_top) order_top begin lift_definition top_MonoTran :: "'a MonoTran" is \ by (fact mono_top) instance by intro_classes (transfer, simp) end instantiation MonoTran :: (lattice) lattice begin lift_definition inf_MonoTran :: "'a MonoTran \ 'a MonoTran \ 'a MonoTran" is inf by (fact mono_inf) lift_definition sup_MonoTran :: "'a MonoTran \ 'a MonoTran \ 'a MonoTran" is sup by (fact mono_sup) instance by intro_classes (transfer, simp)+ end instance MonoTran :: (distrib_lattice) distrib_lattice by intro_classes (transfer, rule sup_inf_distrib1) instantiation MonoTran :: (complete_lattice) complete_lattice begin lift_definition Inf_MonoTran :: "'a MonoTran set \ 'a MonoTran" is Inf by (rule mono_Inf) auto lift_definition Sup_MonoTran :: "'a MonoTran set \ 'a MonoTran" is Sup by (rule mono_Sup) auto instance by intro_classes (transfer, simp add: Inf_lower Sup_upper Inf_greatest Sup_least)+ end context includes lifting_syntax begin lemma [transfer_rule]: "(rel_set A ===> (A ===> pcr_MonoTran HOL.eq) ===> pcr_MonoTran HOL.eq) (\A f. \(f ` A)) (\A f. \(f ` A))" by transfer_prover lemma [transfer_rule]: "(rel_set A ===> (A ===> pcr_MonoTran HOL.eq) ===> pcr_MonoTran HOL.eq) (\A f. \(f ` A)) (\A f. \(f ` A))" by transfer_prover end instance MonoTran :: (complete_distrib_lattice) complete_distrib_lattice proof (intro_classes, transfer) fix A :: "('a \ 'a) set set" assume " \A\A. Ball A mono" from this have [simp]: "{f ` A |f. \Y\A. f Y \ Y} = {x. (\f. (\x. (\x\x. mono x) \ mono (f x)) \ x = f ` A \ (\Y\A. f Y \ Y)) \ (\x\x. mono x)}" apply safe apply (rule_tac x = "\ x . if x \ A then f x else \" in exI) apply (simp add: if_split image_def) by blast+ show " \(Sup ` A) \ \(Inf ` {x. (\f\Collect (pred_fun (\A. Ball A mono) mono). x = f ` A \ (\Y\A. f Y \ Y)) \ Ball x mono})" by (simp add: Inf_Sup) qed definition "dual_fun (f::'a::boolean_algebra \ 'a) = uminus \ f \ uminus" lemma dual_fun_apply [simp]: "dual_fun f p = - f (- p)" by (simp add: dual_fun_def) lemma mono_dual_fun [simp]: "mono f \ mono (dual_fun f)" apply (rule monoI) apply (erule monoE) apply auto done lemma (in order) mono_inf_fun [simp]: fixes x :: "'b::semilattice_inf" shows "mono (inf x)" by (auto intro!: order_class.monoI semilattice_inf_class.inf_mono) lemma (in order) mono_sup_fun [simp]: fixes x :: "'b::semilattice_sup" shows "mono (sup x)" by (auto intro!: order_class.monoI semilattice_sup_class.sup_mono) lemma mono_comp_fun: fixes f :: "'a::order \ 'b::order" shows "mono f \ mono ((\) f)" by (rule monoI) (auto simp add: le_fun_def elim: monoE) definition "Omega_fun f g = inf g \ comp f" lemma Omega_fun_apply [simp]: "Omega_fun f g h p = (g p \ f (h p))" by (simp add: Omega_fun_def) lemma mono_Omega_fun [simp]: "mono f \ mono (Omega_fun f g)" unfolding Omega_fun_def by (auto intro: mono_comp mono_comp_fun) lemma mono_mono_Omega_fun [simp]: fixes f :: "'b::order \ 'a::semilattice_inf" and g :: "'c::semilattice_inf \ 'a" shows "mono f \ mono g \ mono_mono (Omega_fun f g)" apply (auto simp add: mono_mono_def Omega_fun_def) apply (rule mono_comp) apply (rule mono_inf_fun) apply (rule mono_comp_fun) apply assumption done definition "omega_fun f = lfp (Omega_fun f id)" definition "star_fun f = gfp (Omega_fun f id)" lemma mono_omega_fun [simp]: fixes f :: "'a::complete_lattice \ 'a" assumes "mono f" shows "mono (omega_fun f)" proof fix a b :: 'a assume "a \ b" from assms have "mono (lfp (Omega_fun f id))" by (auto intro: mono_mono_Omega_fun) with \a \ b\ show "omega_fun f a \ omega_fun f b" by (auto simp add: omega_fun_def elim: monoE) qed lemma mono_star_fun [simp]: fixes f :: "'a::complete_lattice \ 'a" assumes "mono f" shows "mono (star_fun f)" proof fix a b :: 'a assume "a \ b" from assms have "mono (gfp (Omega_fun f id))" by (auto intro: mono_mono_Omega_fun) with \a \ b\ show "star_fun f a \ star_fun f b" by (auto simp add: star_fun_def elim: monoE) qed lemma lfp_omega_lowerbound: "mono f \ Omega_fun f g A \ A \ omega_fun f \ g \ A" apply (simp add: omega_fun_def) apply (rule_tac P = "\ x . x \ g \ A" and f = "Omega_fun f id" in lfp_ordinal_induct) apply simp_all apply (simp add: le_fun_def o_def inf_fun_def id_def Omega_fun_def) apply auto apply (rule_tac y = "f (A x) \ g x" in order_trans) apply simp_all apply (rule_tac y = "f (S (g x))" in order_trans) apply simp_all apply (simp add: mono_def) apply (auto simp add: ac_simps) apply (unfold Sup_comp_fun) apply (rule SUP_least) by auto lemma gfp_omega_upperbound: "mono f \ A \ Omega_fun f g A \ A \ star_fun f \ g" apply (simp add: star_fun_def) apply (rule_tac P = "\ x . A \ x \ g" and f = "Omega_fun f id" in gfp_ordinal_induct) apply simp_all apply (simp add: le_fun_def o_def inf_fun_def id_def Omega_fun_def) apply auto apply (rule_tac y = "f (A x) \ g x" in order_trans) apply simp_all apply (rule_tac y = "f (A x)" in order_trans) apply simp_all apply (simp add: mono_def) apply (unfold Inf_comp_fun) apply (rule INF_greatest) by auto lemma lfp_omega_greatest: assumes "\u. Omega_fun f g u \ u \ A \ u" shows "A \ omega_fun f \ g" apply (unfold omega_fun_def) apply (simp add: lfp_def) apply (unfold Inf_comp_fun) apply (rule INF_greatest) apply simp apply (rule assms) apply (simp add: le_fun_def) done lemma gfp_star_least: assumes "\u. u \ Omega_fun f g u \ u \ A" shows "star_fun f \ g \ A" apply (unfold star_fun_def) apply (simp add: gfp_def) apply (unfold Sup_comp_fun) apply (rule SUP_least) apply simp apply (rule assms) apply (simp add: le_fun_def) done lemma lfp_omega: "mono f \ omega_fun f \ g = lfp (Omega_fun f g)" apply (rule antisym) apply (rule lfp_omega_lowerbound) apply simp_all apply (simp add: lfp_def) apply (rule Inf_greatest) apply safe apply (rule_tac y = "Omega_fun f g x" in order_trans) apply simp_all apply (rule_tac f = " Omega_fun f g" in monoD) apply simp_all apply (rule Inf_lower) apply simp apply (rule lfp_omega_greatest) apply (simp add: lfp_def) apply (rule Inf_lower) by simp lemma gfp_star: "mono f \ star_fun f \ g = gfp (Omega_fun f g)" apply (rule antisym) apply (rule gfp_star_least) apply (simp add: gfp_def) apply (rule Sup_upper, simp) apply (rule gfp_omega_upperbound) apply simp_all apply (simp add: gfp_def) apply (rule Sup_least) apply safe apply (rule_tac y = "Omega_fun f g x" in order_trans) apply simp_all apply (rule_tac f = " Omega_fun f g" in monoD) apply simp_all apply (rule Sup_upper) by simp definition "assert_fun p q = (p \ q :: 'a::semilattice_inf)" lemma mono_assert_fun [simp]: "mono (assert_fun p)" apply (simp add: assert_fun_def mono_def, safe) by (rule_tac y = x in order_trans, simp_all) lemma assert_fun_le_id [simp]: "assert_fun p \ id" by (simp add: assert_fun_def id_def le_fun_def) lemma assert_fun_disjunctive [simp]: "assert_fun (p::'a::distrib_lattice) \ Apply.disjunctive" by (simp add: assert_fun_def Apply.disjunctive_def inf_sup_distrib) definition "assertion_fun = range assert_fun" lemma assert_cont: "(x :: 'a::boolean_algebra \ 'a) \ id \ x \ Apply.disjunctive \ x = assert_fun (x \)" apply (rule antisym) apply (simp_all add: le_fun_def assert_fun_def, safe) apply (rule_tac f = x in monoD, simp_all) apply (subgoal_tac "x top = sup (x xa) (x (-xa))") apply simp apply (subst inf_sup_distrib) apply simp apply (rule_tac y = "inf (- xa) xa" in order_trans) supply [[simproc del: boolean_algebra_cancel_inf]] apply (simp del: compl_inf_bot) apply (rule_tac y = "x (- xa)" in order_trans) apply simp apply simp apply simp apply (cut_tac x = x and y = xa and z = "-xa" in Apply.disjunctiveD, simp) apply (subst (asm) sup_commute) apply (subst (asm) compl_sup_top) by simp lemma assertion_fun_disj_less_one: "assertion_fun = Apply.disjunctive \ {x::'a::boolean_algebra \ 'a . x \ id}" apply safe apply (simp_all add: assertion_fun_def, auto simp add: image_def) apply (rule_tac x = "x \" in exI) by (rule assert_cont, simp_all) lemma assert_fun_dual: "((assert_fun p) o \) \ (dual_fun (assert_fun p)) = assert_fun p" by (simp add: fun_eq_iff inf_fun_def dual_fun_def o_def assert_fun_def top_fun_def inf_sup_distrib) lemma assertion_fun_dual: "x \ assertion_fun \ (x o \) \ (dual_fun x) = x" by (simp add: assertion_fun_def, safe, simp add: assert_fun_dual) lemma assertion_fun_MonoTran [simp]: "x \ assertion_fun \ mono x" by (unfold assertion_fun_def, auto) lemma assertion_fun_le_one [simp]: "x \ assertion_fun \ x \ id" by (unfold assertion_fun_def, auto) end diff --git a/thys/MonoBoolTranAlgebra/Statements.thy b/thys/MonoBoolTranAlgebra/Statements.thy --- a/thys/MonoBoolTranAlgebra/Statements.thy +++ b/thys/MonoBoolTranAlgebra/Statements.thy @@ -1,437 +1,425 @@ section \Program statements, Hoare and refinement rules\ theory Statements imports Assertion_Algebra begin text \In this section we introduce assume, if, and while program statements as well as Hoare triples, and data refienment. We prove Hoare correctness rules for the program statements and we prove some theorems linking Hoare correctness statement to (data) refinement. Most of the theorems assume a monotonic boolean transformers algebra. The theorem stating the equivalence between a Hoare correctness triple and a refinement statement holds under the assumption that we have a monotonic boolean transformers algebra with post condition statement.\ definition "assume" :: "'a::mbt_algebra Assertion \ 'a" ("[\ _ ]" [0] 1000) where "[\p] = {\p} ^ o" lemma [simp]: "{\p} * \ \ [\p] = {\p}" apply (subgoal_tac "{\p} \ assertion") apply (subst (asm) assertion_def, simp add: assume_def) by simp lemma [simp]: "[\p] * x \ {\-p} * \ = [\p] * x" by (simp add: assume_def uminus_Assertion_def) lemma [simp]: "{\p} * \ \ [\-p] * x = [\-p] * x" by (simp add: assume_def uminus_Assertion_def) lemma assert_sup: "{\p \ q} = {\p} \ {\q}" by (simp add: sup_Assertion_def) lemma assert_inf: "{\p \ q} = {\p} \ {\q}" by (simp add: inf_Assertion_def) lemma assert_neg: "{\-p} = neg_assert {\p}" by (simp add: uminus_Assertion_def) lemma assert_false [simp]: "{\\} = \" by (simp add: bot_Assertion_def) lemma if_Assertion_assumption: "({\p} * x) \ ({\-p} * y) = ([\p] * x) \ ([\-p] * y)" proof - have "({\p} * x) \ {\-p} * y = ({\p} * \ \ [\p]) * x \ ({\-p} * \ \ [\-p]) * y" by simp also have "\ = ({\p} * \ \ ([\p] * x)) \ ({\-p} * \ \ ([\-p] * y))" by (unfold inf_comp, simp) also have "\ = (({\p} * \ \ ([\p] * x)) \ ({\-p} * \)) \ (({\p} * \ \ ([\p] * x)) \ ([\-p] * y))" by (simp add: sup_inf_distrib) also have "\ = (({\p} * \ \ ({\-p} * \)) \ (([\p] * x))) \ (([\-p] * y) \ (([\p] * x) \ ([\-p] * y)))" by (simp add: sup_inf_distrib2) also have "\ = ([\p] * x) \ ([\-p] * y) \ (([\p] * x) \ ([\-p] * y))" apply (simp add: sup_comp [THEN sym] ) by (simp add: assert_sup [THEN sym] inf_assoc) also have "\ = ([\p] * x) \ ([\-p] * y)" by (rule antisym, simp_all add: inf_assoc) finally show ?thesis . qed definition "wp x p = abs_wpt (x * {\p})" lemma wp_assume: "wp [\p] q = -p \ q" apply (simp add: wp_def abs_wpt_def) apply (rule assert_injective) apply simp by (simp add: assert_sup assert_neg assume_def wpt_dual_assertion_comp) lemma assert_commute: "y \ conjunctive \ y * {\p} = {\ wp y p } * y" apply (simp add: wp_def abs_wpt_def) by (rule assertion_commute, simp_all) lemma wp_assert: "wp {\p} q = p \ q" by (simp add: wp_def assertion_inf_comp_eq [THEN sym] assert_inf [THEN sym]) lemma wp_mono [simp]: "mono (wp x)" apply (simp add: le_fun_def wp_def abs_wpt_def less_eq_Assertion_def mono_def) apply (simp add: wpt_def, safe) apply (rule_tac y = " x * {\ xa } * \" in order_trans, simp_all) apply (rule le_comp_right) by (rule le_comp, simp) lemma wp_mono2: "p \ q \ wp x p \ wp x q" apply (cut_tac x = x in wp_mono) apply (unfold mono_def) by blast lemma wp_fun_mono [simp]: "mono wp" apply (simp add: le_fun_def wp_def abs_wpt_def less_eq_Assertion_def mono_def) apply (simp add: wpt_def, safe) apply (rule_tac y = " x * {\ xa } * \" in order_trans, simp_all) apply (rule le_comp_right) by (rule le_comp_right, simp) lemma wp_fun_mono2: "x \ y \ wp x p \ wp y p" apply (cut_tac wp_fun_mono) apply (unfold mono_def) apply (simp add: le_fun_def) by blast lemma wp_comp: "wp (x * y) p = wp x (wp y p)" apply (simp add: wp_def abs_wpt_def) by (unfold wpt_comp_2 [THEN sym] mult.assoc, simp) lemma wp_choice: "wp (x \ y) = wp x \ wp y" apply (simp add: fun_eq_iff wp_def inf_fun_def inf_comp inf_Assertion_def abs_wpt_def) by (simp add: wpt_choice) lemma [simp]: "wp 1 = id" apply (unfold fun_eq_iff, safe) apply (rule assert_injective) by (simp add: wp_def abs_wpt_def) lemma wp_omega_fix: "wp (x ^ \) p = wp x (wp (x ^ \) p) \ p" apply (subst omega_fix) by (simp add: wp_choice wp_comp) lemma wp_omega_least: "(wp x r) \ p \ r \ wp (x ^ \) p \ r" apply (simp add: wp_def abs_wpt_def inf_Assertion_def less_eq_Assertion_def) apply (simp add: wpt_def) apply (rule_tac y = "{\r} * \ \ 1" in order_trans) apply simp apply (rule_tac y = "x ^ \ * {\ p } * \" in order_trans, simp) apply (simp add: mult.assoc) apply (rule omega_least) apply (drule_tac z = \ in le_comp_right) apply (simp add: inf_comp mult.assoc [THEN sym]) by (simp add: assertion_prop) lemma Assertion_wp: "{\wp x p} = (x * {\p} * \) \ 1" apply (simp add: wp_def abs_wpt_def) by (simp add: wpt_def) definition "hoare p S q = (p \ wp S q)" definition "grd x = - (wp x \)" lemma grd_comp: "[\grd x] * x = x" apply (simp add: grd_def wp_def uminus_Assertion_def assume_def neg_assert_def abs_wpt_def dual_sup sup_comp) apply (simp add: wpt_def dual_inf sup_comp dual_comp bot_Assertion_def) by (rule antisym, simp_all) lemma assert_assume: "{\p} * [\p] = {\ p}" by (simp add: assume_def) lemma dual_assume: "[\p] ^ o = {\p}" by (simp add: assume_def) lemma assume_prop: "([\p] * \) \ 1 = [\p]" by (simp add: assume_def dual_assertion_prop) text\An alternative definition of a Hoare triple\ definition "hoare1 p S q = ([\ p ] * S * [\ -q ] = \)" lemma "hoare1 p S q = hoare p S q" apply (simp add: hoare1_def dual_inf dual_comp) apply (simp add: hoare_def wp_def less_eq_Assertion_def abs_wpt_def) apply (simp add: wpt_def) apply safe proof - assume A: "[\ p ] * S * [\ - q ] = \" have "{\p} \ {\p} * \" by simp also have "... \ {\p} * \ * \" by (unfold mult.assoc, simp) also have "... = {\p} * [\ p ] * S * [\ - q ] * \" by (subst A [THEN sym], simp add: mult.assoc) also have "... = {\p} * S * [\ - q ] * \" by (simp add: assert_assume) also have "... \ {\p} * S * {\ q } * \" apply (simp add: mult.assoc) apply (rule le_comp, rule le_comp) apply (simp add: assume_def uminus_Assertion_def) by (simp add: neg_assert_def dual_inf dual_comp sup_comp) also have "... \ S * {\ q } * \" by (simp add: mult.assoc) finally show "{\p} \ S * {\ q } * \" . next assume A: "{\ p } \ S * {\ q } * \" have "\ = ((S * {\q}) ^ o) * \ \ S * {\q} * \" by simp also have "\ \ [\p] * \ \ S * {\q} * \" apply (simp del: dual_neg_top) apply (rule_tac y = "[\p] * \" in order_trans, simp_all) apply (subst dual_le) apply (simp add: dual_comp dual_assume) apply (cut_tac x = "{\p}" and y = "S * {\q} * \" and z = \ in le_comp_right) apply (rule A) by (simp add: mult.assoc) also have "\ = [\p] * S * ({\q} * \)" apply (subst (2) assume_prop [THEN sym]) by (simp_all add: sup_comp mult.assoc) also have "\ \ [\p] * S * ({\q} * \ \ 1)" by (rule le_comp, simp) also have "\ = [\p] * S * [\-q]" apply (simp add: assume_def uminus_Assertion_def) by (simp add: neg_assert_def dual_inf dual_comp) finally show "[\p] * S * [\ - q] = \" by (rule_tac antisym, simp_all) qed lemma hoare_choice: "hoare p (x \ y) q = ((hoare p) x q & (hoare p y q))" apply (unfold hoare_def wp_choice inf_fun_def) by auto definition if_stm:: "'a::mbt_algebra Assertion \ 'a \ 'a \ 'a" ("(If (_)/ then (_)/ else (_))" [0, 0, 10] 10) where "if_stm b x y = (([\ b ] * x) \ ([\ -b ] * y))" lemma if_assertion: "(If p then x else y) = {\p} * x \ {\ -p} * y" by (simp add: if_stm_def if_Assertion_assumption) lemma hoare_if: "hoare p (If b then x else y) q = (hoare (p \ b) x q \ hoare (p \ -b) y q)" by (simp add: hoare_def if_stm_def wp_choice inf_fun_def wp_comp wp_assume sup_neg_inf) lemma hoare_comp: "hoare p (x * y) q = (\ r . (hoare p x r) \ (hoare r y q))" apply (simp add: hoare_def wp_comp) apply safe apply (rule_tac x = "wp y q" in exI, simp) apply (rule_tac y = "wp x r" in order_trans, simp) apply (rule_tac f = "wp x" in monoD) by simp_all lemma hoare_refinement: "hoare p S q = ({\p} * (post {\q}) \ S)" apply (simp add: hoare_def less_eq_Assertion_def Assertion_wp) proof assume A: "{\p} \ S * {\q} * \" have "{\p} * post {\q} = ({\p} * \ \ 1) * post {\q}" by (simp add: assertion_prop) also have "\ = {\p} * \ \ post {\q}" by (simp add: inf_comp) also have "\ \ S * {\q} * \ \ post {\q}" apply simp apply (rule_tac y = "{\p} * \" in order_trans, simp_all) apply (cut_tac x = "{\p}" and y = "S * {\q} * \" and z = \ in le_comp_right) by (rule A, simp) also have "\ \ S" by (simp add: post_2) finally show "{\p} * post {\q} \ S". next assume A: "{\p} * post {\q} \ S" have "{\p} = {\p} * \ \ 1" by (simp add: assertion_prop) also have "\ = {\p} * ((post {\q}) * {\q} * \) \ 1" by (simp add: post_1) also have "\ \ {\p} * ((post {\q}) * {\q} * \)" by simp also have "\ \ S * {\q} * \" apply (cut_tac x = "{\p} * post {\q}" and y = S and z = "{\q} * \" in le_comp_right) apply (simp add: A) by (simp add: mult.assoc) finally show "{\p} \ S * {\q} * \" . qed theorem hoare_fixpoint_mbt: "F x = x \ (!! (w::'a::well_founded) f . (\v. v < w \ hoare (p v) f q) \ hoare (p w) (F f) q) \ hoare (p u) x q" apply (rule less_induct1) proof - fix xa assume A: "\ w f. (\ v . v < w \ hoare (p v) f q) \ hoare (p w) (F f) q" assume B: "F x = x" assume C: "\y . y < xa \ hoare (p y) x q" have D: "hoare (p xa) (F x) q" apply (rule A) by (rule C, simp) show "hoare (p xa) x q" by (cut_tac D, simp add: B) qed lemma hoare_Sup: "hoare (Sup P) x q = (\ p \ P . hoare p x q)" apply (simp add: hoare_def) apply auto apply (rule_tac y = "Sup P" in order_trans, simp_all add: Sup_upper) apply (rule Sup_least) by simp theorem hoare_fixpoint_complete_mbt: "F x = x \ (!! w f . hoare (Sup_less p w) f q \ hoare (p w) (F f) q) \ hoare (Sup (range p)) x q" apply (simp add: hoare_Sup Sup_less_def, safe) apply (rule_tac F = F in hoare_fixpoint_mbt) by auto definition while:: "'a::mbt_algebra Assertion \ 'a \ 'a" ("(While (_)/ do (_))" [0, 10] 10) where "while p x = ([\ p] * x) ^ \ * [\ -p ]" lemma while_false: "(While \ do x) = 1" apply (unfold while_def) apply (subst omega_fix) by (simp_all add: assume_def) lemma while_true: "(While \ do 1) = \" apply (unfold while_def) by (rule antisym, simp_all add: assume_def) lemma hoare_wp [simp]: "hoare (wp x q) x q" by (simp add: hoare_def) lemma hoare_comp_wp: "hoare p (x * y) q = hoare p x (wp y q)" apply (unfold hoare_comp, safe) apply (simp add: hoare_def) apply (rule_tac y = "wp x r" in order_trans, simp) apply (rule wp_mono2, simp) by (rule_tac x = "wp y q" in exI, simp) lemma (in mbt_algebra) hoare_assume: "hoare p [\b] q = (p \ b \ q)" by (simp add: hoare_def wp_assume sup_neg_inf) lemma (in mbt_algebra) hoare_assume_comp: "hoare p ([\b] * x) q = hoare (p \ b) x q" apply (simp add: hoare_comp_wp hoare_assume) by (simp add: hoare_def) lemma hoare_while_mbt: "(\ (w::'b::well_founded) r . (\ v . v < w \ p v \ r) \ hoare ((p w) \ b) x r) \ (\ u . p u \ q) \ hoare (p w) (While b do x) (q \ -b)" apply (unfold while_def) apply (rule_tac F = "\z. [\ b ] * x * z \ [\ - b ]" in hoare_fixpoint_mbt) apply (simp add: mult.assoc [THEN sym]) apply (simp add: omega_comp_fix) apply (unfold hoare_choice) apply safe apply (subst hoare_comp_wp) apply (subst hoare_assume_comp) apply (drule_tac x = w in spec) apply (drule_tac x = "wp f (q \ - b)" in spec) apply (auto simp add: hoare_def) [1] apply (auto simp add: hoare_assume) apply (rule_tac y = "p w" in order_trans) by simp_all lemma hoare_while_complete_mbt: "(\ w::'b::well_founded . hoare ((p w) \ b) x (Sup_less p w)) \ hoare (Sup (range p)) (While b do x) ((Sup (range p)) \ -b)" apply (simp add: hoare_Sup, safe) apply (rule hoare_while_mbt) apply safe apply (drule_tac x = w in spec) apply (simp add: hoare_def) apply (rule_tac y = "wp x (Sup_less p w)" in order_trans, simp_all) apply (rule wp_mono2) apply (simp add: Sup_less_def) apply (rule Sup_least, auto) by (rule SUP_upper, simp) definition "datarefin S S1 D D1 = (D * S \ S1 * D1)" lemma "hoare p S q \ datarefin S S1 D D1 \ hoare (wp D p) S1 (wp D1 q)" apply (simp add: hoare_def datarefin_def) apply (simp add: wp_comp [THEN sym] mult.assoc [THEN sym]) apply (rule_tac y = "wp (D * S) q" in order_trans) apply (subst wp_comp) apply (rule monoD, simp_all) by (rule wp_fun_mono2, simp_all) lemma "hoare p S q \ datarefin ({\p} * S) S1 D D1 \ hoare (wp D p) S1 (wp D1 q)" apply (simp add: hoare_def datarefin_def) apply (rule_tac y = "wp (D * {\p} * S) q" in order_trans) apply (simp add: mult.assoc) apply (subst wp_comp) apply (rule monoD, simp_all) apply (subst wp_comp) apply (unfold wp_assert, simp) apply (unfold wp_comp [THEN sym]) apply (rule wp_fun_mono2) by (simp add: mult.assoc) lemma inf_pres_conj: "x \ conjunctive \ y \ conjunctive \ x \ y \ conjunctive" apply (subst conjunctive_def, safe) apply (simp add: inf_comp conjunctiveD) by (metis (opaque_lifting, no_types) inf_assoc inf_left_commute) lemma sup_pres_disj: "x \ disjunctive \ y \ disjunctive \ x \ y \ disjunctive" apply (subst disjunctive_def, safe) apply (simp add: sup_comp disjunctiveD) by (metis (opaque_lifting, no_types) sup_assoc sup_left_commute) lemma assumption_conjuncive [simp]: "[\p] \ conjunctive" by (simp add: assume_def dual_disjunctive assertion_disjunctive) lemma assumption_disjuncive [simp]: "[\p] \ disjunctive" by (simp add: assume_def dual_conjunctive assertion_conjunctive) lemma if_pres_conj: "x \ conjunctive \ y \ conjunctive \ (If p then x else y) \ conjunctive" apply (unfold if_stm_def) by (simp add: inf_pres_conj comp_pres_conj) lemma if_pres_disj: "x \ disjunctive \ y \ disjunctive \ (If p then x else y) \ disjunctive" apply (unfold if_assertion) by (simp add: sup_pres_disj comp_pres_disj assertion_disjunctive) lemma while_dual_star: "(While p do (x::'a::mbt_algebra)) = (({\ p} * x)^\ * {\ -p })" apply (simp add: while_def) apply (rule antisym) apply (rule omega_least) proof - have "([\ p] * x * (({\ p} * x)^\ * {\-p}) \ [\-p]) = ({\ p} * x * (({\ p} * x)^\ * {\-p})) \ {\-p}" apply (unfold mult.assoc) by (cut_tac p = p and x = "(x * (({\ p } * x)^\ * {\ -p }))" and y = 1 in if_Assertion_assumption, simp) also have "\ = ({\ p} * x)^\ * {\-p}" by (simp add: mult.assoc [THEN sym], simp add: dual_star_comp_fix [THEN sym]) finally show "[\ p ] * x * (({\ p } * x)^\ * {\ - p }) \ [\ - p ] \ ({\ p } * x)^\ * {\ - p }" by simp next show "({\ p } * x)^\ * {\ - p } \ ([\ p ] * x) ^ \ * [\ - p ]" apply (rule dual_star_least) proof - have "{\ p } * x * (([\ p ] * x) ^ \ * [\ - p ]) \ {\ - p } = [\ p ] * x * (([\ p ] * x) ^ \ * [\ - p ]) \ [\ - p ]" apply (unfold mult.assoc) by (cut_tac p = p and x = "(x * (([\p] * x)^\ * [\-p]))" and y = 1 in if_Assertion_assumption, simp) also have "... = ([\ p ] * x) ^ \ * [\ - p ]" apply (simp add: mult.assoc [THEN sym]) by (metis omega_comp_fix) finally show "{\ p } * x * (([\ p ] * x) ^ \ * [\ - p ]) \ {\ - p } \ ([\ p ] * x) ^ \ * [\ - p ] " by simp qed qed lemma while_pres_disj: "(x::'a::mbt_algebra) \ disjunctive \ (While p do x) \ disjunctive" apply (unfold while_dual_star) apply (rule comp_pres_disj) apply (rule dual_star_pres_disj) by (rule comp_pres_disj, simp_all add: assertion_disjunctive) lemma while_pres_conj: "(x::'a::mbt_algebra_fusion) \ conjunctive \ (While p do x) \ conjunctive" apply(unfold while_def) by (simp add: comp_pres_conj omega_pres_conj) -no_notation - bot ("\") and - top ("\") and - inf (infixl "\" 70) and - sup (infixl "\" 65) and - Inf ("\_" [900] 900) and - Sup ("\_" [900] 900) - -no_syntax - "_INF1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) - "_INF" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) - "_SUP1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) - "_SUP" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) +unbundle no_lattice_syntax end diff --git a/thys/Order_Lattice_Props/Order_Lattice_Props_Loc.thy b/thys/Order_Lattice_Props/Order_Lattice_Props_Loc.thy --- a/thys/Order_Lattice_Props/Order_Lattice_Props_Loc.thy +++ b/thys/Order_Lattice_Props/Order_Lattice_Props_Loc.thy @@ -1,575 +1,575 @@ (* Title: Locale-Based Duality Author: Georg Struth Maintainer:Georg Struth *) section \Locale-Based Duality\ theory Order_Lattice_Props_Loc imports Main - "HOL-Library.Lattice_Syntax" +begin -begin +unbundle lattice_syntax text \This section explores order and lattice duality based on locales. Used within the context of a class or locale, this is very effective, though more opaque than the previous approach. Outside of such a context, however, it apparently cannot be used for dualising theorems. Examples are properties of functions between orderings or lattices.\ definition Fix :: "('a \ 'a) \ 'a set" where "Fix f = {x. f x = x}" context ord begin definition min_set :: "'a set \ 'a set" where "min_set X = {y \ X. \x \ X. x \ y \ x = y}" definition max_set :: "'a set \ 'a set" where "max_set X = {x \ X. \y \ X. x \ y \ x = y}" definition directed :: "'a set \ bool" where "directed X = (\Y. finite Y \ Y \ X \ (\x \ X. \y \ Y. y \ x))" definition filtered :: "'a set \ bool" where "filtered X = (\Y. finite Y \ Y \ X \ (\x \ X. \y \ Y. x \ y))" definition downset_set :: "'a set \ 'a set" ("\") where "\X = {y. \x \ X. y \ x}" definition upset_set :: "'a set \ 'a set" ("\") where "\X = {y. \x \ X. x \ y}" definition downset :: "'a \ 'a set" ("\") where "\ = \ \ (\x. {x})" definition upset :: "'a \ 'a set" ("\") where "\ = \ \ (\x. {x})" definition downsets :: "'a set set" where "downsets = Fix \" definition upsets :: "'a set set" where "upsets = Fix \" abbreviation "downset_setp X \ X \ downsets" abbreviation "upset_setp X \ X \ upsets" definition ideals :: "'a set set" where "ideals = {X. X \ {} \ downset_setp X \ directed X}" definition filters :: "'a set set" where "filters = {X. X \ {} \ upset_setp X \ filtered X}" abbreviation "idealp X \ X \ ideals" abbreviation "filterp X \ X \ filters" end abbreviation Sup_pres :: "('a::Sup \ 'b::Sup) \ bool" where "Sup_pres f \ f \ Sup = Sup \ (`) f" abbreviation Inf_pres :: "('a::Inf \ 'b::Inf) \ bool" where "Inf_pres f \ f \ Inf = Inf \ (`) f" abbreviation sup_pres :: "('a::sup \ 'b::sup) \ bool" where "sup_pres f \ (\x y. f (x \ y) = f x \ f y)" abbreviation inf_pres :: "('a::inf \ 'b::inf) \ bool" where "inf_pres f \ (\x y. f (x \ y) = f x \ f y)" abbreviation bot_pres :: "('a::bot \ 'b::bot) \ bool" where "bot_pres f \ f \ = \" abbreviation top_pres :: "('a::top \ 'b::top) \ bool" where "top_pres f \ f \ = \" abbreviation Sup_dual :: "('a::Sup \ 'b::Inf) \ bool" where "Sup_dual f \ f \ Sup = Inf \ (`) f" abbreviation Inf_dual :: "('a::Inf \ 'b::Sup) \ bool" where "Inf_dual f \ f \ Inf = Sup \ (`) f" abbreviation sup_dual :: "('a::sup \ 'b::inf) \ bool" where "sup_dual f \ (\x y. f (x \ y) = f x \ f y)" abbreviation inf_dual :: "('a::inf \ 'b::sup) \ bool" where "inf_dual f \ (\x y. f (x \ y) = f x \ f y)" abbreviation bot_dual :: "('a::bot \ 'b::top) \ bool" where "bot_dual f \ f \ = \" abbreviation top_dual :: "('a::top \ 'b::bot) \ bool" where "top_dual f \ f \ = \" subsection \Duality via Locales\ sublocale ord \ dual_ord: ord "(\)" "(>)" rewrites dual_max_set: "max_set = dual_ord.min_set" and dual_filtered: "filtered = dual_ord.directed" and dual_upset_set: "upset_set = dual_ord.downset_set" and dual_upset: "upset = dual_ord.downset" and dual_upsets: "upsets = dual_ord.downsets" and dual_filters: "filters = dual_ord.ideals" apply unfold_locales unfolding max_set_def ord.min_set_def fun_eq_iff apply blast unfolding filtered_def ord.directed_def apply simp unfolding upset_set_def ord.downset_set_def apply simp apply (simp add: ord.downset_def ord.downset_set_def ord.upset_def ord.upset_set_def) unfolding upsets_def ord.downsets_def apply (metis ord.downset_set_def upset_set_def) unfolding filters_def ord.ideals_def Fix_def ord.downsets_def upsets_def ord.downset_set_def upset_set_def ord.directed_def filtered_def by simp sublocale preorder \ dual_preorder: preorder "(\)" "(>)" apply unfold_locales apply (simp add: less_le_not_le) apply simp using order_trans by blast sublocale order \ dual_order: order "(\)" "(>)" by (unfold_locales, simp) sublocale lattice \ dual_lattice: lattice sup "(\)" "(>)" inf by (unfold_locales, simp_all) sublocale bounded_lattice \ dual_bounded_lattice: bounded_lattice sup "(\)" "(>)" inf \ \ by (unfold_locales, simp_all) sublocale boolean_algebra \ dual_boolean_algebra: boolean_algebra "\x y. x \ -y" uminus sup "(\)" "(>)" inf \ \ by (unfold_locales, simp_all add: inf_sup_distrib1) sublocale complete_lattice \ dual_complete_lattice: complete_lattice Sup Inf sup "(\)" "(>)" inf \ \ rewrites dual_gfp: "gfp = dual_complete_lattice.lfp" proof- show "class.complete_lattice Sup Inf sup (\) (>) inf \ \" by (unfold_locales, simp_all add: Sup_upper Sup_least Inf_lower Inf_greatest) then interpret dual_complete_lattice: complete_lattice Sup Inf sup "(\)" "(>)" inf \ \. show "gfp = dual_complete_lattice.lfp" unfolding gfp_def dual_complete_lattice.lfp_def fun_eq_iff by simp qed context ord begin lemma dual_min_set: "min_set = dual_ord.max_set" by (simp add: dual_ord.dual_max_set) lemma dual_directed: "directed = dual_ord.filtered" by (simp add:dual_ord.dual_filtered) lemma dual_downset: "downset = dual_ord.upset" by (simp add: dual_ord.dual_upset) lemma dual_downset_set: "downset_set = dual_ord.upset_set" by (simp add: dual_ord.dual_upset_set) lemma dual_downsets: "downsets = dual_ord.upsets" by (simp add: dual_ord.dual_upsets) lemma dual_ideals: "ideals = dual_ord.filters" by (simp add: dual_ord.dual_filters) end context complete_lattice begin lemma dual_lfp: "lfp = dual_complete_lattice.gfp" by (simp add: dual_complete_lattice.dual_gfp) end subsection \Properties of Orderings, Again\ context ord begin lemma directed_nonempty: "directed X \ X \ {}" unfolding directed_def by fastforce lemma directed_ub: "directed X \ (\x \ X. \y \ X. \z \ X. x \ z \ y \ z)" by (meson empty_subsetI directed_def finite.emptyI finite_insert insert_subset order_refl) lemma downset_set_prop: "\ = Union \ (`) \" unfolding downset_set_def downset_def fun_eq_iff by fastforce lemma downset_set_prop_var: "\X = (\x \ X. \x)" by (simp add: downset_set_prop) lemma downset_prop: "\x = {y. y \ x}" unfolding downset_def downset_set_def fun_eq_iff comp_def by fastforce end context preorder begin lemma directed_prop: "X \ {} \ (\x \ X. \y \ X. \z \ X. x \ z \ y \ z) \ directed X" proof- assume h1: "X \ {}" and h2: "\x \ X. \y \ X. \z \ X. x \ z \ y \ z" {fix Y have "finite Y \ Y \ X \ (\x \ X. \y \ Y. y \ x)" proof (induct rule: finite_induct) case empty then show ?case using h1 by blast next case (insert x F) then show ?case by (metis h2 insert_iff insert_subset order_trans) qed} thus ?thesis by (simp add: directed_def) qed lemma directed_alt: "directed X = (X \ {} \ (\x \ X. \y \ X. \z \ X. x \ z \ y \ z))" by (metis directed_prop directed_nonempty directed_ub) lemma downset_set_ext: "id \ \" unfolding le_fun_def id_def downset_set_def by auto lemma downset_set_iso: "mono \" unfolding mono_def downset_set_def by blast lemma downset_set_idem [simp]: "\ \ \ = \" unfolding fun_eq_iff downset_set_def comp_def using order_trans by auto lemma downset_faithful: "\x \ \y \ x \ y" by (simp add: downset_prop subset_eq) lemma downset_iso_iff: "(\x \ \y) = (x \ y)" using atMost_iff downset_prop order_trans by blast lemma downset_directed_downset_var [simp]: "directed (\X) = directed X" proof assume h1: "directed X" {fix Y assume h2: "finite Y" and h3: "Y \ \X" hence "\y. \x. y \ Y \ x \ X \ y \ x" by (force simp: downset_set_def) hence "\f. \y. y \ Y \ f y \ X \ y \ f y" by (rule choice) hence "\f. finite (f ` Y) \ f ` Y \ X \ (\y \ Y. y \ f y)" by (metis finite_imageI h2 image_subsetI) hence "\Z. finite Z \ Z \ X \ (\y \ Y. \ z \ Z. y \ z)" by fastforce hence "\Z. finite Z \ Z \ X \ (\y \ Y. \ z \ Z. y \ z) \ (\x \ X. \ z \ Z. z \ x)" by (metis directed_def h1) hence "\x \ X. \y \ Y. y \ x" by (meson order_trans)} thus "directed (\X)" unfolding directed_def downset_set_def by fastforce next assume "directed (\X)" thus "directed X" unfolding directed_def downset_set_def apply clarsimp by (smt Ball_Collect order_refl order_trans subsetCE) qed lemma downset_directed_downset [simp]: "directed \ \ = directed" unfolding fun_eq_iff comp_def by simp lemma directed_downset_ideals: "directed (\X) = (\X \ ideals)" by (metis (mono_tags, lifting) Fix_def comp_apply directed_alt downset_set_idem downsets_def ideals_def mem_Collect_eq) end lemma downset_iso: "mono (\::'a::order \ 'a set)" by (simp add: downset_iso_iff mono_def) context order begin lemma downset_inj: "inj \" by (metis injI downset_iso_iff order.eq_iff) end context lattice begin lemma lat_ideals: "X \ ideals = (X \ {} \ X \ downsets \ (\x \ X. \ y \ X. x \ y \ X))" unfolding ideals_def directed_alt downsets_def Fix_def downset_set_def by (clarsimp, smt sup.cobounded1 sup.orderE sup.orderI sup_absorb2 sup_left_commute mem_Collect_eq) end context bounded_lattice begin lemma bot_ideal: "X \ ideals \ \ \ X" unfolding ideals_def downsets_def Fix_def downset_set_def by fastforce end context complete_lattice begin lemma Sup_downset_id [simp]: "Sup \ \ = id" using Sup_atMost atMost_def downset_prop by fastforce lemma downset_Sup_id: "id \ \ \ Sup" by (simp add: Sup_upper downset_prop le_funI subsetI) lemma Inf_Sup_var: "\(\x \ X. \x) = \X" unfolding downset_prop by (simp add: Collect_ball_eq Inf_eq_Sup) lemma Inf_pres_downset_var: "(\x \ X. \x) = \(\X)" unfolding downset_prop by (safe, simp_all add: le_Inf_iff) end lemma lfp_in_Fix: fixes f :: "'a::complete_lattice \ 'a" shows "mono f \ lfp f \ Fix f" using Fix_def lfp_unfold by fastforce lemma gfp_in_Fix: fixes f :: "'a::complete_lattice \ 'a" shows "mono f \ gfp f \ Fix f" using Fix_def gfp_unfold by fastforce lemma nonempty_Fix: fixes f :: "'a::complete_lattice \ 'a" shows "mono f \ Fix f \ {}" using lfp_in_Fix by fastforce subsection \Dual Properties of Orderings from Locales\ text \These properties can be proved very smoothly overall. But only within the context of a class or locale!\ context ord begin lemma filtered_nonempty: "filtered X \ X \ {}" by (simp add: dual_filtered dual_ord.directed_nonempty) lemma filtered_lb: "filtered X \ (\x \ X. \y \ X. \z \ X. z \ x \ z \ y)" by (simp add: dual_filtered dual_ord.directed_ub) lemma upset_set_prop: "\ = Union \ (`) \" by (simp add: dual_ord.downset_set_prop dual_upset dual_upset_set) lemma upset_set_prop_var: "\X = (\x \ X. \x)" by (simp add: dual_ord.downset_set_prop_var dual_upset dual_upset_set) lemma upset_prop: "\x = {y. x \ y}" by (simp add: dual_ord.downset_prop dual_upset) end context preorder begin lemma filtered_prop: "X \ {} \ (\x \ X. \y \ X. \z \ X. z \ x \ z \ y) \ filtered X" by (simp add: dual_filtered dual_preorder.directed_prop) lemma filtered_alt: "filtered X = (X \ {} \ (\x \ X. \y \ X. \z \ X. z \ x \ z \ y))" by (simp add: dual_filtered dual_preorder.directed_alt) lemma upset_set_ext: "id \ \" by (simp add: dual_preorder.downset_set_ext dual_upset_set) lemma upset_set_anti: "mono \" by (simp add: dual_preorder.downset_set_iso dual_upset_set) lemma up_set_idem [simp]: "\ \ \ = \" by (simp add: dual_upset_set) lemma upset_faithful: "\x \ \y \ y \ x" by (metis dual_preorder.downset_faithful dual_upset) lemma upset_anti_iff: "(\y \ \x) = (x \ y)" by (simp add: dual_preorder.downset_iso_iff dual_upset) lemma upset_filtered_upset [simp]: "filtered \ \ = filtered" by (simp add: dual_filtered dual_upset_set) lemma filtered_upset_filters: "filtered (\X) = (\X \ filters)" using dual_filtered dual_preorder.directed_downset_ideals dual_upset_set ord.dual_filters by fastforce end context order begin lemma upset_inj: "inj \" by (simp add: dual_order.downset_inj dual_upset) end context lattice begin lemma lat_filters: "X \ filters = (X \ {} \ X \ upsets \ (\x \ X. \ y \ X. x \ y \ X))" by (simp add: dual_filters dual_lattice.lat_ideals dual_ord.downsets_def dual_upset_set upsets_def) end context bounded_lattice begin lemma top_filter: "X \ filters \ \ \ X" by (simp add: dual_bounded_lattice.bot_ideal dual_filters) end context complete_lattice begin lemma Inf_upset_id [simp]: "Inf \ \ = id" by (simp add: dual_upset) lemma upset_Inf_id: "id \ \ \ Inf" by (simp add: dual_complete_lattice.downset_Sup_id dual_upset) lemma Sup_Inf_var: " \(\x \ X. \x) = \X" by (simp add: dual_complete_lattice.Inf_Sup_var dual_upset) lemma Sup_dual_upset_var: "(\x \ X. \x) = \(\X)" by (simp add: dual_complete_lattice.Inf_pres_downset_var dual_upset) end subsection \Examples that Do Not Dualise\ lemma upset_anti: "antimono (\::'a::order \ 'a set)" by (simp add: antimono_def upset_anti_iff) context complete_lattice begin lemma fSup_unfold: "(f::nat \ 'a) 0 \ (\n. f (Suc n)) = (\n. f n)" apply (intro order.antisym sup_least) apply (rule Sup_upper, force) apply (rule Sup_mono, force) apply (safe intro!: Sup_least) by (case_tac n, simp_all add: Sup_upper le_supI2) lemma fInf_unfold: "(f::nat \ 'a) 0 \ (\n. f (Suc n)) = (\n. f n)" apply (intro order.antisym inf_greatest) apply (rule Inf_greatest, safe) apply (case_tac n) apply simp_all using Inf_lower inf.coboundedI2 apply force apply (simp add: Inf_lower) by (auto intro: Inf_mono) end lemma fun_isol: "mono f \ mono ((\) f)" by (simp add: le_fun_def mono_def) lemma fun_isor: "mono f \ mono (\x. x \ f)" by (simp add: le_fun_def mono_def) lemma Sup_sup_pres: fixes f :: "'a::complete_lattice \ 'b::complete_lattice" shows "Sup_pres f \ sup_pres f" by (metis (no_types, opaque_lifting) Sup_empty Sup_insert comp_apply image_insert sup_bot.right_neutral) lemma Inf_inf_pres: fixes f :: "'a::complete_lattice \ 'b::complete_lattice" shows"Inf_pres f \ inf_pres f" by (smt INF_insert comp_eq_elim dual_complete_lattice.Sup_empty dual_complete_lattice.Sup_insert inf_top.right_neutral) lemma Sup_bot_pres: fixes f :: "'a::complete_lattice \ 'b::complete_lattice" shows "Sup_pres f \ bot_pres f" by (metis SUP_empty Sup_empty comp_eq_elim) lemma Inf_top_pres: fixes f :: "'a::complete_lattice \ 'b::complete_lattice" shows "Inf_pres f \ top_pres f" by (metis INF_empty comp_eq_elim dual_complete_lattice.Sup_empty) context complete_lattice begin lemma iso_Inf_subdistl: assumes "mono (f::'a \ 'b::complete_lattice)" shows "f \ Inf \ Inf \ (`) f" by (simp add: assms complete_lattice_class.le_Inf_iff le_funI Inf_lower monoD) lemma iso_Sup_supdistl: assumes "mono (f::'a \ 'b::complete_lattice)" shows "Sup \ (`) f \ f \ Sup" by (simp add: assms complete_lattice_class.SUP_le_iff le_funI dual_complete_lattice.Inf_lower monoD) lemma Inf_subdistl_iso: fixes f :: "'a \ 'b::complete_lattice" shows "f \ Inf \ Inf \ (`) f \ mono f" unfolding mono_def le_fun_def comp_def by (metis complete_lattice_class.le_INF_iff Inf_atLeast atLeast_iff) lemma Sup_supdistl_iso: fixes f :: "'a \ 'b::complete_lattice" shows "Sup \ (`) f \ f \ Sup \ mono f" unfolding mono_def le_fun_def comp_def by (metis complete_lattice_class.SUP_le_iff Sup_atMost atMost_iff) lemma supdistl_iso: fixes f :: "'a \ 'b::complete_lattice" shows "(Sup \ (`) f \ f \ Sup) = mono f" using Sup_supdistl_iso iso_Sup_supdistl by force lemma subdistl_iso: fixes f :: "'a \ 'b::complete_lattice" shows "(f \ Inf \ Inf \ (`) f) = mono f" using Inf_subdistl_iso iso_Inf_subdistl by force end lemma fSup_distr: "Sup_pres (\x. x \ f)" unfolding fun_eq_iff comp_def by (smt Inf.INF_cong SUP_apply Sup_apply) lemma fSup_distr_var: "\F \ g = (\f \ F. f \ g)" unfolding fun_eq_iff comp_def by (smt Inf.INF_cong SUP_apply Sup_apply) lemma fInf_distr: "Inf_pres (\x. x \ f)" unfolding fun_eq_iff comp_def by (smt INF_apply Inf.INF_cong Inf_apply) lemma fInf_distr_var: "\F \ g = (\f \ F. f \ g)" unfolding fun_eq_iff comp_def by (smt INF_apply Inf.INF_cong Inf_apply) lemma fSup_subdistl: assumes "mono (f::'a::complete_lattice \ 'b::complete_lattice)" shows "Sup \ (`) ((\) f) \ (\) f \ Sup" using assms by (simp add: SUP_least Sup_upper le_fun_def monoD image_comp) lemma fSup_subdistl_var: fixes f :: "'a::complete_lattice \ 'b::complete_lattice" shows "mono f \ (\g \ G. f \ g) \ f \ \G" by (simp add: SUP_least Sup_upper le_fun_def monoD image_comp) lemma fInf_subdistl: fixes f :: "'a::complete_lattice \ 'b::complete_lattice" shows "mono f \ (\) f \ Inf \ Inf \ (`) ((\) f)" by (simp add: INF_greatest Inf_lower le_fun_def monoD image_comp) lemma fInf_subdistl_var: fixes f :: "'a::complete_lattice \ 'b::complete_lattice" shows "mono f \ f \ \G \ (\g \ G. f \ g)" by (simp add: INF_greatest Inf_lower le_fun_def monoD image_comp) lemma Inf_pres_downset: "Inf_pres (\::'a::complete_lattice \ 'a set)" unfolding downset_prop fun_eq_iff comp_def by (safe, simp_all add: le_Inf_iff) lemma Sup_dual_upset: "Sup_dual (\::'a::complete_lattice \ 'a set)" unfolding upset_prop fun_eq_iff comp_def by (safe, simp_all add: Sup_le_iff) text \This approach could probably be combined with the explicit functor-based one. This may be good for proofs, but seems conceptually rather ugly.\ end \ No newline at end of file diff --git a/thys/Order_Lattice_Props/Order_Lattice_Props_Wenzel.thy b/thys/Order_Lattice_Props/Order_Lattice_Props_Wenzel.thy --- a/thys/Order_Lattice_Props/Order_Lattice_Props_Wenzel.thy +++ b/thys/Order_Lattice_Props/Order_Lattice_Props_Wenzel.thy @@ -1,280 +1,280 @@ (* Title: Duality Based on a Data Type Author: Georg Struth Maintainer:Georg Struth *) section \Duality Based on a Data Type\ theory Order_Lattice_Props_Wenzel imports Main - "HOL-Library.Lattice_Syntax" +begin -begin +unbundle lattice_syntax subsection \Wenzel's Approach Revisited\ text \This approach is similar to, but inferior to the explicit class-based one. The main caveat is that duality is not involutive with this approach, and this allows dualising less theorems.\ text \I copy Wenzel's development \cite{Wenzel} in this subsection and extend it with additional properties. I show only the most important properties.\ datatype 'a dual = dual (un_dual: 'a) ("\") notation un_dual ("\\<^sup>-") lemma dual_inj: "inj \" using injI by fastforce lemma dual_surj: "surj \" using dual.exhaust_sel by blast lemma dual_bij: "bij \" by (simp add: bijI dual_inj dual_surj) text \Dual is not idempotent, and I see no way of imposing this condition. Yet at least an inverse exists --- namely un-dual..\ lemma dual_inv1 [simp]: "\\<^sup>- \ \ = id" by fastforce lemma dual_inv2 [simp]: "\ \ \\<^sup>- = id" by fastforce lemma dual_inv_inj: "inj \\<^sup>-" by (simp add: dual.expand injI) lemma dual_inv_surj: "surj \\<^sup>-" by (metis dual.sel surj_def) lemma dual_inv_bij: "bij \\<^sup>-" by (simp add: bij_def dual_inv_inj dual_inv_surj) lemma dual_iff: "(\ x = y) \ (x = \\<^sup>- y)" by fastforce text \Isabelle data types come with a number of generic functions.\ text \The functor map-dual lifts functions to dual types. Isabelle's generic definition is not straightforward to understand and use. Yet conceptually it can be explained as follows.\ lemma map_dual_def_var [simp]: "(map_dual::('a \ 'b) \ 'a dual \ 'b dual) f = \ \ f \ \\<^sup>-" unfolding fun_eq_iff comp_def by (metis dual.map_sel dual_iff) lemma map_dual_def_var2: "\\<^sup>- \ map_dual f = f \ \\<^sup>-" by (simp add: rewriteL_comp_comp) lemma map_dual_func1: "map_dual (f \ g) = map_dual f \ map_dual g" unfolding fun_eq_iff comp_def by (metis dual.exhaust dual.map) lemma map_dual_func2 : "map_dual id = id" by simp text \The functor map-dual has an inverse functor as well.\ definition map_dual_inv :: "('a dual \ 'b dual) => ('a => 'b)" where "map_dual_inv f = \\<^sup>- \ f \ \" lemma map_dual_inv_func1: "map_dual_inv id = id" by (simp add: map_dual_inv_def) lemma map_dual_inv_func2: "map_dual_inv (f \ g) = map_dual_inv f \ map_dual_inv g" unfolding fun_eq_iff comp_def map_dual_inv_def by (metis dual_iff) lemma map_dual_inv1: "map_dual \ map_dual_inv = id" unfolding fun_eq_iff map_dual_def_var map_dual_inv_def comp_def id_def by (metis dual_iff) lemma map_dual_inv2: "map_dual_inv \ map_dual = id" unfolding fun_eq_iff map_dual_def_var map_dual_inv_def comp_def id_def by (metis dual_iff) text \Hence dual is an isomorphism between categories.\ lemma subset_dual: "(\ ` X = Y) \ (X = \\<^sup>- ` Y)" by (metis dual_inj image_comp image_inv_f_f inv_o_cancel dual_inv2) lemma subset_dual1: "(X \ Y) \ (\ ` X \ \ ` Y)" by (simp add: dual_inj inj_image_subset_iff) lemma dual_ball: "(\x \ X. P (\ x)) \ (\y \ \ ` X. P y)" by simp lemma dual_inv_ball: "(\x \ X. P (\\<^sup>- x)) \ (\y \ \\<^sup>- ` X. P y)" by simp lemma dual_all: "(\x. P (\ x)) \ (\y. P y)" by (metis dual.collapse) lemma dual_inv_all: "(\x. P (\\<^sup>- x)) \ (\y. P y)" by (metis dual_inv_surj surj_def) lemma dual_ex: "(\x. P (\ x)) \ (\y. P y)" by (metis UNIV_I bex_imageD dual_surj) lemma dual_inv_ex: "(\x. P (\\<^sup>- x)) \ (\y. P y)" by (metis dual.sel) lemma dual_Collect: "{\ x |x. P (\ x)} = {y. P y}" by (metis dual.exhaust) lemma dual_inv_Collect: "{\\<^sup>- x |x. P (\\<^sup>- x)} = {y. P y}" by (metis dual.collapse dual.inject) lemma fun_dual1: "(f \ \ = g) \ (f = g \ \\<^sup>-)" by auto lemma fun_dual2: "(\ \ f = g) \ (f = \\<^sup>- \ g)" by auto lemma fun_dual3: "(f \ (`) \ = g) \ (f = g \ (`) \\<^sup>-)" unfolding fun_eq_iff comp_def by (metis subset_dual) lemma fun_dual4: "(f = \\<^sup>- \ g \ (`) \) \ (\ \ f \ (`) \\<^sup>- = g)" by (metis fun_dual2 fun_dual3 o_assoc) text \The next facts show incrementally that the dual of a complete lattice is a complete lattice. This follows once again Wenzel.\ instantiation dual :: (ord) ord begin definition less_eq_dual_def: "(\) = rel_dual (\)" definition less_dual_def: "(<) = rel_dual (>)" instance.. end lemma less_eq_dual_def_var: "(x \ y) = (\\<^sup>- y \ \\<^sup>- x)" apply (rule antisym) apply (simp add: dual.rel_sel less_eq_dual_def) by (simp add: dual.rel_sel less_eq_dual_def) lemma less_dual_def_var: "(x < y) = (\\<^sup>- y < \\<^sup>- x)" by (simp add: dual.rel_sel less_dual_def) instance dual :: (preorder) preorder apply standard apply (simp add: less_dual_def_var less_eq_dual_def_var less_le_not_le) apply (simp add: less_eq_dual_def_var) by (meson less_eq_dual_def_var order_trans) instance dual :: (order) order by (standard, simp add: dual.expand less_eq_dual_def_var) lemma dual_anti: "x \ y \ \ y \ \ x" by (simp add: dual_inj less_eq_dual_def the_inv_f_f) lemma dual_anti_iff: "(x \ y) = (\ y \ \ x)" by (simp add: dual_inj less_eq_dual_def the_inv_f_f) text \map-dual does not map isotone functions to antitone ones. It simply lifts the type!\ lemma "mono f \ mono (map_dual f)" unfolding map_dual_def_var mono_def by (metis comp_apply dual_anti less_eq_dual_def_var) instantiation dual :: (lattice) lattice begin definition inf_dual_def: "x \ y = \ (\\<^sup>- x \ \\<^sup>- y)" definition sup_dual_def: "x \ y = \ (\\<^sup>- x \ \\<^sup>- y)" instance by (standard, simp_all add: dual_inj inf_dual_def sup_dual_def less_eq_dual_def_var the_inv_f_f) end instantiation dual :: (complete_lattice) complete_lattice begin definition Inf_dual_def: "Inf = \ \ Sup \ (`) \\<^sup>-" definition Sup_dual_def: "Sup = \ \ Inf \ (`) \\<^sup>-" definition bot_dual_def: "\ = \ \" definition top_dual_def: "\ = \ \" instance by (standard, simp_all add: Inf_dual_def top_dual_def Sup_dual_def bot_dual_def dual_inj le_INF_iff SUP_le_iff INF_lower SUP_upper less_eq_dual_def_var the_inv_f_f) end text \Next, directed and filtered sets, upsets, downsets, filters and ideals in posets are defined.\ context ord begin definition directed :: "'a set \ bool" where "directed X = (\Y. finite Y \ Y \ X \ (\x \ X. \y \ Y. y \ x))" definition filtered :: "'a set \ bool" where "filtered X = (\Y. finite Y \ Y \ X \ (\x \ X. \y \ Y. x \ y))" definition downset_set :: "'a set \ 'a set" ("\") where "\X = {y. \x \ X. y \ x}" definition upset_set :: "'a set \ 'a set" ("\") where "\X = {y. \x \ X. x \ y}" end subsection \Examples that Do Not Dualise\ text \Filtered and directed sets are dual.\ text \Proofs could be simplified if dual was idempotent.\ lemma filtered_directed_dual: "filtered \ (`) \ = directed" proof- {fix X::"'a set" have "(filtered \ (`) \) X = (\Y. finite (\\<^sup>- ` Y) \ \\<^sup>- ` Y \ X \ (\x \ X.\y \ (\\<^sup>- ` Y). \ x \ \ y))" unfolding filtered_def comp_def by (simp, metis dual_iff finite_subset_image subset_dual subset_dual1) also have "... = (\Y. finite Y \ Y \ X \ (\x \ X.\y \ Y. y \ x))" by (metis dual_anti_iff dual_inv_surj finite_subset_image top.extremum) finally have "(filtered \ (`) \) X = directed X" using directed_def by auto} thus ?thesis unfolding fun_eq_iff by simp qed lemma directed_filtered_dual: "directed \ (`) \ = filtered" proof- {fix X::"'a set" have "(directed \ (`) \) X = (\Y. finite (\\<^sup>- ` Y) \ \\<^sup>- ` Y \ X \ (\x \ X.\y \ (\\<^sup>- ` Y). \ y \ \ x))" unfolding directed_def comp_def by (simp, metis dual_iff finite_subset_image subset_dual subset_dual1) also have "... = (\Y. finite Y \ Y \ X \ (\x \ X.\y \ Y. x \ y))" unfolding dual_anti_iff[symmetric] by (metis dual_inv_surj finite_subset_image top_greatest) finally have "(directed \ (`) \) X = filtered X" using filtered_def by auto} thus ?thesis unfolding fun_eq_iff by simp qed text \This example illustrates the deficiency of the approach. In the class-based approach the second proof is trivial.\ text \The next example shows that this is a systematic problem.\ lemma downset_set_upset_set_dual: "(`) \ \ \ = \ \ (`) \" proof- {fix X::"'a set" have "((`) \ \ \) X = {\ y |y. \x \ X. y \ x}" by (simp add: downset_set_def setcompr_eq_image) also have "... = {\ y |y. \x \ X. \ x \ \ y}" by (meson dual_anti_iff) also have "... = {y. \x \ \ ` X. x \ y}" by (metis (mono_tags, opaque_lifting) dual.exhaust image_iff) finally have "((`) \ \ \) X = (\ \ (`) \) X" by (simp add: upset_set_def)} thus ?thesis unfolding fun_eq_iff by simp qed lemma upset_set_downset_set_dual: "(`) \ \ \ = \ \ (`) \" unfolding downset_set_def upset_set_def fun_eq_iff comp_def apply (safe, force simp: dual_anti) by (metis (mono_tags, lifting) dual.exhaust dual_anti_iff mem_Collect_eq rev_image_eqI) end diff --git a/thys/Order_Lattice_Props/Sup_Lattice.thy b/thys/Order_Lattice_Props/Sup_Lattice.thy --- a/thys/Order_Lattice_Props/Sup_Lattice.thy +++ b/thys/Order_Lattice_Props/Sup_Lattice.thy @@ -1,162 +1,162 @@ (* Title: Sup-Lattices and Other Simplifications Author: Georg Struth Maintainer: Georg Struth *) section \Sup-Lattices and Other Simplifications\ theory Sup_Lattice - imports Main - "HOL-Library.Lattice_Syntax" + imports Main +begin -begin +unbundle lattice_syntax text \Some definitions for orderings and lattices in Isabelle could be simpler. The strict order in in ord could be defined instead of being axiomatised. The function mono could have been defined on ord and not on order---even on a general (di)graph it serves as a morphism. In complete lattices, the supremum---and dually the infimum---suffices to define the other operations (in the Isabelle/HOL-definition infimum, binary supremum and infimum, bottom and top element are axiomatised). This not only increases the number of proof obligations in subclass or sublocale statements, instantiations or interpretations, it also complicates situations where suprema are presented faithfully, e.g. mapped onto suprema in some subalgebra, whereas infima in the subalgebra are different from those in the super-structure.\ text \It would be even nicer to use a class less-eq which dispenses with the strict order symbol in ord. Then one would not have to redefine this symbol in all instantiations or interpretations. At least, it does not carry any proof obligations.\ context ord begin text \ub-set yields the set of all upper bounds of a set; lb-set the set of all lower bounds.\ definition ub_set :: "'a set \ 'a set" where "ub_set X = {y. \x \ X. x \ y}" definition lb_set :: "'a set \ 'a set" where "lb_set X = {y. \x \ X. y \ x}" end definition ord_pres :: "('a::ord \ 'b::ord) \ bool" where "ord_pres f = (\x y. x \ y \ f x \ f y)" lemma ord_pres_mono: fixes f :: "'a::order \ 'b::order" shows "mono f = ord_pres f" by (simp add: mono_def ord_pres_def) class preorder_lean = ord + assumes preorder_refl: "x \ x" and preorder_trans: "x \ y \ y \ z \ x \ z" begin definition le :: "'a \ 'a \ bool" where "le x y = (x \ y \ \ (x \ y))" end sublocale preorder_lean \ prel: preorder "(\)" le by (unfold_locales, auto simp add: le_def preorder_refl preorder_trans) class order_lean = preorder_lean + assumes order_antisym: "x \ y \ x \ y \ x = y" sublocale order_lean \ posl: order "(\)" le by (unfold_locales, simp add: order_antisym) class Sup_lattice = order_lean + Sup + assumes Sups_upper: "x \ X \ x \ \X" and Sups_least: "(\x. x \ X \ x \ z) \ \X \ z" begin definition Infs :: "'a set \ 'a" where "Infs X = \{y. \x \ X. y \ x}" definition sups :: "'a \ 'a \ 'a" where "sups x y = \{x,y}" definition infs :: "'a \ 'a \ 'a" where "infs x y = Infs{x,y}" definition bots :: 'a where "bots = \{}" definition tops :: 'a where "tops = Infs{}" lemma Infs_prop: "Infs = Sup \ lb_set" unfolding fun_eq_iff by (simp add: Infs_def prel.lb_set_def) end class Inf_lattice = order_lean + Inf + assumes Infi_lower: "x \ X \ \X \ x" and Infi_greatest: "(\x. x \ X \ z \ x) \ z \ \X" begin definition Supi :: "'a set \ 'a" where "Supi X = \{y. \x \ X. x \ y}" definition supi :: "'a \ 'a \ 'a" where "supi x y = Supi{x,y}" definition infi :: "'a \ 'a \ 'a" where "infi x y = \{x,y}" definition boti :: 'a where "boti = Supi{}" definition topi :: 'a where "topi = \{}" lemma Supi_prop: "Supi = Inf \ ub_set" unfolding fun_eq_iff by (simp add: Supi_def prel.ub_set_def) end sublocale Inf_lattice \ ldual: Sup_lattice Inf "(\)" rewrites "ldual.Infs = Supi" and "ldual.infs = supi" and "ldual.sups = infi" and "ldual.tops = boti" and "ldual.bots = topi" proof- show "class.Sup_lattice Inf (\)" by (unfold_locales, simp_all add: Infi_lower Infi_greatest preorder_trans) then interpret ldual: Sup_lattice Inf "(\)". show a: "ldual.Infs = Supi" unfolding fun_eq_iff by (simp add: ldual.Infs_def Supi_def) show "ldual.infs = supi" unfolding fun_eq_iff by (simp add: a ldual.infs_def supi_def) show "ldual.sups = infi" unfolding fun_eq_iff by (simp add: ldual.sups_def infi_def) show "ldual.tops = boti" by (simp add: a ldual.tops_def boti_def) show "ldual.bots = topi" by (simp add: ldual.bots_def topi_def) qed sublocale Sup_lattice \ supclat: complete_lattice Infs Sup_class.Sup infs "(\)" le sups bots tops apply unfold_locales unfolding Infs_def infs_def sups_def bots_def tops_def by (simp_all, auto intro: Sups_least, simp_all add: Sups_upper) sublocale Inf_lattice \ infclat: complete_lattice Inf_class.Inf Supi infi "(\)" le supi boti topi by (unfold_locales, simp_all add: ldual.Sups_upper ldual.Sups_least ldual.supclat.Inf_lower ldual.supclat.Inf_greatest) end diff --git a/thys/Polynomial_Factorization/ROOT b/thys/Polynomial_Factorization/ROOT --- a/thys/Polynomial_Factorization/ROOT +++ b/thys/Polynomial_Factorization/ROOT @@ -1,72 +1,71 @@ chapter AFP session "JNF-AFP-Lib" (AFP) in "Lib" = "HOL-Algebra" + description "Theories from HOL-Library and the Archive of Formal Proofs that are used by this entry" options [timeout = 600] sessions Containers "Abstract-Rewriting" Gauss_Jordan Matrix Polynomial_Interpolation Show VectorSpace "HOL-Cardinals" theories Containers.Set_Impl Matrix.Utility Matrix.Ordered_Semiring "Abstract-Rewriting.SN_Order_Carrier" "Abstract-Rewriting.Relative_Rewriting" Show.Show_Instances VectorSpace.VectorSpace Polynomial_Interpolation.Missing_Polynomial Polynomial_Interpolation.Ring_Hom_Poly "HOL-Library.AList" "HOL-Library.Cardinality" "HOL-Library.Char_ord" "HOL-Library.Code_Binary_Nat" "HOL-Library.Code_Target_Numeral" "HOL-Library.DAList" "HOL-Library.DAList_Multiset" "HOL-Library.Infinite_Set" - "HOL-Library.Lattice_Syntax" "HOL-Library.Mapping" "HOL-Library.Monad_Syntax" "HOL-Library.More_List" "HOL-Library.Multiset" "HOL-Library.IArray" "HOL-Library.Phantom_Type" "HOL-Library.Ramsey" "HOL-Library.RBT_Impl" "HOL-Library.Simps_Case_Conv" "HOL-Library.While_Combinator" "HOL-Computational_Algebra.Fundamental_Theorem_Algebra" "HOL-Computational_Algebra.Fraction_Field" "HOL-Computational_Algebra.Polynomial" "HOL-Computational_Algebra.Primes" "HOL-Cardinals.Order_Union" "HOL-Cardinals.Wellorder_Extension" session Polynomial_Factorization (AFP) = "JNF-AFP-Lib" + description "Algebraic Numbers" options [timeout = 600] sessions Partial_Function_MR Polynomial_Interpolation Show Sqrt_Babylonian theories Missing_Multiset Missing_List Precomputation Order_Polynomial Explicit_Roots Dvd_Int_Poly Rational_Root_Test Kronecker_Factorization Square_Free_Factorization Rational_Factorization document_files "root.bib" "root.tex" diff --git a/thys/Residuated_Lattices/Residuated_Boolean_Algebras.thy b/thys/Residuated_Lattices/Residuated_Boolean_Algebras.thy --- a/thys/Residuated_Lattices/Residuated_Boolean_Algebras.thy +++ b/thys/Residuated_Lattices/Residuated_Boolean_Algebras.thy @@ -1,665 +1,667 @@ (* Title: Residuated Boolean Algebras Author: Victor Gomes Maintainer: Georg Struth *) section \Residuated Boolean Algebras\ theory Residuated_Boolean_Algebras imports Residuated_Lattices begin +unbundle lattice_syntax + subsection \Conjugation on Boolean Algebras\ text \ Similarly, as in the previous section, we define the conjugation for arbitrary residuated functions on boolean algebras. \ context boolean_algebra begin lemma inf_bot_iff_le: "x \ y = \ \ x \ -y" by (metis le_iff_inf inf_sup_distrib1 inf_top_right sup_bot.left_neutral sup_compl_top compl_inf_bot inf.assoc inf_bot_right) lemma le_iff_inf_bot: "x \ y \ x \ -y = \" by (metis inf_bot_iff_le compl_le_compl_iff inf_commute) lemma indirect_eq: "(\z. x \ z \ y \ z) \ x = y" by (metis order.eq_iff) text \ Let $B$ be a boolean algebra. The maps $f$ and $g$ on $B$ are a pair of conjugates if and only if for all $x, y \in B$, $f(x) \sqcap y = \bot \Leftrightarrow x \sqcap g(t) = \bot$. \ definition conjugation_pair :: "('a \ 'a) \ ('a \ 'a) \ bool" where "conjugation_pair f g \ \x y. f(x) \ y = \ \ x \ g(y) = \" lemma conjugation_pair_commute: "conjugation_pair f g \ conjugation_pair g f" by (auto simp: conjugation_pair_def inf_commute) lemma conjugate_iff_residuated: "conjugation_pair f g = residuated_pair f (\x. -g(-x))" apply (clarsimp simp: conjugation_pair_def residuated_pair_def inf_bot_iff_le) by (metis double_compl) lemma conjugate_residuated: "conjugation_pair f g \ residuated_pair f (\x. -g(-x))" by (metis conjugate_iff_residuated) lemma residuated_iff_conjugate: "residuated_pair f g = conjugation_pair f (\x. -g(-x))" apply (clarsimp simp: conjugation_pair_def residuated_pair_def inf_bot_iff_le) by (metis double_compl) text \ A map $f$ has a conjugate pair if and only if it is residuated. \ lemma conj_residuatedI1: "\g. conjugation_pair f g \ residuated f" by (metis conjugate_iff_residuated residuated_def) lemma conj_residuatedI2: "\g. conjugation_pair g f \ residuated f" by (metis conj_residuatedI1 conjugation_pair_commute) lemma exist_conjugateI[intro]: "residuated f \ \g. conjugation_pair f g" by (metis residuated_def residuated_iff_conjugate) lemma exist_conjugateI2[intro]: "residuated f \ \g. conjugation_pair g f" by (metis exist_conjugateI conjugation_pair_commute) text \ The conjugate of a residuated function $f$ is unique. \ lemma unique_conjugate[intro]: "residuated f \ \!g. conjugation_pair f g" proof - { fix g h x assume "conjugation_pair f g" and "conjugation_pair f h" hence "g = h" apply (unfold conjugation_pair_def) apply (rule ext) apply (rule order.antisym) by (metis le_iff_inf_bot inf_commute inf_compl_bot)+ } moreover assume "residuated f" ultimately show ?thesis by force qed lemma unique_conjugate2[intro]: "residuated f \ \!g. conjugation_pair g f" by (metis unique_conjugate conjugation_pair_commute) text \ Since the conjugate of a residuated map is unique, we define a conjugate operation. \ definition conjugate :: "('a \ 'a) \ ('a \ 'a)" where "conjugate f \ THE g. conjugation_pair g f" lemma conjugate_iff_def: "residuated f \ f(x) \ y = \ \ x \ conjugate f y = \" apply (clarsimp simp: conjugate_def dest!: unique_conjugate) apply (subgoal_tac "(THE g. conjugation_pair g f) = g") apply (clarsimp simp add: conjugation_pair_def) apply (rule the1_equality) by (auto intro: conjugation_pair_commute) lemma conjugateI1: "residuated f \ f(x) \ y = \ \ x \ conjugate f y = \" by (metis conjugate_iff_def) lemma conjugateI2: "residuated f \ x \ conjugate f y = \ \ f(x) \ y = \" by (metis conjugate_iff_def) text \ Few more lemmas about conjugation follow. \ lemma residuated_conj1: "residuated f \ conjugation_pair f (conjugate f)" using conjugateI1 conjugateI2 conjugation_pair_def by auto lemma residuated_conj2: "residuated f \ conjugation_pair (conjugate f) f" using conjugateI1 conjugateI2 conjugation_pair_def inf_commute by auto lemma conj_residuated: "residuated f \ residuated (conjugate f)" by (force dest!: residuated_conj2 intro: conj_residuatedI1) lemma conj_involution: "residuated f \ conjugate (conjugate f) = f" by (metis conj_residuated residuated_conj1 residuated_conj2 unique_conjugate) lemma residual_conj_eq: "residuated f \ residual (conjugate f) = (\x. -f(-x))" apply (unfold residual_def) apply (rule the1_equality) apply (rule residual_unique) apply (auto intro: conj_residuated conjugate_residuated residuated_conj2) done lemma residual_conj_eq_ext: "residuated f \ residual (conjugate f) x = -f(-x)" by (metis residual_conj_eq) lemma conj_iso: "residuated f \ x \ y \ conjugate f x \ conjugate f y" by (metis conj_residuated res_iso) lemma conjugate_strict: "residuated f \ conjugate f \ = \" by (metis conj_residuated residuated_strict) lemma conjugate_sup: "residuated f \ conjugate f (x \ y) = conjugate f x \ conjugate f y" by (metis conj_residuated residuated_sup) lemma conjugate_subinf: "residuated f \ conjugate f (x \ y) \ conjugate f x \ conjugate f y" by (auto simp: conj_iso) text \ Next we prove some lemmas from Maddux's article. Similar lemmas have been proved in AFP entry for relation algebras. They should be consolidated in the future. \ lemma maddux1: "residuated f \ f(x \ - conjugate f(y)) \ f(x) \ -y" proof - assume assm: "residuated f" hence "f(x \ - conjugate f(y)) \ f x" by (metis inf_le1 res_iso) moreover have "f(x \ - conjugate f (y)) \ y = \" by (metis assm conjugateI2 inf_bot_iff_le inf_le2) ultimately show ?thesis by (metis inf_bot_iff_le le_inf_iff) qed lemma maddux1': "residuated f \ conjugate f(x \ -f(y)) \ conjugate f(x) \ -y" by (metis conj_involution conj_residuated maddux1) lemma maddux2: "residuated f \ f(x) \ y \ f(x \ conjugate f y)" proof - assume resf: "residuated f" obtain z where z_def: "z = f(x \ conjugate f y)" by auto hence "f(x \ conjugate f y) \ -z = \" by (metis inf_compl_bot) hence "x \ conjugate f y \ conjugate f (-z) = \" by (metis conjugate_iff_def resf) hence "x \ conjugate f (y \ -z) = \" apply (subgoal_tac "conjugate f (y \ -z) \ conjugate f y \ conjugate f (-z)") apply (metis (no_types, opaque_lifting) dual_order.trans inf.commute inf_bot_iff_le inf_left_commute) by (metis conj_iso inf_le2 inf_top.left_neutral le_inf_iff resf) hence "f(x) \ y \ -z = \" by (metis conjugateI2 inf.assoc resf) thus ?thesis by (metis double_compl inf_bot_iff_le z_def) qed lemma maddux2': "residuated f \ conjugate f(x) \ y \ conjugate f(x \ f y)" by (metis conj_involution conj_residuated maddux2) lemma residuated_conjugate_ineq: "residuated f \ conjugate f x \ y \ x \ -f(-y)" by (metis conj_residuated residual_galois residual_conj_eq) lemma residuated_comp_closed: "residuated f \ residuated g \ residuated (f o g)" by (auto simp add: residuated_def residuated_pair_def) lemma conjugate_comp: "residuated f \ residuated g \ conjugate (f o g) = conjugate g o conjugate f" proof (rule ext, rule indirect_eq) fix x y assume assms: "residuated f" "residuated g" have "conjugate (f o g) x \ y \ x \ -f(g(-y))" apply (subst residuated_conjugate_ineq) using assms by (auto intro!: residuated_comp_closed) also have "... \ conjugate g (conjugate f x) \ y" using assms by (simp add: residuated_conjugate_ineq) finally show "(conjugate (f \ g) x \ y) = ((conjugate g \ conjugate f) x \ y)" by auto qed lemma conjugate_comp_ext: "residuated f \ residuated g \ conjugate (\x. f (g x)) x = conjugate g (conjugate f x)" using conjugate_comp by (simp add: comp_def) end (* boolean_algebra *) context complete_boolean_algebra begin text \ On a complete boolean algebra, it is possible to give an explicit definition of conjugation. \ lemma conjugate_eq: "residuated f \ conjugate f y = \{x. y \ -f(-x)}" proof - assume assm: "residuated f" obtain g where g_def: "g = conjugate f" by auto have "g y = \{x. x \ g y}" by (auto intro!: order.antisym Inf_lower Inf_greatest) also have "... = \{x. -x \ g y = \}" by (simp add: inf_bot_iff_le) also have "... = \{x. f(-x) \ y = \}" by (metis conjugate_iff_def assm g_def) finally show ?thesis by (simp add: g_def le_iff_inf_bot inf_commute) qed end (* complete_boolean_algebra *) subsection \Residuated Boolean Structures\ text \ In this section, we present various residuated structures based on boolean algebras. The left and right conjugation of the multiplicative operation is defined, and a number of facts is derived. \ class residuated_boolean_algebra = boolean_algebra + residuated_pogroupoid begin subclass residuated_lgroupoid .. definition conjugate_l :: "'a \ 'a \ 'a" (infixl "\" 60) where "x \ y \ -(-x \ y)" definition conjugate_r :: "'a \ 'a \ 'a" (infixl "\" 60) where "x \ y \ -(x \ -y)" lemma residual_conjugate_r: "x \ y = -(x \ -y)" by (metis conjugate_r_def double_compl) lemma residual_conjugate_l: "x \ y = -(-x \ y)" by (metis conjugate_l_def double_compl) lemma conjugation_multl: "x\y \ z = \ \ x \ (z \ y) = \" by (metis conjugate_l_def double_compl le_iff_inf_bot resl_galois) lemma conjugation_multr: "x\y \ z = \ \ y \ (x \ z) = \" by (metis conjugate_r_def inf_bot_iff_le le_iff_inf_bot resr_galois) lemma conjugation_conj: "(x \ y) \ z = \ \ y \ (z \ x) = \" by (metis inf_commute conjugation_multr conjugation_multl) lemma conjugation_pair_multl [simp]: "conjugation_pair (\x. x\y) (\x. x \ y)" by (simp add: conjugation_pair_def conjugation_multl) lemma conjugation_pair_multr [simp]: "conjugation_pair (\x. y\x) (\x. y \ x)" by (simp add: conjugation_pair_def conjugation_multr) lemma conjugation_pair_conj [simp]: "conjugation_pair (\x. y \ x) (\x. x \ y)" by (simp add: conjugation_pair_def conjugation_conj) lemma residuated_conjl1 [simp]: "residuated (\x. x \ y)" by (metis conj_residuatedI2 conjugation_pair_multl) lemma residuated_conjl2 [simp]: "residuated (\x. y \ x)" by (metis conj_residuatedI1 conjugation_pair_conj) lemma residuated_conjr1 [simp]: "residuated (\x. y \ x)" by (metis conj_residuatedI2 conjugation_pair_multr) lemma residuated_conjr2 [simp]: "residuated (\x. x \ y)" by (metis conj_residuatedI2 conjugation_pair_conj) lemma conjugate_multr [simp]: "conjugate (\x. y\x) = (\x. y \ x)" by (metis conjugation_pair_multr residuated_conj1 residuated_multr unique_conjugate) lemma conjugate_conjr1 [simp]: "conjugate (\x. y \ x) = (\x. y\x)" by (metis conjugate_multr conj_involution residuated_multr) lemma conjugate_multl [simp]: "conjugate (\x. x\y) = (\x. x \ y)" by (metis conjugation_pair_multl residuated_conj1 residuated_multl unique_conjugate) lemma conjugate_conjl1 [simp]: "conjugate (\x. x \ y) = (\x. x\y)" proof - have "conjugate (conjugate (\x. x\y)) = conjugate (\x. x \ y)" by simp thus ?thesis by (metis conj_involution[OF residuated_multl]) qed lemma conjugate_conjl2[simp]: "conjugate (\x. y \ x) = (\x. x \ y)" by (metis conjugation_pair_conj unique_conjugate residuated_conj1 residuated_conjl2) lemma conjugate_conjr2[simp]: "conjugate (\x. x \ y) = (\x. y \ x)" proof - have "conjugate (conjugate (\x. y \ x)) = conjugate (\x. x \ y)" by simp thus ?thesis by (metis conj_involution[OF residuated_conjl2]) qed lemma conjl1_iso: "x \ y \ x \ z \ y \ z" by (metis conjugate_l_def compl_mono resl_iso) lemma conjl2_iso: "x \ y \ z \ x \ z \ y" by (metis res_iso residuated_conjl2) lemma conjr1_iso: "x \ y \ z \ x \ z \ y" by (metis res_iso residuated_conjr1) lemma conjr2_iso: "x \ y \ x \ z \ y \ z" by (metis conjugate_r_def compl_mono resr_antitonel) lemma conjl1_sup: "z \ (x \ y) = (z \ x) \ (z \ y)" by (metis conjugate_l_def compl_inf resl_distr) lemma conjl2_sup: "(x \ y) \ z = (x \ z) \ (y \ z)" by (metis (poly_guards_query) residuated_sup residuated_conjl1) lemma conjr1_sup: "z \ (x \ y) = (z \ x) \ (z \ y)" by (metis residuated_sup residuated_conjr1) lemma conjr2_sup: "(x \ y) \ z = (x \ z) \ (y \ z)" by (metis conjugate_r_def compl_inf resr_distl) lemma conjl1_strict: "\ \ x = \" by (metis residuated_strict residuated_conjl1) lemma conjl2_strict: "x \ \ = \" by (metis residuated_strict residuated_conjl2) lemma conjr1_strict: "\ \ x = \" by (metis residuated_strict residuated_conjr2) lemma conjr2_strict: "x \ \ = \" by (metis residuated_strict residuated_conjr1) lemma conjl1_iff: "x \ y \ z \ x \ -(-z\y)" by (metis conjugate_l_def compl_le_swap1 compl_le_swap2 resl_galois) lemma conjl2_iff: "x \ y \ z \ y \ -(-z \ x)" by (metis conjl1_iff conjugate_r_def compl_le_swap2 double_compl resr_galois) lemma conjr1_iff: "x \ y \ z \ y \ -(x\-z)" by (metis conjugate_r_def compl_le_swap1 double_compl resr_galois) lemma conjr2_iff: "x \ y \ z \ x \ -(y \ -z)" by (metis conjugation_conj double_compl inf.commute le_iff_inf_bot) text \ We apply Maddux's lemmas regarding conjugation of an arbitrary residuated function for each of the 6 functions. \ lemma maddux1a: "a\(x \ -(a \ y)) \ a\x" by (insert maddux1 [of "\x. a\x"]) simp lemma maddux1a': "a\(x \ -(a \ y)) \ -y" by (insert maddux1 [of "\x. a\x"]) simp lemma maddux1b: "(x \ -(y \ a))\a \ x\a" by (insert maddux1 [of "\x. x\a"]) simp lemma maddux1b': "(x \ -(y \ a))\a \ -y" by (insert maddux1 [of "\x. x\a"]) simp lemma maddux1c: " a \ x \ -(y \ a) \ a \ x" by (insert maddux1 [of "\x. a \ x"]) simp lemma maddux1c': "a \ x \ -(y \ a) \ -y" by (insert maddux1 [of "\x. a \ x"]) simp lemma maddux1d: "a \ x \ -(a\y) \ a \ x" by (insert maddux1 [of "\x. a \ x"]) simp lemma maddux1d': "a \ x \ -(a\y) \ -y" by (insert maddux1 [of "\x. a \ x"]) simp lemma maddux1e: "x \ -(y\a) \ a \ x \ a" by (insert maddux1 [of "\x. x \ a"]) simp lemma maddux1e': "x \ -(y\a) \ a \ -y" by (insert maddux1 [of "\x. x \ a"]) simp lemma maddux1f: "x \ -(a \ y) \ a \ x \ a" by (insert maddux1 [of "\x. x \ a"]) simp lemma maddux1f': "x \ -(a \ y) \ a \ -y" by (insert maddux1 [of "\x. x \ a"]) simp lemma maddux2a: "a\x \ y \ a\(x \ (a \ y))" by (insert maddux2 [of "\x. a\x"]) simp lemma maddux2b: "x\a \ y \ (x \ (y \ a))\a" by (insert maddux2 [of "\x. x\a"]) simp lemma maddux2c: "(a \ x) \ y \ a \ (x \ (y \ a))" by (insert maddux2 [of "\x. a \ x"]) simp -lemma maddux2d: "(a \ x) \ y \ a \ (x \ a\y)" +lemma maddux2d: "(a \ x) \ y \ a \ (x \ (a\y))" by (insert maddux2 [of "\x. a \ x"]) simp -lemma maddux2e: "(x \ a) \ y \ (x \ y\a) \ a" +lemma maddux2e: "(x \ a) \ y \ (x \ (y\a)) \ a" by (insert maddux2 [of "\x. x \ a"]) simp lemma maddux2f: "(x \ a) \ y \ (x \ (a \ y)) \ a" by (insert maddux2 [of "\x. x \ a"]) simp text \ The multiplicative operation $\cdot$ on a residuated boolean algebra is generally not associative. We prove some equivalences related to associativity. \ lemma res_assoc_iff1: "(\x y z. x\(y\z) = (x\y)\z) \ (\x y z. x \ (y \ z) = y\x \ z)" proof safe fix x y z assume "\x y z. x\(y\z) = (x\y)\z" thus "x \ (y \ z) = y \ x \ z" using conjugate_comp_ext[of "\z. y\z" "\z. x\z"] by auto next fix x y z assume "\x y z. x \ (y \ z) = y\x \ z" thus "x\(y\z) = (x\y)\z" using conjugate_comp_ext[of "\z. y \ z" "\z. x \ z"] by auto qed lemma res_assoc_iff2: "(\x y z. x\(y\z) = (x\y)\z) \ (\x y z. x \ (y \ z) = (x \ z) \ y)" proof safe fix x y z assume "\x y z. x\(y\z) = (x\y)\z" hence "\x y z. (x\y)\z = x\(y\z)" by simp thus "x \ (y \ z) = (x \ z) \ y" using conjugate_comp_ext[of "\x. x\z" "\x. x\y"] by auto next fix x y z assume "\x y z. x \ (y \ z) = (x \ z) \ y" hence "\x y z. (x \ z) \ y = x \ (y \ z)" by simp thus "x\(y\z) = (x\y)\z" using conjugate_comp_ext[of "\z. z \ y" "\x. x \ z"] by auto qed lemma res_assoc_iff3: "(\x y z. x\(y\z) = (x\y)\z) \ (\x y z. (x \ y) \ z = x \ (y \ z))" proof safe fix x y z assume "\x y z. x\(y\z) = (x\y)\z" thus "(x \ y) \ z = x \ (y \ z)" using conjugate_comp_ext[of "\u. x\u" "\u. u\z"] and conjugate_comp_ext[of "\u. u\z" "\u. x\u", symmetric] by auto next fix x y z assume "\x y z. (x \ y) \ z = x \ (y \ z)" thus "x\(y\z) = (x\y)\z" using conjugate_comp_ext[of "\u. x \ u" "\u. u \ z"] and conjugate_comp_ext[of "\u. u \ z" "\u. x \ u", symmetric] by auto qed end (* residuated_boolean_algebra *) class unital_residuated_boolean = residuated_boolean_algebra + one + assumes mult_onel [simp]: "x\1 = x" and mult_oner [simp]: "1\x = x" begin text \ The following equivalences are taken from J{\'o}sson and Tsinakis. \ lemma jonsson1a: "(\f. \x y. x \ y = f(x)\y) \ (\x y. x \ y = (x \ 1)\y)" apply standard apply force apply (rule_tac x="\x. x \ 1" in exI) apply force done lemma jonsson1b: "(\x y. x \ y = (x \ 1)\y) \ (\x y. x\y = (x \ 1) \ y)" proof safe fix x y assume "\x y. x \ y = (x \ 1)\y" hence "conjugate (\y. x \ y) = conjugate (\y. (x \ 1)\y)" by metis thus "x\y = (x \ 1) \ y" by simp next fix x y assume "\x y. x \ y = x \ 1 \ y" thus "x \ y = (x \ 1) \ y" by (metis mult_onel) qed lemma jonsson1c: "(\x y. x \ y = (x \ 1)\y) \ (\x y. y \ x = 1 \ (x \ y))" proof safe fix x y assume "\x y. x \ y = (x \ 1)\y" hence "(\x. x \ y) = (\x. (x \ 1)\y)" by metis hence "(\x. x \ y) = (\x. x\y) o (\x. x \ 1)" by force hence "conjugate (\x. y \ x) = (\x. x\y) o conjugate (\x. 1 \ x)" by simp hence "conjugate (conjugate (\x. y \ x)) = conjugate ((\x. x\y) o conjugate (\x. 1 \ x))" by simp hence "(\x. y \ x) = conjugate ((\x. x\y) o conjugate (\x. 1 \ x))" by simp also have "... = conjugate (conjugate (\x. 1 \ x)) o conjugate (\x. x\y)" by (subst conjugate_comp[symmetric]) simp_all finally show "y \ x = 1 \ (x \ y)" by simp next fix x y assume "\x y. y \ x = 1 \ (x \ y)" hence "(\x. y \ x) = (\x. 1 \ (x \ y))" by metis hence "(\x. y \ x) = (\x. 1 \ x) o conjugate (\x. x\y)" by force hence "conjugate (\x. y \ x) = conjugate ((\x. 1 \ x) o conjugate (\x. x\y))" by metis also have "... = conjugate (conjugate (\x. x\y)) o conjugate (\x. 1 \ x)" by (subst conjugate_comp[symmetric]) simp_all finally have "(\x. x \ y) = (\x. x\y) o (\x. x \ 1)" by simp hence "(\x. x \ y) = (\x. (x \ 1) \ y)" by (simp add: comp_def) thus "x \ y = (x \ 1) \ y" by metis qed lemma jonsson2a: "(\g. \x y. x \ y = x\g(y)) \ (\x y. x \ y = x\(1 \ y))" apply standard apply force apply (rule_tac x="\x. 1 \ x" in exI) apply force done lemma jonsson2b: "(\x y. x \ y = x\(1 \ y)) \ (\x y. x\y = x \ (1 \ y))" proof safe fix x y assume "\x y. x \ y = x\(1 \ y)" hence "conjugate (\x. x \ y) = conjugate (\x. x\(1 \ y))" by metis thus "x\y = x \ (1 \ y)" by simp metis next fix x y assume "\x y. x\y = x \ (1 \ y)" hence "(\x. x\y) = (\x. x \ (1 \ y))" by metis hence "conjugate (\x. x\y) = conjugate (\x. x \ (1 \ y))" by metis thus "x \ y = x \ (1 \ y)" by simp metis qed lemma jonsson2c: "(\x y. x \ y = x\(1 \ y)) \ (\x y. y \ x = (x \ y) \ 1)" proof safe fix x y assume "\x y. x \ y = x\(1 \ y)" hence "(\y. x \ y) = (\y. x\(1 \ y))" by metis hence "(\y. x \ y) = (\y. x\y) o (\y. 1 \ y)" by force hence "conjugate (\y. y \ x) = (\y. x\y) o conjugate (\y. y \ 1)" by force hence "conjugate (conjugate (\y. y \ x)) = conjugate ((\y. x\y) o conjugate (\y. y \ 1))" by metis hence "(\y. y \ x) = conjugate ((\y. x\y) o conjugate (\y. y \ 1))" by simp also have "... = conjugate (conjugate (\y. y \ 1)) o conjugate (\y. x\y)" by (subst conjugate_comp[symmetric]) simp_all finally have "(\y. y \ x) = (\y. x \ y \ 1)" by (simp add: comp_def) thus "y \ x = x \ y \ 1" by metis next fix x y assume "\x y. y \ x = x \ y \ 1" hence "(\y. y \ x) = (\y. x \ y \ 1)" by force hence "(\y. y \ x) = (\y. y \ 1) o conjugate (\y. x\y)" by force hence "conjugate (\y. y \ x) = conjugate ((\y. y \ 1) o conjugate (\y. x\y))" by metis also have "... = conjugate (conjugate (\y. x\y)) o conjugate (\y. y \ 1)" by (subst conjugate_comp[symmetric]) simp_all finally have "(\y. x \ y) = (\y. x\y) o (\y. 1 \ y)" by (metis conjugate_conjr1 conjugate_conjr2 conjugate_multr) thus "x \ y = x \ (1 \ y)" by (simp add: comp_def) qed lemma jonsson3a: "(\x. (x \ 1) \ 1 = x) \ (\x. 1 \ (1 \ x) = x)" proof safe fix x assume "\x. x \ 1 \ 1 = x" thus "1 \ (1 \ x) = x" by (metis compl_le_swap1 compl_le_swap2 conjr2_iff order.eq_iff) next fix x assume "\x. 1 \ (1 \ x) = x" thus "x \ 1 \ 1 = x" by (metis conjugate_l_def conjugate_r_def double_compl jipsen2r) qed lemma jonsson3b: "(\x. (x \ 1) \ 1 = x) \ (x \ y) \ 1 = (x \ 1) \ (y \ 1)" proof (rule order.antisym, auto simp: conjr2_iso) assume assm: "\x. (x \ 1) \ 1 = x" hence "(x \ 1) \ (y \ 1) \ 1 = x \ (((x \ 1) \ (y \ 1) \ 1) \ y)" by (metis (no_types) conjr2_iso inf.cobounded2 inf.commute inf.orderE) hence "(x \ 1) \ (y \ 1) \ 1 \ x \ y" using inf.orderI inf_left_commute by presburger thus "(x \ 1) \ (y \ 1) \ x \ y \ 1" using assm by (metis (no_types) conjr2_iso) qed lemma jonsson3c: "\x. (x \ 1) \ 1 = x \ x \ 1 = 1 \ x" proof (rule indirect_eq) fix z assume assms: "\x. (x \ 1) \ 1 = x" hence "(x \ 1) \ -z = \ \ ((x \ 1) \ -z) \ 1 = \" by (metis compl_sup conjugation_conj double_compl inf_bot_right sup_bot.left_neutral) also have "... \ -z\x \ 1 = \" by (metis assms jonsson3b conjugation_multr) finally have "(x \ 1) \ -z = \ \ (1 \ x) \ -z = \" by (metis conjugation_multl inf.commute) thus "(x \ 1 \ z) \ (1 \ x \ z)" by (metis le_iff_inf_bot) qed end (* unital_residuated_boolean *) class residuated_boolean_semigroup = residuated_boolean_algebra + semigroup_mult begin subclass residuated_boolean_algebra .. text \ The following lemmas hold trivially, since they are equivalent to associativity. \ lemma res_assoc1: "x \ (y \ z) = y\x \ z" by (metis res_assoc_iff1 mult_assoc) lemma res_assoc2: "x \ (y \ z) = (x \ z) \ y" by (metis res_assoc_iff2 mult_assoc) lemma res_assoc3: "(x \ y) \ z = x \ (y \ z)" by (metis res_assoc_iff3 mult_assoc) end (*residuated_boolean_semigroup *) class residuated_boolean_monoid = residuated_boolean_algebra + monoid_mult begin subclass unital_residuated_boolean by standard auto subclass residuated_lmonoid .. lemma jonsson4: "(\x y. x \ y = x\(1 \ y)) \ (\x y. x \ y = (x \ 1)\y)" proof safe fix x y assume assms: "\x y. x \ y = x\(1 \ y)" have "x \ y = (y \ x) \ 1" by (metis assms jonsson2c) also have "... = (y \ ((x \ 1) \ 1)) \ 1" by (metis assms jonsson2b jonsson3a mult_oner) also have "... = (((x \ 1)\y) \ 1) \ 1" by (metis conjugate_r_def double_compl resr3) also have "... = (x \ 1)\y" by (metis assms jonsson2b jonsson3a mult_oner) finally show "x \ y = (x \ 1)\y" . next fix x y assume assms: "\x y. x \ y = (x \ 1)\y" have "y \ x = 1 \ (x \ y)" by (metis assms jonsson1c) also have "... = 1 \ ((1 \ (1 \ x)) \ y)" by (metis assms conjugate_l_def double_compl jonsson1c mult_1_right resl3) also have "... = 1 \ (1 \ (y\(1 \ x)))" by (metis conjugate_l_def double_compl resl3) also have "... = y\(1 \ x)" by (metis assms jonsson1b jonsson1c jonsson3c mult_onel) - finally show "y \ x = y\(1 \ x)". + finally show "y \ x = y\(1 \ x)" . qed end (* residuated_boolean_monoid *) end diff --git a/thys/Residuated_Lattices/Residuated_Lattices.thy b/thys/Residuated_Lattices/Residuated_Lattices.thy --- a/thys/Residuated_Lattices/Residuated_Lattices.thy +++ b/thys/Residuated_Lattices/Residuated_Lattices.thy @@ -1,730 +1,718 @@ (* Title: Residuated Lattices Author: Victor Gomes Maintainer: Georg Struth *) theory Residuated_Lattices imports Kleene_Algebra.Signatures begin notation times (infixl "\" 70) -notation - bot ("\") and - top ("\") and - inf (infixl "\" 65) and - sup (infixl "\" 65) and - Inf ("\_" [900] 900) and - Sup ("\_" [900] 900) +unbundle lattice_syntax -syntax - "_INF1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) - "_INF" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) - "_SUP1" :: "pttrns \ 'b \ 'b" ("(3\_./ _)" [0, 10] 10) - "_SUP" :: "pttrn \ 'a set \ 'b \ 'b" ("(3\_\_./ _)" [0, 0, 10] 10) - section \Residuated Lattices\ subsection \Residuated Functions on a Partial Order\ text \ We follow Galatos and \emph{al.} to define residuated funtions on partial orders. Material from articles by Maddux, and J{\'o}sson and Tsinakis are also considered. This development should in the future be expanded to functions or categories where the sources and targets have different type. Let $P$ be a partial order, or a poset. A map $f : P \to P$ is residuated if there exists a map $g: P \to P$ such that $f(x) \le y \Leftrightarrow x \le g(y)$ for all $x, y \in P$. That is, they are adjoints of a Galois connection. \ context order begin definition residuated_pair :: "('a \ 'a) \ ('a \ 'a) \ bool" where "residuated_pair f g \ \x y. f(x) \ y \ x \ g(y)" theorem residuated_pairI [intro]: assumes "\x y. x \ y \ f x \ f y" and "\x y. x \ y \ g x \ g y" and "\x. x \ (g o f) x" and "\x. (f o g) x \ x" shows "residuated_pair f g" by (metis assms comp_apply local.order_trans residuated_pair_def antisym) definition residuated :: "('a \ 'a) \ bool" where "residuated f \ \g. residuated_pair f g" text \ If a map $f$ is residuated, then its residual $g$ is unique. \ lemma residual_unique: "residuated f \ \!g. residuated_pair f g" unfolding residuated_def residuated_pair_def by (metis ext eq_refl order.antisym) text \ Since the residual of a map $f$ is unique, it makes sense to define a residual operator. \ definition residual :: "('a \ 'a) \ ('a \ 'a)" where "residual f \ THE g. residuated_pair f g" lemma residual_galois: "residuated f \ f(x) \ y \ x \ residual f y" apply (clarsimp simp: residuated_def residuated_pair_def) apply (subgoal_tac "residual f = g") apply (simp_all add: residual_def) apply (rule the1_equality) apply (metis residuated_def residuated_pair_def residual_unique) by (simp add: residuated_pair_def) lemma residual_galoisI1: "residuated f \ f(x) \ y \ x \ residual f y" by (metis residual_galois) lemma residual_galoisI2: "residuated f \ x \ residual f y \ f(x) \ y" by (metis residual_galois) text \ A closure operator on a poset is a map that is expansive, isotone and idempotent. The composition of the residual of a function $f$ with $f$ is a closure operator. \ definition closure :: "('a \ 'a) \ bool" where "closure f \ (\x. x \ f x) \ (\x y. x \ y \ f x \ f y) \ (\x. f(f x) = f x)" lemma res_c1: "residuated f \ x \ residual f (f x)" by (metis local.order.refl residual_galois) lemma res_c2: "residuated f \ x \ y \ residual f (f x) \ residual f (f y)" by (metis local.order.refl local.order_trans residual_galois) lemma res_c3: "residuated f \ residual f (f (residual f (f x))) = residual f (f x)" by (metis order.eq_iff local.order_trans res_c1 residual_galois) lemma res_closure: "residuated f \ closure (residual f o f)" by (auto simp: closure_def intro: res_c1 res_c2 res_c3) text \ Dually, an interior operator on a poset is a map that is contractive, isotone and idempotent. The composition of $f$ with its residual is an interior operator. \ definition interior :: "('a \ 'a) \ bool" where "interior f \ (\x. f x \ x) \ (\x y. x \ y \ f x \ f y) \ (\x. f(f x) = f x)" lemma res_i1: "residuated f \ f (residual f x) \ x" by (metis local.order.refl residual_galois) lemma res_i2: "residuated f \ x \ y \ f (residual f x) \ f (residual f y)" by (metis local.order.refl local.order_trans residual_galois) lemma res_i3: "residuated f \ f (residual f (f (residual f x))) = f (residual f x)" by (metis local.antisym_conv res_c1 res_c3 res_i1 residual_galois) lemma res_interior: "residuated f \ interior (f o residual f)" by (auto simp: interior_def intro: res_i1 res_i2 res_i3) text \Next we show a few basic lemmas about residuated maps.\ lemma res_iso: "residuated f \ x \ y \ f x \ f y" by (metis local.order.trans res_c1 residual_galois) lemma res_residual_iso: "residuated f \ x \ y \ residual f x \ residual f y" by (metis local.order.trans res_i1 residual_galois) lemma res_comp1 [simp]: "residuated f \ f o residual f o f = f" proof - assume resf: "residuated f" { fix x have "(f o residual f o f) x = f x" by (metis resf comp_apply order.eq_iff res_c1 res_i1 res_iso) } thus ?thesis by (metis ext) qed lemma res_comp2 [simp]: "residuated f \ residual f o f o residual f = residual f" proof - assume resf: "residuated f" { fix x have "(residual f o f o residual f) x = residual f x" by (metis resf comp_apply order.eq_iff res_c1 res_i1 res_residual_iso) } thus ?thesis by (metis ext) qed end (* order *) text \ A residuated function $f$ preserves all existing joins. Dually, its residual preserves all existing meets. We restrict our attention to semilattices, where binary joins or meets exist, and to complete lattices, where arbitrary joins and meets exist. \ lemma (in semilattice_sup) residuated_sup: "residuated f \ f (x \ y) = f x \ f y" proof (rule order.antisym) assume assm: "residuated f" thus "f (x \ y) \ f x \ f y" by (metis local.residual_galoisI1 local.residual_galoisI2 local.sup.bounded_iff local.sup_ge1 local.sup_ge2) show "f x \ f y \ f (x \ y)" by (metis assm local.res_iso local.sup.bounded_iff local.sup_ge1 local.sup_ge2) qed lemma (in semilattice_inf) residuated_inf: "residuated f \ residual f (x \ y) = residual f x \ residual f y" proof (rule order.antisym) assume assm: "residuated f" thus "residual f (x \ y) \ residual f x \ residual f y" by (metis local.inf.boundedI local.inf.cobounded1 local.inf.cobounded2 local.res_residual_iso) show "residual f x \ residual f y \ residual f (x \ y)" by (metis assm local.inf.bounded_iff local.inf.cobounded1 local.inf.cobounded2 local.residual_galoisI1 local.residual_galoisI2) qed context bounded_lattice begin lemma residuated_strict: "residuated f \ f \ = \" by (metis local.bot_least local.bot_unique local.res_i1 local.res_iso) lemma res_top: "residuated f \ residual f \ = \" by (metis local.residual_galoisI1 local.top_greatest local.top_unique) end context complete_lattice begin text \ On a complete lattice, a function $f$ is residuated if and only if it preserves arbitrary (possibly infinite) joins. Dually, a function $g$ is a residual of a residuated funtion $f$ if and only if $g$ preserver arbitrary meets. \ lemma residual_eq1: "residuated f \ residual f y = \ {x. f x \ y}" proof (rule order.antisym) assume assm: "residuated f" thus "residual f y \ \{x. f x \ y}" by (auto simp: res_i1 intro!: Sup_upper) show "\{x. f x \ y} \ residual f y" by (auto simp: assm intro!: Sup_least residual_galoisI1) qed lemma residual_eq2: "residuated f \ f x = \ {y. x \ residual f y}" proof (rule order.antisym) assume assm: "residuated f" thus "f x \ \{y. x \ residual f y}" by (auto intro: Inf_greatest residual_galoisI2) show "\{y. x \ residual f y} \ f x" using assm by (auto simp: res_c1 intro!: Inf_lower) qed lemma residuated_Sup: "residuated f \ f(\ X) = \{f x | x. x \ X}" proof (rule order.antisym) assume assm: "residuated f" obtain y where y_def: "y = \{f x |x. x \ X}" by auto hence "\x \ X. f x \ y" by (auto simp: y_def intro: Sup_upper) hence "\x \ X. x \ residual f y" by (auto simp: assm intro!: residual_galoisI1) hence "\X \ residual f y" by (auto intro: Sup_least) thus "f(\ X) \ \{f x |x. x \ X}" by (metis y_def assm residual_galoisI2) qed (clarsimp intro!: Sup_least res_iso Sup_upper) lemma residuated_Inf: "residuated f \ residual f(\ X) = \{residual f x | x. x \ X}" proof (rule order.antisym, clarsimp intro!: Inf_greatest res_residual_iso Inf_lower) assume assm: "residuated f" obtain y where y_def: "y = \{residual f x |x. x \ X}" by auto hence "\x \ X. y \ residual f x" by (auto simp: y_def intro: Inf_lower) hence "\x \ X. f y \ x" by (metis assm residual_galoisI2) hence "f y \ \ X" by (auto intro: Inf_greatest) thus "\{residual f x |x. x \ X} \ residual f (\X)" by (auto simp: assm y_def intro!: residual_galoisI1) qed lemma Sup_sup: "\X. f(\ X) = \{f x | x. x \ X} \ f (x \ y) = f x \ f y" apply (erule_tac x="{x, y}" in allE) by (force intro: Sup_eqI) lemma Sup_residuatedI: "\X. f(\ X) = \{f x | x. x \ X} \ residuated f" proof (unfold residuated_def residuated_pair_def, standard+) fix x y assume "f x \ y" thus "x \ \{x. f x \ y}" by (clarsimp intro!: Sup_upper) next fix x y assume assm: "\X. f (\X) = \{f x |x. x \ X}" hence f_iso: "\x y. x \ y \ f x \ f y" using Sup_sup by (auto simp: le_iff_sup) assume "x \ \{x. f x \ y}" hence "f x \ f(\{x. f x \ y})" by (metis f_iso) also have "... = \{f x | x . f x \ y}" using assm by auto finally show "f x \ y" apply (rule order_trans) by (auto intro!: Sup_least) qed lemma Inf_inf: "\X. f(\ X) = \{f x | x. x \ X} \ f (x \ y) = f x \ f y" apply (erule_tac x="{x, y}" in allE) by (force intro: Inf_eqI) lemma Inf_residuatedI: "\X. \{g x | x. x \ X} = g (\ X) \ \f. residuated_pair f g" proof (unfold residuated_pair_def, standard+) fix x y assume "x \ g y" thus "\{y. x \ g y} \ y" by (clarsimp intro!: Inf_lower) next fix x y assume assm: "\X. \{g x | x. x \ X} = g (\ X)" hence g_iso: "\x y. x \ y \ g x \ g y" using Inf_inf by (auto simp: le_iff_inf) assume "\{y. x \ g y} \ y" hence "g (\{y. x \ g y}) \ g y" by (metis g_iso) hence "(\{g y | y. x \ g y}) \ g y" using assm apply (erule_tac x="{y. x \ g y}" in allE) by (auto intro!: Inf_lower) thus "x \ g y" apply (rule local.dual_order.trans) by (auto intro: Inf_greatest) qed end (* complete lattices *) subsection \Residuated Structures\ text \ In this section, we define residuated algebraic structures, starting from the simplest of all, a \emph{residuated partial ordered groupoid}, to \emph{residuated l-monoids}, which are residuated lattices where the multiplicative operator forms a monoid. \ class pogroupoid = order + times + assumes mult_isor: "x \ y \ x \ z \ y \ z" and mult_isol: "x \ y \ z \ x \ z \ y" text \ A residuated partial ordered groupoid is simply a partial order and a multiplicative groupoid with two extra operators satisfying the residuation laws. It is straighforward to prove that multiplication is compatible with the order, that is, multiplication is isotone. Most of the lemmas below come in pairs; they are related by opposition duality. Formalising this duality is left for future work. \ class residual_r_op = fixes residual_r :: "'a \ 'a \ 'a" (infixr "\" 60) class residual_l_op = fixes residual_l :: "'a \ 'a \ 'a" (infixl "\" 60) class residuated_pogroupoid = order + times + residual_l_op + residual_r_op + assumes resl_galois: "x \ z \ y \ x \ y \ z" and resr_galois: "x \ y \ z \ y \ x \ z" begin lemma reslI [intro]: "x \ y \ z \ x \ z \ y" by (metis resl_galois) lemma resrI [intro]: "x \ y \ z \ y \ x \ z" by (metis resr_galois) lemma residuated_pair_multl [simp]: "residuated_pair (\x. x \ y) (\x. x \ y)" by (auto simp: residuated_pair_def resl_galois) lemma residuated_pair_multr [simp]: "residuated_pair (\y. x \ y) (\y. x \ y)" by (auto simp: residuated_pair_def resr_galois) text \ Multiplication is then obviously residuated. \ lemma residuated_multl [simp]: "residuated (\x. x \ y)" by (metis residuated_def residuated_pair_multl) lemma residuated_multr [simp]: "residuated (\y. x \ y)" by (metis residuated_def residuated_pair_multr) lemma resl_eq [simp]: "residual (\x. x \ y) = (\x. x \ y)" unfolding residual_def apply (rule the1_equality) by (auto simp: intro!: residual_unique) lemma resr_eq [simp]: "residual (\y. x \ y) = (\y. x \ y)" unfolding residual_def apply (rule the1_equality) by (auto simp: intro!: residual_unique) text \ Next we prove a few lemmas, all of which are instantiation of more general facts about residuated functions. \ lemma res_lc1: "x \ x \ y \ y" by auto lemma res_lc2: "x \ y \ x \ z \ z \ y \ z \ z" by (metis local.res_c2 resl_eq residuated_multl) lemma res_lc3 [simp]: "(x \ y \ y) \ y \ y = x \ y \ y" by (metis local.res_c3 resl_eq residuated_multl) lemma res_rc1: "x \ y \ y \ x" by auto lemma res_rc2: "x \ y \ z \ z \ x \ z \ z \ y" by (metis local.res_c2 resr_eq residuated_multr) lemma res_rc3 [simp]: "y \ y \ (y \ y \ x) = y \ y \ x" by (metis local.res_c3 resr_eq residuated_multr) lemma res_li1: "(x \ y) \ y \ x" by (metis local.res_i1 resl_eq residuated_multl) lemma res_li2: "x \ y \ (x \ z) \ z \ (y \ z) \ z" by (metis local.res_i2 resl_eq residuated_multl) lemma res_li3 [simp]: "((x \ y) \ y \ y) \ y = (x \ y) \ y" by (metis local.res_i3 resl_eq residuated_multl) lemma res_ri1: "y \ (y \ x) \ x" by (metis local.res_i1 resr_eq residuated_multr) lemma res_ri2: "x \ y \ z \ (z \ x) \ z \ (z \ y)" by (metis local.res_i2 resr_eq residuated_multr) lemma res_ri3 [simp]: "y \ (y \ y \ (y \ x)) = y \ (y \ x)" by (metis local.res_i3 resr_eq residuated_multr) subclass pogroupoid proof fix x y z show "x \ y \ x \ z \ y \ z" by (metis local.res_iso residuated_multl) show "x \ y \ z \ x \ z \ y" by (metis local.res_iso residuated_multr) qed lemma resl_iso: "x \ y \ x \ z \ y \ z" by (metis res_residual_iso resl_eq residuated_multl) lemma resr_iso: "x \ y \ z \ x \ z \ y" by (metis res_residual_iso resr_eq residuated_multr) lemma resl_comp1 [simp]: "(x \ y \ y) \ y = x \ y" by (metis order.antisym local.mult_isor res_lc1 res_li1) lemma resl_comp2 [simp]: "(x \ y) \ y \ y = x \ y" by (metis order.eq_iff res_lc1 res_li1 resl_iso) lemma resr_comp1 [simp]: "y \ (y \ y \ x) = y \ x" by (metis order.antisym local.mult_isol res_rc1 res_ri1) lemma resr_comp2 [simp]: "y \ y \ (y \ x) = y \ x" by (metis order.eq_iff res_rc1 res_ri1 resr_iso) lemma resl_antitoner: "x \ y \ z \ y \ z \ x" by (metis local.dual_order.trans local.mult_isol res_li1 reslI) lemma resr_antitonel: "x \ y \ y \ z \ x \ z" by (metis local.dual_order.trans local.resl_galois res_ri1 resrI) text \The following lemmas are taken from Galatos and \emph{al.}\ lemma jipsen1l: "x \ y \ (x \ y)" by (metis res_ri1 reslI) lemma jipsen1r: "x \ (y \ x) \ y" by (metis res_li1 resrI) lemma jipsen2l: "(y \ (x \ y)) \ y = x \ y" by (metis jipsen1l jipsen1r order.eq_iff local.resr_antitonel) lemma jipsen2r: "y \ ((y \ x) \ y) = y \ x" by (metis jipsen1l jipsen1r order.eq_iff local.resl_antitoner) end (* residuated_pogroupoid *) text \ In a residuated partial ordered semigroup, the multiplicative operator is associative. \ class residuated_posemigroup = semigroup_mult + residuated_pogroupoid begin lemma resl_trans: "(x \ y) \ (y \ z) \ x \ z" by (metis local.res_li1 local.resl_antitoner local.resl_galois mult_assoc) lemma resr_trans: "(x \ y) \ (y \ z) \ x \ z" by (metis local.res_ri1 local.resr_antitonel local.resr_galois mult_assoc) lemma resr1: "(x \ y) \ z \ (x \ y \ z)" by (metis local.mult_isor local.res_ri1 local.resrI mult_assoc) lemma resr2: "x \ y \ z \ x \ z \ y" by (metis local.mult_isol local.res_ri1 local.resrI mult_assoc) lemma resr3: "x \ y \ z = y \ (x \ z)" by (metis order.eq_iff local.resr_galois mult_assoc) lemma resl1: "z \ (x \ y) \ (z \ x \ y)" by (metis local.mult_isol local.res_li1 local.reslI mult_assoc) lemma resl2: "x \ y \ x \ z \ y \ z" by (metis local.mult_isor local.res_li1 local.reslI mult_assoc) lemma resl3: "x \ y \ z = (x \ z) \ y" by (metis order.eq_iff local.resl_galois mult_assoc) lemma residual_assoc: "x \ (y \ z) = (x \ y) \ z" proof (rule order.antisym) show "x \ (y \ z) \ (x \ y) \ z" by (metis local.res_ri1 local.resl_galois local.resr_galois mult_assoc) show "(x \ y) \ z \ x \ (y \ z)" by (metis local.res_li1 local.resl_galois local.resr_galois mult_assoc) qed end (* residuated_posemigroup *) text \ In a residuated partially ordered monoid, the multiplicative operator has a unit $1$; that is, its reduct forms a monoid. \ class residuated_pomonoid = monoid_mult + residuated_pogroupoid begin subclass residuated_posemigroup .. lemma resl_unit: "x \ 1 = x" by (metis local.mult_1_right local.resl_comp1) lemma resr_unit: "1 \ x = x" by (metis local.mult_1_left local.resr_comp1) lemma mult_one_resl: "x \ (1 \ y) \ x \ y" by (metis local.mult_1_right local.resl1) lemma mult_one_resr: "(x \ 1) \ y \ x \ y" by (metis local.mult_1_left local.resr1) text \More lemmas from Galatos and \emph{al.} follow.\ lemma jipsen3l: "1 \ x \ x" by (metis local.jipsen1l resr_unit) lemma jipsen3r: "1 \ x \ x" by (metis local.jipsen1r resl_unit) lemma jipsen4l [simp]: "(x \ x) \ x = x" by (metis local.mult_1_left local.resl_comp1) lemma jipsen4r [simp]: "x \ (x \ x) = x" by (metis local.mult_1_right local.resr_comp1) lemma jipsen5l [simp]: "(x \ x) \ (x \ x) = x \ x" by (metis jipsen4l local.resl3) lemma jipsen5r [simp]: "(x \ x) \ (x \ x) = x \ x" by (metis jipsen4r local.resr3) lemma jipsen6l: "y \ x \ (y \ z) \ (x \ z)" by (metis local.resl_galois local.resl_trans) lemma jipsen6r: "x \ y \ (z \ x) \ (z \ y)" by (metis local.resr_galois local.resr_trans) lemma jipsen7l: "y \ x \ (z \ y) \ (z \ x)" by (metis local.resr_galois local.resl_trans) lemma jipsen7r: "x \ y \ (x \ z) \ (y \ z)" by (metis local.resl_galois local.resr_trans) end (* residuated_pomonoid *) text \ Residuated partially ordered groupoids can be expanded residuated join semilattice ordered groupoids. They are used as a base for action algebras, which are expansions of Kleene algebras by operations of residuation. Action algebras can be found in the AFP entry for Kleene algebras. \ class residuated_sup_lgroupoid = semilattice_sup + residuated_pogroupoid begin lemma distl: "x \ (y \ z) = x \ y \ x \ z" by (metis local.residuated_multr local.residuated_sup) lemma distr: "(x \ y) \ z = x \ z \ y \ z" by (metis local.residuated_multl local.residuated_sup) lemma resl_subdist_var: "x \ z \ (x \ y) \ z" by (metis local.resl_iso local.sup_ge1) lemma resl_subdist: "(x \ z) \ (y \ z) \ (x \ y) \ z" by (metis local.le_sup_iff local.sup_commute resl_subdist_var) lemma resr_subdist_var: "(x \ y) \ x \ (y \ z)" by (metis local.resr_iso local.sup_ge1) lemma resr_subdist: "(x \ y) \ (x \ z) \ x \ (y \ z)" by (metis sup_commute sup_least resr_subdist_var) lemma resl_superdist_var: "x \ (y \ z) \ x \ y" by (metis local.le_sup_iff local.res_li1 local.reslI local.resr_galois) lemma resr_superdist_var: "(x \ y) \ z \ x \ z" by (metis local.le_sup_iff local.res_ri1 local.resl_galois local.resrI) end (* residuated_sup_lgroupoid *) text \ Similarly, semigroup and monoid variants can be defined. \ class residuated_sup_lsemigroup = semilattice_sup + residuated_posemigroup subclass (in residuated_sup_lsemigroup) residuated_sup_lgroupoid .. class residuated_sup_lmonoid = semilattice_sup + residuated_posemigroup subclass (in residuated_sup_lmonoid) residuated_sup_lsemigroup .. text \ By lattice duality, we define residuated meet semillatice ordered groupoid. \ class residuated_inf_lgroupoid = semilattice_inf + residuated_pogroupoid begin lemma resl_dist: "(x \ y) \ z = (x \ z) \ (y \ z)" by (metis local.resl_eq local.residuated_inf local.residuated_multl) lemma resr_dist: "x \ (y \ z) = (x \ y) \ (x \ z)" by (metis local.resr_eq local.residuated_inf local.residuated_multr) lemma resl_inf_superdist: "x \ y \ x \ (y \ z)" by (metis local.inf_le1 local.resl_antitoner) lemma resr_inf_superdist_var: "x \ y \ (x \ z) \ y" by (metis local.inf_le1 local.resr_antitonel) end (* residuated_inf_lgroupoid *) class residuated_inf_lsemigroup = semilattice_inf + residuated_posemigroup subclass (in residuated_inf_lsemigroup) residuated_inf_lgroupoid .. class residuated_inf_lmonoid = semilattice_inf + residuated_posemigroup subclass (in residuated_inf_lmonoid) residuated_inf_lsemigroup .. text \ Finally, we obtain residuated lattice groupoids. \ class residuated_lgroupoid = lattice + residuated_pogroupoid begin subclass residuated_sup_lgroupoid .. lemma resl_distr: "z \ (x \ y) = (z \ x) \ (z \ y)" proof (rule order.antisym) show "z \ x \ y \ (z \ x) \ (z \ y)" by (metis local.inf.bounded_iff local.resl_superdist_var local.sup_commute) show "(z \ x) \ (z \ y) \ z \ x \ y" by (metis local.inf.cobounded1 local.inf.cobounded2 local.resl_galois local.resr_galois local.sup.bounded_iff) qed lemma resr_distl: "(x \ y) \ z = (x \ z) \ (y \ z)" proof (rule order.antisym) show "x \ y \ z \ (x \ z) \ (y \ z)" by (metis local.inf.bounded_iff local.resr_antitonel local.resr_superdist_var local.sup_ge2) show "(x \ z) \ (y \ z) \ x \ y \ z" by (metis local.inf.cobounded1 local.inf.cobounded2 local.resl_galois local.resr_galois local.sup_least) qed end (* residuated_lgroupoid *) class residuated_lsemigroup = lattice + residuated_sup_lsemigroup subclass (in residuated_lsemigroup) residuated_lgroupoid .. class residuated_lmonoid = lattice + residuated_sup_lmonoid subclass (in residuated_lmonoid) residuated_lsemigroup .. class residuated_blgroupoid = bounded_lattice + residuated_pogroupoid begin lemma multl_strict [simp]: "x \ \ = \" by (metis local.residuated_multr local.residuated_strict) lemma multr_strict [simp]: "\ \ x = \" by (metis local.residuated_multl local.residuated_strict) lemma resl_top [simp]: "\ \ x = \" by (metis local.res_top local.residuated_multl local.resl_eq) lemma resl_impl [simp]: "x \ \ = \" by (metis local.resl_comp2 multl_strict resl_top) lemma resr_top [simp]: "x \ \ = \" by (metis local.resrI local.top_greatest local.top_unique) lemma resr_impl [simp]: "\ \ x = \" by (metis local.resr_comp2 multr_strict resr_top) end (* residuated_blgroupoid *) class residuated_blsemigroup = bounded_lattice + residuated_sup_lsemigroup subclass (in residuated_blsemigroup) residuated_blgroupoid .. class residuated_blmonoid = bounded_lattice + residuated_sup_lmonoid subclass (in residuated_blmonoid) residuated_blsemigroup .. class residuated_clgroupoid = complete_lattice + residuated_pogroupoid begin lemma resl_eq_def: "y \ x = \ {z. z \ x \ y}" by (metis local.residual_eq1 local.residuated_multl local.resl_eq) lemma resr_eq_def: "x \ y = \ {z. x \ z \ y}" by (metis local.residual_eq1 local.residuated_multr local.resr_eq) lemma multl_def: "x \ y = \ {z. x \ z \ y}" proof - have "\{ya. x \ residual (\a. a \ y) ya} = \{z. x \ z \ y}" by simp thus ?thesis by (metis residuated_multl residual_eq2) qed lemma multr_def: "x \ y = \ {z. y \ x \ z}" proof - have "\{ya. y \ residual ((\) x) ya} = \{z. y \ x \ z}" by simp thus ?thesis by (metis residuated_multr residual_eq2) qed end (* residuated_clgroupoid *) class residuated_clsemigroup = complete_lattice + residuated_sup_lsemigroup subclass (in residuated_clsemigroup) residuated_clgroupoid .. class residuated_clmonoid = complete_lattice + residuated_sup_lmonoid subclass (in residuated_clmonoid) residuated_clsemigroup .. end diff --git a/thys/SIFUM_Type_Systems/Preliminaries.thy b/thys/SIFUM_Type_Systems/Preliminaries.thy --- a/thys/SIFUM_Type_Systems/Preliminaries.thy +++ b/thys/SIFUM_Type_Systems/Preliminaries.thy @@ -1,79 +1,81 @@ (* Title: SIFUM-Type-Systems Authors: Sylvia Grewe, Heiko Mantel, Daniel Schoepe *) section \Preliminaries\ theory Preliminaries -imports Main "HOL-Library.Lattice_Syntax" +imports Main begin +unbundle lattice_syntax + text \Possible modes for variables:\ datatype Mode = AsmNoRead | AsmNoWrite | GuarNoRead | GuarNoWrite text \We consider a two-element security lattice:\ datatype Sec = High | Low notation less_eq (infix "\" 50) and less (infix "\" 50) text \@{term Sec} forms a (complete) lattice:\ instantiation Sec :: complete_lattice begin definition top_Sec_def: "\ = High" definition sup_Sec_def: "d1 \ d2 = (if (d1 = High \ d2 = High) then High else Low)" definition inf_Sec_def: "d1 \ d2 = (if (d1 = Low \ d2 = Low) then Low else High)" definition bot_Sec_def: "\ = Low" definition less_eq_Sec_def: "d1 \ d2 = (d1 = d2 \ d1 = Low)" definition less_Sec_def: "d1 < d2 = (d1 = Low \ d2 = High)" definition Sup_Sec_def: "\S = (if (High \ S) then High else Low)" definition Inf_Sec_def: "\S = (if (Low \ S) then Low else High)" instance apply (intro_classes) apply (metis Sec.exhaust Sec.simps(2) less_Sec_def less_eq_Sec_def) apply (metis less_eq_Sec_def) apply (metis less_eq_Sec_def) apply (metis less_eq_Sec_def) apply (metis Sec.exhaust inf_Sec_def less_eq_Sec_def) apply (metis Sec.exhaust inf_Sec_def less_eq_Sec_def) apply (metis Sec.exhaust inf_Sec_def less_eq_Sec_def) apply (metis Sec.exhaust less_eq_Sec_def sup_Sec_def) apply (metis Sec.exhaust less_eq_Sec_def sup_Sec_def) apply (metis Sec.exhaust Sec.simps(2) less_eq_Sec_def sup_Sec_def) apply (metis (full_types) Inf_Sec_def Sec.exhaust less_eq_Sec_def) apply (metis Inf_Sec_def Sec.exhaust less_eq_Sec_def) apply (metis Sec.exhaust Sup_Sec_def less_eq_Sec_def) apply (metis (full_types) Sup_Sec_def less_eq_Sec_def) apply (metis (opaque_lifting, mono_tags) Inf_Sec_def empty_iff top_Sec_def) by (metis (opaque_lifting, mono_tags) Sup_Sec_def bot_Sec_def empty_iff) end text \Memories are mappings from variables to values\ type_synonym ('var, 'val) Mem = "'var \ 'val" text \A mode state maps modes to the set of variables for which the given mode is set.\ type_synonym 'var Mds = "Mode \ 'var set" text \Local configurations:\ type_synonym ('com, 'var, 'val) LocalConf = "('com \ 'var Mds) \ ('var, 'val) Mem" text \Global configurations:\ type_synonym ('com, 'var, 'val) GlobalConf = "('com \ 'var Mds) list \ ('var, 'val) Mem" text \A locale to fix various parametric components in Mantel et. al, and assumptions about them:\ locale sifum_security = fixes dma :: "'Var \ Sec" fixes stop :: "'Com" fixes eval :: "('Com, 'Var, 'Val) LocalConf rel" fixes some_val :: "'Val" fixes some_val' :: "'Val" assumes stop_no_eval: "\ ((((stop, mds), mem), ((c', mds'), mem')) \ eval)" assumes deterministic: "\ (lc, lc') \ eval; (lc, lc'') \ eval \ \ lc' = lc''" assumes finite_memory: "finite {(x::'Var). True}" assumes different_values: "some_val \ some_val'" end