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,523 +1,524 @@ (* 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 Inter Union "(\)" "(\)" "(\)" "(\)" "{}" UNIV minus uminus Id "(O)" - apply (standard, simp_all add: O_assoc) - by blast + +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/MonoBoolTranAlgebra/Statements.thy b/thys/MonoBoolTranAlgebra/Statements.thy --- a/thys/MonoBoolTranAlgebra/Statements.thy +++ b/thys/MonoBoolTranAlgebra/Statements.thy @@ -1,445 +1,437 @@ 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 (in Lattices.boolean_algebra) sup_neg_inf: - "(p \ q \ r) = (p \ -q \ r)" - apply (safe) - apply(cut_tac a = p and c = "q \ r" and b = "-q" and d = "-q" in inf_mono) - apply simp apply simp apply (simp add: inf_sup_distrib2) - apply(cut_tac b = "p \ - q" and d = "r" and a = "q" and c = "q" in sup_mono) - apply simp apply simp by (simp add: sup_inf_distrib) - 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) end diff --git a/thys/Order_Lattice_Props/Order_Duality.thy b/thys/Order_Lattice_Props/Order_Duality.thy --- a/thys/Order_Lattice_Props/Order_Duality.thy +++ b/thys/Order_Lattice_Props/Order_Duality.thy @@ -1,315 +1,315 @@ (* Title: Ad-Hoc Duality for Orderings and Lattices Author: Georg Struth Maintainer: Georg Struth *) section \Ad-Hoc Duality for Orderings and Lattices\ theory Order_Duality imports Sup_Lattice begin text \This component presents an "explicit" formalisation of order and lattice duality. It augments the data type based one used by Wenzel in his lattice components \cite{Wenzel}, and complements the "implicit" formalisation given by locales. It uses a functor dual, supplied within a type class, which is simply a bijection (isomorphism) between types, with the constraint that the dual of a dual object is the original object. In Wenzel's formalisation, by contrast, dual is a bijection, but not idempotent or involutive. In the past, Preoteasa has used a similar approach with Isabelle~\cite{Preoteasa11b}.\ text \Duality is such a fundamental concept in order and lattice theory that it probably deserves to be included in the type classes for these objects, as in this section.\ class dual = fixes dual :: "'a \ 'a" ("\") assumes inj_dual: "inj \" and invol_dual [simp]: "\ \ \ = id" text \This type class allows one to define a type dual. It is actually a dependent type for which dual can be instantiated.\ typedef (overloaded) 'a dual = "range (dual::'a::dual \ 'a)" by fastforce setup_lifting type_definition_dual text \At the moment I have no use for this type.\ context dual begin lemma invol_dual_var [simp]: "\ (\ x) = x" by (simp add: pointfree_idE) lemma surj_dual: "surj \" unfolding surj_def by (metis invol_dual_var) lemma bij_dual: "bij \" by (simp add: bij_def inj_dual surj_dual) lemma inj_dual_iff: "(\ x = \ y) = (x = y)" by (meson inj_dual injD) lemma dual_iff: "(\ x = y) = (x = \ y)" by auto lemma the_inv_dual: "the_inv \ = \" by (metis comp_apply id_def invol_dual_var inj_dual surj_dual surj_fun_eq the_inv_f_o_f_id) end text \In boolean algebras, duality is of course De Morgan duality and can be expressed within the language.\ -sublocale Lattices.boolean_algebra \ ba_dual: dual "uminus" +sublocale boolean_algebra \ ba_dual: dual "uminus" by (unfold_locales, simp_all add: inj_def) definition map_dual:: "('a \ 'b) \ 'a::dual \ 'b::dual" ("\\<^sub>F") where "\\<^sub>F f = \ \ f \ \" lemma map_dual_func1: "\\<^sub>F (f \ g) = \\<^sub>F f \ \\<^sub>F g" by (metis (no_types, lifting) comp_assoc comp_id invol_dual map_dual_def) lemma map_dual_func2 [simp]: "\\<^sub>F id = id" by (simp add: map_dual_def) lemma map_dual_nat_iso: "\\<^sub>F f \ \ = \ \ id f" by (simp add: comp_assoc map_dual_def) lemma map_dual_invol [simp]: "\\<^sub>F \ \\<^sub>F = id" unfolding map_dual_def comp_def fun_eq_iff by simp text \Thus map-dual is naturally isomorphic to the identify functor: The function dual is a natural transformation between map-dual and the identity functor, and, because it has a two-sided inverse --- itself, it is a natural isomorphism.\ text \The generic function set-dual provides another natural transformation (see below). Before introducing it, we introduce useful notation for a widely used function.\ abbreviation "\ \ (\x. {x})" lemma eta_inj: "inj \" by simp definition "set_dual = \ \ \" lemma set_dual_prop: "set_dual (\ x) = {x}" by (metis comp_apply dual_iff set_dual_def) text \The next four lemmas show that (functional) image and preimage are functors (on functions). This does not really belong here, but it is useful for what follows. The interaction between duality and (pre)images is needed in applications.\ lemma image_func1: "(`) (f \ g) = (`) f \ (`) g" unfolding fun_eq_iff by (simp add: image_comp) lemma image_func2: "(`) id = id" by simp lemma vimage_func1: "(-`) (f \ g) = (-`) g \ (-`) f" unfolding fun_eq_iff by (simp add: vimage_comp) lemma vimage_func2: "(-`) id = id" by simp lemma iso_image: "mono ((`) f)" by (simp add: image_mono monoI) lemma iso_preimage: "mono ((-`) f)" by (simp add: monoI vimage_mono) context dual begin lemma image_dual [simp]: "(`) \ \ (`) \ = id" by (metis image_func1 image_func2 invol_dual) lemma vimage_dual [simp]: "(-`) \ \ (-`) \ = id" by (simp add: set.comp) end text \The following natural transformation between the powerset functor (image) and the identity functor is well known.\ lemma power_set_func_nat_trans: "\ \ id f = (`) f \ \" unfolding fun_eq_iff comp_def by simp text \As an instance, set-dual is a natural transformation with built-in type coercion.\ lemma dual_singleton: "(`) \ \ \ = \ \ \" by auto lemma finite_dual [simp]: "finite \ (`) \ = finite" unfolding fun_eq_iff comp_def using inj_dual finite_vimageI inj_vimage_image_eq by fastforce lemma finite_dual_var [simp]: "finite (\ ` X) = finite X" by (metis comp_def finite_dual) lemma subset_dual: "(X = \ ` Y) = (\ ` X = Y)" by (metis image_dual pointfree_idE) lemma subset_dual1: "(X \ Y) = (\ ` X \ \ ` Y)" by (simp add: inj_dual inj_image_subset_iff) lemma dual_empty [simp]: "\ ` {} = {}" by simp lemma dual_UNIV [simp]: "\ ` UNIV = UNIV" by (simp add: surj_dual) lemma fun_dual1: "(f = g \ \) = (f \ \ = g)" by (metis comp_assoc comp_id invol_dual) lemma fun_dual2: "(f = \ \ g) = (\ \ f = g)" by (metis comp_assoc fun.map_id invol_dual) lemma fun_dual3: "(f = g \ (`) \) = (f \ (`) \ = g)" by (metis comp_id image_dual o_assoc) lemma fun_dual4: "(f = (`) \ \ g) = ((`) \ \ f = g)" by (metis comp_assoc id_comp image_dual) lemma fun_dual5: "(f = \ \ g \ \) = (\ \ f \ \ = g)" by (metis comp_assoc fun_dual1 fun_dual2) lemma fun_dual6: "(f = (`) \ \ g \ (`) \) = ((`) \ \ f \ (`) \ = g)" by (simp add: comp_assoc fun_dual3 fun_dual4) lemma fun_dual7: "(f = \ \ g \ (`) \) = (\ \ f \ (`) \ = g)" by (simp add: comp_assoc fun_dual2 fun_dual3) lemma fun_dual8: "(f = (`) \ \ g \ \) = ((`) \ \ f \ \ = g)" by (simp add: comp_assoc fun_dual1 fun_dual4) lemma map_dual_dual: "(\\<^sub>F f = g) = (\\<^sub>F g = f)" by (metis map_dual_invol pointfree_idE) text \The next facts show incrementally that the dual of a complete lattice is a complete lattice.\ class ord_with_dual = dual + ord + assumes ord_dual: "x \ y \ \ y \ \ x" begin lemma dual_dual_ord: "(\ x \ \ y) = (y \ x)" by (metis dual_iff ord_dual) end lemma ord_pres_dual: fixes f :: "'a::ord_with_dual \ 'b::ord_with_dual" shows "ord_pres f \ ord_pres (\\<^sub>F f)" by (simp add: dual_dual_ord map_dual_def ord_pres_def) lemma map_dual_anti: "(f::'a::ord_with_dual \ 'b::ord_with_dual) \ g \ \\<^sub>F g \ \\<^sub>F f" by (simp add: le_fun_def map_dual_def ord_dual) class preorder_with_dual = ord_with_dual + preorder begin lemma less_dual_def_var: "(\ y < \ x) = (x < y)" by (simp add: dual_dual_ord less_le_not_le) end class order_with_dual = preorder_with_dual + order lemma iso_map_dual: fixes f :: "'a::order_with_dual \ 'b::order_with_dual" shows "mono f \ mono (\\<^sub>F f)" by (simp add: ord_pres_dual ord_pres_mono) class lattice_with_dual = lattice + dual + assumes sup_dual_def: "\ (x \ y) = \ x \ \ y" begin subclass order_with_dual by (unfold_locales, metis inf.absorb_iff2 sup_absorb1 sup_commute sup_dual_def) lemma inf_dual: "\ (x \ y) = \ x \ \ y" by (metis invol_dual_var sup_dual_def) lemma inf_to_sup: "x \ y = \ (\ x \ \ y)" using inf_dual dual_iff by fastforce lemma sup_to_inf: "x \ y = \ (\ x \ \ y)" by (simp add: inf_dual) end class bounded_lattice_with_dual = lattice_with_dual + bounded_lattice begin lemma bot_dual: "\ \ = \" by (metis dual_dual_ord dual_iff le_bot top_greatest) lemma top_dual: "\ \ = \" using bot_dual dual_iff by force end class boolean_algebra_with_dual = lattice_with_dual + boolean_algebra -sublocale Lattices.boolean_algebra \ badual: boolean_algebra_with_dual _ _ _ _ _ _ _ _ uminus +sublocale boolean_algebra \ badual: boolean_algebra_with_dual _ _ _ _ _ _ _ _ uminus by unfold_locales simp_all class Sup_lattice_with_dual = Sup_lattice + dual + assumes Sups_dual_def: "\ \ Sup = Infs \ (`) \" class Inf_lattice_with_dual = Inf_lattice + dual + assumes Sups_dual_def: "\ \ Supi = Inf \ (`) \" class complete_lattice_with_dual = complete_lattice + dual + assumes Sups_dual_def: "\ \ Sup = Inf \ (`) \" sublocale Sup_lattice_with_dual \ sclatd: complete_lattice_with_dual Infs Sup infs "(\)" le sups bots tops "\" by (unfold_locales, simp add: Sups_dual_def) sublocale Inf_lattice_with_dual \ iclatd: complete_lattice_with_dual Inf Supi infi "(\)" le supi boti topi "\" by (unfold_locales, simp add: Sups_dual_def) context complete_lattice_with_dual begin lemma Inf_dual: "\ \ Inf = Sup \ (`) \" by (metis comp_assoc comp_id fun.map_id Sups_dual_def image_dual invol_dual) lemma Inf_dual_var: "\ (\X) = \(\ ` X)" using comp_eq_dest Inf_dual by fastforce lemma Inf_to_Sup: "Inf = \ \ Sup \ (`) \" by (auto simp add: Sups_dual_def image_comp) lemma Inf_to_Sup_var: "\X = \ (\(\ ` X))" using Inf_dual_var dual_iff by fastforce lemma Sup_to_Inf: "Sup = \ \ Inf \ (`) \" by (auto simp add: Inf_dual image_comp) lemma Sup_to_Inf_var: "\X = \ (\(\ ` X))" using Sup_to_Inf by force lemma Sup_dual_def_var: "\ (\X) = \ (\ ` X)" using comp_eq_dest Sups_dual_def by fastforce lemma bot_dual_def: "\ \ = \" by (smt Inf_UNIV Sup_UNIV Sups_dual_def surj_dual o_eq_dest) lemma top_dual_def: "\ \ = \" using bot_dual_def dual_iff by blast lemma inf_dual2: "\ (x \ y) = \ x \ \ y" by (smt comp_eq_elim Inf_dual Inf_empty Inf_insert SUP_insert inf_top.right_neutral) lemma sup_dual: "\ (x \ y) = \ x \ \ y" by (metis inf_dual2 dual_iff) subclass lattice_with_dual by (unfold_locales, auto simp: inf_dual sup_dual) subclass bounded_lattice_with_dual.. end end diff --git a/thys/Order_Lattice_Props/Order_Lattice_Props.thy b/thys/Order_Lattice_Props/Order_Lattice_Props.thy --- a/thys/Order_Lattice_Props/Order_Lattice_Props.thy +++ b/thys/Order_Lattice_Props/Order_Lattice_Props.thy @@ -1,1269 +1,1227 @@ (* Title: Properties of Orderings and Lattices Author: Georg Struth Maintainer: Georg Struth *) section \Properties of Orderings and Lattices\ theory Order_Lattice_Props imports Order_Duality begin subsection \Basic Definitions for Orderings and Lattices\ text \The first definition is for order morphisms --- isotone (order-preserving, monotone) functions. An order isomorphism is an order-preserving bijection. This should be defined in the class ord, but mono requires order.\ definition ord_homset :: "('a::order \ 'b::order) set" where "ord_homset = {f::'a::order \ 'b::order. mono f}" definition ord_embed :: "('a::order \ 'b::order) \ bool" where "ord_embed f = (\x y. f x \ f y \ x \ y)" definition ord_iso :: "('a::order \ 'b::order) \ bool" where "ord_iso = bij \ mono \ (mono \ the_inv)" lemma ord_embed_alt: "ord_embed f = (mono f \ (\x y. f x \ f y \ x \ y))" using mono_def ord_embed_def by auto lemma ord_embed_homset: "ord_embed f \ f \ ord_homset" by (simp add: mono_def ord_embed_def ord_homset_def) lemma ord_embed_inj: "ord_embed f \ inj f" unfolding ord_embed_def inj_def by (simp add: eq_iff) lemma ord_iso_ord_embed: "ord_iso f \ ord_embed f" unfolding ord_iso_def ord_embed_def bij_def inj_def mono_def by (clarsimp, metis inj_def the_inv_f_f) lemma ord_iso_alt: "ord_iso f = (ord_embed f \ surj f)" unfolding ord_iso_def ord_embed_def surj_def bij_def inj_def mono_def apply safe by simp_all (metis eq_iff inj_def the_inv_f_f)+ lemma ord_iso_the_inv: "ord_iso f \ mono (the_inv f)" by (simp add: ord_iso_def) lemma ord_iso_inv1: "ord_iso f \ (the_inv f) \ f = id" using ord_embed_inj ord_iso_ord_embed the_inv_into_f_f by fastforce lemma ord_iso_inv2: "ord_iso f \ f \ (the_inv f) = id" using f_the_inv_into_f ord_embed_inj ord_iso_alt by fastforce typedef (overloaded) ('a,'b) ord_homset = "ord_homset::('a::order \ 'b::order) set" by (force simp: ord_homset_def mono_def) setup_lifting type_definition_ord_homset text \The next definition is for the set of fixpoints of a given function. It is important in the context of orders, for instance for proving Tarski's fixpoint theorem, but does not really belong here.\ definition Fix :: "('a \ 'a) \ 'a set" where "Fix f = {x. f x = x}" lemma retraction_prop: "f \ f = f \ f x = x \ x \ range f" by (metis comp_apply f_inv_into_f rangeI) lemma retraction_prop_fix: "f \ f = f \ range f = Fix f" unfolding Fix_def using retraction_prop by fastforce lemma Fix_map_dual: "Fix \ \\<^sub>F = (`) \ \ Fix" unfolding Fix_def map_dual_def comp_def fun_eq_iff by (smt Collect_cong invol_dual pointfree_idE setcompr_eq_image) lemma Fix_map_dual_var: "Fix (\\<^sub>F f) = \ ` (Fix f)" by (metis Fix_map_dual o_def) lemma gfp_dual: "(\::'a::complete_lattice_with_dual \ 'a) \ gfp = lfp \ \\<^sub>F" proof- {fix f:: "'a \ 'a" have "\ (gfp f) = \ (\{u. u \ f u})" by (simp add: gfp_def) also have "... = \(\ ` {u. u \ f u})" by (simp add: Sup_dual_def_var) also have "... = \{\ u |u. u \ f u}" by (simp add: setcompr_eq_image) also have "... = \{u |u. (\\<^sub>F f) u \ u}" by (metis (no_types, opaque_lifting) dual_dual_ord dual_iff map_dual_def o_def) finally have "\ (gfp f) = lfp (\\<^sub>F f)" by (metis lfp_def)} thus ?thesis by auto qed lemma gfp_dual_var: fixes f :: "'a::complete_lattice_with_dual \ 'a" shows "\ (gfp f) = lfp (\\<^sub>F f)" using comp_eq_elim gfp_dual by blast lemma gfp_to_lfp: "gfp = (\::'a::complete_lattice_with_dual \ 'a) \ lfp \ \\<^sub>F" by (simp add: comp_assoc fun_dual2 gfp_dual) lemma gfp_to_lfp_var: fixes f :: "'a::complete_lattice_with_dual \ 'a" shows "gfp f = \ (lfp (\\<^sub>F f))" by (metis gfp_dual_var invol_dual_var) lemma lfp_dual: "(\::'a::complete_lattice_with_dual \ 'a) \ lfp = gfp \ \\<^sub>F" by (simp add: comp_assoc gfp_to_lfp map_dual_invol) lemma lfp_dual_var: fixes f :: "'a::complete_lattice_with_dual \ 'a" shows "\ (lfp f) = gfp (map_dual f)" using comp_eq_dest_lhs lfp_dual by fastforce lemma lfp_to_gfp: "lfp = (\::'a::complete_lattice_with_dual \ 'a) \ gfp \ \\<^sub>F" by (simp add: comp_assoc gfp_dual map_dual_invol) lemma lfp_to_gfp_var: fixes f :: "'a::complete_lattice_with_dual \ 'a" shows "lfp f = \ (gfp (\\<^sub>F f))" by (metis invol_dual_var lfp_dual_var) lemma lfp_in_Fix: fixes f :: "'a::complete_lattice \ 'a" shows "mono f \ lfp f \ Fix f" by (metis (mono_tags, lifting) Fix_def lfp_unfold mem_Collect_eq) lemma gfp_in_Fix: fixes f :: "'a::complete_lattice \ 'a" shows "mono f \ gfp f \ Fix f" by (metis (mono_tags, lifting) Fix_def gfp_unfold mem_Collect_eq) lemma nonempty_Fix: fixes f :: "'a::complete_lattice \ 'a" shows "mono f \ Fix f \ {}" using lfp_in_Fix by fastforce text \Next the minimal and maximal elements of an ordering are defined.\ 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}" end context ord_with_dual begin lemma min_max_set_dual: "(`) \ \ min_set = max_set \ (`) \" unfolding max_set_def min_set_def fun_eq_iff comp_def apply safe using dual_dual_ord inj_dual_iff by auto lemma min_max_set_dual_var: "\ ` (min_set X) = max_set (\ ` X)" using comp_eq_dest min_max_set_dual by fastforce lemma max_min_set_dual: "(`) \ \ max_set = min_set \ (`) \" by (metis (no_types, opaque_lifting) comp_id fun.map_comp id_comp image_dual min_max_set_dual) lemma min_to_max_set: "min_set = (`) \ \ max_set \ (`) \" by (metis comp_id image_dual max_min_set_dual o_assoc) lemma max_min_set_dual_var: "\ ` (max_set X) = min_set (\ ` X)" using comp_eq_dest max_min_set_dual by fastforce lemma min_to_max_set_var: "min_set X = \ ` (max_set (\ ` X))" by (simp add: max_min_set_dual_var pointfree_idE) 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}" definition downset :: "'a \ 'a set" ("\") where "\ = \ \ \" definition upset :: "'a \ 'a set" ("\") where "\ = \ \ \" definition downsets :: "'a set set" where "downsets = Fix \" definition upsets :: "'a set set" where "upsets = Fix \" definition "downclosed_set X = (X \ downsets)" definition "upclosed_set X = (X \ upsets)" definition ideals :: "'a set set" where "ideals = {X. X \ {} \ downclosed_set X \ directed X}" definition filters :: "'a set set" where "filters = {X. X \ {} \ upclosed_set X \ filtered X}" abbreviation "idealp X \ X \ ideals" abbreviation "filterp X \ X \ filters" end text \These notions are pair-wise dual.\ text \Filtered and directed sets are dual.\ context ord_with_dual begin lemma filtered_directed_dual: "filtered \ (`) \ = directed" unfolding filtered_def directed_def fun_eq_iff comp_def apply clarsimp apply safe apply (meson finite_imageI imageI image_mono dual_dual_ord) by (smt finite_subset_image imageE ord_dual) lemma directed_filtered_dual: "directed \ (`) \ = filtered" using filtered_directed_dual by (metis comp_id image_dual o_assoc) lemma filtered_to_directed: "filtered X = directed (\ ` X)" by (metis comp_apply directed_filtered_dual) text \Upsets and downsets are dual.\ lemma downset_set_upset_set_dual: "(`) \ \ \ = \ \ (`) \" unfolding downset_set_def upset_set_def fun_eq_iff comp_def apply safe apply (meson image_eqI ord_dual) by (clarsimp, metis (mono_tags, lifting) dual_iff image_iff mem_Collect_eq ord_dual) lemma upset_set_downset_set_dual: "(`) \ \ \ = \ \ (`) \" using downset_set_upset_set_dual by (metis (no_types, opaque_lifting) comp_id id_comp image_dual o_assoc) lemma upset_set_to_downset_set: "\ = (`) \ \ \ \ (`) \" by (simp add: comp_assoc downset_set_upset_set_dual) lemma upset_set_to_downset_set2: "\ X = \ ` (\ (\ ` X))" by (simp add: upset_set_to_downset_set) lemma downset_upset_dual: "(`) \ \ \ = \ \ \" using downset_def upset_def upset_set_to_downset_set by fastforce lemma upset_to_downset: "(`) \ \ \ = \ \ \" by (metis comp_assoc id_apply ord.downset_def ord.upset_def power_set_func_nat_trans upset_set_downset_set_dual) lemma upset_to_downset2: "\ = (`) \ \ \ \ \" by (simp add: comp_assoc downset_upset_dual) lemma upset_to_downset3: "\ x = \ ` (\ (\ x))" by (simp add: upset_to_downset2) lemma downsets_upsets_dual: "(X \ downsets) = (\ ` X \ upsets)" unfolding downsets_def upsets_def Fix_def by (smt comp_eq_dest downset_set_upset_set_dual image_inv_f_f inj_dual mem_Collect_eq) lemma downset_setp_upset_setp_dual: "upclosed_set \ (`) \ = downclosed_set" unfolding downclosed_set_def upclosed_set_def using downsets_upsets_dual by fastforce lemma upsets_to_downsets: "(X \ upsets) = (\ ` X \ downsets)" by (simp add: downsets_upsets_dual image_comp) lemma upset_setp_downset_setp_dual: "downclosed_set \ (`) \ = upclosed_set" by (metis comp_id downset_setp_upset_setp_dual image_dual o_assoc) text \Filters and ideals are dual.\ lemma ideals_filters_dual: "(X \ ideals) = ((\ ` X) \ filters)" by (smt comp_eq_dest_lhs directed_filtered_dual image_inv_f_f image_is_empty inv_unique_comp filters_def ideals_def inj_dual invol_dual mem_Collect_eq upset_setp_downset_setp_dual) lemma idealp_filterp_dual: "idealp = filterp \ (`) \" unfolding fun_eq_iff by (simp add: ideals_filters_dual) lemma filters_to_ideals: "(X \ filters) = ((\ ` X) \ ideals)" by (simp add: ideals_filters_dual image_comp) lemma filterp_idealp_dual: "filterp = idealp \ (`) \" unfolding fun_eq_iff by (simp add: filters_to_ideals) end subsection \Properties of Orderings\ 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 by fastforce lemma downset_prop2: "y \ x \ y \ \x" by (simp add: downset_prop) lemma ideals_downsets: "X \ ideals \ X \ downsets" by (simp add: downclosed_set_def ideals_def) lemma ideals_directed: "X \ ideals \ directed X" by (simp add: ideals_def) 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_prop_var2: "x \ \X \ y \ x \ y \ \X" unfolding downset_set_def using order_trans by blast lemma downclosed_set_iff: "downclosed_set X = (\x \ X. \y. y \ x \ y \ X)" unfolding downclosed_set_def downsets_def Fix_def downset_set_def by auto lemma downclosed_downset_set: "downclosed_set (\X)" by (simp add: downclosed_set_iff downset_set_prop_var2 downset_def) lemma downclosed_downset: "downclosed_set (\x)" by (simp add: downclosed_downset_set downset_def) 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 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 text \The following proof uses the Axiom of Choice.\ 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 by simp lemma directed_downset_ideals: "directed (\X) = (\X \ ideals)" by (metis (mono_tags, lifting) CollectI Fix_def directed_alt downset_set_idem downclosed_set_def downsets_def ideals_def o_def ord.ideals_directed) lemma downclosed_Fix: "downclosed_set X = (\X = X)" by (metis (mono_tags, lifting) CollectD Fix_def downclosed_downset_set downclosed_set_def downsets_def) end lemma downset_iso: "mono (\::'a::order \ 'a set)" by (simp add: downset_iso_iff mono_def) lemma mono_downclosed: fixes f :: "'a::order \ 'b::order" assumes "mono f" shows "\Y. downclosed_set Y \ downclosed_set (f -` Y)" by (simp add: assms downclosed_set_iff monoD) lemma fixes f :: "'a::order \ 'b::order" assumes "mono f" shows "\Y. downclosed_set X \ downclosed_set (f ` X)" (*nitpick*) oops lemma downclosed_mono: fixes f :: "'a::order \ 'b::order" assumes "\Y. downclosed_set Y \ downclosed_set (f -` Y)" shows "mono f" proof- {fix x y :: "'a::order" assume h: "x \ y" have "downclosed_set (\ (f y))" unfolding downclosed_set_def downsets_def Fix_def downset_set_def downset_def by auto hence "downclosed_set (f -` (\ (f y)))" by (simp add: assms) hence "downclosed_set {z. f z \ f y}" unfolding vimage_def downset_def downset_set_def by auto hence "\z w. (f z \ f y \ w \ z) \ f w \ f y" unfolding downclosed_set_def downclosed_set_def downsets_def Fix_def downset_set_def by force hence "f x \ f y" using h by blast} thus ?thesis.. qed lemma mono_downclosed_iff: "mono f = (\Y. downclosed_set Y \ downclosed_set (f -` Y))" using mono_downclosed downclosed_mono by auto context order begin lemma downset_inj: "inj \" by (metis injI downset_iso_iff order.eq_iff) lemma "(X \ Y) = (\X \ \Y)" (*nitpick*) oops 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 downclosed_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 downclosed_set_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 subsection \Dual Properties of Orderings\ context ord_with_dual begin lemma filtered_nonempty: "filtered X \ X \ {}" using filtered_to_directed ord.directed_nonempty by auto lemma filtered_lb: "filtered X \ (\x \ X. \y \ X. \z \ X. z \ x \ z \ y)" using filtered_to_directed directed_ub dual_dual_ord by fastforce lemma upset_set_prop_var: "\X = (\x \ X. \x)" by (simp add: image_Union downset_set_prop_var upset_set_to_downset_set2 upset_to_downset2) lemma upset_set_prop: "\ = Union \ (`) \" unfolding fun_eq_iff by (simp add: upset_set_prop_var) lemma upset_prop: "\x = {y. x \ y}" unfolding upset_to_downset3 downset_prop image_def using dual_dual_ord by fastforce lemma upset_prop2: "x \ y \ y \ \x" by (simp add: upset_prop) lemma filters_upsets: "X \ filters \ X \ upsets" by (simp add: upclosed_set_def filters_def) lemma filters_filtered: "X \ filters \ filtered X" by (simp add: filters_def) end context preorder_with_dual begin lemma filtered_prop: "X \ {} \ (\x \ X. \y \ X. \z \ X. z \ x \ z \ y) \ filtered X" unfolding filtered_to_directed by (rule directed_prop, blast, metis (full_types) image_iff ord_dual) lemma filtered_alt: "filtered X = (X \ {} \ (\x \ X. \y \ X. \z \ X. z \ x \ z \ y))" by (metis image_empty directed_alt filtered_to_directed filtered_lb filtered_prop) lemma up_set_prop_var2: "x \ \X \ x \ y \ y \ \X" using downset_set_prop_var2 dual_iff ord_dual upset_set_to_downset_set2 by fastforce lemma upclosed_set_iff: "upclosed_set X = (\x \ X. \y. x \ y \ y \ X)" unfolding upclosed_set_def upsets_def Fix_def upset_set_def by auto lemma upclosed_upset_set: "upclosed_set (\X)" using up_set_prop_var2 upclosed_set_iff by blast lemma upclosed_upset: "upclosed_set (\x)" by (simp add: upset_def upclosed_upset_set) lemma upset_set_ext: "id \ \" by (smt comp_def comp_id image_mono le_fun_def downset_set_ext image_dual upset_set_to_downset_set2) lemma upset_set_anti: "mono \" by (metis image_mono downset_set_iso upset_set_to_downset_set2 mono_def) lemma up_set_idem [simp]: "\ \ \ = \" by (metis comp_assoc downset_set_idem upset_set_downset_set_dual upset_set_to_downset_set) lemma upset_faithful: "\x \ \y \ y \ x" by (metis inj_image_subset_iff downset_faithful dual_dual_ord inj_dual upset_to_downset3) lemma upset_anti_iff: "(\y \ \x) = (x \ y)" by (metis downset_iso_iff ord_dual upset_to_downset3 subset_image_iff upset_faithful) lemma upset_filtered_upset [simp]: "filtered \ \ = filtered" by (metis comp_assoc directed_filtered_dual downset_directed_downset upset_set_downset_set_dual) lemma filtered_upset_filters: "filtered (\X) = (\X \ filters)" by (metis comp_apply directed_downset_ideals filtered_to_directed filterp_idealp_dual upset_set_downset_set_dual) lemma upclosed_Fix: "upclosed_set X = (\X = X)" by (simp add: Fix_def upclosed_set_def upsets_def) end lemma upset_anti: "antimono (\::'a::order_with_dual \ 'a set)" by (simp add: antimono_def upset_anti_iff) lemma mono_upclosed: fixes f :: "'a::order_with_dual \ 'b::order_with_dual" assumes "mono f" shows "\Y. upclosed_set Y \ upclosed_set (f -` Y)" by (simp add: assms monoD upclosed_set_iff) lemma mono_upclosed: fixes f :: "'a::order_with_dual \ 'b::order_with_dual" assumes "mono f" shows "\Y. upclosed_set X \ upclosed_set (f ` X)" (*nitpick*) oops lemma upclosed_mono: fixes f :: "'a::order_with_dual \ 'b::order_with_dual" assumes "\Y. upclosed_set Y \ upclosed_set (f -` Y)" shows "mono f" by (metis (mono_tags, lifting) assms dual_order.refl mem_Collect_eq monoI order.trans upclosed_set_iff vimageE vimageI2) lemma mono_upclosed_iff: fixes f :: "'a::order_with_dual \ 'b::order_with_dual" shows "mono f = (\Y. upclosed_set Y \ upclosed_set (f -` Y))" using mono_upclosed upclosed_mono by auto context order_with_dual begin lemma upset_inj: "inj \" by (metis inj_compose inj_on_imageI2 downset_inj inj_dual upset_to_downset) lemma "(X \ Y) = (\Y \ \X)" (*nitpick*) oops end context lattice_with_dual begin lemma lat_filters: "X \ filters = (X \ {} \ X \ upsets \ (\x \ X. \ y \ X. x \ y \ X))" unfolding filters_to_ideals upsets_to_downsets inf_to_sup lat_ideals by (smt image_iff image_inv_f_f image_is_empty inj_image_mem_iff inv_unique_comp inj_dual invol_dual) end context bounded_lattice_with_dual begin lemma top_filter: "X \ filters \ \ \ X" using bot_ideal inj_image_mem_iff inj_dual filters_to_ideals top_dual by fastforce end context complete_lattice_with_dual begin lemma Inf_upset_id [simp]: "Inf \ \ = id" by (metis comp_assoc comp_id Sup_downset_id Sups_dual_def downset_upset_dual invol_dual) lemma upset_Inf_id: "id \ \ \ Inf" by (simp add: Inf_lower le_funI subsetI upset_prop) lemma Sup_Inf_var: " \(\x \ X. \x) = \X" unfolding upset_prop by (simp add: Collect_ball_eq Sup_eq_Inf) lemma Sup_dual_upset_var: "(\x \ X. \x) = \(\X)" unfolding upset_prop by (safe, simp_all add: Sup_le_iff) end -subsection \Shunting Laws\ - -text \The first set of laws supplies so-called shunting laws for boolean algebras. -Such laws rather belong into Isabelle Main.\ - -context Lattices.boolean_algebra -begin - -lemma shunt1: "(x \ y \ z) = (x \ -y \ z)" -proof standard - assume "x \ y \ z" - hence "-y \ (x \ y) \ -y \ z" - using sup.mono by blast - hence "-y \ x \ -y \ z" - by (simp add: sup_inf_distrib1) - thus "x \ -y \ z" - by simp -next - assume "x \ -y \ z" - hence "x \ y \ (-y \ z) \ y" - using inf_mono by auto - thus "x \ y \ z" - using inf.boundedE inf_sup_distrib2 by auto -qed - -lemma shunt2: "(x \ -y \ z) = (x \ y \ z)" - by (simp add: shunt1) - -lemma meet_shunt: "(x \ y = \) = (x \ -y)" - by (simp add: order.eq_iff shunt1) - -lemma join_shunt: "(x \ y = \) = (-x \ y)" - by (metis compl_sup compl_top_eq double_compl meet_shunt) - -lemma meet_shunt_var: "(x - y = \) = (x \ y)" - by (simp add: diff_eq meet_shunt) - -lemma join_shunt_var: "(x \ y = \) = (x \ y)" - by simp - -end - subsection \Properties of Complete Lattices\ definition "Inf_closed_set X = (\Y \ X. \Y \ X)" definition "Sup_closed_set X = (\Y \ X. \Y \ X)" definition "inf_closed_set X = (\x \ X. \y \ X. x \ y \ X)" definition "sup_closed_set X = (\x \ X. \y \ X. x \ y \ X)" text \The following facts about complete lattices add to those in the Isabelle libraries.\ context complete_lattice begin text \The translation between sup and Sup could be improved. The sup-theorems should be direct consequences of Sup-ones. In addition, duality between sup and inf is currently not exploited.\ lemma sup_Sup: "x \ y = \{x,y}" by simp lemma inf_Inf: "x \ y = \{x,y}" by simp text \The next two lemmas are about Sups and Infs of indexed families. These are interesting for iterations and fixpoints.\ 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 Sup_sup_closed: "Sup_closed_set (X::'a::complete_lattice set) \ sup_closed_set X" by (metis Sup_closed_set_def empty_subsetI insert_subsetI sup_Sup sup_closed_set_def) lemma Inf_inf_closed: "Inf_closed_set (X::'a::complete_lattice set) \ inf_closed_set X" by (metis Inf_closed_set_def empty_subsetI inf_Inf inf_closed_set_def insert_subset) subsection \Sup- and Inf-Preservation\ text \Next, important notation for morphism between posets and lattices is introduced: sup-preservation, inf-preservation and related properties.\ 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 \ = \" text \Inf-preservation and sup-preservation relate with duality.\ lemma Inf_pres_map_dual_var: "Inf_pres f = Sup_pres (\\<^sub>F f)" for f :: "'a::complete_lattice_with_dual \ 'b::complete_lattice_with_dual" proof - { fix x :: "'a set" assume "\ (f (\ (\ ` x))) = (\y\x. \ (f (\ y)))" for x then have "\ (f ` \ ` A) = f (\ (\ A))" for A by (metis (no_types) Sup_dual_def_var image_image invol_dual_var subset_dual) then have "\ (f ` x) = f (\ x)" by (metis Sup_dual_def_var subset_dual) } then show ?thesis by (auto simp add: map_dual_def fun_eq_iff Inf_dual_var Sup_dual_def_var image_comp) qed lemma Inf_pres_map_dual: "Inf_pres = Sup_pres \ (\\<^sub>F::('a::complete_lattice_with_dual \ 'b::complete_lattice_with_dual) \ 'a \ 'b)" proof- {fix f::"'a \ 'b" have "Inf_pres f = (Sup_pres \ \\<^sub>F) f" by (simp add: Inf_pres_map_dual_var)} thus ?thesis by force qed lemma Sup_pres_map_dual_var: fixes f :: "'a::complete_lattice_with_dual \ 'b::complete_lattice_with_dual" shows "Sup_pres f = Inf_pres (\\<^sub>F f)" by (metis Inf_pres_map_dual_var fun_dual5 map_dual_def) lemma Sup_pres_map_dual: "Sup_pres = Inf_pres \ (\\<^sub>F::('a::complete_lattice_with_dual \ 'b::complete_lattice_with_dual) \ 'a \ 'b)" by (simp add: Inf_pres_map_dual comp_assoc map_dual_invol) text \The following lemmas relate isotonicity of functions between complete lattices with weak (left) preservation properties of sups and infs.\ 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 Inf_empty Inf_insert comp_eq_elim 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 Inf_empty comp_eq_elim) lemma Sup_sup_dual: fixes f :: "'a::complete_lattice \ 'b::complete_lattice" shows "Sup_dual f \ sup_dual f" by (smt comp_eq_elim image_empty image_insert inf_Inf sup_Sup) lemma Inf_inf_dual: fixes f :: "'a::complete_lattice \ 'b::complete_lattice" shows "Inf_dual f \ inf_dual f" by (smt comp_eq_elim image_empty image_insert inf_Inf sup_Sup) lemma Sup_bot_dual: fixes f :: "'a::complete_lattice \ 'b::complete_lattice" shows "Sup_dual f \ bot_dual f" by (metis INF_empty Sup_empty comp_eq_elim) lemma Inf_top_dual: fixes f :: "'a::complete_lattice \ 'b::complete_lattice" shows "Inf_dual f \ top_dual f" by (metis Inf_empty SUP_empty comp_eq_elim) text \However, Inf-preservation does not imply top-preservation and Sup-preservation does not imply bottom-preservation.\ lemma fixes f :: "'a::complete_lattice \ 'b::complete_lattice" shows "Sup_pres f \ top_pres f" (*nitpick*) oops lemma fixes f :: "'a::complete_lattice \ 'b::complete_lattice" shows "Inf_pres f \ bot_pres f" (*nitpick*) oops context complete_lattice begin lemma iso_Inf_subdistl: fixes f :: "'a \ 'b::complete_lattice" shows "mono f \f \ Inf \ Inf \ (`) f" by (simp add: complete_lattice_class.le_Inf_iff le_funI Inf_lower monoD) lemma iso_Sup_supdistl: fixes f :: "'a \ 'b::complete_lattice" shows "mono f \ Sup \ (`) f \ f \ Sup" by (simp add: complete_lattice_class.Sup_le_iff le_funI Sup_upper 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 ord_iso_Inf_pres: fixes f :: "'a::complete_lattice \ 'b::complete_lattice" shows "ord_iso f \ Inf \ (`) f = f \ Inf" proof- let ?g = "the_inv f" assume h: "ord_iso f" hence a: "mono ?g" by (simp add: ord_iso_the_inv) {fix X :: "'a::complete_lattice set" {fix y :: "'b::complete_lattice" have "(y \ f (\X)) = (?g y \ \X)" by (metis (mono_tags, lifting) UNIV_I f_the_inv_into_f h monoD ord_embed_alt ord_embed_inj ord_iso_alt) also have "... = (\x \ X. ?g y \ x)" by (simp add: le_Inf_iff) also have "... = (\x \ X. y \ f x)" by (metis (mono_tags, lifting) UNIV_I f_the_inv_into_f h monoD ord_embed_alt ord_embed_inj ord_iso_alt) also have "... = (y \ \ (f ` X))" by (simp add: le_INF_iff) finally have "(y \ f (\X)) = (y \ \ (f ` X))".} hence "f (\X) = \ (f ` X)" by (meson dual_order.antisym order_refl)} thus ?thesis unfolding fun_eq_iff by simp qed lemma ord_iso_Sup_pres: fixes f :: "'a::complete_lattice \ 'b::complete_lattice" shows "ord_iso f \ Sup \ (`) f = f \ Sup" proof- let ?g = "the_inv f" assume h: "ord_iso f" hence a: "mono ?g" by (simp add: ord_iso_the_inv) {fix X :: "'a::complete_lattice set" {fix y :: "'b::complete_lattice" have "(f (\X) \ y) = (\X \ ?g y)" by (metis (mono_tags, lifting) UNIV_I f_the_inv_into_f h monoD ord_embed_alt ord_embed_inj ord_iso_alt) also have "... = (\x \ X. x \ ?g y)" by (simp add: Sup_le_iff) also have "... = (\x \ X. f x \ y)" by (metis (mono_tags, lifting) UNIV_I f_the_inv_into_f h monoD ord_embed_alt ord_embed_inj ord_iso_alt) also have "... = (\ (f ` X) \ y)" by (simp add: SUP_le_iff) finally have "(f (\X) \ y) = (\ (f ` X) \ y)".} hence "f (\X) = \ (f ` X)" by (meson dual_order.antisym order_refl)} thus ?thesis unfolding fun_eq_iff by simp qed text \Right preservation of sups and infs is trivial.\ lemma fSup_distr: "Sup_pres (\x. x \ f)" unfolding fun_eq_iff by (simp add: image_comp) lemma fSup_distr_var: "\F \ g = (\f \ F. f \ g)" unfolding fun_eq_iff by (simp add: image_comp) lemma fInf_distr: "Inf_pres (\x. x \ f)" unfolding fun_eq_iff comp_def by (smt INF_apply Inf_fun_def Sup.SUP_cong) lemma fInf_distr_var: "\F \ g = (\f \ F. f \ g)" unfolding fun_eq_iff comp_def by (smt INF_apply INF_cong INF_image Inf_apply image_comp image_def image_image) text \The next set of lemma revisits the preservation properties in the function space.\ lemma fSup_subdistl: assumes "mono (f::'a::complete_lattice \ 'b::complete_lattice)" shows "Sup \ (`) ((\) f) \ (\) f \ Sup" using assms by (simp add: fun_isol supdistl_iso) lemma fSup_subdistl_var: fixes f :: "'a::complete_lattice \ 'b::complete_lattice" shows "mono f \ (\g \ G. f \ g) \ f \ \G" by (simp add: fun_isol mono_Sup) lemma fInf_subdistl: fixes f :: "'a::complete_lattice \ 'b::complete_lattice" shows "mono f \ (\) f \ Inf \ Inf \ (`) ((\) f)" by (simp add: fun_isol subdistl_iso) lemma fInf_subdistl_var: fixes f :: "'a::complete_lattice \ 'b::complete_lattice" shows "mono f \ f \ \G \ (\g \ G. f \ g)" by (simp add: fun_isol mono_Inf) lemma fSup_distl: "Sup_pres f \ Sup_pres ((\) f)" unfolding fun_eq_iff by (simp add: image_comp) lemma fSup_distl_var: "Sup_pres f \ f \ \G = (\g \ G. f \ g)" unfolding fun_eq_iff by (simp add: image_comp) lemma fInf_distl: "Inf_pres f \ Inf_pres ((\) f)" unfolding fun_eq_iff by (simp add: image_comp) lemma fInf_distl_var: "Inf_pres f \ f \ \G = (\g \ G. f \ g)" unfolding fun_eq_iff by (simp add: image_comp) text \Downsets preserve infs whereas upsets preserve sups.\ lemma Inf_pres_downset: "Inf_pres (\::'a::complete_lattice_with_dual \ 'a set)" unfolding downset_prop fun_eq_iff by (safe, simp_all add: le_Inf_iff) lemma Sup_dual_upset: "Sup_dual (\::'a::complete_lattice_with_dual \ 'a set)" unfolding upset_prop fun_eq_iff by (safe, simp_all add: Sup_le_iff) text \Images of Sup-morphisms are closed under Sups and images of Inf-morphisms are closed under Infs.\ lemma Sup_pres_Sup_closed: "Sup_pres f \ Sup_closed_set (range f)" by (metis (mono_tags, lifting) Sup_closed_set_def comp_eq_elim range_eqI subset_image_iff) lemma Inf_pres_Inf_closed: "Inf_pres f \ Inf_closed_set (range f)" by (metis (mono_tags, lifting) Inf_closed_set_def comp_eq_elim range_eqI subset_image_iff) text \It is well known that functions into complete lattices form complete lattices. Here, such results are shown for the subclasses of isotone functions, where additional closure conditions must be respected.\ typedef (overloaded) 'a iso = "{f::'a::order \ 'a::order. mono f}" by (metis Abs_ord_homset_cases ord_homset_def) setup_lifting type_definition_iso instantiation iso :: (complete_lattice) complete_lattice begin lift_definition Inf_iso :: "'a::complete_lattice iso set \ 'a iso" is Sup by (metis (mono_tags, lifting) SUP_subset_mono Sup_apply mono_def subsetI) lift_definition Sup_iso :: "'a::complete_lattice iso set \ 'a iso" is Inf by (smt INF_lower2 Inf_apply le_INF_iff mono_def) lift_definition bot_iso :: "'a::complete_lattice iso" is "\" by (simp add: monoI) lift_definition sup_iso :: "'a::complete_lattice iso \ 'a iso \ 'a iso" is inf by (smt inf_apply inf_mono monoD monoI) lift_definition top_iso :: "'a::complete_lattice iso" is "\" by (simp add: mono_def) lift_definition inf_iso :: "'a::complete_lattice iso \ 'a iso \ 'a iso" is sup by (smt mono_def sup.mono sup_apply) lift_definition less_eq_iso :: "'a::complete_lattice iso \ 'a iso \ bool" is "(\)". lift_definition less_iso :: "'a::complete_lattice iso \ 'a iso \ bool" is "(>)". instance by (intro_classes; transfer, simp_all add: less_fun_def Sup_upper Sup_least Inf_lower Inf_greatest) end text \Duality has been baked into this result because of its relevance for predicate transformers. A proof where Sups are mapped to Sups and Infs to Infs is certainly possible, but two instantiation of the same type and the same classes are unfortunately impossible. Interpretations could be used instead. A corresponding result for Inf-preseving functions and Sup-lattices, is proved in components on transformers, as more advanced properties about Inf-preserving functions are needed.\ subsection \Alternative Definitions for Complete Boolean Algebras\ text \The current definitions of complete boolean algebras deviates from that in most textbooks in that a distributive law with infinite sups and infinite infs is used. There are interesting applications, for instance in topology, where weaker laws are needed --- for instance for frames and locales.\ class complete_heyting_algebra = complete_lattice + assumes ch_dist: "x \ \Y = (\y \ Y. x \ y)" text \Complete Heyting algebras are also known as frames or locales (they differ with respect to their morphisms).\ class complete_co_heyting_algebra = complete_lattice + assumes co_ch_dist: "x \ \Y = (\y \ Y. x \ y)" class complete_boolean_algebra_alt = complete_lattice + boolean_algebra instance set :: (type) complete_boolean_algebra_alt.. context complete_boolean_algebra_alt begin subclass complete_heyting_algebra proof fix x Y {fix t have "(x \ \Y \ t) = (\Y \ -x \ t)" by (simp add: inf.commute shunt1[symmetric]) also have "... = (\y \ Y. y \ -x \ t)" using Sup_le_iff by blast also have "... = (\y \ Y. x \ y \ t)" by (simp add: inf.commute shunt1) finally have "(x \ \Y \ t) = ((\y\Y. x \ y) \ t)" by (simp add: local.SUP_le_iff)} thus "x \ \Y = (\y\Y. x \ y)" using order.eq_iff by blast qed subclass complete_co_heyting_algebra apply unfold_locales apply (rule order.antisym) apply (simp add: INF_greatest Inf_lower2) by (meson eq_refl le_INF_iff le_Inf_iff shunt2) lemma de_morgan1: "-(\X) = (\x \ X. -x)" proof- {fix y have "(y \ -(\X)) = (\X \ -y)" using compl_le_swap1 by blast also have "... = (\x \ X. x \ -y)" by (simp add: Sup_le_iff) also have "... = (\x \ X. y \ -x)" using compl_le_swap1 by blast also have "... = (y \ (\x \ X. -x))" using le_INF_iff by force finally have "(y \ -(\X)) = (y \(\x \ X. -x))".} thus ?thesis using order.antisym by blast qed lemma de_morgan2: "-(\X) = (\x \ X. -x)" by (metis de_morgan1 ba_dual.dual_iff ba_dual.image_dual pointfree_idE) end class complete_boolean_algebra_alt_with_dual = complete_lattice_with_dual + complete_boolean_algebra_alt instantiation set :: (type) complete_boolean_algebra_alt_with_dual begin definition dual_set :: "'a set \ 'a set" where "dual_set = uminus" instance by intro_classes (simp_all add: ba_dual.inj_dual dual_set_def comp_def uminus_Sup id_def) end context complete_boolean_algebra_alt begin sublocale cba_dual: complete_boolean_algebra_alt_with_dual _ _ _ _ _ _ _ _ uminus _ _ by unfold_locales (auto simp: de_morgan2 de_morgan1) end subsection \Atomic Boolean Algebras\ text \Next, atomic boolean algebras are defined.\ context bounded_lattice begin text \Atoms are covers of bottom.\ definition "atom x = (x \ \ \ \(\y. \ < y \ y < x))" definition "atom_map x = {y. atom y \ y \ x}" lemma atom_map_def_var: "atom_map x = \x \ Collect atom" unfolding atom_map_def downset_def downset_set_def comp_def atom_def by fastforce lemma atom_map_atoms: "\(range atom_map) = Collect atom" unfolding atom_map_def atom_def by auto end typedef (overloaded) 'a atoms = "range (atom_map::'a::bounded_lattice \ 'a set)" by blast setup_lifting type_definition_atoms definition at_map :: "'a::bounded_lattice \ 'a atoms" where "at_map = Abs_atoms \ atom_map" class atomic_boolean_algebra = boolean_algebra + assumes atomicity: "x \ \ \ (\y. atom y \ y \ x)" class complete_atomic_boolean_algebra = complete_lattice + atomic_boolean_algebra begin subclass complete_boolean_algebra_alt.. end text \Here are two equivalent definitions for atoms; first in boolean algebras, and then in complete boolean algebras.\ -context Lattices.boolean_algebra +context boolean_algebra begin text \The following two conditions are taken from Koppelberg's book~\cite{Koppelberg89}.\ lemma atom_neg: "atom x \ x \ \ \ (\y z. x \ y \ x \ -y)" - by (metis atom_def dual_order.order_iff_strict inf.cobounded1 inf.commute meet_shunt) + by (auto simp add: atom_def) (metis local.dual_order.not_eq_order_implies_strict local.inf.cobounded1 local.inf.cobounded2 local.inf_shunt) lemma atom_sup: "(\y. x \ y \ x \ -y) \ (\y z. (x \ y \ x \ z) = (x \ y \ z))" by (metis inf.orderE le_supI1 shunt2) lemma sup_atom: "x \ \ \ (\y z. (x \ y \ x \ z) = (x \ y \ z)) \ atom x" - unfolding atom_def apply clarsimp by (metis bot_less inf.absorb2 less_le_not_le meet_shunt sup_compl_top) + by (auto simp add: atom_def) (metis (full_types) local.inf.boundedI local.inf.cobounded2 local.inf_shunt local.inf_sup_ord(4) local.le_iff_sup local.shunt1 local.sup.absorb1 local.sup.strict_order_iff) lemma atom_sup_iff: "atom x = (x \ \ \ (\y z. (x \ y \ x \ z) = (x \ y \ z)))" - by (standard, auto simp add: atom_neg atom_sup sup_atom) + by rule (auto simp add: atom_neg atom_sup sup_atom) lemma atom_neg_iff: "atom x = (x \ \ \ (\y z. x \ y \ x \ -y))" - by (standard, auto simp add: atom_neg atom_sup sup_atom) + by rule (auto simp add: atom_neg atom_sup sup_atom) lemma atom_map_bot_pres: "atom_map \ = {}" using atom_def atom_map_def le_bot by auto lemma atom_map_top_pres: "atom_map \ = Collect atom" using atom_map_def by auto end context complete_boolean_algebra_alt begin lemma atom_Sup: "\Y. x \ \ \ (\y. x \ y \ x \ -y) \ ((\y \ Y. x \ y) = (x \ \Y))" - by (metis Sup_least Sup_upper2 compl_le_swap1 le_iff_inf meet_shunt) + by (metis Sup_least Sup_upper2 compl_le_swap1 le_iff_inf inf_shunt) lemma Sup_atom: "x \ \ \ (\Y. (\y \ Y. x \ y) = (x \ \Y)) \ atom x" proof- assume h1: "x \ \" and h2: "\Y. (\y \ Y. x \ y) = (x \ \Y)" hence "\y z. (x \ y \ x \ z) = (x \ y \ z)" by (smt insert_iff sup_Sup sup_bot.right_neutral) thus "atom x" by (simp add: h1 sup_atom) qed lemma atom_Sup_iff: "atom x = (x \ \ \ (\Y. (\y \ Y. x \ y) = (x \ \Y)))" by standard (auto simp: atom_neg atom_Sup Sup_atom) end 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 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 Lattices.boolean_algebra \ dual_boolean_algebra: Lattices.boolean_algebra "\x y. x \ -y" uminus sup "(\)" "(>)" inf \ \ +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/Representations.thy b/thys/Order_Lattice_Props/Representations.thy --- a/thys/Order_Lattice_Props/Representations.thy +++ b/thys/Order_Lattice_Props/Representations.thy @@ -1,620 +1,620 @@ (* Title: Representation Theorems for Orderings and Lattices Author: Georg Struth Maintainer: Georg Struth *) section \Representation Theorems for Orderings and Lattices\ theory Representations imports Order_Lattice_Props begin subsection \Representation of Posets\ text \The isomorphism between partial orders and downsets with set inclusion is well known. It forms the basis of Priestley and Stone duality. I show it not only for objects, but also order morphisms, hence establish equivalences and isomorphisms between categories.\ typedef (overloaded) 'a downset = "range (\::'a::ord \ 'a set)" by fastforce setup_lifting type_definition_downset text \The map ds yields the isomorphism between the set and the powerset level if its range is restricted to downsets.\ definition ds :: "'a::ord \ 'a downset" where "ds = Abs_downset \ \" text \In a complete lattice, its inverse is Sup.\ definition SSup :: "'a::complete_lattice downset \ 'a" where "SSup = Sup \ Rep_downset" lemma ds_SSup_inv: "ds \ SSup = (id::'a::complete_lattice downset \ 'a downset)" unfolding ds_def SSup_def by (smt Rep_downset Rep_downset_inverse cSup_atMost eq_id_iff imageE o_def ord_class.atMost_def ord_class.downset_prop) lemma SSup_ds_inv: "SSup \ ds = (id::'a::complete_lattice \ 'a)" unfolding ds_def SSup_def fun_eq_iff id_def comp_def by (simp add: Abs_downset_inverse pointfree_idE) instantiation downset :: (ord) order begin lift_definition less_eq_downset :: "'a downset \ 'a downset \ bool" is "(\X Y. Rep_downset X \ Rep_downset Y)" . lift_definition less_downset :: "'a downset \ 'a downset \ bool" is "(\X Y. Rep_downset X \ Rep_downset Y)" . instance by (intro_classes, transfer, auto simp: Rep_downset_inject less_eq_downset_def) end lemma ds_iso: "mono ds" unfolding mono_def ds_def fun_eq_iff comp_def by (metis Abs_downset_inverse downset_iso_iff less_eq_downset.rep_eq rangeI) lemma ds_faithful: "ds x \ ds y \ x \ (y::'a::order)" by (simp add: Abs_downset_inverse downset_faithful ds_def less_eq_downset.rep_eq) lemma ds_inj: "inj (ds::'a::order \ 'a downset)" by (simp add: ds_faithful dual_order.antisym injI) lemma ds_surj: "surj ds" by (metis (no_types, opaque_lifting) Rep_downset Rep_downset_inverse ds_def image_iff o_apply surj_def) lemma ds_bij: "bij (ds::'a::order \ 'a downset)" by (simp add: bijI ds_inj ds_surj) lemma ds_ord_iso: "ord_iso ds" unfolding ord_iso_def comp_def inf_bool_def by (smt UNIV_I ds_bij ds_faithful ds_inj ds_iso ds_surj f_the_inv_into_f inf1I mono_def) text \The morphishms between orderings and downsets are isotone functions. One can define functors mapping back and forth between these.\ definition map_ds :: "('a::complete_lattice \ 'b::complete_lattice) \ ('a downset \ 'b downset)" where "map_ds f = ds \ f \ SSup" text \This definition is actually contrived. We have shown that a function f between posets P and Q is isotone if and only if the inverse image of f maps downclosed sets in Q to downclosed sets in P. There is the following duality: ds is a natural transformation between the identity functor and the preimage functor as a contravariant functor from P to Q. Hence orderings with isotone maps and downsets with downset-preserving maps are dual, which is a first step towards Stone duality. I don't see a way of proving this with Isabelle, as the types of the preimage of f are the wrong way and I don't see how I could capture opposition with what I have.\ (*lemma "mono (f::'a::complete_lattice \ 'b::complete_lattimap_ds f = Abs_downset \ (-`) f \ Rep_downset" doesn't work! *) lemma map_ds_prop: fixes f :: "'a::complete_lattice \ 'b::complete_lattice" shows "map_ds f \ ds = ds \ f" unfolding map_ds_def by (simp add: SSup_ds_inv comp_assoc) lemma map_ds_prop2: fixes f :: "'a::complete_lattice \ 'b::complete_lattice" shows "map_ds f \ ds = ds \ id f" unfolding map_ds_def by (simp add: SSup_ds_inv comp_assoc) text \This is part of showing that map-ds is naturally isomorphic to the identity functor, ds being the natural isomorphism.\ definition map_SSup :: "('a downset \ 'b downset) \ ('a::complete_lattice \ 'b::complete_lattice)" where "map_SSup F = SSup \ F \ ds" lemma map_ds_iso_pres: fixes f :: "'a::complete_lattice \ 'b::complete_lattice" shows "mono f \ mono (map_ds f)" unfolding fun_eq_iff mono_def map_ds_def ds_def SSup_def comp_def by (metis Abs_downset_inverse Sup_subset_mono downset_iso_iff less_eq_downset.rep_eq rangeI) lemma map_SSup_iso_pres: fixes F :: "'a::complete_lattice downset \ 'b::complete_lattice downset" shows "mono F \ mono (map_SSup F)" unfolding fun_eq_iff mono_def map_SSup_def ds_def SSup_def comp_def by (metis Abs_downset_inverse Sup_subset_mono downset_iso_iff less_eq_downset.rep_eq rangeI) lemma map_SSup_prop: fixes F :: "'a::complete_lattice downset \ 'b::complete_lattice downset" shows "ds \ map_SSup F = F \ ds" unfolding map_SSup_def by (metis ds_SSup_inv fun.map_id0 id_def rewriteL_comp_comp) lemma map_SSup_prop2: fixes F :: "'a::complete_lattice downset \ 'b::complete_lattice downset" shows "ds \ map_SSup F = id F \ ds" by (simp add: map_SSup_prop) lemma map_ds_func1: "map_ds id = (id::'a::complete_lattice downset\ 'a downset)" by (simp add: ds_SSup_inv map_ds_def) lemma map_ds_func2: fixes g :: "'a::complete_lattice \ 'b::complete_lattice" shows "map_ds (f \ g) = map_ds f \ map_ds g" unfolding map_ds_def fun_eq_iff comp_def ds_def SSup_def by (metis Abs_downset_inverse Sup_atMost atMost_def downset_prop rangeI) lemma map_SSup_func1: "map_SSup (id::'a::complete_lattice downset\ 'a downset) = id" by (simp add: SSup_ds_inv map_SSup_def) lemma map_SSup_func2: fixes F :: "'c::complete_lattice downset \ 'b::complete_lattice downset" and G :: "'a::complete_lattice downset \ 'c downset" shows "map_SSup (F \ G) = map_SSup F \ map_SSup G" unfolding map_SSup_def fun_eq_iff comp_def id_def ds_def by (metis comp_apply ds_SSup_inv ds_def id_apply) lemma map_SSup_map_ds_inv: "map_SSup \ map_ds = (id::('a::complete_lattice \ 'b::complete_lattice) \ ('a \ 'b))" by (metis (no_types, opaque_lifting) SSup_ds_inv comp_def eq_id_iff fun.map_comp fun.map_id0 id_apply map_SSup_prop map_ds_prop) lemma map_ds_map_SSup_inv: "map_ds \ map_SSup = (id::('a::complete_lattice downset \ 'b::complete_lattice downset) \ ('a downset \ 'b downset))" unfolding map_SSup_def map_ds_def SSup_def ds_def id_def comp_def fun_eq_iff by (metis (no_types, lifting) Rep_downset Rep_downset_inverse Sup_downset_id image_iff pointfree_idE) lemma inj_map_ds: "inj (map_ds::('a::complete_lattice \ 'b::complete_lattice) \ ('a downset \ 'b downset))" by (metis (no_types, lifting) SSup_ds_inv fun.map_id0 id_comp inj_def map_ds_prop rewriteR_comp_comp2) lemma inj_map_SSup: "inj (map_SSup::('a::complete_lattice downset \ 'b::complete_lattice downset) \ ('a \ 'b))" by (metis inj_on_id inj_on_imageI2 map_ds_map_SSup_inv) lemma map_ds_map_SSup_iff: fixes g :: "'a::complete_lattice \ 'b::complete_lattice" shows "(f = map_ds g) = (map_SSup f = g)" by (metis inj_eq inj_map_ds map_ds_map_SSup_inv pointfree_idE) text \This gives an isomorphism between categories.\ lemma surj_map_ds: "surj (map_ds::('a::complete_lattice \ 'b::complete_lattice) \ ('a downset \ 'b downset))" by (simp add: map_ds_map_SSup_iff surj_def) lemma surj_map_SSup: "surj (map_SSup::('a::complete_lattice_with_dual downset \ 'b::complete_lattice_with_dual downset) \ ('a \ 'b))" by (metis map_ds_map_SSup_iff surjI) text \There is of course a dual result for upsets with the reverse inclusion ordering. Once again, it seems impossible to capture the "real" duality that uses the inverse image functor.\ typedef (overloaded) 'a upset = "range (\::'a::ord \ 'a set)" by fastforce setup_lifting type_definition_upset definition us :: "'a::ord \ 'a upset" where "us = Abs_upset \ \" definition IInf :: "'a::complete_lattice upset \ 'a" where "IInf = Inf \ Rep_upset" lemma us_ds: "us = Abs_upset \ (`) \ \ Rep_downset \ ds \ (\::'a::ord_with_dual \ 'a)" unfolding us_def ds_def fun_eq_iff comp_def by (simp add: Abs_downset_inverse upset_to_downset2) lemma IInf_SSup: "IInf = \ \ SSup \ Abs_downset \ (`) (\::'a::complete_lattice_with_dual \ 'a) \ Rep_upset" unfolding IInf_def SSup_def fun_eq_iff comp_def by (metis (no_types, opaque_lifting) Abs_downset_inverse Rep_upset Sup_dual_def_var image_iff rangeI subset_dual upset_to_downset3) lemma us_IInf_inv: "us \ IInf = (id::'a::complete_lattice_with_dual upset \ 'a upset)" unfolding us_def IInf_def fun_eq_iff id_def comp_def by (metis (no_types, lifting) Inf_upset_id Rep_upset Rep_upset_inverse f_the_inv_into_f pointfree_idE upset_inj) lemma IInf_us_inv: "IInf \ us = (id::'a::complete_lattice_with_dual \ 'a)" unfolding us_def IInf_def fun_eq_iff id_def comp_def by (metis Abs_upset_inverse Sup_Inf_var Sup_atLeastAtMost Sup_dual_upset_var order_refl rangeI) instantiation upset :: (ord) order begin lift_definition less_eq_upset :: "'a upset \ 'a upset \ bool" is "(\X Y. Rep_upset X \ Rep_upset Y)" . lift_definition less_upset :: "'a upset \ 'a upset \ bool" is "(\X Y. Rep_upset X \ Rep_upset Y)" . instance by (intro_classes, transfer, simp_all add: less_le_not_le less_eq_upset.rep_eq Rep_upset_inject) end lemma us_iso: "x \ y \ us x \ us (y::'a::order_with_dual)" by (simp add: Abs_upset_inverse less_eq_upset.rep_eq upset_anti_iff us_def) lemma us_faithful: "us x \ us y \ x \ (y::'a::order_with_dual)" by (simp add: Abs_upset_inverse upset_faithful us_def less_eq_upset.rep_eq) lemma us_inj: "inj (us::'a::order_with_dual \ 'a upset)" unfolding inj_def by (simp add: us_faithful dual_order.antisym) lemma us_surj: "surj (us::'a::order_with_dual \ 'a upset)" unfolding surj_def by (metis Rep_upset Rep_upset_inverse comp_def image_iff us_def) lemma us_bij: "bij (us::'a::order_with_dual \ 'a upset)" by (simp add: bij_def us_surj us_inj) lemma us_ord_iso: "ord_iso (us::'a::order_with_dual \ 'a upset)" unfolding ord_iso_def by (simp, metis (no_types, lifting) UNIV_I f_the_inv_into_f monoI us_iso us_bij us_faithful us_inj us_surj) definition map_us :: "('a::complete_lattice \ 'b::complete_lattice) \ ('a upset \ 'b upset)" where "map_us f = us \ f \ IInf" lemma map_us_prop: "map_us f \ (us::'a::complete_lattice_with_dual \ 'a upset) = us \ id f" unfolding map_us_def by (simp add: IInf_us_inv comp_assoc) definition map_IInf :: "('a upset \ 'b upset) \ ('a::complete_lattice \ 'b::complete_lattice)" where "map_IInf F = IInf \ F \ us" lemma map_IInf_prop: "(us::'a::complete_lattice_with_dual \ 'a upset) \ map_IInf F = id F \ us" proof- have "us \ map_IInf F = (us \ IInf) \ F \ us" by (simp add: fun.map_comp map_IInf_def) thus ?thesis by (simp add: us_IInf_inv) qed lemma map_us_func1: "map_us id = (id::'a::complete_lattice_with_dual upset \ 'a upset)" unfolding map_us_def fun_eq_iff comp_def us_def id_def IInf_def by (metis (no_types, lifting) Inf_upset_id Rep_upset Rep_upset_inverse image_iff pointfree_idE) lemma map_us_func2: fixes f :: "'c::complete_lattice_with_dual \ 'b::complete_lattice_with_dual" and g :: "'a::complete_lattice_with_dual \ 'c" shows "map_us (f \ g) = map_us f \ map_us g" unfolding map_us_def fun_eq_iff comp_def us_def IInf_def by (metis Abs_upset_inverse Sup_Inf_var Sup_atLeastAtMost Sup_dual_upset_var order_refl rangeI) lemma map_IInf_func1: "map_IInf id = (id::'a::complete_lattice_with_dual \ 'a)" unfolding map_IInf_def fun_eq_iff comp_def id_def us_def IInf_def by (simp add: Abs_upset_inverse pointfree_idE) lemma map_IInf_func2: fixes F :: "'c::complete_lattice_with_dual upset \ 'b::complete_lattice_with_dual upset" and G :: "'a::complete_lattice_with_dual upset \ 'c upset" shows "map_IInf (F \ G) = map_IInf F \ map_IInf G" unfolding map_IInf_def fun_eq_iff comp_def id_def us_def by (metis comp_apply id_apply us_IInf_inv us_def) lemma map_IInf_map_us_inv: "map_IInf \ map_us = (id::('a::complete_lattice_with_dual \ 'b::complete_lattice_with_dual) \ ('a \ 'b))" unfolding map_IInf_def map_us_def IInf_def us_def id_def comp_def fun_eq_iff by (simp add: Abs_upset_inverse pointfree_idE) lemma map_us_map_IInf_inv: "map_us \ map_IInf = (id::('a::complete_lattice_with_dual upset \ 'b::complete_lattice_with_dual upset) \ ('a upset \ 'b upset))" unfolding map_IInf_def map_us_def IInf_def us_def id_def comp_def fun_eq_iff by (metis (no_types, lifting) Inf_upset_id Rep_upset Rep_upset_inverse image_iff pointfree_idE) lemma inj_map_us: "inj (map_us::('a::complete_lattice_with_dual \ 'b::complete_lattice_with_dual) \ ('a upset \ 'b upset))" unfolding map_us_def us_def IInf_def inj_def comp_def fun_eq_iff by (metis (no_types, opaque_lifting) Abs_upset_inverse Inf_upset_id pointfree_idE rangeI) lemma inj_map_IInf: "inj (map_IInf::('a::complete_lattice_with_dual upset \ 'b::complete_lattice_with_dual upset) \ ('a \ 'b))" unfolding map_IInf_def fun_eq_iff inj_def comp_def IInf_def us_def by (metis (no_types, opaque_lifting) Inf_upset_id Rep_upset Rep_upset_inverse image_iff pointfree_idE) lemma map_us_map_IInf_iff: fixes g :: "'a::complete_lattice_with_dual \ 'b::complete_lattice_with_dual" shows "(f = map_us g) = (map_IInf f = g)" by (metis inj_eq inj_map_us map_us_map_IInf_inv pointfree_idE) lemma map_us_mono_pres: fixes f :: "'a::complete_lattice_with_dual \ 'b::complete_lattice_with_dual" shows "mono f \ mono (map_us f)" unfolding mono_def map_us_def comp_def us_def IInf_def by (metis Abs_upset_inverse Inf_superset_mono less_eq_upset.rep_eq rangeI upset_anti_iff) lemma map_IInf_mono_pres: fixes F :: "'a::complete_lattice_with_dual upset \ 'b::complete_lattice_with_dual upset" shows "mono F \ mono (map_IInf F)" unfolding mono_def map_IInf_def comp_def us_def IInf_def oops lemma surj_map_us: "surj (map_us::('a::complete_lattice_with_dual \ 'b::complete_lattice_with_dual) \ ('a upset \ 'b upset))" by (simp add: map_us_map_IInf_iff surj_def) lemma surj_map_IInf: "surj (map_IInf::('a::complete_lattice_with_dual upset \ 'b::complete_lattice_with_dual upset) \ ('a \ 'b))" by (metis map_us_map_IInf_iff surjI) text \Hence we have again an isomorphism --- or rather equivalence --- between categories. Here, however, duality is not consistently picked up.\ subsection \Stone's Theorem in the Presence of Atoms\ text \Atom-map is a boolean algebra morphism.\ -context Lattices.boolean_algebra +context boolean_algebra begin lemma atom_map_compl_pres: "atom_map (-x) = Collect atom - atom_map x" proof- {fix y have "(y \ atom_map (-x)) = (atom y \ y \ -x)" by (simp add: atom_map_def) also have "... = (atom y \ \(y \ x))" - by (metis atom_sup_iff inf.orderE meet_shunt sup_compl_top top.ordering_top_axioms ordering_top.extremum) + by (metis atom_sup_iff inf.orderE inf_shunt sup_compl_top top.ordering_top_axioms ordering_top.extremum) also have "... = (y \ Collect atom - atom_map x)" using atom_map_def by auto finally have "(y \ atom_map (-x)) = (y \ Collect atom - atom_map x)".} thus ?thesis by blast qed lemma atom_map_sup_pres: "atom_map (x \ y) = atom_map x \ atom_map y" proof- {fix z have "(z \ atom_map (x \ y)) = (atom z \ z \ x \ y)" by (simp add: atom_map_def) also have "... = (atom z \ (z \ x \ z \ y))" using atom_sup_iff by auto also have "... = (z \ atom_map x \ atom_map y)" using atom_map_def by auto finally have "(z \ atom_map (x \ y)) = (z \ atom_map x \ atom_map y)" by blast} thus ?thesis by blast qed lemma atom_map_inf_pres: "atom_map (x \ y) = atom_map x \ atom_map y" by (smt Diff_Un atom_map_compl_pres atom_map_sup_pres compl_inf double_compl) lemma atom_map_minus_pres: "atom_map (x - y) = atom_map x - atom_map y" using atom_map_compl_pres atom_map_def diff_eq by auto end text \The homomorphic images of boolean algebras under atom-map are boolean algebras --- in fact powerset boolean algebras.\ -instantiation atoms :: (Lattices.boolean_algebra) Lattices.boolean_algebra +instantiation atoms :: (boolean_algebra) boolean_algebra begin lift_definition minus_atoms :: "'a atoms \ 'a atoms \ 'a atoms" is "\x y. Abs_atoms (Rep_atoms x - Rep_atoms y)". lift_definition uminus_atoms :: "'a atoms \ 'a atoms" is "\x. Abs_atoms (Collect atom - Rep_atoms x)". lift_definition bot_atoms :: "'a atoms" is "Abs_atoms {}". lift_definition sup_atoms :: "'a atoms \ 'a atoms \ 'a atoms" is "\x y. Abs_atoms (Rep_atoms x \ Rep_atoms y)". lift_definition top_atoms :: "'a atoms" is "Abs_atoms (Collect atom)". lift_definition inf_atoms :: "'a atoms \ 'a atoms \ 'a atoms" is "\x y. Abs_atoms (Rep_atoms x \ Rep_atoms y)". lift_definition less_eq_atoms :: "'a atoms \ 'a atoms \ bool" is "(\x y. Rep_atoms x \ Rep_atoms y)". lift_definition less_atoms :: "'a atoms \ 'a atoms \ bool" is "(\x y. Rep_atoms x \ Rep_atoms y)". instance apply intro_classes apply (transfer, simp add: less_le_not_le) apply (transfer, simp) apply (transfer, blast) apply (simp add: Rep_atoms_inject less_eq_atoms.abs_eq) apply (transfer, smt Abs_atoms_inverse Rep_atoms atom_map_inf_pres image_iff inf_le1 rangeI) apply (transfer, smt Abs_atoms_inverse Rep_atoms atom_map_inf_pres image_iff inf_le2 rangeI) apply (transfer, smt Abs_atoms_inverse Rep_atoms atom_map_inf_pres image_iff le_iff_sup rangeI sup_inf_distrib1) apply (transfer, smt Abs_atoms_inverse Rep_atoms atom_map_sup_pres image_iff image_iff inf.orderE inf_sup_aci(6) le_iff_sup order_refl rangeI rangeI) apply (transfer, smt Abs_atoms_inverse Rep_atoms atom_map_sup_pres image_iff inf_sup_aci(6) le_iff_sup rangeI sup.left_commute sup.right_idem) apply (transfer, subst Abs_atoms_inverse, metis (no_types, lifting) Rep_atoms atom_map_sup_pres image_iff rangeI, simp) apply transfer using Abs_atoms_inverse atom_map_bot_pres apply blast apply (transfer, metis Abs_atoms_inverse Rep_atoms atom_map_compl_pres atom_map_top_pres diff_eq double_compl inf_le1 rangeE rangeI) apply (transfer, smt Abs_atoms_inverse Rep_atoms atom_map_inf_pres atom_map_sup_pres image_iff rangeI sup_inf_distrib1) apply (transfer, metis (no_types, opaque_lifting) Abs_atoms_inverse Diff_disjoint Rep_atoms atom_map_compl_pres rangeE rangeI) apply (transfer, smt Abs_atoms_inverse uminus_atoms.abs_eq Rep_atoms Un_Diff_cancel atom_map_compl_pres atom_map_inf_pres atom_map_minus_pres atom_map_sup_pres atom_map_top_pres diff_eq double_compl inf_compl_bot_right rangeE rangeI sup_commute sup_compl_top) by transfer (smt Abs_atoms_inverse Rep_atoms atom_map_compl_pres atom_map_inf_pres atom_map_minus_pres diff_eq rangeE rangeI) end text \The homomorphism atom-map can then be restricted in its output type to the powerset boolean algebra.\ lemma at_map_bot_pres: "at_map \ = \" by (simp add: at_map_def atom_map_bot_pres bot_atoms.transfer) lemma at_map_top_pres: "at_map \ = \" by (simp add: at_map_def atom_map_top_pres top_atoms.transfer) lemma at_map_compl_pres: "at_map \ uminus = uminus \ at_map" unfolding fun_eq_iff by (simp add: Abs_atoms_inverse at_map_def atom_map_compl_pres uminus_atoms.abs_eq) lemma at_map_sup_pres: "at_map (x \ y) = at_map x \ at_map y" unfolding at_map_def comp_def by (metis (mono_tags, lifting) Abs_atoms_inverse atom_map_sup_pres rangeI sup_atoms.transfer) lemma at_map_inf_pres: "at_map (x \ y) = at_map x \ at_map y" unfolding at_map_def comp_def by (metis (mono_tags, lifting) Abs_atoms_inverse atom_map_inf_pres inf_atoms.transfer rangeI) lemma at_map_minus_pres: "at_map (x - y) = at_map x - at_map y" unfolding at_map_def comp_def by (simp add: Abs_atoms_inverse atom_map_minus_pres minus_atoms.abs_eq) context atomic_boolean_algebra begin text \In atomic boolean algebras, atom-map is an embedding that maps atoms of the boolean algebra to those of the powerset boolean algebra. Analogous properties hold for at-map.\ lemma inj_atom_map: "inj atom_map" proof- {fix x y ::'a assume "x \ y" hence "x \ -y \ \ \ -x \ y \ \" - by (auto simp: meet_shunt) + by (auto simp: inf_shunt) hence "\z. atom z \ (z \ x \ -y \ z \ -x \ y)" using atomicity by blast hence "\z. atom z \ ((z \ atom_map x \ \(z \ atom_map y)) \ (\(z \ atom_map x) \ z \ atom_map y))" - unfolding atom_def atom_map_def by (clarsimp, metis diff_eq inf.orderE meet_shunt_var) + unfolding atom_def atom_map_def by (clarsimp, metis diff_eq inf.orderE diff_shunt_var) hence "atom_map x \ atom_map y" by blast} thus ?thesis by (meson injI) qed lemma atom_map_atom_pres: "atom x \ atom_map x = {x}" unfolding atom_def atom_map_def by (auto simp: bot_less dual_order.order_iff_strict) lemma atom_map_atom_pres2: "atom x \ atom (atom_map x)" proof- assume "atom x" hence "atom_map x = {x}" by (simp add: atom_map_atom_pres) thus "atom (atom_map x)" using bounded_lattice_class.atom_def by auto qed end lemma inj_at_map: "inj (at_map::'a::atomic_boolean_algebra \ 'a atoms)" unfolding at_map_def comp_def by (metis (no_types, lifting) Abs_atoms_inverse inj_atom_map inj_def rangeI) lemma at_map_atom_pres: "atom (x::'a::atomic_boolean_algebra) \ at_map x = Abs_atoms {x}" unfolding at_map_def comp_def by (simp add: atom_map_atom_pres) lemma at_map_atom_pres2: "atom (x::'a::atomic_boolean_algebra) \ atom (at_map x)" unfolding at_map_def comp_def by (metis Abs_atoms_inverse atom_def atom_map_atom_pres2 atom_map_bot_pres bot_atoms.abs_eq less_atoms.abs_eq rangeI) text \Homomorphic images of atomic boolean algebras under atom-map are therefore atomic (rather obviously).\ instance atoms :: (atomic_boolean_algebra) atomic_boolean_algebra proof intro_classes fix x::"'a atoms" assume "x \ \" hence "\y. x = at_map y \ x \ \" unfolding at_map_def comp_def by (metis Abs_atoms_cases rangeE) hence "\y. x = at_map y \ (\z. atom z \ z \ y)" using at_map_bot_pres atomicity by force hence "\y. x = at_map y \ (\z. atom (at_map z) \ at_map z \ at_map y)" by (metis at_map_atom_pres2 at_map_sup_pres sup.orderE sup_ge2) thus "\y. atom y \ y \ x" by fastforce qed context complete_boolean_algebra_alt begin text \In complete boolean algebras, atom-map is surjective; more precisely it is the left inverse of Sup, at least for sets of atoms. Below, this statement is made more explicit for at-map.\ lemma surj_atom_map: "Y \ Collect atom \ Y = atom_map (\Y)" proof assume "Y \ Collect atom" thus "Y \ atom_map (\Y)" using Sup_upper atom_map_def by force next assume "Y \ Collect atom" hence a: "\y. y \ Y \ atom y" by blast {fix z assume h: "z \ Collect atom - Y" hence "\y \ Y. y \ z = \" - by (metis DiffE a h atom_def dual_order.not_eq_order_implies_strict inf.absorb_iff2 inf_le2 meet_shunt mem_Collect_eq) + by (metis DiffE a h atom_def dual_order.not_eq_order_implies_strict inf.absorb_iff2 inf_le2 inf_shunt mem_Collect_eq) hence "\Y \ z = \" - using Sup_least meet_shunt by simp + using Sup_least inf_shunt by simp hence "z \ atom_map (\Y)" using atom_map_bot_pres atom_map_def by force} thus "atom_map (\Y) \ Y" using atom_map_def by force qed text \In this setting, atom-map is a complete boolean algebra morphism.\ lemma atom_map_Sup_pres: "atom_map (\X) = (\x \ X. atom_map x)" proof- {fix z have "(z \ atom_map (\X)) = (atom z \ z \ \X)" by (simp add: atom_map_def) also have "... = (atom z \ (\x \ X. z \ x))" using atom_Sup_iff by auto also have "... = (z \ (\x \ X. atom_map x))" using atom_map_def by auto finally have "(z \ atom_map (\X)) = (z \ (\x \ X. atom_map x))" by blast} thus ?thesis by blast qed lemma atom_map_Sup_pres_var: "atom_map \ Sup = Sup \ (`) atom_map" unfolding fun_eq_iff comp_def by (simp add: atom_map_Sup_pres) text \For Inf-preservation, it is important that Infs are restricted to homomorphic images; hence they need to be pushed into the set of all atoms.\ lemma atom_map_Inf_pres: "atom_map (\X) = Collect atom \ (\x \ X. atom_map x)" proof- have "atom_map (\X) = atom_map (-(\x \ X. -x))" by (smt Collect_cong SUP_le_iff atom_map_def compl_le_compl_iff compl_le_swap1 le_Inf_iff) also have "... = Collect atom - atom_map (\x \ X. -x)" using atom_map_compl_pres by blast also have "... = Collect atom - (\x \ X. atom_map (-x))" by (simp add: atom_map_Sup_pres) also have "... = Collect atom - (\x \ X. Collect atom - atom_map (x))" using atom_map_compl_pres by blast also have "... = Collect atom \ (\x \ X. atom_map x)" by blast finally show ?thesis. qed end text \It follows that homomorphic images of complete boolean algebras under atom-map form complete boolean algebras.\ instantiation atoms :: (complete_boolean_algebra_alt) complete_boolean_algebra_alt begin lift_definition Inf_atoms :: "'a::complete_boolean_algebra_alt atoms set \ 'a::complete_boolean_algebra_alt atoms" is "\X. Abs_atoms (Collect atom \ Inter ((`) Rep_atoms X))". lift_definition Sup_atoms :: "'a::complete_boolean_algebra_alt atoms set \ 'a::complete_boolean_algebra_alt atoms" is "\X. Abs_atoms (Union ((`) Rep_atoms X))". instance apply (intro_classes; transfer) apply (metis (no_types, opaque_lifting) Abs_atoms_inverse image_iff inf_le1 le_Inf_iff le_infI2 order_refl rangeI surj_atom_map) apply (metis (no_types, lifting) Abs_atoms_inverse Int_subset_iff Rep_atoms Sup_upper atom_map_atoms inf_le1 le_INF_iff rangeI surj_atom_map) apply (metis Abs_atoms_inverse Rep_atoms SUP_least SUP_upper Sup_upper atom_map_atoms rangeI surj_atom_map) apply (metis Abs_atoms_inverse Rep_atoms SUP_least Sup_upper atom_map_atoms rangeI surj_atom_map) by simp_all end text \Once more, properties proved above can now be restricted to at-map.\ lemma surj_at_map_var: "at_map \ Sup \ Rep_atoms = (id::'a::complete_boolean_algebra_alt atoms \ 'a atoms)" unfolding at_map_def comp_def fun_eq_iff id_def by (metis Rep_atoms Rep_atoms_inverse Sup_upper atom_map_atoms surj_atom_map) lemma surj_at_map: "surj (at_map::'a::complete_boolean_algebra_alt \ 'a atoms)" unfolding surj_def at_map_def comp_def by (metis Rep_atoms Rep_atoms_inverse image_iff) lemma at_map_Sup_pres: "at_map \ Sup = Sup \ (`) (at_map::'a::complete_boolean_algebra_alt \ 'a atoms)" unfolding fun_eq_iff at_map_def comp_def atom_map_Sup_pres by (smt Abs_atoms_inverse Sup.SUP_cong Sup_atoms.transfer UN_extend_simps(10) rangeI) lemma at_map_Sup_pres_var: "at_map (\X) = (\(x::'a::complete_boolean_algebra_alt) \ X. (at_map x))" using at_map_Sup_pres comp_eq_elim by blast lemma at_map_Inf_pres: "at_map (\X) = Abs_atoms (Collect atom \ (\x \ X. (Rep_atoms (at_map (x::'a::complete_boolean_algebra_alt)))))" unfolding at_map_def comp_def by (metis (no_types, lifting) Abs_atoms_inverse Sup.SUP_cong atom_map_Inf_pres rangeI) lemma at_map_Inf_pres_var: "at_map \ Inf = Inf \ (`) (at_map::'a::complete_boolean_algebra_alt \ 'a atoms)" unfolding fun_eq_iff comp_def by (metis Inf_atoms.abs_eq at_map_Inf_pres image_image) text \Finally, on complete atomic boolean algebras (CABAs), at-map is an isomorphism, that is, a bijection that preserves the complete boolean algebra operations. Thus every CABA is isomorphic to a powerset boolean algebra and every powerset boolean algebra is a CABA. The bijective pair is given by at-map and Sup (defined on the powerset algebra). This theorem is a little version of Stone's theorem. In the general case, ultrafilters play the role of atoms.\ lemma "Sup \ atom_map = (id::'a::complete_atomic_boolean_algebra \ 'a)" unfolding fun_eq_iff comp_def id_def by (metis Union_upper atom_map_atoms inj_atom_map inj_def rangeI surj_atom_map) lemma inj_at_map_var: "Sup \ Rep_atoms \ at_map = (id ::'a::complete_atomic_boolean_algebra \ 'a)" unfolding at_map_def comp_def fun_eq_iff id_def by (metis Abs_atoms_inverse Union_upper atom_map_atoms inj_atom_map inj_def rangeI surj_atom_map) lemma bij_at_map: "bij (at_map::'a::complete_atomic_boolean_algebra \ 'a atoms)" unfolding bij_def by (simp add: inj_at_map surj_at_map) instance atoms :: (complete_atomic_boolean_algebra) complete_atomic_boolean_algebra.. text \A full consideration of Stone duality is left for future work.\ (* Failed attempt to prove Tarski's fixpoint theorem: The problem is that we want to use mono, but this has two type parameters. It doesn't work inside of the one-type-parameter typedef. Yet isotonicity is needed to prove inhabitance of the type. I could develop a theory of isotone endos and prove the existence of lfps and gfps, duplicating the more general facts for mono. But that's not the point. Because of this I see no direct way of proving Tarski's fixpoint theorem. Any way out? class complete_lattice_with_iso = complete_lattice + fixes f :: "'a \ 'a" (* assumes isof: "x \ y \ f x \ f y"*) typedef (overloaded) 'a Fix = "Fix (f::'a::complete_lattice_with_iso \ 'a)" setup_lifting type_definition_Fix *) end diff --git a/thys/PSemigroupsConvolution/Binary_Modalities.thy b/thys/PSemigroupsConvolution/Binary_Modalities.thy --- a/thys/PSemigroupsConvolution/Binary_Modalities.thy +++ b/thys/PSemigroupsConvolution/Binary_Modalities.thy @@ -1,438 +1,434 @@ (* Title: Binary Modalities Author: Brijesh Dongol, Victor Gomes, Ian J Hayes, Georg Struth Maintainer: Victor Gomes Georg Struth *) section \Binary Modalities and Relational Convolution\ theory Binary_Modalities imports Quantales begin subsection \Auxiliary Properties\ lemma SUP_is_Sup: "(SUP f\F. f y) = \{(f::'a \ 'b::proto_near_quantale) y |f. f \ F}" proof (rule antisym) fix f::"'a \ 'b::proto_near_quantale" have "f \ F \ f y \ {f y |f. f \ F}" by (simp add: Setcompr_eq_image) hence "f \ F \ f y \ \{f y |f. f \ F}" by (simp add: Sup_upper) thus "(SUP f\F. f y) \ \{(f::'a \ 'b::proto_near_quantale) y |f. f \ F}" by (simp add: Setcompr_eq_image) next fix x have "x \ {f y |f. f \ F} \ x \ (SUP f\F. f y) " using mk_disjoint_insert by force thus "Sup {(f::'a \ 'b::proto_near_quantale) y |f. f \ F} \ (SUP f\F. f y)" by (simp add: Setcompr_eq_image) qed lemma bmod_auxl: "{x \ g z |x. \f. x = f y \ f \ F} = {f y \ g z |f. f \ F}" by force lemma bmod_auxr: "{f y \ x |x. \g. x = g z \ g \ G} = {f y \ g z |g. g \ G}" by force lemma bmod_assoc_aux1: "\{\{(f :: 'a \ 'b::proto_near_quantale) u \ g v \ h w |u v. R y u v} |y w. R x y w} = \{(f u \ g v) \ h w |u v y w. R y u v \ R x y w}" apply (rule antisym) apply (rule Sup_least, safe) apply (intro Sup_least Sup_upper, force) apply (rule Sup_least, safe) by (rule Sup_upper2, auto)+ lemma bmod_assoc_aux2: "\{\{(f::'a \ 'b::proto_near_quantale) u \ g v \ h w |v w. R y v w} |u y. R x u y} = \{f u \ g v \ h w |u v w y. R y v w \ R x u y}" apply (rule antisym) apply (rule Sup_least, safe) apply (intro Sup_least Sup_upper, force) apply (rule Sup_least, safe) by (rule Sup_upper2, auto)+ subsection \Binary Modalities\ text \Most of the development in the papers mentioned in the introduction generalises to proto-near-quantales. Binary modalities are interesting for various substructural logics over ternary Kripke frames. They also arise, e.g., as chop modalities in interval logics or as separation conjunction in separation logic. Binary modalities can be understood as a convolution operation parametrised by a ternary operation. Our development yields a unifying framework.\ text \We would prefer a notation that is more similar to our articles, that is, $f\ast_R g$, but we don' know how we could index an infix operator by a variable in Isabelle.\ definition bmod_comp :: "('a \ 'b \ 'c \ bool) \ ('b \ 'd::proto_near_quantale) \ ('c \ 'd) \ 'a \ 'd" ("\") where "\ R f g x = \{f y \ g z |y z. R x y z}" definition bmod_bres :: "('c \ 'b \ 'a \ bool) \ ('b \ 'd::proto_near_quantale) \ ('c \ 'd) \ 'a \ 'd" ("\") where "\ R f g x = \{(f y) \ (g z) |y z. R z y x}" definition bmod_fres :: "('b \ 'a \ 'c \ bool) \ ('b \ 'd::proto_near_quantale) \ ('c \ 'd) \ 'a \ 'd" ("\")where "\ R f g x = \{(f y) \ (g z) |y z. R y x z}" lemma bmod_un_rel: "\ (R \ S) = \ R \ \ S" apply (clarsimp simp: fun_eq_iff bmod_comp_def Sup_union_distrib[symmetric] Collect_disj_eq[symmetric]) by (metis (no_types, lifting)) lemma bmod_Un_rel: "\ (\\) f g x = \{\ R f g x |R. R \ \}" apply (simp add: bmod_comp_def) apply (rule antisym) apply (rule Sup_least, safe) apply (rule Sup_upper2, force) apply (rule Sup_upper, force) apply (rule Sup_least, safe)+ by (metis (mono_tags, lifting) Sup_upper mem_Collect_eq) lemma bmod_sup_fun1: "\ R (f \ g) = \ R f \ \ R g" apply (clarsimp simp add: fun_eq_iff bmod_comp_def sup_distr) apply (rule antisym) apply (intro Sup_least, safe) apply (rule sup_least) apply (intro le_supI1 Sup_upper, force) apply (intro le_supI2 Sup_upper, force) apply (rule sup_least) by (intro Sup_least, safe, rule Sup_upper2, force, simp)+ lemma bmod_Sup_fun1: "\ R (\\) g x = \{\ R f g x |f. f \ \}" proof - have "\ R (\{f. f \ \}) g x = \{\{f y |f. f \ \} \ g z |y z. R x y z}" by (simp add: bmod_comp_def SUP_is_Sup) also have "... = \{\{f y \ g z |f. f \ \} |y z. R x y z}" by (simp add: Sup_distr bmod_auxl) also have "... = \{\{f y \ g z |y z. R x y z} |f. f \ \}" apply (rule antisym) by ((rule Sup_least, safe)+ , (rule Sup_upper2, force, rule Sup_upper, force))+ finally show ?thesis by (simp add: bmod_comp_def) qed lemma bmod_sup_fun2: "\ R (f::'a \ 'b::weak_proto_quantale) (g \ h) = \ R f g \ \ R f h" apply (clarsimp simp add: fun_eq_iff bmod_comp_def sup_distl) apply (rule antisym) apply (intro Sup_least, safe) apply (rule sup_least) apply (intro le_supI1 Sup_upper, force) apply (intro le_supI2 Sup_upper, force) apply (rule sup_least) by (intro Sup_least, safe, rule Sup_upper2, force, simp)+ lemma bmod_Sup_fun2_weak: assumes "\ \ {}" shows "\ R f (\\) x = \{\ R f (g::'a \ 'b::weak_proto_quantale) x |g. g \ \}" proof - have set: "\z. {g z |g::'a \ 'b. g \ \} \ {}" using assms by blast have "\ R f (\{g. g \ \}) x = \{f y \ \{g z |g. g \ \} |y z. R x y z}" by (simp add: bmod_comp_def SUP_is_Sup) also have "... = \{\{f y \ g z |g. g \ \} |y z. R x y z}" by (simp add: weak_Sup_distl[OF set] bmod_auxr) also have "... = \{\{f y \ g z |y z. R x y z} |g. g \ \}" apply (rule antisym) by ((rule Sup_least, safe)+, (rule Sup_upper2, force, rule Sup_upper, force))+ finally show ?thesis by (auto simp: bmod_comp_def) qed lemma bmod_Sup_fun2: "\ R f (\\) x = \{\ R f (g::'a \ 'b::proto_quantale) x |g. g \ \}" proof - have "\ R f (\{g. g \ \}) x = \{f y \ \{g z |g. g \ \} |y z. R x y z}" by (simp add: bmod_comp_def SUP_is_Sup) also have "... = \{\{f y \ g z |g. g \ \} |y z. R x y z}" by (simp add: Sup_distl bmod_auxr) also have "... = \{\{f y \ g z |y z. R x y z} |g. g \ \}" apply (rule antisym) by ((rule Sup_least, safe)+, (rule Sup_upper2, force, rule Sup_upper, force))+ finally show ?thesis by (auto simp: bmod_comp_def) qed lemma bmod_comp_bres_galois: "(\x. \ R f g x \ h x) \ (\x. g x \ \ R f h x)" (* nitpick[expect=genuine] *) oops text \The following Galois connection requires functions into proto-quantales.\ lemma bmod_comp_bres_galois: "(\x. \ R (f::'a \ 'b::proto_quantale) g x \ h x) \ (\x. g x \ \ R f h x)" proof - have "(\x. \ R f g x \ h x) \ (\x y z. R x y z \ (f y) \ (g z) \ h x)" apply (simp add: bmod_comp_def, standard, safe) apply (metis (mono_tags, lifting) CollectI Sup_le_iff) by (rule Sup_least, force) also have "... \ (\x y z. R x y z \ g z \ (f y) \ (h x))" by (simp add: bres_galois) finally show ?thesis apply (simp add: fun_eq_iff bmod_bres_def) apply standard using le_Inf_iff apply fastforce by (metis (mono_tags, lifting) CollectI le_Inf_iff) qed lemma bmod_comp_fres_galois: "(\x. \ R f g x \ h x) \ (\x. f x \ \ R h g x)" proof - have "(\x. \ R f g x \ h x) \ (\x y z. R x y z \ (f y) \ (g z) \ h x)" apply (simp add: bmod_comp_def, standard, safe) apply (metis (mono_tags, lifting) CollectI Sup_le_iff) by (rule Sup_least, force) also have "... \ (\x y z. R x y z \ f y \ (h x) \ (g z))" by (simp add: fres_galois) finally show ?thesis apply (simp add: bmod_fres_def fun_eq_iff) apply standard using le_Inf_iff apply fastforce by (metis (mono_tags, lifting) CollectI le_Inf_iff) qed subsection \Relational Convolution and Correspondence Theory\ text\We now fix a ternary relation $\rho$ and can then hide the parameter in a convolution-style notation.\ class rel_magma = fixes \ :: "'a \ 'a \ 'a \ bool" begin definition times_rel_fun :: "('a \ 'b::proto_near_quantale) \ ('a \ 'b) \ 'a \ 'b" (infix "\" 70) where "f \ g = \ \ f g" lemma rel_fun_Sup_distl_weak: "G \ {} \ (f::'a \ 'b::weak_proto_quantale) \ \G = \{f \ g |g. g \ G}" proof- fix f :: "'a \ 'b" and G :: "('a \ 'b) set" assume a: "G \ {}" show "f \ \G = \{f \ g |g. g \ G}" apply (clarsimp simp: fun_eq_iff times_rel_fun_def bmod_Sup_fun2_weak[OF a]) apply (rule antisym) apply (rule Sup_least, safe) apply (rule SUP_upper2, force+) apply (rule SUP_least, safe) by (rule Sup_upper2, force+) qed lemma rel_fun_Sup_distl: "(f::'a \ 'b::proto_quantale) \ \G = \{f \ g |g. g \ G}" apply (clarsimp simp: fun_eq_iff times_rel_fun_def bmod_Sup_fun2) apply (rule antisym) apply (rule Sup_least, safe) apply (rule SUP_upper2, force+) apply (rule SUP_least, safe) by (rule Sup_upper2, force+) lemma rel_fun_Sup_distr: "\G \ (f::'a \ 'b::proto_near_quantale) = \{g \ f |g. g \ G}" apply (clarsimp simp: fun_eq_iff times_rel_fun_def bmod_Sup_fun1) apply (rule antisym) apply (rule Sup_least, safe) apply (rule SUP_upper2, force+) apply (rule SUP_least, safe) by (rule Sup_upper2, force+) end class rel_semigroup = rel_magma + assumes rel_assoc: "(\y. \ y u v \ \ x y w) \ (\z. \ z v w \ \ x u z)" begin text \Nitpick produces counterexamples even for weak quantales. Hence one cannot generally lift functions into weak quantales to weak quantales.\ lemma bmod_assoc: "\ \ (\ \ (f::'a \ 'b::weak_quantale) g) h x = \ \ f (\ \ g h) x" (*nitpick[show_all,expect=genuine]*) oops lemma bmod_assoc: "\ \ (\ \ (f::'a \ 'b::quantale) g) h x = \ \ f (\ \ g h) x" proof - have "\ \ (\ \ f g) h x = \{\{f u \ g v \ h z |u v. \ y u v} |y z. \ x y z}" apply (simp add: bmod_comp_def Sup_distr) apply (rule antisym) by (intro Sup_least Sup_upper, safe, intro exI conjI, simp_all, rule_tac f = Sup and g = Sup in cong, auto)+ also have "... = \{f u \ g v \ h z |u v y z. \ y u v \ \ x y z}" by (simp add: bmod_assoc_aux1) also have "... = \{f u \ g v \ h z |u v z y. \ y v z \ \ x u y}" apply (rule antisym) apply (rule Sup_least, rule Sup_upper, safe) using rel_assoc apply force apply (rule Sup_least, rule Sup_upper, safe) using rel_assoc by blast also have "... = \{\{f u \ g v \ h z |v z. \ y v z} |u y. \ x u y}" by (simp add: bmod_assoc_aux2) also have "... = \{f u \ \{g v \ h z |v z. \ y v z} |u y. \ x u y}" apply (simp add: Sup_distl mult.assoc) apply (rule antisym) by (intro Sup_least Sup_upper, safe, intro exI conjI, simp_all, rule_tac f = Sup and g = Sup in cong, auto)+ finally show ?thesis by (auto simp: bmod_comp_def) qed lemma rel_fun_assoc: "((f :: 'a \ 'b::quantale) \ g) \ h = f \ (g \ h)" by (simp add: times_rel_fun_def fun_eq_iff bmod_assoc[symmetric]) end lemma "\ R (\ R f f) f x = \ R f (\ R f f) x" (*apply (simp add: bmod_comp_def) nitpick[expect=genuine] *) oops class rel_monoid = rel_semigroup + fixes \ :: "'a set" assumes unitl_ex: "\e \ \. \ x e x" and unitr_ex: "\e \ \. \ x x e" and unitl_eq: "e \ \ \ \ x e y \ x = y" and unitr_eq: "e \ \ \ \ x y e \ x = y" begin lemma xi_prop: "e1 \ \ \ e2 \ \ \ e1 \ e2 \ \ \ x e1 e2 \ \ \ x e2 e1" using unitl_eq unitr_eq by blast definition pid :: "'a \ 'b::unital_weak_quantale" ("\") where "\ x = (if x \ \ then 1 else \)" text \Due to the absence of right annihilation, the right unit law fails for functions into weak quantales.\ lemma bmod_onel: "\ \ f (\::'a \ 'b::unital_weak_quantale) x = f x" (*nitpick[expect=genuine]*) oops text \A unital quantale is required for this lifting.\ lemma bmod_onel: "\ \ f (\::'a \ 'b::unital_quantale) x = f x" apply (simp add: bmod_comp_def pid_def) apply (rule antisym) apply (rule Sup_least, safe) apply (simp add: bres_galois) using unitr_eq apply fastforce apply (metis bot.extremum) by (metis (mono_tags, lifting) Sup_upper mem_Collect_eq unitr_ex) lemma bmod_oner: "\ \ \ f x = f x" apply (simp add: bmod_comp_def pid_def) apply (rule antisym) apply (rule Sup_least, safe) apply (simp add: fres_galois) using unitl_eq apply fastforce apply (metis bot.extremum) by (metis (mono_tags, lifting) Sup_upper mem_Collect_eq unitl_ex) lemma pid_unitl [simp]: "\ \ f = f" by (simp add: fun_eq_iff times_rel_fun_def bmod_oner) lemma pid_unitr [simp]: "f \ (\::'a \ 'b::unital_quantale) = f" by (simp add: fun_eq_iff times_rel_fun_def bmod_onel) lemma bmod_assoc_weak_aux: "f u \ \{g v \ h z |v z. \ y v z} = \{(f::'a \ 'b::weak_quantale) u \ g v \ h z |v z. \ y v z}" apply (subst weak_Sup_distl) using unitl_ex apply force apply simp by (metis (no_types, lifting) mult.assoc) lemma bmod_assoc_weak: "\ \ (\ \ (f::'a \ 'b::weak_quantale) g) h x = \ \ f (\ \ g h) x" proof - have "\ \ (\ \ f g) h x = \{\{f u \ g v \ h z |u v. \ y u v} |y z. \ x y z}" apply (simp add: bmod_comp_def Sup_distr) apply (rule antisym) by (intro Sup_least Sup_upper, safe, intro exI conjI, simp_all, rule_tac f = Sup and g = Sup in cong, auto)+ also have "... = \{f u \ g v \ h z |u v y z. \ y u v \ \ x y z}" by (simp add: bmod_assoc_aux1) also have "... = \{f u \ g v \ h z |u v z y. \ y v z \ \ x u y}" apply (rule antisym) apply (rule Sup_least, rule Sup_upper, safe) using rel_assoc apply force apply (rule Sup_least, rule Sup_upper, safe) using rel_assoc by blast also have "... = \{\{f u \ g v \ h z |v z. \ y v z} |u y. \ x u y}" by (simp add: bmod_assoc_aux2) also have "... = \{f u \ \{g v \ h z |v z. \ y v z} |u y. \ x u y}" by (simp add: bmod_assoc_weak_aux) finally show ?thesis by (auto simp: bmod_comp_def) qed lemma rel_fun_assoc_weak: "((f :: 'a \ 'b::weak_quantale) \ g) \ h = f \ (g \ h)" by (simp add: times_rel_fun_def fun_eq_iff bmod_assoc_weak[symmetric]) end lemma (in rel_semigroup) "\id. \f x. (\ \ f id x = f x \ \ \ id f x = f x)" (* apply (simp add: bmod_comp_def) nitpick [expect=genuine] *) oops class rel_ab_semigroup = rel_semigroup + assumes rel_comm: "\ x y z \ \ x z y" begin lemma bmod_comm: "\ \ (f::'a \ 'b::ab_quantale) g = \ \ g f" by (simp add: fun_eq_iff bmod_comp_def mult.commute, meson rel_comm) lemma "\ \ f g = \ \ g f" (* nitpick [expect=genuine] *) oops lemma bmod_bres_fres_eq: "\ \ (f::'a \ 'b::ab_quantale) g = \ \ g f" by (simp add: fun_eq_iff bmod_bres_def bmod_fres_def bres_fres_eq, meson rel_comm) lemma rel_fun_comm: "(f :: 'a \ 'b::ab_quantale) \ g = g \ f" by (simp add: times_rel_fun_def bmod_comm) end class rel_ab_monoid = rel_ab_semigroup + rel_monoid subsection \Lifting to Function Spaces\ text \We lift by interpretation, since we need sort instantiations to be used for functions from PAM's to Quantales. Both instantiations cannot be used in Isabelle at the same time.\ interpretation rel_fun: weak_proto_quantale Inf Sup inf less_eq less sup bot "top :: 'a::rel_magma \ 'b::weak_proto_quantale" times_rel_fun by standard (simp_all add: rel_fun_Sup_distr rel_fun_Sup_distl_weak) interpretation rel_fun: proto_quantale Inf Sup inf less_eq less sup bot "top :: 'a::rel_magma \ 'b::proto_quantale" times_rel_fun by standard (simp add: rel_fun_Sup_distl) text \Nitpick shows that the lifting of weak quantales to weak quantales does not work for relational semigroups, because associativity fails.\ interpretation rel_fun: weak_quantale times_rel_fun Inf Sup inf less_eq less sup bot "top::'a::rel_semigroup \ 'b::weak_quantale" (*apply standard nitpick[expect=genuine]*) oops text \Associativity is obtained when lifting from relational monoids, but the right unit law doesn't hold in the quantale on the function space, according to our above results. Hence the lifting results into a non-unital quantale, in which only the left unit law holds, as shown above. We don't provide a special class for such quantales. Hence we lift only to non-unital quantales.\ interpretation rel_fun: weak_quantale times_rel_fun Inf Sup inf less_eq less sup bot "top::'a::rel_monoid \ 'b::unital_weak_quantale" by standard (simp_all add: rel_fun_assoc_weak) interpretation rel_fun2: quantale times_rel_fun Inf Sup inf less_eq less sup bot "top::'a::rel_semigroup \ 'b::quantale" by standard (simp add: rel_fun_assoc) interpretation rel_fun2: distrib_quantale Inf Sup inf less_eq less sup bot "top::'a::rel_semigroup \ 'b::distrib_quantale" times_rel_fun .. -interpretation rel_fun2: bool_quantale Inf Sup inf less_eq less sup bot "top::'a::rel_semigroup \ 'b::bool_quantale" minus uminus times_rel_fun .. +interpretation rel_fun2: bool_quantale minus uminus inf less_eq less sup bot \top::'a::rel_semigroup \ 'b::bool_quantale\ Inf Sup times_rel_fun .. interpretation rel_fun2: unital_quantale pid times_rel_fun Inf Sup inf less_eq less sup bot "top::'a::rel_monoid \ 'b::unital_quantale" by (standard; simp) interpretation rel_fun2: distrib_unital_quantale Inf Sup inf less_eq less sup bot "top::'a::rel_monoid \ 'b::distrib_unital_quantale" pid times_rel_fun .. -interpretation rel_fun2: bool_unital_quantale Inf Sup inf less_eq less sup bot "top::'a::rel_monoid \ 'b::bool_unital_quantale" minus uminus pid times_rel_fun .. +interpretation rel_fun2: bool_unital_quantale minus uminus inf less_eq less sup bot \top::'a::rel_monoid \ 'b::bool_unital_quantale\ Inf Sup pid times_rel_fun .. interpretation rel_fun: ab_quantale times_rel_fun Inf Sup inf less_eq less sup bot "top::'a::rel_ab_semigroup \ 'b::ab_quantale" by standard (simp add: rel_fun_comm) interpretation rel_fun: ab_unital_quantale times_rel_fun Inf Sup inf less_eq less sup bot "top::'a::rel_ab_monoid \ 'b::ab_unital_quantale" pid .. interpretation rel_fun2: distrib_ab_unital_quantale Inf Sup inf less_eq less sup bot "top::'a::rel_ab_monoid \ 'b::distrib_ab_unital_quantale" times_rel_fun pid .. interpretation rel_fun2: bool_ab_unital_quantale times_rel_fun Inf Sup inf less_eq less sup bot "top::'a::rel_ab_monoid \ 'b::bool_ab_unital_quantale" minus uminus pid .. end - - - - diff --git a/thys/Relation_Algebra/More_Boolean_Algebra.thy b/thys/Relation_Algebra/More_Boolean_Algebra.thy --- a/thys/Relation_Algebra/More_Boolean_Algebra.thy +++ b/thys/Relation_Algebra/More_Boolean_Algebra.thy @@ -1,293 +1,293 @@ (* Title: Relation Algebra Author: Alasdair Armstrong, Simon Foster, Georg Struth, Tjark Weber Maintainer: Georg Struth Tjark Weber *) section \(More) Boolean Algebra\ theory More_Boolean_Algebra imports Main begin subsection \Laws of Boolean Algebra\ text \The following laws of Boolean algebra support relational proofs. We might add laws for the binary minus since that would make certain theorems look more nicely. These are currently not so well supported.\ -context Lattices.boolean_algebra +context boolean_algebra begin no_notation times (infixl "\" 70) and plus (infixl "+" 65) and Groups.zero_class.zero ("0") and Groups.one_class.one ("1") notation inf (infixl "\" 70) and sup (infixl "+" 65) and bot ("0") and top ("1") lemma meet_assoc: "x \ (y \ z) = (x \ y) \ z" by (metis inf_assoc) lemma aux4 [simp]: "x \ y + x \ -y = x" by (metis inf_sup_distrib1 inf_top_right sup_compl_top) lemma aux4_comm [simp]: "x \ -y + x \ y = x" by (metis aux4 sup.commute) lemma aux6 [simp]: "(x + y) \ -x = y \ -x" by (metis inf_compl_bot inf_sup_distrib2 sup_bot_left) lemma aux6_var [simp]: "(-x + y) \ x = x \ y" by (metis compl_inf_bot inf_commute inf_sup_distrib2 sup_bot_left) lemma aux9 [simp]: "x + -x \ y = x + y" by (metis aux4 aux6 inf.commute inf_sup_absorb) lemma join_iso: "x \ y \ x + z \ y + z" by (metis eq_refl sup_mono) lemma join_isol: "x \ y \ z + x \ z + y" by (metis join_iso sup.commute) lemma join_double_iso: "x \ y \ w + x + z \ w + y + z" by (metis le_iff_inf sup_inf_distrib1 sup_inf_distrib2) lemma comp_anti: "x \ y \ -y \ -x" by (metis compl_le_swap2 double_compl) lemma meet_iso: "x \ y \ x \ z \ y \ z" by (metis eq_refl inf_mono) lemma meet_isor: "x \ y \ z \ x \ z \ y" by (metis inf.commute meet_iso) lemma meet_double_iso: "x \ y \ w \ x \ z \ w \ y \ z" by (metis meet_iso meet_isor) lemma de_morgan_3 [simp]: "-(-x \ -y) = x + y" by (metis compl_sup double_compl) lemma subdist_2_var: "x + y \ z \ x + y" by (metis eq_refl inf_le1 sup_mono) lemma dist_alt: "\x + z = y + z; x \ z = y \ z\ \ x = y" by (metis aux4 aux6 sup.commute) text \Finally we prove the Galois connections for complementation.\ lemma galois_aux: "x \ y = 0 \ x \ -y" by (metis aux6 compl_sup double_compl inf.commute le_iff_inf sup_bot_right sup_compl_top) lemma galois_aux2: "x \ -y = 0 \ x \ y" by (metis double_compl galois_aux) lemma galois_1: "x \ -y \ z \ x \ y + z" apply (rule iffI) apply (metis inf_le2 join_iso le_iff_sup le_supE join_isol aux4) apply (metis meet_iso aux6 le_infE) done lemma galois_2: "x \ y + -z \ x \ z \ y" apply (rule iffI) apply (metis compl_sup double_compl galois_1 inf.commute) apply (metis inf.commute order_trans subdist_2_var aux4 join_iso) done lemma galois_aux3: "x + y = 1 \ -x \ y" by (metis galois_1 inf_top_left top_unique) lemma galois_aux4: "-x + y = 1 \ x \ y" by (metis double_compl galois_aux3) subsection \Boolean Algebras with Operators\ text \We follow J\'onsson and Tarski to define pairs of conjugate functions on Boolean algebras. We also consider material from Maddux's article. This gives rise to a Galois connection and the notion of Boolean algebras with operators. We do not explicitly define families of functions over Boolean algebras as a type class. This development should certainly be expanded do deal with complete Boolean algebras one the one hand and other lattices on the other hand. Boolean algebras with operators and their variants can be applied in various ways. The prime example are relation algebras. The modular laws, for instance, can be derived by instantiation. Other applications are antidomain semirings where modal operators satisfy conjugations and Galois connections, and algebras of predicate transformers.\ text\We define conjugation as a predicate which holds if a pair of functions are conjugates.\ definition is_conjugation :: "('a \ 'a) \ ('a \ 'a) \ bool" where "is_conjugation f g \ (\x y . f x \ y = 0 \ x \ g y = 0)" text \We now prove the standard lemmas. First we show that conjugation is symmetric and that conjugates are uniqely defined.\ lemma is_conjugation_sym: "is_conjugation f g \ is_conjugation g f" by (metis inf.commute is_conjugation_def) lemma is_conjugation_unique: "\is_conjugation f g; is_conjugation f h\ \ g = h" by (metis galois_aux inf.commute double_compl order.eq_iff ext is_conjugation_def) text \Next we show that conjugates give rise to adjoints in a Galois connection.\ lemma conj_galois_1: assumes "is_conjugation f g" shows "f x \ y \ x \ -g (-y)" by (metis assms is_conjugation_def double_compl galois_aux) lemma conj_galois_2: assumes "is_conjugation f g" shows "g x \ y \ x \ -f (-y)" by (metis assms is_conjugation_sym conj_galois_1) text \Now we prove some of the standard properties of adjoints and conjugates. In fact, conjugate functions even distribute over all existing suprema. We display the next proof in detail because it is elegant.\ lemma f_pre_additive: assumes "is_conjugation f g" shows "f (x + y) \ z \ f x + f y \ z" proof - have "f (x + y) \ z \ x + y \ -g (-z)" by (metis assms conj_galois_1) also have "... \ x \ -g (-z) \ y \ -g (-z)" by (metis le_sup_iff) also have "... \ f x \ z \ f y \ z" by (metis assms conj_galois_1) thus ?thesis by (metis le_sup_iff calculation) qed lemma f_additive: assumes "is_conjugation f g" shows "f (sup x y) = sup (f x) (f y)" by (metis assms order.eq_iff f_pre_additive) lemma g_pre_additive: assumes "is_conjugation f g" shows "g (sup x y) \ z \ sup (g x) (g y) \ z" by (metis assms is_conjugation_sym f_pre_additive) lemma g_additive: assumes "is_conjugation f g" shows "g (sup x y) = sup (g x) (g y)" by (metis assms is_conjugation_sym f_additive) text \Additivity of adjoints obviously implies their isotonicity.\ lemma f_iso: assumes "is_conjugation f g" shows "x \ y \ f x \ f y" by (metis assms f_additive le_iff_sup) lemma g_iso: assumes "is_conjugation f g" shows "x \ y \ g x \ g y" by (metis assms is_conjugation_sym f_iso) lemma f_subdist: assumes "is_conjugation f g" shows "f (x \ y) \ f x" by (metis assms f_iso inf_le1) lemma g_subdist: assumes "is_conjugation f g" shows "g (x \ y) \ g x" by (metis assms g_iso inf_le1) text \Next we prove cancellation and strictness laws.\ lemma cancellation_1: assumes "is_conjugation f g" shows "f (-g x) \ -x" by (metis assms conj_galois_1 double_compl eq_refl) lemma cancellation_2: assumes "is_conjugation f g" shows "g (-f x) \ -x" by (metis assms is_conjugation_sym cancellation_1) lemma f_strict: assumes "is_conjugation f g" shows "f 0 = 0" by (metis assms inf.idem inf_bot_left is_conjugation_def) lemma g_strict: assumes "is_conjugation f g" shows "g 0 = 0" by (metis assms is_conjugation_sym f_strict) text \The following variants of modular laws have more concrete counterparts in relation algebra.\ lemma modular_1_aux: assumes "is_conjugation f g" shows "f (x \ -g y) \ y = 0" by (metis assms galois_aux inf_le2 is_conjugation_def) lemma modular_2_aux: assumes "is_conjugation f g" shows "g (x \ -f y) \ y = 0" by (metis assms is_conjugation_sym modular_1_aux) lemma modular_1: assumes "is_conjugation f g" shows "f x \ y = f (x \ g y) \ y" proof - have "f x \ y = f (x \ g y + x \ -g y) \ y" by (metis aux4) hence "f x \ y = (f (x \ g y) + f (x \ -g y)) \ y" by (metis assms f_additive) hence "f x \ y = f (x \ g y) \ y + f (x \ -g y) \ y" by (metis inf.commute inf_sup_distrib1) thus ?thesis by (metis assms modular_1_aux sup_bot_right) qed lemma modular_2: assumes "is_conjugation f g" shows "g x \ y = g (x \ f y) \ y" by (metis assms is_conjugation_sym modular_1) lemma conjugate_eq_aux: "is_conjugation f g \ f (x \ -g y) \ f x \ -y" by (metis f_subdist galois_aux le_inf_iff modular_1_aux) lemma conjugate_eq: "is_conjugation f g \ (\x y. f (x \ -g y) \ f x \ -y \ g (y \ -f x) \ g y \ -x)" (is "?l \ ?r") proof assume ?l thus ?r by (metis is_conjugation_sym conjugate_eq_aux) next assume r: ?r have "\x y. f x \ y = 0 \ x \ g y = 0" by (metis aux4 inf.left_commute inf_absorb1 inf_compl_bot inf_left_idem sup_bot_left r) hence "\x y. x \ g y = 0 \ f x \ y = 0" by (metis aux4 inf.commute inf.left_commute inf_absorb1 inf_compl_bot sup_commute sup_inf_absorb r) thus "is_conjugation f g" by (metis is_conjugation_def) qed lemma conjugation_prop1: "is_conjugation f g \ f y \ z \ f (y \ g z)" by (metis le_infE modular_1 order_refl) lemma conjugation_prop2: "is_conjugation f g \ g z \ y \ g (z \ f y)" by (metis is_conjugation_sym conjugation_prop1) end (* boolean_algebra *) end diff --git a/thys/Relation_Algebra/Relation_Algebra_Models.thy b/thys/Relation_Algebra/Relation_Algebra_Models.thy --- a/thys/Relation_Algebra/Relation_Algebra_Models.thy +++ b/thys/Relation_Algebra/Relation_Algebra_Models.thy @@ -1,149 +1,149 @@ (* Title: Relation Algebra Author: Alasdair Armstrong, Simon Foster, Georg Struth, Tjark Weber Maintainer: Georg Struth Tjark Weber *) section \Models of Relation Algebra\ theory Relation_Algebra_Models imports Relation_Algebra Kleene_Algebra.Inf_Matrix begin text \We formalise two models. First we show the obvious: binary relations form a relation algebra. Then we show that infinite matrices (which we formalised originally for Kleene algebras) form models of relation algebra if we restrict their element type to @{typ bool}.\ subsection \Binary Relations\ text \Since Isabelle's libraries for binary relations are very well developed, the proof for this model is entirely trivial.\ interpretation rel_relation_algebra: relation_algebra "(-)" uminus "(\)" "(\)" "(\)" "(\)" "{}" UNIV "(O)" Relation.converse Id by unfold_locales auto subsection \Infinite Boolean Matrices\ text \Next we consider infinite Boolean matrices. We define the maximal Boolean matrix (all of its entries are @{const True}), the converse or transpose of a matrix, the intersection of two Boolean matrices and the complement of a Boolean matrix.\ definition mat_top :: "('a, 'b, bool) matrix" ("\") where "\ i j \ True" definition mat_transpose :: "('a, 'b, 'c) matrix \ ('b, 'a, 'c) matrix" ("_\<^sup>\" [101] 100) where "f\<^sup>\ \ (\i j. f j i)" definition mat_inter :: "('a, 'b, bool) matrix \ ('a, 'b, bool) matrix \ ('a, 'b, bool) matrix" (infixl "\" 70) where "f \ g \ (\i j. f i j \ g i j)" definition mat_complement :: "('a, 'b, bool) matrix \ ('a, 'b, bool) matrix" ("_\<^sup>c" [101] 100) where "f\<^sup>c = (\i j. - f i j)" text \Next we show that the Booleans form a dioid. We state this as an \emph{instantiation} result. The Kleene algebra files contain an \emph{interpretation} proof, which is not sufficient for our purposes.\ instantiation bool :: dioid_one_zero begin definition zero_bool_def: "zero_bool \ False" definition one_bool_def: "one_bool \ True" definition times_bool_def: "times_bool \ (\)" definition plus_bool_def: "plus_bool \ (\)" instance by standard (auto simp: plus_bool_def times_bool_def one_bool_def zero_bool_def) end text \We now show that infinite Boolean matrices form a Boolean algebra.\ lemma le_funI2: "(\i j. f i j \ g i j) \ f \ g" by (metis le_funI) -interpretation matrix_ba: Lattices.boolean_algebra "\f g. f \ g\<^sup>c" mat_complement "(\)" "(\)" "(<)" mat_add mat_zero mat_top +interpretation matrix_ba: boolean_algebra "\f g. f \ g\<^sup>c" mat_complement "(\)" "(\)" "(<)" mat_add mat_zero mat_top by standard (force intro!: le_funI simp: mat_inter_def plus_bool_def mat_add_def mat_zero_def zero_bool_def mat_top_def mat_complement_def)+ text \We continue working towards the main result of this section, that infinite Boolean matrices form a relation algebra.\ lemma mat_mult_var: "(f \ g) = (\i j. \ {(f i k) * (g k j) | k. k \ UNIV})" by (rule ext)+ (simp add: mat_mult_def) text \The following fact is related to proving the last relation algebra axiom in the matrix model. It is more complicated than necessary since finite infima are not well developed in Isabelle. Instead we translate properties of finite infima into properties of finite suprema by using Boolean algebra. For finite suprema we have developed special-purpose theorems in the Kleene algebra files.\ lemma mat_res_pointwise: fixes i j :: "'a::finite" and x :: "('a, 'a, bool) matrix" shows "(x\<^sup>\ \ (x \ y)\<^sup>c) i j \ (y\<^sup>c) i j" proof - have "\{(x\<^sup>\) i k \ ((x \ y)\<^sup>c) k j |k. k \ UNIV} \ (y\<^sup>c) i j \ (\k. ((x\<^sup>\) i k \ ((x \ y)\<^sup>c) k j) \ (y\<^sup>c) i j)" by (subst sum_sup) auto also have "\ \ (\k. ((x\<^sup>\) i k \ - (x \ y) k j) \ (y\<^sup>c) i j)" by (simp only: mat_complement_def) also have "\ \ (\k. (x\<^sup>\) i k \ ((y\<^sup>c) i j \ (x \ y) k j))" by auto also have "\ \ (\k. (x\<^sup>\) i k \ (- y i j \ (x \ y) k j))" by (simp only: mat_complement_def) also have "\ \ (\k. ((x\<^sup>\) i k \ y i j) \ (x \ y) k j)" by auto also have "\ \ (\k. (x k i \ y i j) \ (x \ y) k j)" by (simp add: mat_transpose_def) also have "\ \ (\k. (x k i \ y i j) \ \{x k l \ y l j |l. l \ UNIV})" by (simp add: mat_mult_def times_bool_def) also have "\ \ (\k. \{x k i \ y i j} \ \{x k l \ y l j |l. l \ UNIV})" by simp also have "\ \ True" by (intro iffI TrueI allI sum_intro[rule_format]) auto moreover have "(x\<^sup>\ \ (x \ y)\<^sup>c) i j = \{(x\<^sup>\) i k \ ((x \ y)\<^sup>c) k j |k. k \ UNIV}" by (subst mat_mult_def) (simp add: times_bool_def) ultimately show ?thesis by auto qed text \Finally the main result of this section.\ interpretation matrix_ra: relation_algebra "\f g. f \ g\<^sup>c" mat_complement "(\)" "(\)" "(<)" "(\)" "\i j. False" \ "(\)" mat_transpose \ proof fix x y z :: "'a::finite \ 'a \ bool" show "(\(i::'a) j::'a. False) \ x" by (metis predicate2I) show "x \ x\<^sup>c = (\i j. False)" by (metis matrix_ba.bot.extremum matrix_ba.inf_compl_bot rev_predicate2D) show "x \ x\<^sup>c = \" by (fact matrix_ba.sup_compl_top) show "x \ y\<^sup>c = x \ y\<^sup>c" by (fact refl) show "x \ y \ z = x \ (y \ z)" by (metis mat_mult_assoc) show "x \ \ = x" by (fact mat_oner) show "x \ y \ z = (x \ z) \ (y \ z)" by (fact mat_distr) show "(x\<^sup>\)\<^sup>\ = x" by (simp add: mat_transpose_def) show "(x \ y)\<^sup>\ = x\<^sup>\ \ y\<^sup>\" by (simp add: mat_transpose_def mat_add_def) show "(x \ y)\<^sup>\ = y\<^sup>\ \ x\<^sup>\" by (simp add: mat_transpose_def mat_mult_var times_bool_def conj_commute) show "x\<^sup>\ \ (x \ y)\<^sup>c \ y\<^sup>c" by (metis le_funI2 mat_res_pointwise) qed end 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,665 @@ (* Title: Residuated Boolean Algebras Author: Victor Gomes Maintainer: Georg Struth *) section \Residuated Boolean Algebras\ theory Residuated_Boolean_Algebras imports Residuated_Lattices begin subsection \Conjugation on Boolean Algebras\ text \ Similarly, as in the previous section, we define the conjugation for arbitrary residuated functions on boolean algebras. \ -context Lattices.boolean_algebra +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)" by (insert maddux2 [of "\x. a \ x"]) simp 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)". qed end (* residuated_boolean_monoid *) end diff --git a/thys/Residuated_Lattices/Residuated_Relation_Algebra.thy b/thys/Residuated_Lattices/Residuated_Relation_Algebra.thy --- a/thys/Residuated_Lattices/Residuated_Relation_Algebra.thy +++ b/thys/Residuated_Lattices/Residuated_Relation_Algebra.thy @@ -1,118 +1,118 @@ (* Title: Residuated Relation Algebras Author: Victor Gomes Maintainer: Georg Struth *) section \Residuated Relation Algebras\ theory Residuated_Relation_Algebra imports Residuated_Boolean_Algebras Relation_Algebra.Relation_Algebra begin -context Lattices.boolean_algebra begin +context boolean_algebra begin text \ The notation used in the relation algebra AFP entry differs a little from ours. \ notation times (infixl "\" 70) and plus (infixl "+" 65) and Groups.zero_class.zero ("0") and Groups.one_class.one ("1") no_notation inf (infixl "\" 70) and sup (infixl "+" 65) and bot ("0") and top ("1") end text \ We prove that a unital residuated boolean algebra enriched with two simple equalities form a non-associative relation algebra, that is, a relation algebra where the associativity law does not hold. \ class nra = unital_residuated_boolean + assumes conv1: "x \ y = (x \ 1)\y" and conv2: "x \ y = x\(1 \ y)" begin text \ When the converse operation is set to be $\lambda x.\ x \rhd 1$, a unital residuated boolean algebra forms a non associative relation algebra. \ lemma conv_invol: "x \ 1 \ 1 = x" by (metis local.conv1 local.jonsson1b local.mult_onel) lemma conv_add: "x \ y \ 1 = (x \ 1) \ (y \ 1)" by (metis local.conjr2_sup) lemma conv_contrav: "x\y \ 1 = (y \ 1)\(x \ 1)" by (metis local.conv1 local.conv2 local.jonsson1b local.jonsson2c) lemma conv_res: "(x \ 1)\- (x\y) \ - y" by (metis local.comp_anti local.conjugate_r_def local.conv1 local.double_compl local.res_rc1) text \ Similarly, for $x^\smile = 1 \lhd x$, since $x \rhd 1 = 1 \lhd x$ when $x \rhd 1 \rhd 1 = x$ holds. \ lemma conv_invol': "1 \ (1 \ x) = x" by (metis local.conv_invol local.jonsson3c) lemma conv_add': "1 \ (x \ y) = (1 \ x) \ (1 \ y)" by (metis local.conjl1_sup) lemma conv_contrav': "1 \ x\y = (1 \ y)\(1 \ x)" by (metis local.conv_contrav local.conv_invol local.jonsson3c) lemma conv_res': "(1 \ x)\- (x\y) \ - y" by (metis conv_res local.conv_invol local.jonsson3c) end (* nra *) text \ Since the previous axioms are equivalent when multiplication is associative in a residuated boolean monoid, one of them are sufficient to derive a relation algebra. \ class residuated_ra = residuated_boolean_monoid + assumes conv: "x \ y = (x \ 1)\y" begin subclass nra proof (standard, fact conv) fix x y show "x \ y = x\(1 \ y)" by (metis conv jonsson4) qed sublocale relation_algebra where composition = "(\)" and unit = 1 and converse = "\x. x \ 1" proof fix x y z show "x\y\z = x\(y\z)" by (metis local.mult_assoc) show "x\1 = x" by (metis local.mult_onel) show "(x \ y)\z = x\z \ y\z" by (metis local.distr) show "x \ 1 \ 1 = x" by (metis local.conv_invol) show "x \ y \ 1 = (x \ 1) \ (y \ 1)" by (metis local.conv_add) show "x\y \ 1 = (y \ 1)\(x \ 1)" by (metis local.conv_contrav) show "(x \ 1)\- (x\y) \ - y" by (metis local.conv_res) qed end (* residuated_ra *) end diff --git a/thys/Robbins-Conjecture/Robbins_Conjecture.thy b/thys/Robbins-Conjecture/Robbins_Conjecture.thy --- a/thys/Robbins-Conjecture/Robbins_Conjecture.thy +++ b/thys/Robbins-Conjecture/Robbins_Conjecture.thy @@ -1,985 +1,985 @@ section \Robbins Conjecture\ theory Robbins_Conjecture imports Main begin text \The document gives a formalization of the proof of the Robbins conjecture, following A. Mann, \emph{A Complete Proof of the Robbins Conjecture}, 2003, DOI 10.1.1.6.7838\ section \Axiom Systems\ text \The following presents several axiom systems that shall be under study. The first axiom sets common systems that underly all of the systems we shall be looking at. The second system is a reformulation of Boolean algebra. We shall follow pages 7--8 in S. Koppelberg. \emph{General Theory of Boolean Algebras}, Volume 1 of \emph{Handbook of Boolean Algebras}. North Holland, 1989. Note that our formulation deviates slightly from this, as we only provide one distribution axiom, as the dual is redundant. The third system is Huntington's algebra and the fourth system is Robbins' algebra. Apart from the common system, all of these systems are demonstrated to be equivalent to the library formulation of Boolean algebra, under appropriate interpretation.\ subsection \Common Algebras\ class common_algebra = uminus + fixes inf :: "'a \ 'a \ 'a" (infixl "\" 70) fixes sup :: "'a \ 'a \ 'a" (infixl "\" 65) fixes bot :: "'a" ("\") fixes top :: "'a" ("\") assumes sup_assoc: "x \ (y \ z) = (x \ y) \ z" assumes sup_comm: "x \ y = y \ x" context common_algebra begin definition less_eq :: "'a \ 'a \ bool" (infix "\" 50) where "x \ y = (x \ y = y)" definition less :: "'a \ 'a \ bool" (infix "\" 50) where "x \ y = (x \ y \ \ y \ x)" definition minus :: "'a \ 'a \ 'a" (infixl "-" 65) where "minus x y = (x \ - y)" (* We shall need some object in order to define falsum and verum *) definition secret_object1 :: "'a" ("\") where "\ = (SOME x. True)" end class ext_common_algebra = common_algebra + assumes inf_eq: "x \ y = -(- x \ - y)" assumes top_eq: "\ = \ \ - \" assumes bot_eq: "\ = -(\ \ - \)" subsection \Boolean Algebra\ class boolean_algebra_II = common_algebra + assumes inf_comm: "x \ y = y \ x" assumes inf_assoc: "x \ (y \ z) = (x \ y) \ z" assumes sup_absorb: "x \ (x \ y) = x" assumes inf_absorb: "x \ (x \ y) = x" assumes sup_inf_distrib1: "x \ y \ z = (x \ y) \ (x \ z)" assumes sup_compl: "x \ - x = \" assumes inf_compl: "x \ - x = \" subsection \Huntington's Algebra\ class huntington_algebra = ext_common_algebra + assumes huntington: "- (-x \ -y) \ - (-x \ y) = x" subsection \Robbins' Algebra\ class robbins_algebra = ext_common_algebra + assumes robbins: "- (- (x \ y) \ - (x \ -y)) = x" section \Equivalence\ text \With our axiom systems defined, we turn to providing equivalence results between them. We shall begin by illustrating equivalence for our formulation and the library formulation of Boolean algebra.\ subsection \Boolean Algebra\ text \The following provides the canonical definitions for order and relative complementation for Boolean algebras. These are necessary since the Boolean algebras presented in the Isabelle/HOL library have a lot of structure, while our formulation is considerably simpler. Since our formulation of Boolean algebras is considerably simple, it is easy to show that the library instantiates our axioms.\ context boolean_algebra_II begin lemma boolean_II_is_boolean: "class.boolean_algebra minus uminus (\) (\) (\) (\) \ \" apply unfold_locales apply (metis inf_absorb inf_assoc inf_comm inf_compl less_def less_eq_def minus_def sup_absorb sup_assoc sup_comm sup_compl sup_inf_distrib1 sup_absorb inf_comm)+ done end -context Lattices.boolean_algebra begin +context boolean_algebra begin lemma boolean_is_boolean_II: "class.boolean_algebra_II uminus inf sup bot top" apply unfold_locales apply (metis sup_assoc sup_commute sup_inf_absorb sup_compl_top inf_assoc inf_commute inf_sup_absorb inf_compl_bot sup_inf_distrib1)+ done end subsection \Huntington Algebra\ text \We shall illustrate here that all Boolean algebra using our formulation are Huntington algebras, and illustrate that every Huntington algebra may be interpreted as a Boolean algebra. Since the Isabelle/HOL library has good automation, it is convenient to first show that the library instances Huntington algebras to exploit previous results, and then use our previously derived correspondence.\ -context Lattices.boolean_algebra begin +context boolean_algebra begin lemma boolean_is_huntington: "class.huntington_algebra uminus inf sup bot top" apply unfold_locales apply (metis double_compl inf_sup_distrib1 inf_top_right compl_inf inf_commute inf_compl_bot compl_sup sup_commute sup_compl_top sup_compl_top sup_assoc)+ done end context boolean_algebra_II begin lemma boolean_II_is_huntington: "class.huntington_algebra uminus (\) (\) \ \" proof - interpret boolean: - Lattices.boolean_algebra minus uminus "(\)" "(\)" "(\)" "(\)" \ \ + boolean_algebra minus uminus "(\)" "(\)" "(\)" "(\)" \ \ by (fact boolean_II_is_boolean) show ?thesis by (simp add: boolean.boolean_is_huntington) qed end context huntington_algebra begin lemma huntington_id: "x \ -x = -x \ -(-x)" proof - from huntington have "x \ -x = -(-x \ -(-(-x))) \ -(-x \ -(-x)) \ (-(-(-x) \ -(-(-x))) \ -(-(-x) \ -(-x)))" by simp also from sup_comm have "\ = -(-(-x) \ -(-x)) \ -(-(-x) \ -(-(-x))) \ (-(-(-x) \ -x) \ -(-(-(-x)) \ -x))" by simp also from sup_assoc have "\ = -(-(-x) \ -(-x)) \ (-(-(-x) \ -(-(-x))) \ -(-(-x) \ -x)) \ -(-(-(-x)) \ -x)" by simp also from sup_comm have "\ = -(-(-x) \ -(-x)) \ (-(-(-x) \ -x) \ -(-(-x) \ -(-(-x)))) \ -(-(-(-x)) \ -x)" by simp also from sup_assoc have "\ = -(-(-x) \ -(-x)) \ -(-(-x) \ -x) \ (-(-(-x) \ -(-(-x))) \ -(-(-(-x)) \ -x))" by simp also from sup_comm have "\ = -(-(-x) \ -(-x)) \ -(-(-x) \ -x) \ (-(-(-(-x)) \ -(-x)) \ -(-(-(-x)) \ -x))" by simp also from huntington have "\ = -x \ -(-x)" by simp finally show ?thesis by simp qed lemma dbl_neg: "- (-x) = x" apply (metis huntington huntington_id sup_comm) done lemma towards_sup_compl: "x \ -x = y \ -y" proof - from huntington have "x \ -x = -(-x \ -(-y)) \ -(-x \ -y) \ (-(-(-x) \ -(-y)) \ -(-(-x) \ -y))" by simp also from sup_comm have "\ = -(-(-y) \ -x) \ -(-y \ -x) \ (-(-y \ -(-x)) \ -(-(-y) \ -(-x)))" by simp also from sup_assoc have "\ = -(-(-y) \ -x) \ (-(-y \ -x) \ -(-y \ -(-x))) \ -(-(-y) \ -(-x))" by simp also from sup_comm have "\ = -(-y \ -(-x)) \ -(-y \ -x) \ -(-(-y) \ -x) \ -(-(-y) \ -(-x))" by simp also from sup_assoc have "\ = -(-y \ -(-x)) \ -(-y \ -x) \ (-(-(-y) \ -x) \ -(-(-y) \ -(-x)))" by simp also from sup_comm have "\ = -(-y \ -(-x)) \ -(-y \ -x) \ (-(-(-y) \ -(-x)) \ -(-(-y) \ -x))" by simp also from huntington have "y \ -y = \" by simp finally show ?thesis by simp qed lemma sup_compl: "x \ -x = \" by (simp add: top_eq towards_sup_compl) lemma towards_inf_compl: "x \ -x = y \ -y" by (metis dbl_neg inf_eq sup_comm sup_compl) lemma inf_compl: "x \ -x = \" by (metis dbl_neg sup_comm bot_eq towards_inf_compl inf_eq) lemma towards_idem: "\ = \ \ \" by (metis dbl_neg huntington inf_compl inf_eq sup_assoc sup_comm sup_compl) lemma sup_ident: "x \ \ = x" by (metis dbl_neg huntington inf_compl inf_eq sup_assoc sup_comm sup_compl towards_idem) lemma inf_ident: "x \ \ = x" by (metis dbl_neg inf_compl inf_eq sup_ident sup_comm sup_compl) lemma sup_idem: "x \ x = x" by (metis dbl_neg huntington inf_compl inf_eq sup_ident sup_comm sup_compl) lemma inf_idem: "x \ x = x" by (metis dbl_neg inf_eq sup_idem) lemma sup_nil: "x \ \ = \" by (metis sup_idem sup_assoc sup_comm sup_compl) lemma inf_nil: "x \ \ = \" by (metis dbl_neg inf_compl inf_eq sup_nil sup_comm sup_compl) lemma sup_absorb: "x \ x \ y = x" by (metis huntington inf_eq sup_idem sup_assoc sup_comm) lemma inf_absorb: "x \ (x \ y) = x" by (metis dbl_neg inf_eq sup_absorb) lemma partition: "x \ y \ x \ -y = x" by (metis dbl_neg huntington inf_eq sup_comm) lemma demorgans1: "-(x \ y) = -x \ -y" by (metis dbl_neg inf_eq) lemma demorgans2: "-(x \ y) = -x \ -y" by (metis dbl_neg inf_eq) lemma inf_comm: "x \ y = y \ x" by (metis inf_eq sup_comm) lemma inf_assoc: "x \ (y \ z) = x \ y \ z" by (metis dbl_neg inf_eq sup_assoc) lemma inf_sup_distrib1: "x \ (y \ z) = (x \ y) \ (x \ z)" proof - from partition have "x \ (y \ z) = x \ (y \ z) \ y \ x \ (y \ z) \ -y" .. also from inf_assoc have "\ = x \ ((y \ z) \ y) \ x \ (y \ z) \ -y" by simp also from inf_comm have "\ = x \ (y \ (y \ z)) \ x \ (y \ z) \ -y" by simp also from inf_absorb have "\ = (x \ y) \ (x \ (y \ z) \ -y)" by simp also from partition have "\ = ((x \ y \ z) \ (x \ y \ -z)) \ ((x \ (y \ z) \ -y \ z) \ (x \ (y \ z) \ -y \ -z))" by simp also from inf_assoc have "\ = ((x \ y \ z) \ (x \ y \ -z)) \ ((x \ ((y \ z) \ (-y \ z))) \ (x \ ((y \ z) \ (-y \ -z))))" by simp also from demorgans2 have "\ = ((x \ y \ z) \ (x \ y \ -z)) \ ((x \ ((y \ z) \ (-y \ z))) \ (x \ ((y \ z) \ -(y \ z))))" by simp also from inf_compl have "\ = ((x \ y \ z) \ (x \ y \ -z)) \ ((x \ ((y \ z) \ (-y \ z))) \ (x \ \))" by simp also from inf_nil have "\ = ((x \ y \ z) \ (x \ y \ -z)) \ ((x \ ((y \ z) \ (-y \ z))) \ \)" by simp also from sup_idem have "\ = ((x \ y \ z) \ (x \ y \ z) \ (x \ y \ -z)) \ ((x \ ((y \ z) \ (-y \ z))) \ \)" by simp also from sup_ident have "\ = ((x \ y \ z) \ (x \ y \ z) \ (x \ y \ -z)) \ (x \ ((y \ z) \ (-y \ z)))" by simp also from inf_comm have "\ = ((x \ y \ z) \ (x \ y \ z) \ (x \ y \ -z)) \ (x \ ((-y \ z) \ (y \ z)))" by simp also from sup_comm have "\ = ((x \ y \ z) \ (x \ y \ z) \ (x \ y \ -z)) \ (x \ ((-y \ z) \ (z \ y)))" by simp also from inf_assoc have "\ = ((x \ y \ z) \ (x \ (y \ z)) \ (x \ y \ -z)) \ (x \ (-y \ (z \ (z \ y))))" by simp also from inf_absorb have "\ = ((x \ y \ z) \ (x \ (y \ z)) \ (x \ y \ -z)) \ (x \ (-y \ z))" by simp also from inf_comm have "\ = ((x \ y \ z) \ (x \ (z \ y)) \ (x \ y \ -z)) \ (x \ (z \ -y))" by simp also from sup_assoc have "\ = ((x \ y \ z) \ ((x \ (z \ y)) \ (x \ y \ -z))) \ (x \ (z \ -y))" by simp also from sup_comm have "\ = ((x \ y \ z) \ ((x \ y \ -z) \ (x \ (z \ y)))) \ (x \ (z \ -y))" by simp also from sup_assoc have "\ = ((x \ y \ z) \ (x \ y \ -z)) \ ((x \ (z \ y)) \ (x \ (z \ -y)))" by simp also from inf_assoc have "\ = ((x \ y \ z) \ (x \ y \ -z)) \ ((x \ z \ y) \ (x \ z \ -y))" by simp also from partition have "\ = (x \ y) \ (x \ z)" by simp finally show ?thesis by simp qed lemma sup_inf_distrib1: "x \ (y \ z) = (x \ y) \ (x \ z)" proof - from dbl_neg have "x \ (y \ z) = -(-(-(-x) \ (y \ z)))" by simp also from inf_eq have "\ = -(-x \ (-y \ -z))" by simp also from inf_sup_distrib1 have "\ = -((-x \ -y) \ (-x \ -z))" by simp also from demorgans2 have "\ = -(-x \ -y) \ -(-x \ -z)" by simp also from demorgans1 have "\ = (-(-x) \ -(-y)) \ (-(-x) \ -(-z))" by simp also from dbl_neg have "\ = (x \ y) \ (x \ z)" by simp finally show ?thesis by simp qed lemma huntington_is_boolean_II: "class.boolean_algebra_II uminus (\) (\) \ \" apply unfold_locales apply (metis inf_comm inf_assoc sup_absorb inf_absorb sup_inf_distrib1 sup_compl inf_compl)+ done lemma huntington_is_boolean: "class.boolean_algebra minus uminus (\) (\) (\) (\) \ \" proof - interpret boolean_II: boolean_algebra_II uminus "(\)" "(\)" \ \ by (fact huntington_is_boolean_II) show ?thesis by (simp add: boolean_II.boolean_II_is_boolean) qed end subsection \Robbins' Algebra\ -context Lattices.boolean_algebra begin +context boolean_algebra begin lemma boolean_is_robbins: "class.robbins_algebra uminus inf sup bot top" apply unfold_locales apply (metis sup_assoc sup_commute compl_inf double_compl sup_compl_top inf_compl_bot diff_eq sup_bot_right sup_inf_distrib1)+ done end context boolean_algebra_II begin lemma boolean_II_is_robbins: "class.robbins_algebra uminus inf sup bot top" proof - interpret boolean: - Lattices.boolean_algebra minus uminus "(\)" "(\)" "(\)" "(\)" \ \ + boolean_algebra minus uminus "(\)" "(\)" "(\)" "(\)" \ \ by (fact boolean_II_is_boolean) show ?thesis by (simp add: boolean.boolean_is_robbins) qed end context huntington_algebra begin lemma huntington_is_robbins: "class.robbins_algebra uminus inf sup bot top" proof - interpret boolean: - Lattices.boolean_algebra minus uminus "(\)" "(\)" "(\)" "(\)" \ \ + boolean_algebra minus uminus "(\)" "(\)" "(\)" "(\)" \ \ by (fact huntington_is_boolean) show ?thesis by (simp add: boolean.boolean_is_robbins) qed end text \Before diving into the proof that the Robbins algebra is Boolean, we shall present some shorthand machinery\ context common_algebra begin (* Iteration Machinery/Shorthand *) primrec copyp :: "nat \ 'a \ 'a" (infix "\" 80) where copyp_0: "0 \ x = x" | copyp_Suc: "(Suc k) \ x = (k \ x) \ x" no_notation Product_Type.Times (infixr "\" 80) primrec copy :: "nat \ 'a \ 'a" (infix "\" 85) where "0 \ x = x" | "(Suc k) \ x = k \ x" (* Theorems for translating shorthand into syntax *) lemma one: "1 \ x = x" proof - have "1 = Suc(0)" by arith hence "1 \ x = Suc(0) \ x" by metis also have "\ = x" by simp finally show ?thesis by simp qed lemma two: "2 \ x = x \ x" proof - have "2 = Suc(Suc(0))" by arith hence "2 \ x = Suc(Suc(0)) \ x" by metis also have "\ = x \ x" by simp finally show ?thesis by simp qed lemma three: "3 \ x = x \ x \ x" proof - have "3 = Suc(Suc(Suc(0)))" by arith hence "3 \ x = Suc(Suc(Suc(0))) \ x" by metis also have "\ = x \ x \ x" by simp finally show ?thesis by simp qed lemma four: "4 \ x = x \ x \ x \ x" proof - have "4 = Suc(Suc(Suc(Suc(0))))" by arith hence "4 \ x = Suc(Suc(Suc(Suc(0)))) \ x" by metis also have "\ = x \ x \ x \ x" by simp finally show ?thesis by simp qed lemma five: "5 \ x = x \ x \ x \ x \ x" proof - have "5 = Suc(Suc(Suc(Suc(Suc(0)))))" by arith hence "5 \ x = Suc(Suc(Suc(Suc(Suc(0))))) \ x" by metis also have "\ = x \ x \ x \ x \ x" by simp finally show ?thesis by simp qed lemma six: "6 \ x = x \ x \ x \ x \ x \ x" proof - have "6 = Suc(Suc(Suc(Suc(Suc(Suc(0))))))" by arith hence "6 \ x = Suc(Suc(Suc(Suc(Suc(Suc(0)))))) \ x" by metis also have "\ = x \ x \ x \ x \ x \ x" by simp finally show ?thesis by simp qed (* Distribution Laws *) lemma copyp_distrib: "k \ (x \ y) = (k \ x) \ (k \ y)" proof (induct k) case 0 show ?case by simp case Suc thus ?case by (simp, metis sup_assoc sup_comm) qed corollary copy_distrib: "k \ (x \ y) = (k \ x) \ (k \ y)" by (induct k, (simp add: sup_assoc sup_comm copyp_distrib)+) lemma copyp_arith: "(k + l + 1) \ x = (k \ x) \ (l \ x)" proof (induct l) case 0 have "k + 0 + 1 = Suc(k)" by arith thus ?case by simp case (Suc l) note ind_hyp = this have "k + Suc(l) + 1 = Suc(k + l + 1)" by arith+ hence "(k + Suc(l) + 1) \ x = (k + l + 1) \ x \ x" by (simp add: ind_hyp) also from ind_hyp have "\ = (k \ x) \ (l \ x) \ x" by simp also note sup_assoc finally show ?case by simp qed lemma copy_arith: assumes "k \ 0" and "l \ 0" shows "(k + l) \ x = (k \ x) \ (l \ x)" using assms proof - from assms have "\ k'. Suc(k') = k" and "\ l'. Suc(l') = l" by arith+ from this obtain k' l' where A: "Suc(k') = k" and B: "Suc(l') = l" by fast+ from this have A1: "k \ x = k' \ x" and B1: "l \ x = l' \ x" by fastforce+ from A B have "k + l = Suc(k' + l' + 1)" by arith hence "(k + l) \ x = (k' + l' + 1) \ x" by simp also from copyp_arith have "\ = k' \ x \ l' \ x" by fast also from A1 B1 have "\ = k \ x \ l \ x" by fastforce finally show ?thesis by simp qed end text \The theorem asserting all Robbins algebras are Boolean comes in 6 movements. First: The Winker identity is proved. Second: Idempotence for a particular object is proved. Note that falsum is defined in terms of this object. Third: An identity law for falsum is derived. Fourth: Idempotence for supremum is derived. Fifth: The double negation law is proven Sixth: Robbin's algebras are proven to be Huntington Algebras.\ context robbins_algebra begin definition secret_object2 :: "'a" ("\") where "\ = -(-(\ \ \ \ \) \ \)" definition secret_object3 :: "'a" ("\") where "\ = \ \ \" definition secret_object4 :: "'a" ("\") where "\ = \ \ (-(\ \ -\) \ -(\ \ -\))" definition secret_object5 :: "'a" ("\") where "\ = \ \ -(\ \ -\)" definition winker_object :: "'a" ("\") where "\ = \ \ \ \ \" definition fake_bot :: "'a" ("\\") where "\\ = -(\ \ -\)" (* Towards Winker's Identity *) (* These lemmas are due to Alan Mann *) lemma robbins2: "y = -(-(-x \ y) \ -(x \ y))" by (metis robbins sup_comm) lemma mann0: "-(x \ y) = -(-(-(x \ y) \ -x \ y) \ y)" by (metis robbins sup_comm sup_assoc) lemma mann1: "-(-x \ y) = -(-(-(-x \ y) \ x \ y) \ y)" by (metis robbins sup_comm sup_assoc) lemma mann2: "y = -(-(-(-x \ y) \ x \ y \ y) \ -(-x \ y))" by (metis mann1 robbins sup_comm sup_assoc) lemma mann3: "z = -(-(-(-(-x \ y) \ x \ y \ y) \ -(-x \ y) \ z) \ -(y \ z))" proof - let ?w = "-(-(-x \ y) \ x \ y \ y) \ -(-x \ y)" from robbins[where x="z" and y="?w"] sup_comm mann2 have "z = -(-(y \ z) \ -(?w \ z))" by metis thus ?thesis by (metis sup_comm) qed lemma mann4: "-(y \ z) = -(-(-(-(-x \ y) \ x \ y \ y) \ -(-x \ y) \ -(y \ z) \ z) \ z)" proof - from robbins2[where x="-(-(-x \ y) \ x \ y \ y) \ -(-x \ y) \ z" and y="-(y \ z)"] mann3[where x="x" and y="y" and z="z"] have "-(y \ z) = -(z \ -(-(-(-x \ y) \ x \ y \ y) \ -(-x \ y) \ z \ -(y \ z)))" by metis with sup_comm sup_assoc show ?thesis by metis qed lemma mann5: "u = -(-(-(-(-(-x \ y) \ x \ y \ y) \ -(-x \ y) \ - (y \ z) \ z) \ z \ u) \ -(-(y \ z) \ u))" using robbins2[where x="-(-(-(-x \ y) \ x \ y \ y) \ -(-x \ y) \ -(y \ z) \ z) \ z" and y="u"] mann4[where x=x and y=y and z=z] sup_comm by metis lemma mann6: "-(- 3\x \ x) = -(-(-(- 3\x \ x) \ - 3\x) \ -(-(- 3\x \ x) \ 5\x))" proof - have "3+2=(5::nat)" and "3\(0::nat)" and "2\(0::nat)" by arith+ with copy_arith have \: "3\x \ 2\x = 5\x" by metis let ?p = "-(- 3\x \ x)" { fix q from sup_comm have "-(q \ 5\x) = -(5\x \ q)" by metis also from \ mann0[where x="3\x" and y="q \ 2\x"] sup_assoc sup_comm have "\ = -(-(-(3\x \ (q \ 2\x)) \ - 3\x \ (q \ 2\x)) \ (q \ 2\x))" by metis also from sup_assoc have "\ = -(-(-((3\x \ q) \ 2\x) \ - 3\x \ (q \ 2\x)) \ (q \ 2\x))" by metis also from sup_comm have "\ = -(-(-((q \ 3\x) \ 2\x) \ - 3\x \ (q \ 2\x)) \ (q \ 2\x))" by metis also from sup_assoc have "\ = -(-(-(q \ (3\x \ 2\x)) \ - 3\x \ (q \ 2\x)) \ (q \ 2\x))" by metis also from \ have "\ = -(-(-(q \ 5\x) \ - 3\x \ (q \ 2\x)) \ (q \ 2\x))" by metis also from sup_assoc have "\ = -(-(-(q \ 5\x) \ (- 3\x \ q) \ 2\x) \ (q \ 2\x))" by metis also from sup_comm have "\ = -(-(-(q \ 5\x) \ (q \ - 3\x) \ 2\x) \ (2\x \ q))" by metis also from sup_assoc have "\ = -(-(-(q \ 5\x) \ q \ - 3\x \ 2\x) \ 2\x \ q)" by metis finally have "-(q \ 5\x) = -(-(-(q \ 5\x) \ q \ - 3\x \ 2\x) \ 2\x \ q)" by simp } hence \: "-(?p \ 5\x) = -(-(-(?p \ 5\x) \ ?p \ - 3\x \ 2\x) \ 2\x \ ?p)" by simp from mann5[where x="3\x" and y="x" and z="2\x" and u="?p"] sup_assoc three[where x=x] five[where x=x] have "?p = -(-(-(-(?p \ 5\x) \ ?p \ -(x \ 2\x) \ 2\x) \ 2\x \ ?p) \ -(-(x \ 2\x) \ ?p))" by metis also from sup_comm have "\ = -(-(-(-(?p \ 5\x) \ ?p \ -(2\x \ x) \ 2\x) \ 2\x \ ?p) \ -(-(2\x \ x) \ ?p))" by metis also from two[where x=x] three[where x=x] have "\ = -(-(-(-(?p \ 5\x) \ ?p \ - 3\x \ 2\x) \ 2\x \ ?p) \ -(- 3\x \ ?p))" by metis also from \ have "\ = -(-(?p \ 5\x) \ -(- 3\x \ ?p))" by simp also from sup_comm have "\ = -(-(?p \ 5\x) \ -(?p \ - 3\x))" by simp also from sup_comm have "\ = -(-(?p \ - 3\x) \ -(?p \ 5\x))" by simp finally show ?thesis . qed lemma mann7: "- 3\x = -(-(- 3\x \ x) \ 5\x)" proof - let ?p = "-(- 3\x \ x)" let ?q = "?p \ - 3\x" let ?r = "-(?p \ 5\x)" from robbins2[where x="?q" and y="?r"] mann6[where x=x] have "?r = - (?p \ - (?q \ ?r))" by simp also from sup_comm have "\ = - (- (?q \ ?r) \ ?p)" by simp also from sup_comm have "\ = - (- (?r \ ?q) \ ?p)" by simp finally have \: "?r = - (- (?r \ ?q) \ ?p)" . from mann3[where x="3\x" and y="x" and z="- 3\x"] sup_comm have "- 3\x = -(-(-(?p \ 3\x \ x \ x) \ ?p \ - 3\x) \ ?p)" by metis also from sup_assoc have "\ = -(-(-(?p \ (3\x \ x \ x)) \ ?q) \ ?p)" by metis also from three[where x=x] five[where x=x] have "\ = -(-(?r \ ?q) \ ?p)" by metis finally have "- 3\x = -(-(?r \ ?q) \ ?p)" by metis with \ show ?thesis by simp qed lemma mann8: "-(- 3\x \ x) \ 2\x = -(-(-(- 3\x \ x) \ - 3\x \ 2\x) \ - 3\x)" (is "?lhs = ?rhs") proof - let ?p = "-(- 3\x \ x)" let ?q = "?p \ 2\x" let ?r = "3\x" have "3+2=(5::nat)" and "3\(0::nat)" and "2\(0::nat)" by arith+ with copy_arith have \: "3\x \ 2\x = 5\x" by metis from robbins2[where x="?r" and y="?q"] and sup_assoc have "?q = -(-(- 3\x \ ?q) \ -(3\x \ ?p \ 2\x))" by metis also from sup_comm have "\ = -(-(?q \ - 3\x) \ -(?p \ 3\x \ 2\x))" by metis also from \ sup_assoc have "\ = -(-(?q \ - 3\x) \ -(?p \ 5\x))" by metis also from mann7[where x=x] have "\ = -(-(?q \ - 3\x) \ - 3\x)" by metis also from sup_assoc have "\ = -(-(?p \ (2\x \ - 3\x)) \ - 3\x)" by metis also from sup_comm have "\ = -(-(?p \ (- 3\x \ 2\x)) \ - 3\x)" by metis also from sup_assoc have "\ = ?rhs" by metis finally show ?thesis by simp qed lemma mann9: "x = -(-(- 3\x \ x) \ - 3\x )" proof - let ?p = "-(- 3\x \ x)" let ?q = "?p \ 4\x" have "4+1=(5::nat)" and "1\(0::nat)" and "4\(0::nat)" by arith+ with copy_arith one have \: "4\x \ x = 5\x" by metis with sup_assoc robbins2[where y=x and x="?q"] have "x = -(-(-?q \ x) \ -(?p \ 5\x))" by metis with mann7 have "x = -(-(-?q \ x) \ - 3\x)" by metis moreover have "3+1=(4::nat)" and "1\(0::nat)" and "3\(0::nat)" by arith+ with copy_arith one have \: "3\x \ x = 4\x" by metis with mann1[where x="3\x" and y="x"] sup_assoc have "-(-?q \ x) = ?p" by metis ultimately show ?thesis by simp qed lemma mann10: "y = -(-(-(- 3\x \ x) \ - 3\x \ y) \ -(x \ y))" using robbins2[where x="-(- 3\x \ x) \ - 3\x" and y=y] mann9[where x=x] sup_comm by metis theorem mann: "2\x = -(- 3\x \ x) \ 2\x" using mann10[where x=x and y="2\x"] mann8[where x=x] two[where x=x] three[where x=x] sup_comm by metis corollary winkerr: "\ \ \ = \" using mann secret_object2_def secret_object3_def two three by metis corollary winker: "\ \ \ = \" by (metis winkerr sup_comm) corollary multi_winkerp: "\ \ k \ \ = \" by (induct k, (simp add: winker sup_comm sup_assoc)+) corollary multi_winker: "\ \ k \ \ = \" by (induct k, (simp add: multi_winkerp winker sup_comm sup_assoc)+) (* Towards Idempotence *) lemma less_eq_introp: "-(x \ -(y \ z)) = -(x \ y \ -z) \ y \ x" by (metis robbins sup_assoc less_eq_def sup_comm[where x=x and y=y]) corollary less_eq_intro: "-(x \ -(y \ z)) = -(x \ y \ -z) \ x \ y = x" by (metis less_eq_introp less_eq_def sup_comm) lemma eq_intro: "-(x \ -(y \ z)) = -(y \ -(x \ z)) \ x = y" by (metis robbins sup_assoc sup_comm) lemma copyp0: assumes "-(x \ -y) = z" shows "-(x \ -(y \ k \ (x \ z))) = z" using assms proof (induct k) case 0 show ?case by (simp, metis assms robbins sup_assoc sup_comm) case Suc note ind_hyp = this show ?case by (simp, metis ind_hyp robbins sup_assoc sup_comm) qed lemma copyp1: assumes "-(-(x \ -y) \ -y) = x" shows "-(y \ k \ (x \ -(x \ -y))) = -y" using assms proof - let ?z = "-(x \ - y)" let ?ky = "y \ k \ (x \ ?z)" have "-(x \ -?ky) = ?z" by (simp add: copyp0) hence "-(-?ky \ -(-y \ ?z)) = ?z" by (metis assms sup_comm) also have "-(?z \ -?ky) = x" by (metis assms copyp0 sup_comm) hence "?z = -(-y \ -(-?ky \ ?z))" by (metis sup_comm) finally show ?thesis by (metis eq_intro) qed corollary copyp2: assumes "-(x \ y) = -y" shows "-(y \ k \ (x \ -(x \ -y))) = -y" by (metis assms robbins sup_comm copyp1) lemma two_threep: assumes "-(2 \ x \ y) = -y" and "-(3 \ x \ y) = -y" shows "2 \ x \ y = 3 \ x \ y" using assms proof - from assms two three have A: "-(x \ x \ y) = -y" and B: "-(x \ x \ x \ y) = -y" by simp+ with sup_assoc copyp2[where x="x" and y="x \ x \ y" and k="0"] have "-(x \ x \ y \ x \ -(x \ -y)) = -y" by simp moreover from sup_comm sup_assoc A B copyp2[where x="x \ x" and y="y" and k="0"] have "-(y \ x \ x \ -(x \ x \ -y)) = -y" by fastforce with sup_comm sup_assoc have "-(x \ x \ y \ -(x \ (x \ -y))) = -y" by metis ultimately have "-(x \ x \ y \ -(x \ (x \ -y))) = -(x \ x \ y \ x \ -(x \ -y))" by simp with less_eq_intro have "x \ x \ y = x \ x \ y \ x" by metis with sup_comm sup_assoc two three show ?thesis by metis qed lemma two_three: assumes "-(x \ y) = -y \ -(-(x \ -y) \ -y) = x" shows "y \ 2 \ (x \ -(x \ -y)) = y \ 3 \ (x \ -(x \ -y))" (is "y \ ?z2 = y \ ?z3") using assms proof assume "-(x \ y) = -y" with copyp2[where k="Suc(0)"] copyp2[where k="Suc(Suc(0))"] two three have "-(y \ ?z2) = -y" and "-(y \ ?z3) = -y" by simp+ with two_threep sup_comm show ?thesis by metis next assume "-(-(x \ -y) \ -y) = x" with copyp1[where k="Suc(0)"] copyp1[where k="Suc(Suc(0))"] two three have "-(y \ ?z2) = -y" and "-(y \ ?z3) = -y" by simp+ with two_threep sup_comm show ?thesis by metis qed lemma sup_idem: "\ \ \ = \" proof - from winkerr two copyp2[where x="\" and y="\" and k="Suc(0)"] have "-\ = -(\ \ 2 \ (\ \ -(\ \ -\)))" by simp also from copy_distrib sup_assoc have "\ = -(\ \ 2 \ \ \ 2 \ (-(\ \ -\)))" by simp also from sup_assoc secret_object4_def two multi_winker[where k="2"] have "\ = -\" by metis finally have "-\ = -\" by simp with secret_object4_def sup_assoc three have "\ \ -(\ \ -\) = \ \ 3 \ (-(\ \ -\))" by simp also from copy_distrib[where k="3"] multi_winker[where k="3"] sup_assoc have "\ = \ \ 3 \ (\ \ -(\ \ -\))" by metis also from winker sup_comm two_three[where x="\" and y="\"] have "\ = \ \ 2 \ (\ \ -(\ \ -\))" by fastforce also from copy_distrib[where k="2"] multi_winker[where k="2"] sup_assoc two secret_object4_def have "\ = \" by metis finally have \: "\ \ -(\ \ -\) = \" by simp from secret_object4_def winkerr sup_assoc have "\ \ \ = \" by metis hence "\ \ \ = \" by (metis sup_comm) hence "-(-(\ \ -\) \ -\) = -(-(\ \ (\ \ -\)) \ -\)" by (metis sup_assoc) also from \ have "\ = -(-(\ \ (\ \ -\)) \ -(\ \ -(\ \ -\)))" by metis also from robbins have "\ = \" by metis finally have "-(-(\ \ -\) \ -\) = \" by simp with two_three[where x="\" and y="\"] secret_object5_def sup_comm have "3 \ \ \ \ = 2 \ \ \ \" by fastforce with secret_object5_def sup_assoc sup_comm have "3 \ \ \ \ = 2 \ \ \ \" by fastforce with two three four five six have "6 \ \ = 3 \ \" by simp moreover have "3 + 3 = (6::nat)" and "3 \ (0::nat)" by arith+ moreover note copy_arith[where k="3" and l="3" and x="\"] winker_object_def three ultimately show ?thesis by simp qed (* Idempotence implies the identity law *) lemma sup_ident: "x \ \\ = x" proof - have I: "\ = -(-\ \ \\)" by (metis fake_bot_def inf_eq robbins sup_comm sup_idem) { fix x have "x = -(-(x \ -\ \ \\) \ -(x \ \))" by (metis I robbins sup_assoc) } note II = this have III: "-\ = -(-(\ \ -\ \ -\) \ \)" by (metis robbins[where x="-\" and y="\ \ -\"] I sup_comm fake_bot_def) hence "\ = -(-(\ \ -\ \ -\) \ -\)" by (metis robbins[where x="\" and y="\ \ -\ \ -\"] sup_comm[where x="\" and y="-(\ \ -\ \ -\)"] sup_assoc sup_idem) hence "-(\ \ -\ \ -\) = \\" by (metis robbins[where x="-(\ \ -\ \ -\)" and y="\"] III sup_comm fake_bot_def) hence "-\ = -(\ \ \\)" by (metis III sup_comm) hence "\ = -(-(\ \ \\) \ -(\ \ \\ \ -\))" by (metis II sup_idem sup_comm sup_assoc) moreover have "\ \ \\ = -(-(\ \ \\) \ -(\ \ \\ \ -\))" by (metis robbins[where x="\ \ \\" and y="\"] sup_comm[where y="\"] sup_assoc sup_idem) ultimately have "\ = \ \ \\" by auto hence "x \ \\ = -(-(x \ \) \ -(x \ \\ \ -\))" by (metis robbins[where x="x \ \\" and y=\] sup_comm[where x="\\" and y=\] sup_assoc) thus ?thesis by (metis sup_assoc sup_comm II) qed (* The identity law implies double negation *) lemma dbl_neg: "- (-x) = x" proof - { fix x have "\\ = -(-x \ -(-x))" by (metis robbins sup_comm sup_ident) } note I = this { fix x have "-x = -(-(-x \ -(-(-x))))" by (metis I robbins sup_comm sup_ident) } note II = this { fix x have "-(-(-x)) = -(-(-x \ -(-(-x))))" by (metis I II robbins sup_assoc sup_comm sup_ident) } note III = this show ?thesis by (metis II III robbins) qed (* Double negation implies Huntington's axiom, hence Boolean*) theorem robbins_is_huntington: "class.huntington_algebra uminus (\) (\) \ \" apply unfold_locales apply (metis dbl_neg robbins sup_comm) done theorem robbins_is_boolean_II: "class.boolean_algebra_II uminus (\) (\) \ \" proof - interpret huntington: huntington_algebra uminus "(\)" "(\)" \ \ by (fact robbins_is_huntington) show ?thesis by (simp add: huntington.huntington_is_boolean_II) qed theorem robbins_is_boolean: "class.boolean_algebra minus uminus (\) (\) (\) (\) \ \" proof - interpret huntington: huntington_algebra uminus "(\)" "(\)" \ \ by (fact robbins_is_huntington) show ?thesis by (simp add: huntington.huntington_is_boolean) qed end no_notation secret_object1 ("\") and secret_object2 ("\") and secret_object3 ("\") and secret_object4 ("\") and secret_object5 ("\") and winker_object ("\") and less_eq (infix "\" 50) and less (infix "\" 50) and inf (infixl "\" 70) and sup (infixl "\" 65) and top ("\") and bot ("\") and copyp (infix "\" 80) and copy (infix "\" 85) notation Product_Type.Times (infixr "\" 80) end diff --git a/thys/Stone_Algebras/Lattice_Basics.thy b/thys/Stone_Algebras/Lattice_Basics.thy --- a/thys/Stone_Algebras/Lattice_Basics.thy +++ b/thys/Stone_Algebras/Lattice_Basics.thy @@ -1,561 +1,561 @@ (* Title: Lattice Basics Author: Walter Guttmann Maintainer: Walter Guttmann *) section \Lattice Basics\ text \ This theory provides notations, basic definitions and facts of lattice-related structures used throughout the subsequent development. \ theory Lattice_Basics imports Main begin subsection \General Facts and Notations\ text \ The following results extend basic Isabelle/HOL facts. \ lemma imp_as_conj: assumes "P x \ Q x" shows "P x \ Q x \ P x" using assms by auto lemma if_distrib_2: "f (if c then x else y) (if c then z else w) = (if c then f x z else f y w)" by simp lemma left_invertible_inj: "(\x . g (f x) = x) \ inj f" by (metis injI) lemma invertible_bij: assumes "\x . g (f x) = x" and "\y . f (g y) = y" shows "bij f" by (metis assms bijI') lemma finite_ne_subset_induct [consumes 3, case_names singleton insert]: assumes "finite F" and "F \ {}" and "F \ S" and singleton: "\x . P {x}" and insert: "\x F . finite F \ F \ {} \ F \ S \ x \ S \ x \ F \ P F \ P (insert x F)" shows "P F" using assms(1-3) apply (induct rule: finite_ne_induct) apply (simp add: singleton) by (simp add: insert) lemma finite_set_of_finite_funs_pred: assumes "finite { x::'a . True }" and "finite { y::'b . P y }" shows "finite { f . (\x::'a . P (f x)) }" using assms finite_set_of_finite_funs by force text \ We use the following notations for the join, meet and complement operations. Changing the precedence of the unary complement allows us to write terms like \--x\ instead of \-(-x)\. \ context sup begin notation sup (infixl "\" 65) definition additive :: "('a \ 'a) \ bool" where "additive f \ \x y . f (x \ y) = f x \ f y" end context inf begin notation inf (infixl "\" 67) end context uminus begin no_notation uminus ("- _" [81] 80) notation uminus ("- _" [80] 80) end subsection \Orders\ text \ We use the following definition of monotonicity for operations defined in classes. The standard \mono\ places a sort constraint on the target type. We also give basic properties of Galois connections and lift orders to functions. \ context ord begin definition isotone :: "('a \ 'a) \ bool" where "isotone f \ \x y . x \ y \ f x \ f y" definition galois :: "('a \ 'a) \ ('a \ 'a) \ bool" where "galois l u \ \x y . l x \ y \ x \ u y" definition lifted_less_eq :: "('a \ 'a) \ ('a \ 'a) \ bool" ("(_ \\ _)" [51, 51] 50) where "f \\ g \ \x . f x \ g x" end context order begin lemma order_lesseq_imp: "(\z . x \ z \ y \ z) \ y \ x" using order_trans by blast lemma galois_char: "galois l u \ (\x . x \ u (l x)) \ (\x . l (u x) \ x) \ isotone l \ isotone u" apply (rule iffI) apply (metis (full_types) galois_def isotone_def order_refl order_trans) using galois_def isotone_def order_trans by blast lemma galois_closure: "galois l u \ l x = l (u (l x)) \ u x = u (l (u x))" by (simp add: galois_char isotone_def order.antisym) lemma lifted_reflexive: "f = g \ f \\ g" by (simp add: lifted_less_eq_def) lemma lifted_transitive: "f \\ g \ g \\ h \ f \\ h" using lifted_less_eq_def order_trans by blast lemma lifted_antisymmetric: "f \\ g \ g \\ f \ f = g" by (rule ext, rule order.antisym) (simp_all add: lifted_less_eq_def) text \ If the image of a finite non-empty set under \f\ is a totally ordered, there is an element that minimises the value of \f\. \ lemma finite_set_minimal: assumes "finite s" and "s \ {}" and "\x\s . \y\s . f x \ f y \ f y \ f x" shows "\m\s . \z\s . f m \ f z" apply (rule finite_ne_subset_induct[where S=s]) apply (rule assms(1)) apply (rule assms(2)) apply simp apply simp by (metis assms(3) insert_iff order_trans subsetD) end subsection \Semilattices\ text \ The following are basic facts in semilattices. \ context semilattice_sup begin lemma sup_left_isotone: "x \ y \ x \ z \ y \ z" using sup.mono by blast lemma sup_right_isotone: "x \ y \ z \ x \ z \ y" using sup.mono by blast lemma sup_left_divisibility: "x \ y \ (\z . x \ z = y)" using sup.absorb2 sup.cobounded1 by blast lemma sup_right_divisibility: "x \ y \ (\z . z \ x = y)" by (metis sup.cobounded2 sup.orderE) lemma sup_same_context: "x \ y \ z \ y \ x \ z \ x \ z = y \ z" by (simp add: le_iff_sup sup_left_commute) lemma sup_relative_same_increasing: "x \ y \ x \ z = x \ w \ y \ z = y \ w" using sup.assoc sup_right_divisibility by auto end text \ Every bounded semilattice is a commutative monoid. Finite sums defined in commutative monoids are available via the following sublocale. \ context bounded_semilattice_sup_bot begin sublocale sup_monoid: comm_monoid_add where plus = sup and zero = bot apply unfold_locales apply (simp add: sup_assoc) apply (simp add: sup_commute) by simp end context semilattice_inf begin lemma inf_same_context: "x \ y \ z \ y \ x \ z \ x \ z = y \ z" using order.antisym by auto end text \ The following class requires only the existence of upper bounds, which is a property common to bounded semilattices and (not necessarily bounded) lattices. We use it in our development of filters. \ class directed_semilattice_inf = semilattice_inf + assumes ub: "\z . x \ z \ y \ z" text \ We extend the \inf\ sublocale, which dualises the order in semilattices, to bounded semilattices. \ context bounded_semilattice_inf_top begin subclass directed_semilattice_inf apply unfold_locales using top_greatest by blast sublocale inf: bounded_semilattice_sup_bot where sup = inf and less_eq = greater_eq and less = greater and bot = top by unfold_locales (simp_all add: less_le_not_le) end subsection \Lattices\ context lattice begin subclass directed_semilattice_inf apply unfold_locales using sup_ge1 sup_ge2 by blast definition dual_additive :: "('a \ 'a) \ bool" where "dual_additive f \ \x y . f (x \ y) = f x \ f y" end text \ Not every bounded lattice has complements, but two elements might still be complements of each other as captured in the following definition. In this situation we can apply, for example, the shunting property shown below. We introduce most definitions using the \abbreviation\ command. \ context bounded_lattice begin abbreviation "complement x y \ x \ y = top \ x \ y = bot" lemma complement_symmetric: "complement x y \ complement y x" by (simp add: inf.commute sup.commute) definition conjugate :: "('a \ 'a) \ ('a \ 'a) \ bool" where "conjugate f g \ \x y . f x \ y = bot \ x \ g y = bot" end class dense_lattice = bounded_lattice + assumes bot_meet_irreducible: "x \ y = bot \ x = bot \ y = bot" context distrib_lattice begin lemma relative_equality: "x \ z = y \ z \ x \ z = y \ z \ x = y" by (metis inf.commute inf_sup_absorb inf_sup_distrib2) end text \ Distributive lattices with a greatest element are widely used in the construction theorem for Stone algebras. \ class distrib_lattice_bot = bounded_lattice_bot + distrib_lattice class distrib_lattice_top = bounded_lattice_top + distrib_lattice class bounded_distrib_lattice = bounded_lattice + distrib_lattice begin subclass distrib_lattice_bot .. subclass distrib_lattice_top .. lemma complement_shunting: assumes "complement z w" shows "z \ x \ y \ x \ w \ y" proof assume 1: "z \ x \ y" have "x = (z \ w) \ x" by (simp add: assms) also have "... \ y \ (w \ x)" using 1 sup.commute sup.left_commute inf_sup_distrib2 sup_right_divisibility by fastforce also have "... \ w \ y" by (simp add: inf.coboundedI1) finally show "x \ w \ y" . next assume "x \ w \ y" hence "z \ x \ z \ (w \ y)" using inf.sup_right_isotone by auto also have "... = z \ y" by (simp add: assms inf_sup_distrib1) also have "... \ y" by simp finally show "z \ x \ y" . qed end subsection \Linear Orders\ text \ We next consider lattices with a linear order structure. In such lattices, join and meet are selective operations, which give the maximum and the minimum of two elements, respectively. Moreover, the lattice is automatically distributive. \ class bounded_linorder = linorder + order_bot + order_top class linear_lattice = lattice + linorder begin lemma max_sup: "max x y = x \ y" by (metis max.boundedI max.cobounded1 max.cobounded2 sup_unique) lemma min_inf: "min x y = x \ y" by (simp add: inf.absorb1 inf.absorb2 min_def) lemma sup_inf_selective: "(x \ y = x \ x \ y = y) \ (x \ y = y \ x \ y = x)" by (meson inf.absorb1 inf.absorb2 le_cases sup.absorb1 sup.absorb2) lemma sup_selective: "x \ y = x \ x \ y = y" using sup_inf_selective by blast lemma inf_selective: "x \ y = x \ x \ y = y" using sup_inf_selective by blast subclass distrib_lattice apply standard apply (rule order.antisym) apply (auto simp add: le_supI2) apply (metis inf_selective inf.coboundedI1 inf.coboundedI2 order.eq_iff) done lemma sup_less_eq: "x \ y \ z \ x \ y \ x \ z" by (metis le_supI1 le_supI2 sup_selective) lemma inf_less_eq: "x \ y \ z \ x \ z \ y \ z" by (metis inf.coboundedI1 inf.coboundedI2 inf_selective) lemma sup_inf_sup: "x \ y = (x \ y) \ (x \ y)" by (metis sup_commute sup_inf_absorb sup_left_commute) end text \ The following class derives additional properties if the linear order of the lattice has a least and a greatest element. \ class linear_bounded_lattice = bounded_lattice + linorder begin subclass linear_lattice .. subclass bounded_linorder .. subclass bounded_distrib_lattice .. lemma sup_dense: "x \ top \ y \ top \ x \ y \ top" by (metis sup_selective) lemma inf_dense: "x \ bot \ y \ bot \ x \ y \ bot" by (metis inf_selective) lemma sup_not_bot: "x \ bot \ x \ y \ bot" by simp lemma inf_not_top: "x \ top \ x \ y \ top" by simp subclass dense_lattice apply unfold_locales using inf_dense by blast end text \ Every bounded linear order can be expanded to a bounded lattice. Join and meet are maximum and minimum, respectively. \ class linorder_lattice_expansion = bounded_linorder + sup + inf + assumes sup_def [simp]: "x \ y = max x y" assumes inf_def [simp]: "x \ y = min x y" begin subclass linear_bounded_lattice apply unfold_locales by auto end subsection \Non-trivial Algebras\ text \ Some results, such as the existence of certain filters, require that the algebras are not trivial. This is not an assumption of the order and lattice classes that come with Isabelle/HOL; for example, \bot = top\ may hold in bounded lattices. \ class non_trivial = assumes consistent: "\x y . x \ y" class non_trivial_order = non_trivial + order class non_trivial_order_bot = non_trivial_order + order_bot class non_trivial_bounded_order = non_trivial_order_bot + order_top begin lemma bot_not_top: "bot \ top" proof - from consistent obtain x y :: 'a where "x \ y" by auto thus ?thesis by (metis bot_less top.extremum_strict) qed end subsection \Homomorphisms\ text \ This section gives definitions of lattice homomorphisms and isomorphisms and basic properties. \ class sup_inf_top_bot_uminus = sup + inf + top + bot + uminus class sup_inf_top_bot_uminus_ord = sup_inf_top_bot_uminus + ord -context Lattices.boolean_algebra +context boolean_algebra begin subclass sup_inf_top_bot_uminus_ord . end abbreviation sup_homomorphism :: "('a::sup \ 'b::sup) \ bool" where "sup_homomorphism f \ \x y . f (x \ y) = f x \ f y" abbreviation inf_homomorphism :: "('a::inf \ 'b::inf) \ bool" where "inf_homomorphism f \ \x y . f (x \ y) = f x \ f y" abbreviation bot_homomorphism :: "('a::bot \ 'b::bot) \ bool" where "bot_homomorphism f \ f bot = bot" abbreviation top_homomorphism :: "('a::top \ 'b::top) \ bool" where "top_homomorphism f \ f top = top" abbreviation minus_homomorphism :: "('a::minus \ 'b::minus) \ bool" where "minus_homomorphism f \ \x y . f (x - y) = f x - f y" abbreviation uminus_homomorphism :: "('a::uminus \ 'b::uminus) \ bool" where "uminus_homomorphism f \ \x . f (-x) = -f x" abbreviation sup_inf_homomorphism :: "('a::{sup,inf} \ 'b::{sup,inf}) \ bool" where "sup_inf_homomorphism f \ sup_homomorphism f \ inf_homomorphism f" abbreviation sup_inf_top_homomorphism :: "('a::{sup,inf,top} \ 'b::{sup,inf,top}) \ bool" where "sup_inf_top_homomorphism f \ sup_inf_homomorphism f \ top_homomorphism f" abbreviation sup_inf_top_bot_homomorphism :: "('a::{sup,inf,top,bot} \ 'b::{sup,inf,top,bot}) \ bool" where "sup_inf_top_bot_homomorphism f \ sup_inf_top_homomorphism f \ bot_homomorphism f" abbreviation bounded_lattice_homomorphism :: "('a::bounded_lattice \ 'b::bounded_lattice) \ bool" where "bounded_lattice_homomorphism f \ sup_inf_top_bot_homomorphism f" abbreviation sup_inf_top_bot_uminus_homomorphism :: "('a::sup_inf_top_bot_uminus \ 'b::sup_inf_top_bot_uminus) \ bool" where "sup_inf_top_bot_uminus_homomorphism f \ sup_inf_top_bot_homomorphism f \ uminus_homomorphism f" abbreviation sup_inf_top_bot_uminus_ord_homomorphism :: "('a::sup_inf_top_bot_uminus_ord \ 'b::sup_inf_top_bot_uminus_ord) \ bool" where "sup_inf_top_bot_uminus_ord_homomorphism f \ sup_inf_top_bot_uminus_homomorphism f \ (\x y . x \ y \ f x \ f y)" abbreviation sup_inf_top_isomorphism :: "('a::{sup,inf,top} \ 'b::{sup,inf,top}) \ bool" where "sup_inf_top_isomorphism f \ sup_inf_top_homomorphism f \ bij f" abbreviation bounded_lattice_top_isomorphism :: "('a::bounded_lattice_top \ 'b::bounded_lattice_top) \ bool" where "bounded_lattice_top_isomorphism f \ sup_inf_top_isomorphism f" abbreviation sup_inf_top_bot_uminus_isomorphism :: "('a::sup_inf_top_bot_uminus \ 'b::sup_inf_top_bot_uminus) \ bool" where "sup_inf_top_bot_uminus_isomorphism f \ sup_inf_top_bot_uminus_homomorphism f \ bij f" abbreviation boolean_algebra_isomorphism :: "('a::boolean_algebra \ 'b::boolean_algebra) \ bool" where "boolean_algebra_isomorphism f \ sup_inf_top_bot_uminus_isomorphism f \ minus_homomorphism f" lemma sup_homomorphism_mono: "sup_homomorphism (f::'a::semilattice_sup \ 'b::semilattice_sup) \ mono f" by (metis le_iff_sup monoI) lemma sup_isomorphism_ord_isomorphism: assumes "sup_homomorphism (f::'a::semilattice_sup \ 'b::semilattice_sup)" and "bij f" shows "x \ y \ f x \ f y" proof assume "x \ y" thus "f x \ f y" by (metis assms(1) le_iff_sup) next assume "f x \ f y" hence "f (x \ y) = f y" by (simp add: assms(1) le_iff_sup) hence "x \ y = y" by (metis injD bij_is_inj assms(2)) thus "x \ y" by (simp add: le_iff_sup) qed lemma minus_homomorphism_default: assumes "\x y::'a::{inf,minus,uminus} . x - y = x \ -y" and "\x y::'b::{inf,minus,uminus} . x - y = x \ -y" and "inf_homomorphism (f::'a \ 'b)" and "uminus_homomorphism f" shows "minus_homomorphism f" by (simp add: assms) end diff --git a/thys/Stone_Algebras/P_Algebras.thy b/thys/Stone_Algebras/P_Algebras.thy --- a/thys/Stone_Algebras/P_Algebras.thy +++ b/thys/Stone_Algebras/P_Algebras.thy @@ -1,1328 +1,1328 @@ (* Title: Pseudocomplemented Algebras Author: Walter Guttmann Maintainer: Walter Guttmann *) section \Pseudocomplemented Algebras\ text \ This theory expands lattices with a pseudocomplement operation. In particular, we consider the following algebraic structures: \begin{itemize} \item pseudocomplemented lattices (p-algebras) \item pseudocomplemented distributive lattices (distributive p-algebras) \item Stone algebras \item Heyting semilattices \item Heyting lattices \item Heyting algebras \item Heyting-Stone algebras \item Brouwer algebras \item Boolean algebras \end{itemize} Most of these structures and many results in this theory are discussed in \cite{BalbesDwinger1974,Birkhoff1967,Blyth2005,Curry1977,Graetzer1971,Maddux1996}. \ theory P_Algebras imports Lattice_Basics begin subsection \P-Algebras\ text \ In this section we add a pseudocomplement operation to lattices and to distributive lattices. \ subsubsection \Pseudocomplemented Lattices\ text \ The pseudocomplement of an element \y\ is the greatest element whose meet with \y\ is the least element of the lattice. \ class p_algebra = bounded_lattice + uminus + assumes pseudo_complement: "x \ y = bot \ x \ -y" begin subclass sup_inf_top_bot_uminus_ord . text \ Regular elements and dense elements are frequently used in pseudocomplemented algebras. \ abbreviation "regular x \ x = --x" abbreviation "dense x \ -x = bot" abbreviation "complemented x \ \y . x \ y = bot \ x \ y = top" abbreviation "in_p_image x \ \y . x = -y" abbreviation "selection s x \ s = --s \ x" abbreviation "dense_elements \ { x . dense x }" abbreviation "regular_elements \ { x . in_p_image x }" lemma p_bot [simp]: "-bot = top" using inf_top.left_neutral pseudo_complement top_unique by blast lemma p_top [simp]: "-top = bot" by (metis eq_refl inf_top.comm_neutral pseudo_complement) text \ The pseudocomplement satisfies the following half of the requirements of a complement. \ lemma inf_p [simp]: "x \ -x = bot" using inf.commute pseudo_complement by fastforce lemma p_inf [simp]: "-x \ x = bot" by (simp add: inf_commute) lemma pp_inf_p: "--x \ -x = bot" by simp text \ The double complement is a closure operation. \ lemma pp_increasing: "x \ --x" using inf_p pseudo_complement by blast lemma ppp [simp]: "---x = -x" by (metis order.antisym inf.commute order_trans pseudo_complement pp_increasing) lemma pp_idempotent: "----x = --x" by simp lemma regular_in_p_image_iff: "regular x \ in_p_image x" by auto lemma pseudo_complement_pp: "x \ y = bot \ --x \ -y" by (metis inf_commute pseudo_complement ppp) lemma p_antitone: "x \ y \ -y \ -x" by (metis inf_commute order_trans pseudo_complement pp_increasing) lemma p_antitone_sup: "-(x \ y) \ -x" by (simp add: p_antitone) lemma p_antitone_inf: "-x \ -(x \ y)" by (simp add: p_antitone) lemma p_antitone_iff: "x \ -y \ y \ -x" using order_lesseq_imp p_antitone pp_increasing by blast lemma pp_isotone: "x \ y \ --x \ --y" by (simp add: p_antitone) lemma pp_isotone_sup: "--x \ --(x \ y)" by (simp add: p_antitone) lemma pp_isotone_inf: "--(x \ y) \ --x" by (simp add: p_antitone) text \ One of De Morgan's laws holds in pseudocomplemented lattices. \ lemma p_dist_sup [simp]: "-(x \ y) = -x \ -y" apply (rule order.antisym) apply (simp add: p_antitone) using inf_le1 inf_le2 le_sup_iff p_antitone_iff by blast lemma p_supdist_inf: "-x \ -y \ -(x \ y)" by (simp add: p_antitone) lemma pp_dist_pp_sup [simp]: "--(--x \ --y) = --(x \ y)" by simp lemma p_sup_p [simp]: "-(x \ -x) = bot" by simp lemma pp_sup_p [simp]: "--(x \ -x) = top" by simp lemma dense_pp: "dense x \ --x = top" by (metis p_bot p_top ppp) lemma dense_sup_p: "dense (x \ -x)" by simp lemma regular_char: "regular x \ (\y . x = -y)" by auto lemma pp_inf_bot_iff: "x \ y = bot \ --x \ y = bot" by (simp add: pseudo_complement_pp) text \ Weak forms of the shunting property hold. Most require a pseudocomplemented element on the right-hand side. \ lemma p_shunting_swap: "x \ y \ -z \ x \ z \ -y" by (metis inf_assoc inf_commute pseudo_complement) lemma pp_inf_below_iff: "x \ y \ -z \ --x \ y \ -z" by (simp add: inf_commute p_shunting_swap) lemma p_inf_pp [simp]: "-(x \ --y) = -(x \ y)" apply (rule order.antisym) apply (simp add: inf.coboundedI2 p_antitone pp_increasing) using inf_commute p_antitone_iff pp_inf_below_iff by auto lemma p_inf_pp_pp [simp]: "-(--x \ --y) = -(x \ y)" by (simp add: inf_commute) lemma regular_closed_inf: "regular x \ regular y \ regular (x \ y)" by (metis p_dist_sup ppp) lemma regular_closed_p: "regular (-x)" by simp lemma regular_closed_pp: "regular (--x)" by simp lemma regular_closed_bot: "regular bot" by simp lemma regular_closed_top: "regular top" by simp lemma pp_dist_inf [simp]: "--(x \ y) = --x \ --y" by (metis p_dist_sup p_inf_pp_pp ppp) lemma inf_import_p [simp]: "x \ -(x \ y) = x \ -y" apply (rule order.antisym) using p_shunting_swap apply fastforce using inf.sup_right_isotone p_antitone by auto text \ Pseudocomplements are unique. \ lemma p_unique: "(\x . x \ y = bot \ x \ z) \ z = -y" using inf.order_eq_iff pseudo_complement by auto lemma maddux_3_5: "x \ x = x \ -(y \ -y)" by simp lemma shunting_1_pp: "x \ --y \ x \ -y = bot" by (simp add: pseudo_complement) lemma pp_pp_inf_bot_iff: "x \ y = bot \ --x \ --y = bot" by (simp add: pseudo_complement_pp) lemma inf_pp_semi_commute: "x \ --y \ --(x \ y)" using inf.eq_refl p_antitone_iff p_inf_pp by presburger lemma inf_pp_commute: "--(--x \ y) = --x \ --y" by simp lemma sup_pp_semi_commute: "x \ --y \ --(x \ y)" by (simp add: p_antitone_iff) lemma regular_sup: "regular z \ (x \ z \ y \ z \ --(x \ y) \ z)" apply (rule iffI) apply (metis le_supI pp_isotone) using dual_order.trans sup_ge2 pp_increasing pp_isotone_sup by blast lemma dense_closed_inf: "dense x \ dense y \ dense (x \ y)" by (simp add: dense_pp) lemma dense_closed_sup: "dense x \ dense y \ dense (x \ y)" by simp lemma dense_closed_pp: "dense x \ dense (--x)" by simp lemma dense_closed_top: "dense top" by simp lemma dense_up_closed: "dense x \ x \ y \ dense y" using dense_pp top_le pp_isotone by auto lemma regular_dense_top: "regular x \ dense x \ x = top" using p_bot by blast lemma selection_char: "selection s x \ (\y . s = -y \ x)" by (metis inf_import_p inf_commute regular_closed_p) lemma selection_closed_inf: "selection s x \ selection t x \ selection (s \ t) x" by (metis inf_assoc inf_commute inf_idem pp_dist_inf) lemma selection_closed_pp: "regular x \ selection s x \ selection (--s) x" by (metis pp_dist_inf) lemma selection_closed_bot: "selection bot x" by simp lemma selection_closed_id: "selection x x" using inf.le_iff_sup pp_increasing by auto text \ Conjugates are usually studied for Boolean algebras, however, some of their properties generalise to pseudocomplemented algebras. \ lemma conjugate_unique_p: assumes "conjugate f g" and "conjugate f h" shows "uminus \ g = uminus \ h" proof - have "\x y . x \ g y = bot \ x \ h y = bot" using assms conjugate_def inf.commute by simp hence "\x y . x \ -(g y) \ x \ -(h y)" using inf.commute pseudo_complement by simp hence "\y . -(g y) = -(h y)" using order.eq_iff by blast thus ?thesis by auto qed lemma conjugate_symmetric: "conjugate f g \ conjugate g f" by (simp add: conjugate_def inf_commute) lemma additive_isotone: "additive f \ isotone f" by (metis additive_def isotone_def le_iff_sup) lemma dual_additive_antitone: assumes "dual_additive f" shows "isotone (uminus \ f)" proof - have "\x y . f (x \ y) \ f x" using assms dual_additive_def by simp hence "\x y . x \ y \ f y \ f x" by (metis sup_absorb2) hence "\x y . x \ y \ -(f x) \ -(f y)" by (simp add: p_antitone) thus ?thesis by (simp add: isotone_def) qed lemma conjugate_dual_additive: assumes "conjugate f g" shows "dual_additive (uminus \ f)" proof - have 1: "\x y z . -z \ -(f (x \ y)) \ -z \ -(f x) \ -z \ -(f y)" proof (intro allI) fix x y z have "(-z \ -(f (x \ y))) = (f (x \ y) \ -z = bot)" by (simp add: p_antitone_iff pseudo_complement) also have "... = ((x \ y) \ g(-z) = bot)" using assms conjugate_def by auto also have "... = (x \ y \ -(g(-z)))" by (simp add: pseudo_complement) also have "... = (x \ -(g(-z)) \ y \ -(g(-z)))" by (simp add: le_sup_iff) also have "... = (x \ g(-z) = bot \ y \ g(-z) = bot)" by (simp add: pseudo_complement) also have "... = (f x \ -z = bot \ f y \ -z = bot)" using assms conjugate_def by auto also have "... = (-z \ -(f x) \ -z \ -(f y))" by (simp add: p_antitone_iff pseudo_complement) finally show "-z \ -(f (x \ y)) \ -z \ -(f x) \ -z \ -(f y)" by simp qed have "\x y . -(f (x \ y)) = -(f x) \ -(f y)" proof (intro allI) fix x y have "-(f x) \ -(f y) = --(-(f x) \ -(f y))" by simp hence "-(f x) \ -(f y) \ -(f (x \ y))" using 1 by (metis inf_le1 inf_le2) thus "-(f (x \ y)) = -(f x) \ -(f y)" using 1 order.antisym by fastforce qed thus ?thesis using dual_additive_def by simp qed lemma conjugate_isotone_pp: "conjugate f g \ isotone (uminus \ uminus \ f)" by (simp add: comp_assoc conjugate_dual_additive dual_additive_antitone) lemma conjugate_char_1_pp: "conjugate f g \ (\x y . f(x \ -(g y)) \ --f x \ -y \ g(y \ -(f x)) \ --g y \ -x)" proof assume 1: "conjugate f g" show "\x y . f(x \ -(g y)) \ --f x \ -y \ g(y \ -(f x)) \ --g y \ -x" proof (intro allI) fix x y have 2: "f(x \ -(g y)) \ -y" using 1 by (simp add: conjugate_def pseudo_complement) have "f(x \ -(g y)) \ --f(x \ -(g y))" by (simp add: pp_increasing) also have "... \ --f x" using 1 conjugate_isotone_pp isotone_def by simp finally have 3: "f(x \ -(g y)) \ --f x \ -y" using 2 by simp have 4: "isotone (uminus \ uminus \ g)" using 1 conjugate_isotone_pp conjugate_symmetric by auto have 5: "g(y \ -(f x)) \ -x" using 1 by (metis conjugate_def inf.cobounded2 inf_commute pseudo_complement) have "g(y \ -(f x)) \ --g(y \ -(f x))" by (simp add: pp_increasing) also have "... \ --g y" using 4 isotone_def by auto finally have "g(y \ -(f x)) \ --g y \ -x" using 5 by simp thus "f(x \ -(g y)) \ --f x \ -y \ g(y \ -(f x)) \ --g y \ -x" using 3 by simp qed next assume 6: "\x y . f(x \ -(g y)) \ --f x \ -y \ g(y \ -(f x)) \ --g y \ -x" hence 7: "\x y . f x \ y = bot \ x \ g y = bot" by (metis inf.le_iff_sup inf.le_sup_iff inf_commute pseudo_complement) have "\x y . x \ g y = bot \ f x \ y = bot" using 6 by (metis inf.le_iff_sup inf.le_sup_iff inf_commute pseudo_complement) thus "conjugate f g" using 7 conjugate_def by auto qed lemma conjugate_char_1_isotone: "conjugate f g \ isotone f \ isotone g \ f(x \ -(g y)) \ f x \ -y \ g(y \ -(f x)) \ g y \ -x" by (simp add: conjugate_char_1_pp ord.isotone_def) lemma dense_lattice_char_1: "(\x y . x \ y = bot \ x = bot \ y = bot) \ (\x . x \ bot \ dense x)" by (metis inf_top.left_neutral p_bot p_inf pp_inf_bot_iff) lemma dense_lattice_char_2: "(\x y . x \ y = bot \ x = bot \ y = bot) \ (\x . regular x \ x = bot \ x = top)" by (metis dense_lattice_char_1 inf_top.left_neutral p_inf regular_closed_p regular_closed_top) lemma restrict_below_Rep_eq: "x \ --y \ z \ x \ y = x \ z \ y" by (metis inf.absorb2 inf.commute inf.left_commute pp_increasing) (* lemma p_inf_sup_below: "-x \ (x \ y) \ y" nitpick [expect=genuine] oops lemma complement_p: "x \ y = bot \ x \ y = top \ -x = y" nitpick [expect=genuine] oops lemma complemented_regular: "complemented x \ regular x" nitpick [expect=genuine] oops *) end text \ The following class gives equational axioms for the pseudocomplement operation. \ class p_algebra_eq = bounded_lattice + uminus + assumes p_bot_eq: "-bot = top" and p_top_eq: "-top = bot" and inf_import_p_eq: "x \ -(x \ y) = x \ -y" begin lemma inf_p_eq: "x \ -x = bot" by (metis inf_bot_right inf_import_p_eq inf_top_right p_top_eq) subclass p_algebra apply unfold_locales apply (rule iffI) apply (metis inf.orderI inf_import_p_eq inf_top.right_neutral p_bot_eq) by (metis (full_types) inf.left_commute inf.orderE inf_bot_right inf_commute inf_p_eq) end subsubsection \Pseudocomplemented Distributive Lattices\ text \ We obtain further properties if we assume that the lattice operations are distributive. \ class pd_algebra = p_algebra + bounded_distrib_lattice begin lemma p_inf_sup_below: "-x \ (x \ y) \ y" by (simp add: inf_sup_distrib1) lemma pp_inf_sup_p [simp]: "--x \ (x \ -x) = x" using inf.absorb2 inf_sup_distrib1 pp_increasing by auto lemma complement_p: "x \ y = bot \ x \ y = top \ -x = y" by (metis pseudo_complement inf.commute inf_top.left_neutral sup.absorb_iff1 sup.commute sup_bot.right_neutral sup_inf_distrib2 p_inf) lemma complemented_regular: "complemented x \ regular x" using complement_p inf.commute sup.commute by fastforce lemma regular_inf_dense: "\y z . regular y \ dense z \ x = y \ z" by (metis pp_inf_sup_p dense_sup_p ppp) lemma maddux_3_12 [simp]: "(x \ -y) \ (x \ y) = x" by (metis p_inf sup_bot_right sup_inf_distrib1) lemma maddux_3_13 [simp]: "(x \ y) \ -x = y \ -x" by (simp add: inf_sup_distrib2) lemma maddux_3_20: "((v \ w) \ (-v \ x)) \ -((v \ y) \ (-v \ z)) = (v \ w \ -y) \ (-v \ x \ -z)" proof - have "v \ w \ -(v \ y) \ -(-v \ z) = v \ w \ -(v \ y)" by (meson inf.cobounded1 inf_absorb1 le_infI1 p_antitone_iff) also have "... = v \ w \ -y" using inf.sup_relative_same_increasing inf_import_p inf_le1 by blast finally have 1: "v \ w \ -(v \ y) \ -(-v \ z) = v \ w \ -y" . have "-v \ x \ -(v \ y) \ -(-v \ z) = -v \ x \ -(-v \ z)" by (simp add: inf.absorb1 le_infI1 p_antitone_inf) also have "... = -v \ x \ -z" by (simp add: inf.assoc inf_left_commute) finally have 2: "-v \ x \ -(v \ y) \ -(-v \ z) = -v \ x \ -z" . have "((v \ w) \ (-v \ x)) \ -((v \ y) \ (-v \ z)) = (v \ w \ -(v \ y) \ -(-v \ z)) \ (-v \ x \ -(v \ y) \ -(-v \ z))" by (simp add: inf_assoc inf_sup_distrib2) also have "... = (v \ w \ -y) \ (-v \ x \ -z)" using 1 2 by simp finally show ?thesis . qed lemma order_char_1: "x \ y \ x \ y \ -x" by (metis inf.sup_left_isotone inf_sup_absorb le_supI1 maddux_3_12 sup_commute) lemma order_char_2: "x \ y \ x \ -x \ y \ -x" using order_char_1 by auto lemma half_shunting: "x \ y \ z \ x \ -z \ y" by (metis inf.sup_right_isotone inf_commute inf_sup_distrib1 sup.boundedE maddux_3_12) (* lemma pp_dist_sup [simp]: "--(x \ y) = --x \ --y" nitpick [expect=genuine] oops lemma regular_closed_sup: "regular x \ regular y \ regular (x \ y)" nitpick [expect=genuine] oops lemma regular_complemented_iff: "regular x \ complemented x" nitpick [expect=genuine] oops lemma selection_closed_sup: "selection s x \ selection t x \ selection (s \ t) x" nitpick [expect=genuine] oops lemma stone [simp]: "-x \ --x = top" nitpick [expect=genuine] oops *) end subsection \Stone Algebras\ text \ A Stone algebra is a distributive lattice with a pseudocomplement that satisfies the following equation. We thus obtain the other half of the requirements of a complement at least for the regular elements. \ class stone_algebra = pd_algebra + assumes stone [simp]: "-x \ --x = top" begin text \ As a consequence, we obtain both De Morgan's laws for all elements. \ lemma p_dist_inf [simp]: "-(x \ y) = -x \ -y" proof (rule p_unique[THEN sym], rule allI, rule iffI) fix w assume "w \ (x \ y) = bot" hence "w \ --x \ y = bot" using inf_commute inf_left_commute pseudo_complement by auto hence 1: "w \ --x \ -y" by (simp add: pseudo_complement) have "w = (w \ -x) \ (w \ --x)" using distrib_imp2 sup_inf_distrib1 by auto thus "w \ -x \ -y" using 1 by (metis inf_le2 sup.mono) next fix w assume "w \ -x \ -y" thus "w \ (x \ y) = bot" using order_trans p_supdist_inf pseudo_complement by blast qed lemma pp_dist_sup [simp]: "--(x \ y) = --x \ --y" by simp lemma regular_closed_sup: "regular x \ regular y \ regular (x \ y)" by simp text \ The regular elements are precisely the ones having a complement. \ lemma regular_complemented_iff: "regular x \ complemented x" by (metis inf_p stone complemented_regular) lemma selection_closed_sup: "selection s x \ selection t x \ selection (s \ t) x" by (simp add: inf_sup_distrib2) lemma huntington_3_pp [simp]: "-(-x \ -y) \ -(-x \ y) = --x" by (metis p_dist_inf p_inf sup.commute sup_bot_left sup_inf_distrib1) lemma maddux_3_3 [simp]: "-(x \ y) \ -(x \ -y) = -x" by (simp add: sup_commute sup_inf_distrib1) lemma maddux_3_11_pp: "(x \ -y) \ (x \ --y) = x" by (metis inf_sup_distrib1 inf_top_right stone) lemma maddux_3_19_pp: "(-x \ y) \ (--x \ z) = (--x \ y) \ (-x \ z)" proof - have "(--x \ y) \ (-x \ z) = (--x \ z) \ (y \ -x) \ (y \ z)" by (simp add: inf.commute inf_sup_distrib1 sup.assoc) also have "... = (--x \ z) \ (y \ -x) \ (y \ z \ (-x \ --x))" by simp also have "... = (--x \ z) \ ((y \ -x) \ (y \ -x \ z)) \ (y \ z \ --x)" using inf_sup_distrib1 sup_assoc inf_commute inf_assoc by presburger also have "... = (--x \ z) \ (y \ -x) \ (y \ z \ --x)" by simp also have "... = ((--x \ z) \ (--x \ z \ y)) \ (y \ -x)" by (simp add: inf_assoc inf_commute sup.left_commute sup_commute) also have "... = (--x \ z) \ (y \ -x)" by simp finally show ?thesis by (simp add: inf_commute sup_commute) qed lemma compl_inter_eq_pp: "--x \ y = --x \ z \ -x \ y = -x \ z \ y = z" by (metis inf_commute inf_p inf_sup_distrib1 inf_top.right_neutral p_bot p_dist_inf) lemma maddux_3_21_pp [simp]: "--x \ (-x \ y) = --x \ y" by (simp add: sup.commute sup_inf_distrib1) lemma shunting_2_pp: "x \ --y \ -x \ --y = top" by (metis inf_top_left p_bot p_dist_inf pseudo_complement) lemma shunting_p: "x \ y \ -z \ x \ -z \ -y" by (metis inf.assoc p_dist_inf p_shunting_swap pseudo_complement) text \ The following weak shunting property is interesting as it does not require the element \z\ on the right-hand side to be regular. \ lemma shunting_var_p: "x \ -y \ z \ x \ z \ --y" proof assume "x \ -y \ z" hence "z \ --y = --y \ (z \ x \ -y)" by (simp add: sup.absorb1 sup.commute) thus "x \ z \ --y" by (metis inf_commute maddux_3_21_pp sup.commute sup.left_commute sup_left_divisibility) next assume "x \ z \ --y" thus "x \ -y \ z" by (metis inf.mono maddux_3_12 sup_ge2) qed (* Whether conjugate_char_2_pp can be proved in pd_algebra or in p_algebra is unknown. *) lemma conjugate_char_2_pp: "conjugate f g \ f bot = bot \ g bot = bot \ (\x y . f x \ y \ --(f(x \ --(g y))) \ g y \ x \ --(g(y \ --(f x))))" proof assume 1: "conjugate f g" hence 2: "dual_additive (uminus \ g)" using conjugate_symmetric conjugate_dual_additive by auto show "f bot = bot \ g bot = bot \ (\x y . f x \ y \ --(f(x \ --(g y))) \ g y \ x \ --(g(y \ --(f x))))" proof (intro conjI) show "f bot = bot" using 1 by (metis conjugate_def inf_idem inf_bot_left) next show "g bot = bot" using 1 by (metis conjugate_def inf_idem inf_bot_right) next show "\x y . f x \ y \ --(f(x \ --(g y))) \ g y \ x \ --(g(y \ --(f x)))" proof (intro allI) fix x y have 3: "y \ -(f(x \ -(g y)))" using 1 by (simp add: conjugate_def pseudo_complement inf_commute) have 4: "x \ -(g(y \ -(f x)))" using 1 conjugate_def inf.commute pseudo_complement by fastforce have "y \ -(f(x \ --(g y))) = y \ -(f(x \ -(g y))) \ -(f(x \ --(g y)))" using 3 by (simp add: inf.le_iff_sup inf_commute) also have "... = y \ -(f((x \ -(g y)) \ (x \ --(g y))))" using 1 conjugate_dual_additive dual_additive_def inf_assoc by auto also have "... = y \ -(f x)" by (simp add: maddux_3_11_pp) also have "... \ -(f x)" by simp finally have 5: "f x \ y \ --(f(x \ --(g y)))" by (simp add: inf_commute p_shunting_swap) have "x \ -(g(y \ --(f x))) = x \ -(g(y \ -(f x))) \ -(g(y \ --(f x)))" using 4 by (simp add: inf.le_iff_sup inf_commute) also have "... = x \ -(g((y \ -(f x)) \ (y \ --(f x))))" using 2 by (simp add: dual_additive_def inf_assoc) also have "... = x \ -(g y)" by (simp add: maddux_3_11_pp) also have "... \ -(g y)" by simp finally have "g y \ x \ --(g(y \ --(f x)))" by (simp add: inf_commute p_shunting_swap) thus "f x \ y \ --(f(x \ --(g y))) \ g y \ x \ --(g(y \ --(f x)))" using 5 by simp qed qed next assume "f bot = bot \ g bot = bot \ (\x y . f x \ y \ --(f(x \ --(g y))) \ g y \ x \ --(g(y \ --(f x))))" thus "conjugate f g" by (unfold conjugate_def, metis inf_commute le_bot pp_inf_bot_iff regular_closed_bot) qed lemma conjugate_char_2_pp_additive: assumes "conjugate f g" and "additive f" and "additive g" shows "f x \ y \ f(x \ --(g y)) \ g y \ x \ g(y \ --(f x))" proof - have "f x \ y = f ((x \ --g y) \ (x \ -g y)) \ y" by (simp add: sup.commute sup_inf_distrib1) also have "... = (f (x \ --g y) \ y) \ (f (x \ -g y) \ y)" using assms(2) additive_def inf_sup_distrib2 by auto also have "... = f (x \ --g y) \ y" by (metis assms(1) conjugate_def inf_le2 pseudo_complement sup_bot.right_neutral) finally have 2: "f x \ y \ f (x \ --g y)" by simp have "g y \ x = g ((y \ --f x) \ (y \ -f x)) \ x" by (simp add: sup.commute sup_inf_distrib1) also have "... = (g (y \ --f x) \ x) \ (g (y \ -f x) \ x)" using assms(3) additive_def inf_sup_distrib2 by auto also have "... = g (y \ --f x) \ x" by (metis assms(1) conjugate_def inf.cobounded2 pseudo_complement sup_bot.right_neutral inf_commute) finally have "g y \ x \ g (y \ --f x)" by simp thus ?thesis using 2 by simp qed (* lemma compl_le_swap2_iff: "-x \ y \ -y \ x" nitpick [expect=genuine] oops lemma huntington_3: "x = -(-x \ -y) \ -(-x \ y)" nitpick [expect=genuine] oops lemma maddux_3_1: "x \ -x = y \ -y" nitpick [expect=genuine] oops lemma maddux_3_4: "x \ (y \ -y) = z \ -z" nitpick [expect=genuine] oops lemma maddux_3_11: "x = (x \ y) \ (x \ -y)" nitpick [expect=genuine] oops lemma maddux_3_19: "(-x \ y) \ (x \ z) = (x \ y) \ (-x \ z)" nitpick [expect=genuine] oops lemma compl_inter_eq: "x \ y = x \ z \ -x \ y = -x \ z \ y = z" nitpick [expect=genuine] oops lemma maddux_3_21: "x \ y = x \ (-x \ y)" nitpick [expect=genuine] oops lemma shunting_1: "x \ y \ x \ -y = bot" nitpick [expect=genuine] oops lemma shunting_2: "x \ y \ -x \ y = top" nitpick [expect=genuine] oops lemma conjugate_unique: "conjugate f g \ conjugate f h \ g = h" nitpick [expect=genuine] oops lemma conjugate_isotone_pp: "conjugate f g \ isotone f" nitpick [expect=genuine] oops lemma conjugate_char_1: "conjugate f g \ (\x y . f(x \ -(g y)) \ f x \ -y \ g(y \ -(f x)) \ g y \ -x)" nitpick [expect=genuine] oops lemma conjugate_char_2: "conjugate f g \ f bot = bot \ g bot = bot \ (\x y . f x \ y \ f(x \ g y) \ g y \ x \ g(y \ f x))" nitpick [expect=genuine] oops lemma shunting: "x \ y \ z \ x \ z \ -y" nitpick [expect=genuine] oops lemma shunting_var: "x \ -y \ z \ x \ z \ y" nitpick [expect=genuine] oops lemma sup_compl_top: "x \ -x = top" nitpick [expect=genuine] oops lemma selection_closed_p: "selection s x \ selection (-s) x" nitpick [expect=genuine] oops lemma selection_closed_pp: "selection s x \ selection (--s) x" nitpick [expect=genuine] oops *) end abbreviation stone_algebra_isomorphism :: "('a::stone_algebra \ 'b::stone_algebra) \ bool" where "stone_algebra_isomorphism f \ sup_inf_top_bot_uminus_isomorphism f" text \ Every bounded linear order can be expanded to a Stone algebra. The pseudocomplement takes \bot\ to the \top\ and every other element to \bot\. \ class linorder_stone_algebra_expansion = linorder_lattice_expansion + uminus + assumes uminus_def [simp]: "-x = (if x = bot then top else bot)" begin subclass stone_algebra apply unfold_locales using bot_unique min_def top_le by auto text \ The regular elements are the least and greatest elements. All elements except the least element are dense. \ lemma regular_bot_top: "regular x \ x = bot \ x = top" by simp lemma not_bot_dense: "x \ bot \ --x = top" by simp end subsection \Heyting Algebras\ text \ In this section we add a relative pseudocomplement operation to semilattices and to lattices. \ subsubsection \Heyting Semilattices\ text \ The pseudocomplement of an element \y\ relative to an element \z\ is the least element whose meet with \y\ is below \z\. This can be stated as a Galois connection. Specialising \z = bot\ gives (non-relative) pseudocomplements. Many properties can already be shown if the underlying structure is just a semilattice. \ class implies = fixes implies :: "'a \ 'a \ 'a" (infixl "\" 65) class heyting_semilattice = semilattice_inf + implies + assumes implies_galois: "x \ y \ z \ x \ y \ z" begin lemma implies_below_eq [simp]: "y \ (x \ y) = y" using implies_galois inf.absorb_iff1 inf.cobounded1 by blast lemma implies_increasing: "x \ y \ x" by (simp add: inf.orderI) lemma implies_galois_swap: "x \ y \ z \ y \ x \ z" by (metis implies_galois inf_commute) lemma implies_galois_var: "x \ y \ z \ y \ x \ z" by (simp add: implies_galois_swap implies_galois) lemma implies_galois_increasing: "x \ y \ (x \ y)" using implies_galois by blast lemma implies_galois_decreasing: "(y \ x) \ y \ x" using implies_galois by blast lemma implies_mp_below: "x \ (x \ y) \ y" using implies_galois_decreasing inf_commute by auto lemma implies_isotone: "x \ y \ z \ x \ z \ y" using implies_galois order_trans by blast lemma implies_antitone: "x \ y \ y \ z \ x \ z" by (meson implies_galois_swap order_lesseq_imp) lemma implies_isotone_inf: "x \ (y \ z) \ x \ y" by (simp add: implies_isotone) lemma implies_antitone_inf: "x \ z \ (x \ y) \ z" by (simp add: implies_antitone) lemma implies_curry: "x \ (y \ z) = (x \ y) \ z" by (metis implies_galois_decreasing implies_galois inf_assoc order.antisym) lemma implies_curry_flip: "x \ (y \ z) = y \ (x \ z)" by (simp add: implies_curry inf_commute) lemma triple_implies [simp]: "((x \ y) \ y) \ y = x \ y" using implies_antitone implies_galois_swap order.eq_iff by auto lemma implies_mp_eq [simp]: "x \ (x \ y) = x \ y" by (metis implies_below_eq implies_mp_below inf_left_commute inf.absorb2) lemma implies_dist_implies: "x \ (y \ z) \ (x \ y) \ (x \ z)" using implies_curry implies_curry_flip by auto lemma implies_import_inf [simp]: "x \ ((x \ y) \ (x \ z)) = x \ (y \ z)" by (metis implies_curry implies_mp_eq inf_commute) lemma implies_dist_inf: "x \ (y \ z) = (x \ y) \ (x \ z)" proof - have "(x \ y) \ (x \ z) \ x \ y \ z" by (simp add: implies_galois) hence "(x \ y) \ (x \ z) \ x \ (y \ z)" using implies_galois by blast thus ?thesis by (simp add: implies_isotone order.eq_iff) qed lemma implies_itself_top: "y \ x \ x" by (simp add: implies_galois_swap implies_increasing) lemma inf_implies_top: "z \ (x \ y) \ x" using implies_galois_var le_infI1 by blast lemma inf_inf_implies [simp]: "z \ ((x \ y) \ x) = z" by (simp add: inf_implies_top inf_absorb1) lemma le_implies_top: "x \ y \ z \ x \ y" using implies_antitone implies_itself_top order.trans by blast lemma le_iff_le_implies: "x \ y \ x \ x \ y" using implies_galois inf_idem by force lemma implies_inf_isotone: "x \ y \ (x \ z) \ (y \ z)" by (metis implies_curry implies_galois_increasing implies_isotone) lemma implies_transitive: "(x \ y) \ (y \ z) \ x \ z" using implies_dist_implies implies_galois_var implies_increasing order_lesseq_imp by blast lemma implies_inf_absorb [simp]: "x \ (x \ y) = x \ y" using implies_dist_inf implies_itself_top inf.absorb_iff2 by auto lemma implies_implies_absorb [simp]: "x \ (x \ y) = x \ y" by (simp add: implies_curry) lemma implies_inf_identity: "(x \ y) \ y = y" by (simp add: inf_commute) lemma implies_itself_same: "x \ x = y \ y" by (simp add: le_implies_top order.eq_iff) end text \ The following class gives equational axioms for the relative pseudocomplement operation (inequalities can be written as equations). \ class heyting_semilattice_eq = semilattice_inf + implies + assumes implies_mp_below: "x \ (x \ y) \ y" and implies_galois_increasing: "x \ y \ (x \ y)" and implies_isotone_inf: "x \ (y \ z) \ x \ y" begin subclass heyting_semilattice apply unfold_locales apply (rule iffI) apply (metis implies_galois_increasing implies_isotone_inf inf.absorb2 order_lesseq_imp) by (metis implies_mp_below inf_commute order_trans inf_mono order_refl) end text \ The following class allows us to explicitly give the pseudocomplement of an element relative to itself. \ class bounded_heyting_semilattice = bounded_semilattice_inf_top + heyting_semilattice begin lemma implies_itself [simp]: "x \ x = top" using implies_galois inf_le2 top_le by blast lemma implies_order: "x \ y \ x \ y = top" by (metis implies_galois inf_top.left_neutral top_unique) lemma inf_implies [simp]: "(x \ y) \ x = top" using implies_order inf_le1 by blast lemma top_implies [simp]: "top \ x = x" by (metis implies_mp_eq inf_top.left_neutral) end subsubsection \Heyting Lattices\ text \ We obtain further properties if the underlying structure is a lattice. In particular, the lattice operations are automatically distributive in this case. \ class heyting_lattice = lattice + heyting_semilattice begin lemma sup_distrib_inf_le: "(x \ y) \ (x \ z) \ x \ (y \ z)" proof - have "x \ z \ y \ (x \ (y \ z))" using implies_galois_var implies_increasing sup.bounded_iff sup.cobounded2 by blast hence "x \ y \ (x \ z) \ (x \ (y \ z))" using implies_galois_swap implies_increasing le_sup_iff by blast thus ?thesis by (simp add: implies_galois) qed subclass distrib_lattice apply unfold_locales using distrib_sup_le order.eq_iff sup_distrib_inf_le by auto lemma implies_isotone_sup: "x \ y \ x \ (y \ z)" by (simp add: implies_isotone) lemma implies_antitone_sup: "(x \ y) \ z \ x \ z" by (simp add: implies_antitone) lemma implies_sup: "x \ z \ (y \ z) \ ((x \ y) \ z)" proof - have "(x \ z) \ (y \ z) \ y \ z" by (simp add: implies_galois) hence "(x \ z) \ (y \ z) \ (x \ y) \ z" using implies_galois_swap implies_galois_var by fastforce thus ?thesis by (simp add: implies_galois) qed lemma implies_dist_sup: "(x \ y) \ z = (x \ z) \ (y \ z)" apply (rule order.antisym) apply (simp add: implies_antitone) by (simp add: implies_sup implies_galois) lemma implies_antitone_isotone: "(x \ y) \ (x \ y) \ x \ y" by (simp add: implies_antitone_sup implies_dist_inf le_infI2) lemma implies_antisymmetry: "(x \ y) \ (y \ x) = (x \ y) \ (x \ y)" by (metis implies_dist_sup implies_inf_absorb inf.commute) lemma sup_inf_implies [simp]: "(x \ y) \ (x \ y) = y" by (simp add: inf_sup_distrib2 sup.absorb2) lemma implies_subdist_sup: "(x \ y) \ (x \ z) \ x \ (y \ z)" by (simp add: implies_isotone) lemma implies_subdist_inf: "(x \ z) \ (y \ z) \ (x \ y) \ z" by (simp add: implies_antitone) lemma implies_sup_absorb: "(x \ y) \ z \ (x \ z) \ (y \ z)" by (metis implies_dist_sup implies_isotone_sup implies_increasing inf_inf_implies le_sup_iff sup_inf_implies) lemma sup_below_implies_implies: "x \ y \ (x \ y) \ y" by (simp add: implies_dist_sup implies_galois_swap implies_increasing) end class bounded_heyting_lattice = bounded_lattice + heyting_lattice begin subclass bounded_heyting_semilattice .. lemma implies_bot [simp]: "bot \ x = top" using implies_galois top_unique by fastforce end subsubsection \Heyting Algebras\ text \ The pseudocomplement operation can be defined in Heyting algebras, but it is typically not part of their signature. We add the definition as an axiom so that we can use the class hierarchy, for example, to inherit results from the class \pd_algebra\. \ class heyting_algebra = bounded_heyting_lattice + uminus + assumes uminus_eq: "-x = x \ bot" begin subclass pd_algebra apply unfold_locales using bot_unique implies_galois uminus_eq by auto lemma boolean_implies_below: "-x \ y \ x \ y" by (simp add: implies_increasing implies_isotone uminus_eq) lemma negation_implies: "-(x \ y) = --x \ -y" proof (rule order.antisym) show "-(x \ y) \ --x \ -y" using boolean_implies_below p_antitone by auto next have "x \ -y \ (x \ y) = bot" by (metis implies_mp_eq inf_p inf_bot_left inf_commute inf_left_commute) hence "--x \ -y \ (x \ y) = bot" using pp_inf_bot_iff inf_assoc by auto thus "--x \ -y \ -(x \ y)" by (simp add: pseudo_complement) qed lemma double_negation_dist_implies: "--(x \ y) = --x \ --y" apply (rule order.antisym) apply (metis pp_inf_below_iff implies_galois_decreasing implies_galois negation_implies ppp) by (simp add: p_antitone_iff negation_implies) (* lemma stone: "-x \ --x = top" nitpick [expect=genuine] oops *) end text \ The following class gives equational axioms for Heyting algebras. \ class heyting_algebra_eq = bounded_lattice + implies + uminus + assumes implies_mp_eq: "x \ (x \ y) = x \ y" and implies_import_inf: "x \ ((x \ y) \ (x \ z)) = x \ (y \ z)" and inf_inf_implies: "z \ ((x \ y) \ x) = z" and uminus_eq_eq: "-x = x \ bot" begin subclass heyting_algebra apply unfold_locales apply (rule iffI) apply (metis implies_import_inf inf.sup_left_divisibility inf_inf_implies le_iff_inf) apply (metis implies_mp_eq inf.commute inf.le_sup_iff inf.sup_right_isotone) by (simp add: uminus_eq_eq) end text \ A relative pseudocomplement is not enough to obtain the Stone equation, so we add it in the following class. \ class heyting_stone_algebra = heyting_algebra + assumes heyting_stone: "-x \ --x = top" begin subclass stone_algebra by unfold_locales (simp add: heyting_stone) (* lemma pre_linear: "(x \ y) \ (y \ x) = top" nitpick [expect=genuine] oops *) end subsubsection \Brouwer Algebras\ text \ Brouwer algebras are dual to Heyting algebras. The dual pseudocomplement of an element \y\ relative to an element \x\ is the least element whose join with \y\ is above \x\. We can now use the binary operation provided by Boolean algebras in Isabelle/HOL because it is compatible with dual relative pseudocomplements (not relative pseudocomplements). \ class brouwer_algebra = bounded_lattice + minus + uminus + assumes minus_galois: "x \ y \ z \ x - y \ z" and uminus_eq_minus: "-x = top - x" begin sublocale brouwer: heyting_algebra where inf = sup and less_eq = greater_eq and less = greater and sup = inf and bot = top and top = bot and implies = "\x y . y - x" apply unfold_locales apply simp apply simp apply simp apply simp apply (metis minus_galois sup_commute) by (simp add: uminus_eq_minus) lemma curry_minus: "x - (y \ z) = (x - y) - z" by (simp add: brouwer.implies_curry sup_commute) lemma minus_subdist_sup: "(x - z) \ (y - z) \ (x \ y) - z" by (simp add: brouwer.implies_dist_inf) lemma inf_sup_minus: "(x \ y) \ (x - y) = x" by (simp add: inf.absorb1 brouwer.inf_sup_distrib2) end subsection \Boolean Algebras\ text \ This section integrates Boolean algebras in the above hierarchy. In particular, we strengthen several results shown above. \ -context Lattices.boolean_algebra +context boolean_algebra begin text \ Every Boolean algebra is a Stone algebra, a Heyting algebra and a Brouwer algebra. \ subclass stone_algebra apply unfold_locales apply (rule iffI) apply (metis compl_sup_top inf.orderI inf_bot_right inf_sup_distrib1 inf_top_right sup_inf_absorb) using inf.commute inf.sup_right_divisibility apply fastforce by simp sublocale heyting: heyting_algebra where implies = "\x y . -x \ y" apply unfold_locales apply (rule iffI) using shunting_var_p sup_commute apply fastforce using shunting_var_p sup_commute apply force by simp subclass brouwer_algebra apply unfold_locales apply (simp add: diff_eq shunting_var_p sup.commute) by (simp add: diff_eq) lemma huntington_3 [simp]: "-(-x \ -y) \ -(-x \ y) = x" using huntington_3_pp by auto lemma maddux_3_1: "x \ -x = y \ -y" by simp lemma maddux_3_4: "x \ (y \ -y) = z \ -z" by simp lemma maddux_3_11 [simp]: "(x \ y) \ (x \ -y) = x" using brouwer.maddux_3_12 sup.commute by auto lemma maddux_3_19: "(-x \ y) \ (x \ z) = (x \ y) \ (-x \ z)" using maddux_3_19_pp by auto lemma compl_inter_eq: "x \ y = x \ z \ -x \ y = -x \ z \ y = z" by (metis inf_commute maddux_3_11) lemma maddux_3_21 [simp]: "x \ (-x \ y) = x \ y" by (simp add: sup_inf_distrib1) lemma shunting_1: "x \ y \ x \ -y = bot" by (simp add: pseudo_complement) lemma uminus_involutive: "uminus \ uminus = id" by auto lemma uminus_injective: "uminus \ f = uminus \ g \ f = g" by (metis comp_assoc id_o minus_comp_minus) lemma conjugate_unique: "conjugate f g \ conjugate f h \ g = h" using conjugate_unique_p uminus_injective by blast lemma dual_additive_additive: "dual_additive (uminus \ f) \ additive f" by (metis additive_def compl_eq_compl_iff dual_additive_def p_dist_sup o_def) lemma conjugate_additive: "conjugate f g \ additive f" by (simp add: conjugate_dual_additive dual_additive_additive) lemma conjugate_isotone: "conjugate f g \ isotone f" by (simp add: conjugate_additive additive_isotone) lemma conjugate_char_1: "conjugate f g \ (\x y . f(x \ -(g y)) \ f x \ -y \ g(y \ -(f x)) \ g y \ -x)" by (simp add: conjugate_char_1_pp) lemma conjugate_char_2: "conjugate f g \ f bot = bot \ g bot = bot \ (\x y . f x \ y \ f(x \ g y) \ g y \ x \ g(y \ f x))" by (simp add: conjugate_char_2_pp) lemma shunting: "x \ y \ z \ x \ z \ -y" by (simp add: heyting.implies_galois sup.commute) lemma shunting_var: "x \ -y \ z \ x \ z \ y" by (simp add: shunting) end class non_trivial_stone_algebra = non_trivial_bounded_order + stone_algebra class non_trivial_boolean_algebra = non_trivial_stone_algebra + boolean_algebra end diff --git a/thys/Stone_Relation_Algebras/Matrix_Relation_Algebras.thy b/thys/Stone_Relation_Algebras/Matrix_Relation_Algebras.thy --- a/thys/Stone_Relation_Algebras/Matrix_Relation_Algebras.thy +++ b/thys/Stone_Relation_Algebras/Matrix_Relation_Algebras.thy @@ -1,613 +1,613 @@ (* Title: Matrix Relation Algebras Author: Walter Guttmann Maintainer: Walter Guttmann *) section \Matrix Relation Algebras\ text \ This theory gives matrix models of Stone relation algebras and more general structures. We consider only square matrices. The main result is that matrices over Stone relation algebras form a Stone relation algebra. We use the monoid structure underlying semilattices to provide finite sums, which are necessary for defining the composition of two matrices. See \cite{ArmstrongFosterStruthWeber2016,ArmstrongGomesStruthWeber2016} for similar liftings to matrices for semirings and relation algebras. A technical difference is that those theories are mostly based on semirings whereas our hierarchy is mostly based on lattices (and our semirings directly inherit from semilattices). Relation algebras have both a semiring and a lattice structure such that semiring addition and lattice join coincide. In particular, finite sums and finite suprema coincide. Isabelle/HOL has separate theories for semirings and lattices, based on separate addition and join operations and different operations for finite sums and finite suprema. Reusing results from both theories is beneficial for relation algebras, but not always easy to realise. \ theory Matrix_Relation_Algebras imports Relation_Algebras begin subsection \Finite Suprema\ text \ We consider finite suprema in idempotent semirings and Stone relation algebras. We mostly use the first of the following notations, which denotes the supremum of expressions \t(x)\ over all \x\ from the type of \x\. For finite types, this is implemented in Isabelle/HOL as the repeated application of binary suprema. \ syntax "_sum_sup_monoid" :: "idt \ 'a::bounded_semilattice_sup_bot \ 'a" ("(\\<^sub>_ _)" [0,10] 10) "_sum_sup_monoid_bounded" :: "idt \ 'b set \ 'a::bounded_semilattice_sup_bot \ 'a" ("(\\<^bsub>_\_\<^esub> _)" [0,51,10] 10) translations "\\<^sub>x t" => "XCONST sup_monoid.sum (\x . t) { x . CONST True }" "\\<^bsub>x\X\<^esub> t" => "XCONST sup_monoid.sum (\x . t) X" context idempotent_semiring begin text \ The following induction principles are useful for comparing two suprema. The first principle works because types are not empty. \ lemma one_sup_induct [case_names one sup]: fixes f g :: "'b::finite \ 'a" assumes one: "\i . P (f i) (g i)" and sup: "\j I . j \ I \ P (\\<^bsub>i\I\<^esub> f i) (\\<^bsub>i\I\<^esub> g i) \ P (f j \ (\\<^bsub>i\I\<^esub> f i)) (g j \ (\\<^bsub>i\I\<^esub> g i))" shows "P (\\<^sub>k f k) (\\<^sub>k g k)" proof - let ?X = "{ k::'b . True }" have "finite ?X" and "?X \ {}" by auto thus ?thesis proof (induct rule: finite_ne_induct) case (singleton i) thus ?case using one by simp next case (insert j I) thus ?case using sup by simp qed qed lemma bot_sup_induct [case_names bot sup]: fixes f g :: "'b::finite \ 'a" assumes bot: "P bot bot" and sup: "\j I . j \ I \ P (\\<^bsub>i\I\<^esub> f i) (\\<^bsub>i\I\<^esub> g i) \ P (f j \ (\\<^bsub>i\I\<^esub> f i)) (g j \ (\\<^bsub>i\I\<^esub> g i))" shows "P (\\<^sub>k f k) (\\<^sub>k g k)" apply (induct rule: one_sup_induct) using bot sup apply fastforce using sup by blast text \ Now many properties of finite suprema follow by simple applications of the above induction rules. In particular, we show distributivity of composition, isotonicity and the upper-bound property. \ lemma comp_right_dist_sum: fixes f :: "'b::finite \ 'a" shows "(\\<^sub>k f k * x) = (\\<^sub>k f k) * x" proof (induct rule: one_sup_induct) case one show ?case by simp next case (sup j I) thus ?case using mult_right_dist_sup by auto qed lemma comp_left_dist_sum: fixes f :: "'b::finite \ 'a" shows "(\\<^sub>k x * f k) = x * (\\<^sub>k f k)" proof (induct rule: one_sup_induct) case one show ?case by simp next case (sup j I) thus ?case by (simp add: mult_left_dist_sup) qed lemma leq_sum: fixes f g :: "'b::finite \ 'a" shows "(\k . f k \ g k) \ (\\<^sub>k f k) \ (\\<^sub>k g k)" proof (induct rule: one_sup_induct) case one thus ?case by simp next case (sup j I) thus ?case using sup_mono by blast qed lemma ub_sum: fixes f :: "'b::finite \ 'a" shows "f i \ (\\<^sub>k f k)" proof - have "i \ { k . True }" by simp thus "f i \ (\\<^sub>k f (k::'b))" by (metis finite_code sup_monoid.sum.insert sup_ge1 mk_disjoint_insert) qed lemma lub_sum: fixes f :: "'b::finite \ 'a" assumes "\k . f k \ x" shows "(\\<^sub>k f k) \ x" proof (induct rule: one_sup_induct) case one show ?case by (simp add: assms) next case (sup j I) thus ?case using assms le_supI by blast qed lemma lub_sum_iff: fixes f :: "'b::finite \ 'a" shows "(\k . f k \ x) \ (\\<^sub>k f k) \ x" using order.trans ub_sum lub_sum by blast end context stone_relation_algebra begin text \ In Stone relation algebras, we can also show that converse, double complement and meet distribute over finite suprema. \ lemma conv_dist_sum: fixes f :: "'b::finite \ 'a" shows "(\\<^sub>k (f k)\<^sup>T) = (\\<^sub>k f k)\<^sup>T" proof (induct rule: one_sup_induct) case one show ?case by simp next case (sup j I) thus ?case by (simp add: conv_dist_sup) qed lemma pp_dist_sum: fixes f :: "'b::finite \ 'a" shows "(\\<^sub>k --f k) = --(\\<^sub>k f k)" proof (induct rule: one_sup_induct) case one show ?case by simp next case (sup j I) thus ?case by simp qed lemma inf_right_dist_sum: fixes f :: "'b::finite \ 'a" shows "(\\<^sub>k f k \ x) = (\\<^sub>k f k) \ x" by (rule comp_inf.comp_right_dist_sum) end subsection \Square Matrices\ text \ Because our semiring and relation algebra type classes only work for homogeneous relations, we only look at square matrices. \ type_synonym ('a,'b) square = "'a \ 'a \ 'b" text \ We use standard matrix operations. The Stone algebra structure is lifted componentwise. Composition is matrix multiplication using given composition and supremum operations. Its unit lifts given zero and one elements into an identity matrix. Converse is matrix transpose with an additional componentwise transpose. \ definition less_eq_matrix :: "('a,'b::ord) square \ ('a,'b) square \ bool" (infix "\" 50) where "f \ g = (\e . f e \ g e)" definition less_matrix :: "('a,'b::ord) square \ ('a,'b) square \ bool" (infix "\" 50) where "f \ g = (f \ g \ \ g \ f)" definition sup_matrix :: "('a,'b::sup) square \ ('a,'b) square \ ('a,'b) square" (infixl "\" 65) where "f \ g = (\e . f e \ g e)" definition inf_matrix :: "('a,'b::inf) square \ ('a,'b) square \ ('a,'b) square" (infixl "\" 67) where "f \ g = (\e . f e \ g e)" definition minus_matrix :: "('a,'b::{uminus,inf}) square \ ('a,'b) square \ ('a,'b) square" (infixl "\" 65) where "f \ g = (\e . f e \ -g e)" definition implies_matrix :: "('a,'b::implies) square \ ('a,'b) square \ ('a,'b) square" (infixl "\" 65) where "f \ g = (\e . f e \ g e)" definition times_matrix :: "('a,'b::{times,bounded_semilattice_sup_bot}) square \ ('a,'b) square \ ('a,'b) square" (infixl "\" 70) where "f \ g = (\(i,j) . \\<^sub>k f (i,k) * g (k,j))" definition uminus_matrix :: "('a,'b::uminus) square \ ('a,'b) square" ("\ _" [80] 80) where "\f = (\e . -f e)" definition conv_matrix :: "('a,'b::conv) square \ ('a,'b) square" ("_\<^sup>t" [100] 100) where "f\<^sup>t = (\(i,j) . (f (j,i))\<^sup>T)" definition bot_matrix :: "('a,'b::bot) square" ("mbot") where "mbot = (\e . bot)" definition top_matrix :: "('a,'b::top) square" ("mtop") where "mtop = (\e . top)" definition one_matrix :: "('a,'b::{one,bot}) square" ("mone") where "mone = (\(i,j) . if i = j then 1 else bot)" subsection \Stone Algebras\ text \ We first lift the Stone algebra structure. Because all operations are componentwise, this also works for infinite matrices. \ interpretation matrix_order: order where less_eq = less_eq_matrix and less = "less_matrix :: ('a,'b::order) square \ ('a,'b) square \ bool" apply unfold_locales apply (simp add: less_matrix_def) apply (simp add: less_eq_matrix_def) apply (meson less_eq_matrix_def order_trans) by (meson less_eq_matrix_def antisym ext) interpretation matrix_semilattice_sup: semilattice_sup where sup = sup_matrix and less_eq = less_eq_matrix and less = "less_matrix :: ('a,'b::semilattice_sup) square \ ('a,'b) square \ bool" apply unfold_locales apply (simp add: sup_matrix_def less_eq_matrix_def) apply (simp add: sup_matrix_def less_eq_matrix_def) by (simp add: sup_matrix_def less_eq_matrix_def) interpretation matrix_semilattice_inf: semilattice_inf where inf = inf_matrix and less_eq = less_eq_matrix and less = "less_matrix :: ('a,'b::semilattice_inf) square \ ('a,'b) square \ bool" apply unfold_locales apply (simp add: inf_matrix_def less_eq_matrix_def) apply (simp add: inf_matrix_def less_eq_matrix_def) by (simp add: inf_matrix_def less_eq_matrix_def) interpretation matrix_bounded_semilattice_sup_bot: bounded_semilattice_sup_bot where sup = sup_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix :: ('a,'b::bounded_semilattice_sup_bot) square" apply unfold_locales by (simp add: bot_matrix_def less_eq_matrix_def) interpretation matrix_bounded_semilattice_inf_top: bounded_semilattice_inf_top where inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and top = "top_matrix :: ('a,'b::bounded_semilattice_inf_top) square" apply unfold_locales by (simp add: less_eq_matrix_def top_matrix_def) interpretation matrix_lattice: lattice where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = "less_matrix :: ('a,'b::lattice) square \ ('a,'b) square \ bool" .. interpretation matrix_distrib_lattice: distrib_lattice where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = "less_matrix :: ('a,'b::distrib_lattice) square \ ('a,'b) square \ bool" apply unfold_locales by (simp add: sup_inf_distrib1 sup_matrix_def inf_matrix_def) interpretation matrix_bounded_lattice: bounded_lattice where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix :: ('a,'b::bounded_lattice) square" and top = top_matrix .. interpretation matrix_bounded_distrib_lattice: bounded_distrib_lattice where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix :: ('a,'b::bounded_distrib_lattice) square" and top = top_matrix .. interpretation matrix_p_algebra: p_algebra where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix :: ('a,'b::p_algebra) square" and top = top_matrix and uminus = uminus_matrix apply unfold_locales apply (unfold inf_matrix_def bot_matrix_def less_eq_matrix_def uminus_matrix_def) by (meson pseudo_complement) interpretation matrix_pd_algebra: pd_algebra where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix :: ('a,'b::pd_algebra) square" and top = top_matrix and uminus = uminus_matrix .. text \ In particular, matrices over Stone algebras form a Stone algebra. \ interpretation matrix_stone_algebra: stone_algebra where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix :: ('a,'b::stone_algebra) square" and top = top_matrix and uminus = uminus_matrix by unfold_locales (simp add: sup_matrix_def uminus_matrix_def top_matrix_def) interpretation matrix_heyting_stone_algebra: heyting_stone_algebra where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix :: ('a,'b::heyting_stone_algebra) square" and top = top_matrix and uminus = uminus_matrix and implies = implies_matrix apply unfold_locales apply (unfold inf_matrix_def sup_matrix_def bot_matrix_def top_matrix_def less_eq_matrix_def uminus_matrix_def implies_matrix_def) apply (simp add: implies_galois) apply (simp add: uminus_eq) by simp -interpretation matrix_boolean_algebra: Lattices.boolean_algebra where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix :: ('a,'b::boolean_algebra) square" and top = top_matrix and uminus = uminus_matrix and minus = minus_matrix +interpretation matrix_boolean_algebra: boolean_algebra where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix :: ('a,'b::boolean_algebra) square" and top = top_matrix and uminus = uminus_matrix and minus = minus_matrix apply unfold_locales apply simp apply (simp add: sup_matrix_def uminus_matrix_def top_matrix_def) by (simp add: inf_matrix_def uminus_matrix_def minus_matrix_def) subsection \Semirings\ text \ Next, we lift the semiring structure. Because of composition, this requires a restriction to finite matrices. \ interpretation matrix_monoid: monoid_mult where times = times_matrix and one = "one_matrix :: ('a::finite,'b::idempotent_semiring) square" proof fix f g h :: "('a,'b) square" show "(f \ g) \ h = f \ (g \ h)" proof (rule ext, rule prod_cases) fix i j have "((f \ g) \ h) (i,j) = (\\<^sub>l (f \ g) (i,l) * h (l,j))" by (simp add: times_matrix_def) also have "... = (\\<^sub>l (\\<^sub>k f (i,k) * g (k,l)) * h (l,j))" by (simp add: times_matrix_def) also have "... = (\\<^sub>l \\<^sub>k (f (i,k) * g (k,l)) * h (l,j))" by (metis (no_types) comp_right_dist_sum) also have "... = (\\<^sub>l \\<^sub>k f (i,k) * (g (k,l) * h (l,j)))" by (simp add: mult.assoc) also have "... = (\\<^sub>k \\<^sub>l f (i,k) * (g (k,l) * h (l,j)))" using sup_monoid.sum.swap by auto also have "... = (\\<^sub>k f (i,k) * (\\<^sub>l g (k,l) * h (l,j)))" by (metis (no_types) comp_left_dist_sum) also have "... = (\\<^sub>k f (i,k) * (g \ h) (k,j))" by (simp add: times_matrix_def) also have "... = (f \ (g \ h)) (i,j)" by (simp add: times_matrix_def) finally show "((f \ g) \ h) (i,j) = (f \ (g \ h)) (i,j)" . qed next fix f :: "('a,'b) square" show "mone \ f = f" proof (rule ext, rule prod_cases) fix i j have "(mone \ f) (i,j) = (\\<^sub>k mone (i,k) * f (k,j))" by (simp add: times_matrix_def) also have "... = (\\<^sub>k (if i = k then 1 else bot) * f (k,j))" by (simp add: one_matrix_def) also have "... = (\\<^sub>k if i = k then 1 * f (k,j) else bot * f (k,j))" by (metis (full_types, opaque_lifting)) also have "... = (\\<^sub>k if i = k then f (k,j) else bot)" by (meson mult_left_one mult_left_zero) also have "... = f (i,j)" by simp finally show "(mone \ f) (i,j) = f (i,j)" . qed next fix f :: "('a,'b) square" show "f \ mone = f" proof (rule ext, rule prod_cases) fix i j have "(f \ mone) (i,j) = (\\<^sub>k f (i,k) * mone (k,j))" by (simp add: times_matrix_def) also have "... = (\\<^sub>k f (i,k) * (if k = j then 1 else bot))" by (simp add: one_matrix_def) also have "... = (\\<^sub>k if k = j then f (i,k) * 1 else f (i,k) * bot)" by (metis (full_types, opaque_lifting)) also have "... = (\\<^sub>k if k = j then f (i,k) else bot)" by (meson mult.right_neutral semiring.mult_zero_right) also have "... = f (i,j)" by simp finally show "(f \ mone) (i,j) = f (i,j)" . qed qed interpretation matrix_idempotent_semiring: idempotent_semiring where sup = sup_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix :: ('a::finite,'b::idempotent_semiring) square" and one = one_matrix and times = times_matrix proof fix f g h :: "('a,'b) square" show "f \ g \ f \ h \ f \ (g \ h)" proof (unfold less_eq_matrix_def, rule allI, rule prod_cases) fix i j have "(f \ g \ f \ h) (i,j) = (f \ g) (i,j) \ (f \ h) (i,j)" by (simp add: sup_matrix_def) also have "... = (\\<^sub>k f (i,k) * g (k,j)) \ (\\<^sub>k f (i,k) * h (k,j))" by (simp add: times_matrix_def) also have "... = (\\<^sub>k f (i,k) * g (k,j) \ f (i,k) * h (k,j))" by (simp add: sup_monoid.sum.distrib) also have "... = (\\<^sub>k f (i,k) * (g (k,j) \ h (k,j)))" by (simp add: mult_left_dist_sup) also have "... = (\\<^sub>k f (i,k) * (g \ h) (k,j))" by (simp add: sup_matrix_def) also have "... = (f \ (g \ h)) (i,j)" by (simp add: times_matrix_def) finally show "(f \ g \ f \ h) (i,j) \ (f \ (g \ h)) (i,j)" by simp qed next fix f g h :: "('a,'b) square" show "(f \ g) \ h = f \ h \ g \ h" proof (rule ext, rule prod_cases) fix i j have "((f \ g) \ h) (i,j) = (\\<^sub>k (f \ g) (i,k) * h (k,j))" by (simp add: times_matrix_def) also have "... = (\\<^sub>k (f (i,k) \ g (i,k)) * h (k,j))" by (simp add: sup_matrix_def) also have "... = (\\<^sub>k f (i,k) * h (k,j) \ g (i,k) * h (k,j))" by (meson mult_right_dist_sup) also have "... = (\\<^sub>k f (i,k) * h (k,j)) \ (\\<^sub>k g (i,k) * h (k,j))" by (simp add: sup_monoid.sum.distrib) also have "... = (f \ h) (i,j) \ (g \ h) (i,j)" by (simp add: times_matrix_def) also have "... = (f \ h \ g \ h) (i,j)" by (simp add: sup_matrix_def) finally show "((f \ g) \ h) (i,j) = (f \ h \ g \ h) (i,j)" . qed next fix f :: "('a,'b) square" show "mbot \ f = mbot" proof (rule ext, rule prod_cases) fix i j have "(mbot \ f) (i,j) = (\\<^sub>k mbot (i,k) * f (k,j))" by (simp add: times_matrix_def) also have "... = (\\<^sub>k bot * f (k,j))" by (simp add: bot_matrix_def) also have "... = bot" by simp also have "... = mbot (i,j)" by (simp add: bot_matrix_def) finally show "(mbot \ f) (i,j) = mbot (i,j)" . qed next fix f :: "('a,'b) square" show "mone \ f = f" by simp next fix f :: "('a,'b) square" show "f \ f \ mone" by simp next fix f g h :: "('a,'b) square" show "f \ (g \ h) = f \ g \ f \ h" proof (rule ext, rule prod_cases) fix i j have "(f \ (g \ h)) (i,j) = (\\<^sub>k f (i,k) * (g \ h) (k,j))" by (simp add: times_matrix_def) also have "... = (\\<^sub>k f (i,k) * (g (k,j) \ h (k,j)))" by (simp add: sup_matrix_def) also have "... = (\\<^sub>k f (i,k) * g (k,j) \ f (i,k) * h (k,j))" by (meson mult_left_dist_sup) also have "... = (\\<^sub>k f (i,k) * g (k,j)) \ (\\<^sub>k f (i,k) * h (k,j))" by (simp add: sup_monoid.sum.distrib) also have "... = (f \ g) (i,j) \ (f \ h) (i,j)" by (simp add: times_matrix_def) also have "... = (f \ g \ f \ h) (i,j)" by (simp add: sup_matrix_def) finally show "(f \ (g \ h)) (i,j) = (f \ g \ f \ h) (i,j)" . qed next fix f :: "('a,'b) square" show "f \ mbot = mbot" proof (rule ext, rule prod_cases) fix i j have "(f \ mbot) (i,j) = (\\<^sub>k f (i,k) * mbot (k,j))" by (simp add: times_matrix_def) also have "... = (\\<^sub>k f (i,k) * bot)" by (simp add: bot_matrix_def) also have "... = bot" by simp also have "... = mbot (i,j)" by (simp add: bot_matrix_def) finally show "(f \ mbot) (i,j) = mbot (i,j)" . qed qed interpretation matrix_bounded_idempotent_semiring: bounded_idempotent_semiring where sup = sup_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix :: ('a::finite,'b::bounded_idempotent_semiring) square" and top = top_matrix and one = one_matrix and times = times_matrix proof fix f :: "('a,'b) square" show "f \ mtop = mtop" proof fix e have "(f \ mtop) e = f e \ mtop e" by (simp add: sup_matrix_def) also have "... = f e \ top" by (simp add: top_matrix_def) also have "... = top" by simp also have "... = mtop e" by (simp add: top_matrix_def) finally show "(f \ mtop) e = mtop e" . qed qed subsection \Stone Relation Algebras\ text \ Finally, we show that matrices over Stone relation algebras form a Stone relation algebra. \ interpretation matrix_stone_relation_algebra: stone_relation_algebra where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix :: ('a::finite,'b::stone_relation_algebra) square" and top = top_matrix and uminus = uminus_matrix and one = one_matrix and times = times_matrix and conv = conv_matrix proof fix f g h :: "('a,'b) square" show "(f \ g) \ h = f \ (g \ h)" by (simp add: matrix_monoid.mult_assoc) next fix f g h :: "('a,'b) square" show "(f \ g) \ h = f \ h \ g \ h" by (simp add: matrix_idempotent_semiring.mult_right_dist_sup) next fix f :: "('a,'b) square" show "mbot \ f = mbot" by simp next fix f :: "('a,'b) square" show "mone \ f = f" by simp next fix f :: "('a,'b) square" show "f\<^sup>t\<^sup>t = f" proof (rule ext, rule prod_cases) fix i j have "(f\<^sup>t\<^sup>t) (i,j) = ((f\<^sup>t) (j,i))\<^sup>T" by (simp add: conv_matrix_def) also have "... = f (i,j)" by (simp add: conv_matrix_def) finally show "(f\<^sup>t\<^sup>t) (i,j) = f (i,j)" . qed next fix f g :: "('a,'b) square" show "(f \ g)\<^sup>t = f\<^sup>t \ g\<^sup>t" proof (rule ext, rule prod_cases) fix i j have "((f \ g)\<^sup>t) (i,j) = ((f \ g) (j,i))\<^sup>T" by (simp add: conv_matrix_def) also have "... = (f (j,i) \ g (j,i))\<^sup>T" by (simp add: sup_matrix_def) also have "... = (f\<^sup>t) (i,j) \ (g\<^sup>t) (i,j)" by (simp add: conv_matrix_def conv_dist_sup) also have "... = (f\<^sup>t \ g\<^sup>t) (i,j)" by (simp add: sup_matrix_def) finally show "((f \ g)\<^sup>t) (i,j) = (f\<^sup>t \ g\<^sup>t) (i,j)" . qed next fix f g :: "('a,'b) square" show "(f \ g)\<^sup>t = g\<^sup>t \ f\<^sup>t" proof (rule ext, rule prod_cases) fix i j have "((f \ g)\<^sup>t) (i,j) = ((f \ g) (j,i))\<^sup>T" by (simp add: conv_matrix_def) also have "... = (\\<^sub>k f (j,k) * g (k,i))\<^sup>T" by (simp add: times_matrix_def) also have "... = (\\<^sub>k (f (j,k) * g (k,i))\<^sup>T)" by (metis (no_types) conv_dist_sum) also have "... = (\\<^sub>k (g (k,i))\<^sup>T * (f (j,k))\<^sup>T)" by (simp add: conv_dist_comp) also have "... = (\\<^sub>k (g\<^sup>t) (i,k) * (f\<^sup>t) (k,j))" by (simp add: conv_matrix_def) also have "... = (g\<^sup>t \ f\<^sup>t) (i,j)" by (simp add: times_matrix_def) finally show "((f \ g)\<^sup>t) (i,j) = (g\<^sup>t \ f\<^sup>t) (i,j)" . qed next fix f g h :: "('a,'b) square" show "(f \ g) \ h \ f \ (g \ (f\<^sup>t \ h))" proof (unfold less_eq_matrix_def, rule allI, rule prod_cases) fix i j have "((f \ g) \ h) (i,j) = (f \ g) (i,j) \ h (i,j)" by (simp add: inf_matrix_def) also have "... = (\\<^sub>k f (i,k) * g (k,j)) \ h (i,j)" by (simp add: times_matrix_def) also have "... = (\\<^sub>k f (i,k) * g (k,j) \ h (i,j))" by (metis (no_types) inf_right_dist_sum) also have "... \ (\\<^sub>k f (i,k) * (g (k,j) \ (f (i,k))\<^sup>T * h (i,j)))" by (rule leq_sum, meson dedekind_1) also have "... = (\\<^sub>k f (i,k) * (g (k,j) \ (f\<^sup>t) (k,i) * h (i,j)))" by (simp add: conv_matrix_def) also have "... \ (\\<^sub>k f (i,k) * (g (k,j) \ (\\<^sub>l (f\<^sup>t) (k,l) * h (l,j))))" by (rule leq_sum, rule allI, rule comp_right_isotone, rule inf.sup_right_isotone, rule ub_sum) also have "... = (\\<^sub>k f (i,k) * (g (k,j) \ (f\<^sup>t \ h) (k,j)))" by (simp add: times_matrix_def) also have "... = (\\<^sub>k f (i,k) * (g \ (f\<^sup>t \ h)) (k,j))" by (simp add: inf_matrix_def) also have "... = (f \ (g \ (f\<^sup>t \ h))) (i,j)" by (simp add: times_matrix_def) finally show "((f \ g) \ h) (i,j) \ (f \ (g \ (f\<^sup>t \ h))) (i,j)" . qed next fix f g :: "('a,'b) square" show "\\(f \ g) = \\f \ \\g" proof (rule ext, rule prod_cases) fix i j have "(\\(f \ g)) (i,j) = --((f \ g) (i,j))" by (simp add: uminus_matrix_def) also have "... = --(\\<^sub>k f (i,k) * g (k,j))" by (simp add: times_matrix_def) also have "... = (\\<^sub>k --(f (i,k) * g (k,j)))" by (metis (no_types) pp_dist_sum) also have "... = (\\<^sub>k --(f (i,k)) * --(g (k,j)))" by (meson pp_dist_comp) also have "... = (\\<^sub>k (\\f) (i,k) * (\\g) (k,j))" by (simp add: uminus_matrix_def) also have "... = (\\f \ \\g) (i,j)" by (simp add: times_matrix_def) finally show "(\\(f \ g)) (i,j) = (\\f \ \\g) (i,j)" . qed next let ?o = "mone :: ('a,'b) square" show "\\?o = ?o" proof (rule ext, rule prod_cases) fix i j have "(\\?o) (i,j) = --(?o (i,j))" by (simp add: uminus_matrix_def) also have "... = --(if i = j then 1 else bot)" by (simp add: one_matrix_def) also have "... = (if i = j then --1 else --bot)" by simp also have "... = (if i = j then 1 else bot)" by auto also have "... = ?o (i,j)" by (simp add: one_matrix_def) finally show "(\\?o) (i,j) = ?o (i,j)" . qed qed end diff --git a/thys/Subset_Boolean_Algebras/Subset_Boolean_Algebras.thy b/thys/Subset_Boolean_Algebras/Subset_Boolean_Algebras.thy --- a/thys/Subset_Boolean_Algebras/Subset_Boolean_Algebras.thy +++ b/thys/Subset_Boolean_Algebras/Subset_Boolean_Algebras.thy @@ -1,3608 +1,3608 @@ (* Title: Subset Boolean Algebras Authors: Walter Guttmann, Bernhard Möller Maintainer: Walter Guttmann *) theory Subset_Boolean_Algebras imports Stone_Algebras.P_Algebras begin section \Boolean Algebras\ text \ We show that Isabelle/HOL's \boolean_algebra\ class is equivalent to Huntington's axioms \cite{Huntington1933}. See \cite{WamplerDoty2016} for related results. \ subsection \Huntington's Axioms\ text \Definition 1\ class huntington = sup + uminus + assumes associative: "x \ (y \ z) = (x \ y) \ z" assumes commutative: "x \ y = y \ x" assumes huntington: "x = -(-x \ y) \ -(-x \ -y)" begin lemma top_unique: "x \ -x = y \ -y" proof - have "x \ -x = y \ -(--y \ -x) \ -(--y \ --x)" by (smt associative commutative huntington) thus ?thesis by (metis associative huntington) qed end subsection \Equivalence to \boolean_algebra\ Class\ text \Definition 2\ class extended = sup + inf + minus + uminus + bot + top + ord + assumes top_def: "top = (THE x . \y . x = y \ -y)" (* define without imposing uniqueness *) assumes bot_def: "bot = -(THE x . \y . x = y \ -y)" assumes inf_def: "x \ y = -(-x \ -y)" assumes minus_def: "x - y = -(-x \ y)" assumes less_eq_def: "x \ y \ x \ y = y" assumes less_def: "x < y \ x \ y = y \ \ (y \ x = x)" class huntington_extended = huntington + extended begin lemma top_char: "top = x \ -x" using top_def top_unique by auto lemma bot_char: "bot = -top" by (simp add: bot_def top_def) subclass boolean_algebra proof show 1: "\x y. (x < y) = (x \ y \ \ y \ x)" by (simp add: less_def less_eq_def) show 2: "\x. x \ x" proof - fix x have "x \ top = top \ --x" by (metis (full_types) associative top_char) thus "x \ x" by (metis (no_types) associative huntington less_eq_def top_char) qed show 3: "\x y z. x \ y \ y \ z \ x \ z" by (metis associative less_eq_def) show 4: "\x y. x \ y \ y \ x \ x = y" by (simp add: commutative less_eq_def) show 5: "\x y. x \ y \ x" using 2 by (metis associative huntington inf_def less_eq_def) show 6: "\x y. x \ y \ y" using 5 commutative inf_def by fastforce show 8: "\x y. x \ x \ y" using 2 associative less_eq_def by auto show 9: "\y x. y \ x \ y" using 8 commutative by fastforce show 10: "\y x z. y \ x \ z \ x \ y \ z \ x" by (metis associative less_eq_def) show 11: "\x. bot \ x" using 8 by (metis bot_char huntington top_char) show 12: "\x. x \ top" using 6 11 by (metis huntington bot_def inf_def less_eq_def top_def) show 13: "\x y z. x \ y \ z = (x \ y) \ (x \ z)" proof - have 2: "\x y z . x \ (y \ z) = (x \ y) \ z" by (simp add: associative) have 3: "\x y z . (x \ y) \ z = x \ (y \ z)" using 2 by metis have 4: "\x y . x \ y = y \ x" by (simp add: commutative) have 5: "\x y . x = - (- x \ y) \ - (- x \ - y)" by (simp add: huntington) have 6: "\x y . - (- x \ y) \ - (- x \ - y) = x" using 5 by metis have 7: "\x y . x \ y = - (- x \ - y)" by (simp add: inf_def) have 10: "\x y z . x \ (y \ z) = y \ (x \ z)" using 3 4 by metis have 11: "\x y z . - (- x \ y) \ (- (- x \ - y) \ z) = x \ z" using 3 6 by metis have 12: "\x y . - (x \ - y) \ - (- y \ - x) = y" using 4 6 by metis have 13: "\x y . - (- x \ y) \ - (- y \ - x) = x" using 4 6 by metis have 14: "\x y . - x \ - (- (- x \ y) \ - - (- x \ - y)) = - x \ y" using 6 by metis have 18: "\x y z . - (x \ - y) \ (- (- y \ - x) \ z) = y \ z" using 3 12 by metis have 20: "\x y . - (- x \ - y) \ - (y \ - x) = x" using 4 12 by metis have 21: "\x y . - (x \ - y) \ - (- x \ - y) = y" using 4 12 by metis have 22: "\x y . - x \ - (- (y \ - x) \ - - (- x \ - y)) = y \ - x" using 6 12 by metis have 23: "\x y . - x \ - (- x \ (- y \ - (y \ - x))) = y \ - x" using 3 4 6 12 by metis have 24: "\x y . - x \ - (- (- x \ - y) \ - - (- x \ y)) = - x \ - y" using 6 12 by metis have 28: "\x y . - (- x \ - y) \ - (- y \ x) = y" using 4 13 by metis have 30: "\x y . - x \ - (- y \ (- x \ - (- x \ y))) = - x \ y" using 3 4 6 13 by metis have 32: "\x y z . - (- x \ y) \ (z \ - (- y \ - x)) = z \ x" using 10 13 by metis have 37: "\x y z . - (- x \ - y) \ (- (y \ - x) \ z) = x \ z" using 3 20 by metis have 39: "\x y z . - (- x \ - y) \ (z \ - (y \ - x)) = z \ x" using 10 20 by metis have 40: "\x y z . - (x \ - y) \ (- (- x \ - y) \ z) = y \ z" using 3 21 by metis have 43: "\x y . - x \ - (- y \ (- x \ - (y \ - x))) = y \ - x" using 3 4 6 21 by metis have 47: "\x y z . - (x \ y) \ - (- (- x \ z) \ - (- (- x \ - z) \ y)) = - x \ z" using 6 11 by metis have 55: "\x y . x \ - (- y \ - - x) = y \ - (- x \ y)" using 4 11 12 by metis have 58: "\x y . x \ - (- - y \ - x) = x \ - (- x \ y)" using 4 11 13 by metis have 63: "\x y . x \ - (- - x \ - y) = y \ - (- x \ y)" using 4 11 21 by metis have 71: "\x y . x \ - (- y \ x) = y \ - (- x \ y)" using 4 11 28 by metis have 75: "\x y . x \ - (- y \ x) = y \ - (y \ - x)" using 4 71 by metis have 78: "\x y . - x \ (y \ - (- x \ (y \ - - (- x \ - y)))) = - x \ - (- x \ - y)" using 3 4 6 71 by metis have 86: "\x y . - (- x \ - (- y \ x)) \ - (y \ - (- x \ y)) = - y \ x" using 4 20 71 by metis have 172: "\x y . - x \ - (- x \ - y) = y \ - (- - x \ y)" using 14 75 by metis have 201: "\x y . x \ - (- y \ - - x) = y \ - (y \ - x)" using 4 55 by metis have 236: "\x y . x \ - (- - y \ - x) = x \ - (y \ - x)" using 4 58 by metis have 266: "\x y . - x \ - (- (- x \ - (y \ - - x)) \ - - (- x \ - - (- - x \ y))) = - x \ - (- - x \ y)" using 14 58 236 by metis have 678: "\x y z . - (- x \ - (- y \ x)) \ (- (y \ - (- x \ y)) \ z) = - y \ (x \ z)" using 3 4 37 71 by smt have 745: "\x y z . - (- x \ - (- y \ x)) \ (z \ - (y \ - (- x \ y))) = z \ (- y \ x)" using 4 39 71 by metis have 800: "\x y . - - x \ (- y \ (- (y \ - - x) \ - (- x \ (- - x \ (- y \ - (y \ - - x)))))) = x \ - (y \ - - x)" using 3 23 63 by metis have 944: "\x y . x \ - (x \ - - (- (- x \ - y) \ - - (- x \ y))) = - (- x \ - y) \ - (- (- x \ - y) \ - - (- x \ y))" using 4 24 71 by metis have 948: "\x y . - x \ - (- (y \ - (y \ - - x)) \ - - (- x \ (- y \ - x))) = - x \ - (- y \ - x)" using 24 75 by metis have 950: "\x y . - x \ - (- (y \ - (- - x \ y)) \ - - (- x \ (- x \ - y))) = - x \ - (- x \ - y)" using 24 75 by metis have 961: "\x y . - x \ - (- (y \ - (- - x \ y)) \ - - (- x \ (- - - x \ - y))) = y \ - (- - x \ y)" using 24 63 by metis have 966: "\x y . - x \ - (- (y \ - (y \ - - x)) \ - - (- x \ (- y \ - - - x))) = y \ - (y \ - - x)" using 24 201 by metis have 969: "\x y . - x \ - (- (- x \ - (y \ - - x)) \ - - (- x \ (- - y \ - - x))) = - x \ - (y \ - - x)" using 24 236 by metis have 1096: "\x y z . - x \ (- (- x \ - y) \ z) = y \ (- (- - x \ y) \ z)" using 3 172 by metis have 1098: "\x y z . - x \ (y \ - (- x \ - z)) = y \ (z \ - (- - x \ z))" using 10 172 by metis have 1105: "\x y . x \ - x = y \ - y" using 4 10 12 32 172 by metis have 1109: "\x y z . x \ (- x \ y) = z \ (- z \ y)" using 3 1105 by metis have 1110: "\x y z . x \ - x = y \ (z \ - (y \ z))" using 3 1105 by metis have 1114: "\x y . - (- x \ - - x) = - (y \ - y)" using 7 1105 by metis have 1115: "\x y z . x \ (y \ - y) = z \ (x \ - z)" using 10 1105 by metis have 1117: "\x y . - (x \ - - x) \ - (y \ - y) = - x" using 4 13 1105 by metis have 1121: "\x y . - (x \ - x) \ - (y \ - - y) = - y" using 4 28 1105 by metis have 1122: "\x . - - x = x" using 4 28 1105 1117 by metis have 1134: "\x y z . - (x \ - y) \ (z \ - z) = y \ (- y \ - x)" using 18 1105 1122 by metis have 1140: "\x . - x \ - (x \ (x \ - x)) = - x \ - x" using 4 22 1105 1122 1134 by metis have 1143: "\x y . x \ (- x \ y) = y \ (x \ - y)" using 37 1105 1122 1134 by metis have 1155: "\x y . - (x \ - x) \ - (y \ y) = - y" using 1121 1122 by metis have 1156: "\x y . - (x \ x) \ - (y \ - y) = - x" using 1117 1122 by metis have 1157: "\x y . - (x \ - x) = - (y \ - y)" using 4 1114 1122 by metis have 1167: "\x y z . - x \ (y \ - (- x \ - z)) = y \ (z \ - (x \ z))" using 1098 1122 by metis have 1169: "\x y z . - x \ (- (- x \ - y) \ z) = y \ (- (x \ y) \ z)" using 1096 1122 by metis have 1227: "\x y . - x \ - (- x \ (y \ (x \ - (- x \ - (y \ x))))) = - x \ - (y \ x)" using 3 4 969 1122 by smt have 1230: "\x y . - x \ - (- x \ (- y \ (- x \ - (y \ - (y \ x))))) = y \ - (y \ x)" using 3 4 966 1122 by smt have 1234: "\x y . - x \ - (- x \ (- x \ (- y \ - (y \ - (x \ y))))) = y \ - (x \ y)" using 3 4 961 1122 by metis have 1239: "\x y . - x \ - (- x \ - y) = y \ - (x \ y)" using 3 4 950 1122 1234 by metis have 1240: "\x y . - x \ - (- y \ - x) = y \ - (y \ x)" using 3 4 948 1122 1230 by metis have 1244: "\x y . x \ - (x \ (y \ (y \ - (x \ y)))) = - (- x \ - y) \ - (y \ (y \ - (x \ y)))" using 3 4 944 1122 1167 by metis have 1275: "\x y . x \ (- y \ (- (y \ x) \ - (x \ (- x \ (- y \ - (y \ x)))))) = x \ - (y \ x)" using 10 800 1122 by metis have 1346: "\x y . - x \ - (x \ (y \ (y \ (x \ - (x \ (y \ x)))))) = - x \ - (x \ y)" using 3 4 10 266 1122 1167 by smt have 1377: "\x y . - x \ (y \ - (- x \ (y \ (- x \ - y)))) = y \ - (x \ y)" using 78 1122 1239 by metis have 1394: "\x y . - (- x \ - y) \ - (y \ (y \ (- x \ - (x \ y)))) = x" using 3 4 10 20 30 1122 1239 by smt have 1427: "\x y . - (- x \ - y) \ - (y \ - (x \ (x \ - (x \ y)))) = x \ (x \ - (x \ y))" using 3 4 30 40 1240 by smt have 1436: "\x . - x \ - (x \ (x \ (- x \ - x))) = - x \ (- x \ - (x \ - x))" using 3 4 30 1140 1239 by smt have 1437: "\x y . - (x \ y) \ - (x \ - y) = - x" using 6 1122 by metis have 1438: "\x y . - (x \ y) \ - (y \ - x) = - y" using 12 1122 by metis have 1439: "\x y . - (x \ y) \ - (- y \ x) = - x" using 13 1122 by metis have 1440: "\x y . - (x \ - y) \ - (y \ x) = - x" using 20 1122 by metis have 1441: "\x y . - (x \ y) \ - (- x \ y) = - y" using 21 1122 by metis have 1568: "\x y . x \ (- y \ - x) = y \ (- y \ x)" using 10 1122 1143 by metis have 1598: "\x . - x \ - (x \ (x \ (x \ - x))) = - x \ (- x \ - (x \ - x))" using 4 1436 1568 by metis have 1599: "\x y . - x \ (y \ - (x \ (- x \ (- x \ y)))) = y \ - (x \ y)" using 10 1377 1568 by smt have 1617: "\x . x \ (- x \ (- x \ - (x \ - x))) = x \ - x" using 3 4 10 71 1122 1155 1568 1598 by metis have 1632: "\x y z . - (x \ - x) \ - (- y \ (- (z \ - z) \ - (y \ - (x \ - x)))) = y \ - (x \ - x)" using 43 1157 by metis have 1633: "\x y z . - (x \ - x) \ - (- y \ (- (x \ - x) \ - (y \ - (z \ - z)))) = y \ - (x \ - x)" using 43 1157 by metis have 1636: "\x y . x \ - (y \ (- y \ - (x \ x))) = x \ x" using 43 1109 1122 by metis have 1645: "\x y . x \ - x = y \ (y \ - y)" using 3 1110 1156 by metis have 1648: "\x y z . - (x \ (y \ (- y \ - x))) \ - (z \ - z) = - (y \ - y)" using 3 1115 1156 by metis have 1657: "\x y z . x \ - x = y \ (z \ - z)" using 1105 1645 by metis have 1664: "\x y z . x \ - x = y \ (z \ - y)" using 1115 1645 by metis have 1672: "\x y z . x \ - x = y \ (- y \ z)" using 3 4 1657 by metis have 1697: "\x y z . - x \ (y \ x) = z \ - z" using 1122 1664 by metis have 1733: "\x y z . - (x \ y) \ - (- (z \ - z) \ - (- (- x \ - x) \ y)) = x \ - x" using 4 47 1105 1122 by metis have 1791: "\x y z . x \ - (y \ (- y \ z)) = x \ - (x \ - x)" using 4 71 1122 1672 by metis have 1818: "\x y z . x \ - (- y \ (z \ y)) = x \ - (x \ - x)" using 4 71 1122 1697 by metis have 1861: "\x y z . - (x \ - x) \ - (y \ - (z \ - z)) = - y" using 1437 1657 by metis have 1867: "\x y z . - (x \ - x) \ - (- y \ - (z \ y)) = y" using 1122 1437 1697 by metis have 1868: "\x y . x \ - (y \ - y) = x" using 1122 1155 1633 1861 by metis have 1869: "\x y z . - (x \ - x) \ - (- y \ (- (z \ - z) \ - y)) = y" using 1632 1868 by metis have 1870: "\x y . - (x \ - x) \ - y = - y" using 1861 1868 by metis have 1872: "\x y z . x \ - (- y \ (z \ y)) = x" using 1818 1868 by metis have 1875: "\x y z . x \ - (y \ (- y \ z)) = x" using 1791 1868 by metis have 1883: "\x y . - (x \ (y \ (- y \ - x))) = - (y \ - y)" using 1648 1868 by metis have 1885: "\x . x \ (x \ - x) = x \ - x" using 4 1568 1617 1868 by metis have 1886: "\x . - x \ - x = - x" using 1598 1868 1885 by metis have 1890: "\x . - (x \ x) = - x" using 1156 1868 by metis have 1892: "\x y . - (x \ - x) \ y = y" using 1122 1869 1870 1886 by metis have 1893: "\x y . - (- x \ - (y \ x)) = x" using 1867 1892 by metis have 1902: "\x y . x \ (y \ - (x \ y)) = x \ - x" using 3 4 1122 1733 1886 1892 by metis have 1908: "\x . x \ x = x" using 1636 1875 1890 by metis have 1910: "\x y . x \ - (y \ x) = - y \ x" using 1599 1875 by metis have 1921: "\x y . x \ (- y \ - (y \ x)) = - y \ x" using 1275 1875 1910 by metis have 1951: "\x y . - x \ - (y \ x) = - x" using 1227 1872 1893 1908 by metis have 1954: "\x y z . x \ (y \ - (x \ z)) = y \ (- z \ x)" using 745 1122 1910 1951 by metis have 1956: "\x y z . x \ (- (x \ y) \ z) = - y \ (x \ z)" using 678 1122 1910 1951 by metis have 1959: "\x y . x \ - (x \ y) = - y \ x" using 86 1122 1910 1951 by metis have 1972: "\x y . x \ (- x \ y) = x \ - x" using 1902 1910 by metis have 2000: "\x y . - (- x \ - y) \ - (y \ (- x \ y)) = x \ - (y \ (- x \ y))" using 4 1244 1910 1959 by metis have 2054: "\x y . x \ - (y \ (- x \ y)) = x" using 1394 1921 2000 by metis have 2057: "\x y . - (x \ (y \ - y)) = - (y \ - y)" using 1883 1972 by metis have 2061: "\x y . x \ (- y \ x) = x \ - y" using 4 1122 1427 1910 1959 2054 by metis have 2090: "\x y z . x \ (- (y \ x) \ z) = x \ (- y \ z)" using 1122 1169 1956 by metis have 2100: "\x y . - x \ - (x \ y) = - x" using 4 1346 1868 1885 1910 1959 1972 2057 by metis have 2144: "\x y . x \ - (y \ - x) = x" using 1122 1440 2000 2061 by metis have 2199: "\x y . x \ (x \ y) = x \ y" using 3 1908 by metis have 2208: "\x y z . x \ (- (y \ - x) \ z) = x \ z" using 3 2144 by metis have 2349: "\x y z . - (x \ y) \ - (x \ (y \ z)) = - (x \ y)" using 3 2100 by metis have 2432: "\x y z . - (x \ (y \ z)) \ - (y \ (z \ - x)) = - (y \ z)" using 3 1438 by metis have 2530: "\x y z . - (- (x \ y) \ z) = - (y \ (- x \ z)) \ - (- y \ z)" using 4 1122 1439 2090 2208 by smt have 3364: "\x y z . - (- x \ y) \ (z \ - (x \ y)) = z \ - y" using 3 4 1122 1441 1910 1954 2199 by metis have 5763: "\x y z . - (x \ y) \ - (- x \ (y \ z)) = - (x \ y) \ - (y \ z)" using 4 2349 3364 by metis have 6113: "\x y z . - (x \ (y \ z)) \ - (z \ - x) = - (y \ z) \ - (z \ - x)" using 4 2432 3364 5763 by metis show "\x y z. x \ y \ z = (x \ y) \ (x \ z)" proof - fix x y z have "- (y \ z \ x) = - (- (- y \ z) \ - (- y \ - z) \ x) \ - (x \ - - z)" using 1437 2530 6113 by (smt commutative inf_def) thus "x \ y \ z = (x \ y) \ (x \ z)" using 12 1122 by (metis commutative inf_def) qed qed show 14: "\x. x \ - x = bot" proof - fix x have "(bot \ x) \ (bot \ -x) = bot" using huntington bot_def inf_def by auto thus "x \ -x = bot" using 11 less_eq_def by force qed show 15: "\x. x \ - x = top" using 5 14 by (metis (no_types, lifting) huntington bot_def less_eq_def top_def) show 16: "\x y. x - y = x \ - y" using 15 by (metis commutative huntington inf_def minus_def) show 7: "\x y z. x \ y \ x \ z \ x \ y \ z" by (simp add: 13 less_eq_def) qed end -context Lattices.boolean_algebra +context boolean_algebra begin sublocale ba_he: huntington_extended proof show "\x y z. x \ (y \ z) = x \ y \ z" by (simp add: sup_assoc) show "\x y. x \ y = y \ x" by (simp add: sup_commute) show "\x y. x = - (- x \ y) \ - (- x \ - y)" by simp show "top = (THE x. \y. x = y \ - y)" by auto show "bot = - (THE x. \y. x = y \ - y)" by auto show "\x y. x \ y = - (- x \ - y)" by simp show "\x y. x - y = - (- x \ y)" by (simp add: diff_eq) show "\x y. (x \ y) = (x \ y = y)" by (simp add: le_iff_sup) show "\x y. (x < y) = (x \ y = y \ y \ x \ x)" using sup.strict_order_iff sup_commute by auto qed end subsection \Stone Algebras\ text \ We relate Stone algebras to Boolean algebras. \ class stone_algebra_extended = stone_algebra + minus + assumes stone_minus_def[simp]: "x - y = x \ -y" class regular_stone_algebra = stone_algebra_extended + assumes double_complement[simp]: "--x = x" begin subclass boolean_algebra proof show "\x. x \ - x = bot" by simp show "\x. x \ - x = top" using regular_dense_top by fastforce show "\x y. x - y = x \ - y" by simp qed end -context Lattices.boolean_algebra +context boolean_algebra begin sublocale ba_rsa: regular_stone_algebra proof show "\x y. x - y = x \ - y" by (simp add: diff_eq) show "\x. - - x = x" by simp qed end section \Alternative Axiomatisations of Boolean Algebras\ text \ We consider four axiomatisations of Boolean algebras based only on join and complement. The first three are from the literature and the fourth, a version using equational axioms, is new. The motivation for Byrne's and the new axiomatisation is that the axioms are easier to understand than Huntington's third axiom. We also include Meredith's axiomatisation. \ subsection \Lee Byrne's Formulation A\ text \ The following axiomatisation is from \cite[Formulation A]{Byrne1946}; see also \cite{Frink1941}. \ text \Theorem 3\ class boolean_algebra_1 = sup + uminus + assumes ba1_associative: "x \ (y \ z) = (x \ y) \ z" assumes ba1_commutative: "x \ y = y \ x" assumes ba1_complement: "x \ -y = z \ -z \ x \ y = x" begin subclass huntington proof show 1: "\x y z. x \ (y \ z) = x \ y \ z" by (simp add: ba1_associative) show "\x y. x \ y = y \ x" by (simp add: ba1_commutative) show "\x y. x = - (- x \ y) \ - (- x \ - y)" proof - have 2: "\x y. y \ (y \ x) = y \ x" using 1 by (metis ba1_complement) hence "\x. --x = x" by (smt ba1_associative ba1_commutative ba1_complement) hence "\x y. y \ -(y \ -x) = y \ x" by (smt ba1_associative ba1_commutative ba1_complement) thus "\x y. x = -(-x \ y) \ -(-x \ - y)" using 2 by (smt ba1_commutative ba1_complement) qed qed end context huntington begin sublocale h_ba1: boolean_algebra_1 proof show "\x y z. x \ (y \ z) = x \ y \ z" by (simp add: associative) show "\x y. x \ y = y \ x" by (simp add: commutative) show "\x y z. (x \ - y = z \ - z) = (x \ y = x)" proof fix x y z have 1: "\x y z. -(-x \ y) \ (-(-x \ -y) \ z) = x \ z" using associative huntington by force have 2: "\x y. -(x \ -y) \ -(-y \ -x) = y" by (metis commutative huntington) show "x \ - y = z \ - z \ x \ y = x" by (metis 1 2 associative commutative top_unique) show "x \ y = x \ x \ - y = z \ - z" by (metis associative huntington commutative top_unique) qed qed end subsection \Lee Byrne's Formulation B\ text \ The following axiomatisation is from \cite[Formulation B]{Byrne1946}. \ text \Theorem 4\ class boolean_algebra_2 = sup + uminus + assumes ba2_associative_commutative: "(x \ y) \ z = (y \ z) \ x" assumes ba2_complement: "x \ -y = z \ -z \ x \ y = x" begin subclass boolean_algebra_1 proof show "\x y z. x \ (y \ z) = x \ y \ z" by (smt ba2_associative_commutative ba2_complement) show "\x y. x \ y = y \ x" by (metis ba2_associative_commutative ba2_complement) show "\x y z. (x \ - y = z \ - z) = (x \ y = x)" by (simp add: ba2_complement) qed end context boolean_algebra_1 begin sublocale ba1_ba2: boolean_algebra_2 proof show "\x y z. x \ y \ z = y \ z \ x" using ba1_associative commutative by force show "\x y z. (x \ - y = z \ - z) = (x \ y = x)" by (simp add: ba1_complement) qed end subsection \Meredith's Equational Axioms\ text \ The following axiomatisation is from \cite[page 221 (1) \{A,N\}]{MeredithPrior1968}. \ class boolean_algebra_mp = sup + uminus + assumes ba_mp_1: "-(-x \ y) \ x = x" assumes ba_mp_2: "-(-x \ y) \ (z \ y) = y \ (z \ x)" begin subclass huntington proof show "\x y z. x \ (y \ z) = x \ y \ z" by (metis ba_mp_1 ba_mp_2) show "\x y. x \ y = y \ x" by (metis ba_mp_1 ba_mp_2) show "\x y. x = - (- x \ y) \ - (- x \ - y)" by (metis ba_mp_1 ba_mp_2) qed end context huntington begin sublocale mp_h: boolean_algebra_mp proof show 1: "\x y. - (- x \ y) \ x = x" by (metis h_ba1.ba1_associative h_ba1.ba1_complement huntington) show "\x y z. - (- x \ y) \ (z \ y) = y \ (z \ x)" proof - fix x y z have "y = -(-x \ -y) \ y" using 1 h_ba1.ba1_commutative by auto thus "-(-x \ y) \ (z \ y) = y \ (z \ x)" by (metis h_ba1.ba1_associative h_ba1.ba1_commutative huntington) qed qed end subsection \An Equational Axiomatisation based on Semilattices\ text \ The following version is an equational axiomatisation based on semilattices. We add the double complement rule and that \top\ is unique. The final axiom \ba3_export\ encodes the logical statement $P \vee Q = P \vee (\neg P \wedge Q)$. Its dual appears in \cite{BalbesHorn1970}. \ text \Theorem 5\ class boolean_algebra_3 = sup + uminus + assumes ba3_associative: "x \ (y \ z) = (x \ y) \ z" assumes ba3_commutative: "x \ y = y \ x" assumes ba3_idempotent[simp]: "x \ x = x" assumes ba3_double_complement[simp]: "--x = x" assumes ba3_top_unique: "x \ -x = y \ -y" assumes ba3_export: "x \ -(x \ y) = x \ -y" begin subclass huntington proof show "\x y z. x \ (y \ z) = x \ y \ z" by (simp add: ba3_associative) show "\x y. x \ y = y \ x" by (simp add: ba3_commutative) show "\x y. x = - (- x \ y) \ - (- x \ - y)" by (metis ba3_commutative ba3_double_complement ba3_export ba3_idempotent ba3_top_unique) qed end context huntington begin sublocale h_ba3: boolean_algebra_3 proof show "\x y z. x \ (y \ z) = x \ y \ z" by (simp add: h_ba1.ba1_associative) show "\x y. x \ y = y \ x" by (simp add: h_ba1.ba1_commutative) show 3: "\x. x \ x = x" using h_ba1.ba1_complement by blast show 4: "\x. - - x = x" by (metis h_ba1.ba1_commutative huntington top_unique) show "\x y. x \ - x = y \ - y" by (simp add: top_unique) show "\x y. x \ - (x \ y) = x \ - y" using 3 4 by (smt h_ba1.ba1_ba2.ba2_associative_commutative h_ba1.ba1_complement) qed end section \Subset Boolean Algebras\ text \ We apply Huntington's axioms to the range of a unary operation, which serves as complement on the range. This gives a Boolean algebra structure on the range without imposing any further constraints on the set. The obtained structure is used as a reference in the subsequent development and to inherit the results proved here. This is taken from \cite{Guttmann2012c,GuttmannStruthWeber2011b} and follows the development of Boolean algebras in \cite{Maddux1996}. \ text \Definition 6\ class subset_boolean_algebra = sup + uminus + assumes sub_associative: "-x \ (-y \ -z) = (-x \ -y) \ -z" assumes sub_commutative: "-x \ -y = -y \ -x" assumes sub_complement: "-x = -(--x \ -y) \ -(--x \ --y)" assumes sub_sup_closed: "-x \ -y = --(-x \ -y)" begin text \uniqueness of \top\, resulting in the lemma \top_def\ to replace the assumption \sub_top_def\\ lemma top_unique: "-x \ --x = -y \ --y" by (metis sub_associative sub_commutative sub_complement) text \consequences for join and complement\ lemma double_negation[simp]: "---x = -x" by (metis sub_complement sub_sup_closed) lemma complement_1: "--x = -(-x \ -y) \ -(-x \ --y)" by (metis double_negation sub_complement) lemma sup_right_zero_var: "-x \ (-y \ --y) = -z \ --z" by (smt complement_1 sub_associative sub_sup_closed top_unique) lemma sup_right_unit_idempotent: "-x \ -x = -x \ -(-y \ --y)" by (metis complement_1 double_negation sub_sup_closed sup_right_zero_var) lemma sup_idempotent[simp]: "-x \ -x = -x" by (smt complement_1 double_negation sub_associative sup_right_unit_idempotent) lemma complement_2: "-x = -(-(-x \ -y) \ -(-x \ --y))" using complement_1 by auto lemma sup_eq_cases: "-x \ -y = -x \ -z \ --x \ -y = --x \ -z \ -y = -z" by (metis complement_2 sub_commutative) lemma sup_eq_cases_2: "-y \ -x = -z \ -x \ -y \ --x = -z \ --x \ -y = -z" using sub_commutative sup_eq_cases by auto end text \Definition 7\ class subset_extended = sup + inf + minus + uminus + bot + top + ord + assumes sub_top_def: "top = (THE x . \y . x = -y \ --y)" (* define without imposing uniqueness *) assumes sub_bot_def: "bot = -(THE x . \y . x = -y \ --y)" assumes sub_inf_def: "-x \ -y = -(--x \ --y)" assumes sub_minus_def: "-x - -y = -(--x \ -y)" assumes sub_less_eq_def: "-x \ -y \ -x \ -y = -y" assumes sub_less_def: "-x < -y \ -x \ -y = -y \ \ (-y \ -x = -x)" class subset_boolean_algebra_extended = subset_boolean_algebra + subset_extended begin lemma top_def: "top = -x \ --x" using sub_top_def top_unique by blast text \consequences for meet\ lemma inf_closed: "-x \ -y = --(-x \ -y)" by (simp add: sub_inf_def) lemma inf_associative: "-x \ (-y \ -z) = (-x \ -y) \ -z" using sub_associative sub_inf_def sub_sup_closed by auto lemma inf_commutative: "-x \ -y = -y \ -x" by (simp add: sub_commutative sub_inf_def) lemma inf_idempotent[simp]: "-x \ -x = -x" by (simp add: sub_inf_def) lemma inf_absorb[simp]: "(-x \ -y) \ -x = -x" by (metis complement_1 sup_idempotent sub_inf_def sub_associative sub_sup_closed) lemma sup_absorb[simp]: "-x \ (-x \ -y) = -x" by (metis sub_associative sub_complement sub_inf_def sup_idempotent) lemma inf_demorgan: "-(-x \ -y) = --x \ --y" using sub_inf_def sub_sup_closed by auto lemma sub_sup_demorgan: "-(-x \ -y) = --x \ --y" by (simp add: sub_inf_def) lemma sup_cases: "-x = (-x \ -y) \ (-x \ --y)" by (metis inf_closed inf_demorgan sub_complement) lemma inf_cases: "-x = (-x \ -y) \ (-x \ --y)" by (metis complement_2 sub_sup_closed sub_sup_demorgan) lemma inf_complement_intro: "(-x \ -y) \ --x = -y \ --x" proof - have "(-x \ -y) \ --x = (-x \ -y) \ (--x \ -y) \ --x" by (metis inf_absorb inf_associative sub_sup_closed) also have "... = -y \ --x" by (metis inf_cases sub_commutative) finally show ?thesis . qed lemma sup_complement_intro: "-x \ -y = -x \ (--x \ -y)" by (metis inf_absorb inf_commutative inf_complement_intro sub_sup_closed sup_cases) lemma inf_left_dist_sup: "-x \ (-y \ -z) = (-x \ -y) \ (-x \ -z)" proof - have "-x \ (-y \ -z) = (-x \ (-y \ -z) \ -y) \ (-x \ (-y \ -z) \ --y)" by (metis sub_inf_def sub_sup_closed sup_cases) also have "... = (-x \ -y) \ (-x \ -z \ --y)" by (metis inf_absorb inf_associative inf_complement_intro sub_sup_closed) also have "... = (-x \ -y) \ ((-x \ -y \ -z) \ (-x \ -z \ --y))" using sub_associative sub_inf_def sup_absorb by auto also have "... = (-x \ -y) \ ((-x \ -z \ -y) \ (-x \ -z \ --y))" by (metis inf_associative inf_commutative) also have "... = (-x \ -y) \ (-x \ -z)" by (metis sub_inf_def sup_cases) finally show ?thesis . qed lemma sup_left_dist_inf: "-x \ (-y \ -z) = (-x \ -y) \ (-x \ -z)" proof - have "-x \ (-y \ -z) = -(--x \ (--y \ --z))" by (metis sub_inf_def sub_sup_closed sub_sup_demorgan) also have "... = (-x \ -y) \ (-x \ -z)" by (metis inf_left_dist_sup sub_sup_closed sub_sup_demorgan) finally show ?thesis . qed lemma sup_right_dist_inf: "(-y \ -z) \ -x = (-y \ -x) \ (-z \ -x)" using sub_commutative sub_inf_def sup_left_dist_inf by auto lemma inf_right_dist_sup: "(-y \ -z) \ -x = (-y \ -x) \ (-z \ -x)" by (metis inf_commutative inf_left_dist_sup sub_sup_closed) lemma case_duality: "(--x \ -y) \ (-x \ -z) = (-x \ -y) \ (--x \ -z)" proof - have 1: "-(--x \ --y) \ ----x = --x \ -y" using inf_commutative inf_complement_intro sub_sup_closed sub_sup_demorgan by auto have 2: "-(----x \ -(--x \ -z)) = -----x \ ---z" by (metis (no_types) double_negation sup_complement_intro sub_sup_demorgan) have 3: "-(--x \ --y) \ -x = -x" using inf_commutative inf_left_dist_sup sub_sup_closed sub_sup_demorgan by auto hence "-(--x \ --y) = -x \ -y" using sub_sup_closed sub_sup_demorgan by auto thus ?thesis by (metis double_negation 1 2 3 inf_associative inf_left_dist_sup sup_complement_intro) qed lemma case_duality_2: "(-x \ -y) \ (--x \ -z) = (-x \ -z) \ (--x \ -y)" using case_duality sub_commutative sub_inf_def by auto lemma complement_cases: "((-v \ -w) \ (--v \ -x)) \ -((-v \ -y) \ (--v \ -z)) = (-v \ -w \ --y) \ (--v \ -x \ --z)" proof - have 1: "(--v \ -w) = --(--v \ -w) \ (-v \ -x) = --(-v \ -x) \ (--v \ --y) = --(--v \ --y) \ (-v \ --z) = --(-v \ --z)" using sub_inf_def sub_sup_closed by auto have 2: "(-v \ (-x \ --z)) = --(-v \ (-x \ --z))" using sub_inf_def sub_sup_closed by auto have "((-v \ -w) \ (--v \ -x)) \ -((-v \ -y) \ (--v \ -z)) = ((-v \ -w) \ (--v \ -x)) \ (-(-v \ -y) \ -(--v \ -z))" using sub_inf_def by auto also have "... = ((-v \ -w) \ (--v \ -x)) \ ((--v \ --y) \ (-v \ --z))" using inf_demorgan by auto also have "... = (--v \ -w) \ (-v \ -x) \ ((--v \ --y) \ (-v \ --z))" by (metis case_duality double_negation) also have "... = (--v \ -w) \ ((-v \ -x) \ ((--v \ --y) \ (-v \ --z)))" by (metis 1 inf_associative sub_inf_def) also have "... = (--v \ -w) \ ((-v \ -x) \ (--v \ --y) \ (-v \ --z))" by (metis 1 inf_associative) also have "... = (--v \ -w) \ ((--v \ --y) \ (-v \ -x) \ (-v \ --z))" by (metis 1 inf_commutative) also have "... = (--v \ -w) \ ((--v \ --y) \ ((-v \ -x) \ (-v \ --z)))" by (metis 1 inf_associative) also have "... = (--v \ -w) \ ((--v \ --y) \ (-v \ (-x \ --z)))" by (simp add: sup_left_dist_inf) also have "... = (--v \ -w) \ (--v \ --y) \ (-v \ (-x \ --z))" using 1 2 by (metis inf_associative) also have "... = (--v \ (-w \ --y)) \ (-v \ (-x \ --z))" by (simp add: sup_left_dist_inf) also have "... = (-v \ (-w \ --y)) \ (--v \ (-x \ --z))" by (metis case_duality complement_1 complement_2 sub_inf_def) also have "... = (-v \ -w \ --y) \ (--v \ -x \ --z)" by (simp add: inf_associative) finally show ?thesis . qed lemma inf_cases_2: "--x = -(-x \ -y) \ -(-x \ --y)" using sub_inf_def sup_cases by auto text \consequences for \top\ and \bot\\ lemma sup_complement[simp]: "-x \ --x = top" using top_def by auto lemma inf_complement[simp]: "-x \ --x = bot" by (metis sub_bot_def sub_inf_def sub_top_def top_def) lemma complement_bot[simp]: "-bot = top" using inf_complement inf_demorgan sup_complement by fastforce lemma complement_top[simp]: "-top = bot" using sub_bot_def sub_top_def by blast lemma sup_right_zero[simp]: "-x \ top = top" using sup_right_zero_var by auto lemma sup_left_zero[simp]: "top \ -x = top" by (metis complement_bot sub_commutative sup_right_zero) lemma inf_right_unit[simp]: "-x \ bot = bot" by (metis complement_bot complement_top double_negation sub_sup_demorgan sup_right_zero) lemma inf_left_unit[simp]: "bot \ -x = bot" by (metis complement_top inf_commutative inf_right_unit) lemma sup_right_unit[simp]: "-x \ bot = -x" using sup_right_unit_idempotent by auto lemma sup_left_unit[simp]: "bot \ -x = -x" by (metis complement_top sub_commutative sup_right_unit) lemma inf_right_zero[simp]: "-x \ top = -x" by (metis inf_left_dist_sup sup_cases top_def) lemma sub_inf_left_zero[simp]: "top \ -x = -x" using inf_absorb top_def by fastforce lemma bot_double_complement[simp]: "--bot = bot" by simp lemma top_double_complement[simp]: "--top = top" by simp text \consequences for the order\ lemma reflexive: "-x \ -x" by (simp add: sub_less_eq_def) lemma transitive: "-x \ -y \ -y \ -z \ -x \ -z" by (metis sub_associative sub_less_eq_def) lemma antisymmetric: "-x \ -y \ -y \ -x \ -x = -y" by (simp add: sub_commutative sub_less_eq_def) lemma sub_bot_least: "bot \ -x" using sup_left_unit complement_top sub_less_eq_def by blast lemma top_greatest: "-x \ top" using complement_bot sub_less_eq_def sup_right_zero by blast lemma upper_bound_left: "-x \ -x \ -y" by (metis sub_associative sub_less_eq_def sub_sup_closed sup_idempotent) lemma upper_bound_right: "-y \ -x \ -y" using sub_commutative upper_bound_left by fastforce lemma sub_sup_left_isotone: assumes "-x \ -y" shows "-x \ -z \ -y \ -z" proof - have "-x \ -y = -y" by (meson assms sub_less_eq_def) thus ?thesis by (metis (full_types) sub_associative sub_commutative sub_sup_closed upper_bound_left) qed lemma sub_sup_right_isotone: "-x \ -y \ -z \ -x \ -z \ -y" by (simp add: sub_commutative sub_sup_left_isotone) lemma sup_isotone: assumes "-p \ -q" and "-r \ -s" shows "-p \ -r \ -q \ -s" proof - have "\x y. \ -x \ -y \ -r \ -x \ -y \ -s" by (metis (full_types) assms(2) sub_sup_closed sub_sup_right_isotone transitive) thus ?thesis by (metis (no_types) assms(1) sub_sup_closed sub_sup_left_isotone) qed lemma sub_complement_antitone: "-x \ -y \ --y \ --x" by (metis inf_absorb inf_demorgan sub_less_eq_def) lemma less_eq_inf: "-x \ -y \ -x \ -y = -x" by (metis inf_absorb inf_commutative sub_less_eq_def upper_bound_right sup_absorb) lemma inf_complement_left_antitone: "-x \ -y \ -(-y \ -z) \ -(-x \ -z)" by (simp add: sub_complement_antitone inf_demorgan sub_sup_left_isotone) lemma sub_inf_left_isotone: "-x \ -y \ -x \ -z \ -y \ -z" using sub_complement_antitone inf_closed inf_complement_left_antitone by fastforce lemma sub_inf_right_isotone: "-x \ -y \ -z \ -x \ -z \ -y" by (simp add: inf_commutative sub_inf_left_isotone) lemma inf_isotone: assumes "-p \ -q" and "-r \ -s" shows "-p \ -r \ -q \ -s" proof - have "\w x y z. (-w \ -x \ -y \ \ -w \ -x \ -z) \ \ -z \ -y" by (metis (no_types) inf_closed sub_inf_right_isotone transitive) thus ?thesis by (metis (no_types) assms inf_closed sub_inf_left_isotone) qed lemma least_upper_bound: "-x \ -z \ -y \ -z \ -x \ -y \ -z" by (metis sub_sup_closed transitive upper_bound_right sup_idempotent sup_isotone upper_bound_left) lemma lower_bound_left: "-x \ -y \ -x" by (metis sub_inf_def upper_bound_right sup_absorb) lemma lower_bound_right: "-x \ -y \ -y" using inf_commutative lower_bound_left by fastforce lemma greatest_lower_bound: "-x \ -y \ -x \ -z \ -x \ -y \ -z" by (metis inf_closed sub_inf_left_isotone less_eq_inf transitive lower_bound_left lower_bound_right) lemma less_eq_sup_top: "-x \ -y \ --x \ -y = top" by (metis complement_1 inf_commutative inf_complement_intro sub_inf_left_zero less_eq_inf sub_complement sup_complement_intro top_def) lemma less_eq_inf_bot: "-x \ -y \ -x \ --y = bot" by (metis complement_bot complement_top double_negation inf_demorgan less_eq_sup_top sub_inf_def) lemma shunting: "-x \ -y \ -z \ -y \ --x \ -z" proof (cases "--x \ (-z \ --y) = top") case True have "\v w. -v \ -w \ -w \ --v \ top" using less_eq_sup_top sub_commutative by blast thus ?thesis by (metis True sub_associative sub_commutative sub_inf_def sub_sup_closed) next case False hence "--x \ (-z \ --y) \ top \ \ -y \ -z \ --x" by (metis (no_types) less_eq_sup_top sub_associative sub_commutative sub_sup_closed) thus ?thesis using less_eq_sup_top sub_associative sub_commutative sub_inf_def sub_sup_closed by auto qed lemma shunting_right: "-x \ -y \ -z \ -x \ -z \ --y" by (metis inf_commutative sub_commutative shunting) lemma sup_less_eq_cases: assumes "-z \ -x \ -y" and "-z \ --x \ -y" shows "-z \ -y" proof - have "-z \ (-x \ -y) \ (--x \ -y)" by (metis assms greatest_lower_bound sub_sup_closed) also have "... = -y" by (metis inf_cases sub_commutative) finally show ?thesis . qed lemma sup_less_eq_cases_2: "-x \ -y \ -x \ -z \ --x \ -y \ --x \ -z \ -y \ -z" by (metis least_upper_bound sup_less_eq_cases sub_sup_closed) lemma sup_less_eq_cases_3: "-y \ -x \ -z \ -x \ -y \ --x \ -z \ --x \ -y \ -z" by (simp add: sup_less_eq_cases_2 sub_commutative) lemma inf_less_eq_cases: "-x \ -y \ -z \ --x \ -y \ -z \ -y \ -z" by (simp add: shunting sup_less_eq_cases) lemma inf_less_eq_cases_2: "-x \ -y \ -x \ -z \ --x \ -y \ --x \ -z \ -y \ -z" by (metis greatest_lower_bound inf_closed inf_less_eq_cases) lemma inf_less_eq_cases_3: "-y \ -x \ -z \ -x \ -y \ --x \ -z \ --x \ -y \ -z" by (simp add: inf_commutative inf_less_eq_cases_2) lemma inf_eq_cases: "-x \ -y = -x \ -z \ --x \ -y = --x \ -z \ -y = -z" by (metis inf_commutative sup_cases) lemma inf_eq_cases_2: "-y \ -x = -z \ -x \ -y \ --x = -z \ --x \ -y = -z" using inf_commutative inf_eq_cases by auto lemma wnf_lemma_1: "((-x \ -y) \ (--x \ -z)) \ -x = -x \ -y" proof - have "\u v w. (-u \ (-v \ --w)) \ -w = -u \ -w" by (metis inf_right_zero sub_associative sub_sup_closed sup_complement sup_idempotent sup_right_dist_inf) thus ?thesis by (metis (no_types) sub_associative sub_commutative sub_sup_closed sup_idempotent) qed lemma wnf_lemma_2: "((-x \ -y) \ (-z \ --y)) \ -y = -x \ -y" using sub_commutative wnf_lemma_1 by fastforce lemma wnf_lemma_3: "((-x \ -z) \ (--x \ -y)) \ --x = --x \ -y" by (metis case_duality case_duality_2 double_negation sub_commutative wnf_lemma_2) lemma wnf_lemma_4: "((-z \ -y) \ (-x \ --y)) \ --y = -x \ --y" using sub_commutative wnf_lemma_3 by auto end class subset_boolean_algebra' = sup + uminus + assumes sub_associative': "-x \ (-y \ -z) = (-x \ -y) \ -z" assumes sub_commutative': "-x \ -y = -y \ -x" assumes sub_complement': "-x = -(--x \ -y) \ -(--x \ --y)" assumes sub_sup_closed': "\z . -x \ -y = -z" begin subclass subset_boolean_algebra proof show "\x y z. - x \ (- y \ - z) = - x \ - y \ - z" by (simp add: sub_associative') show "\x y. - x \ - y = - y \ - x" by (simp add: sub_commutative') show "\x y. - x = - (- - x \ - y) \ - (- - x \ - - y)" by (simp add: sub_complement') show "\x y. - x \ - y = - - (- x \ - y)" proof - fix x y have "\x y. -y \ (-(--y \ -x) \ -(---x \ -y)) = -y \ --x" by (metis (no_types) sub_associative' sub_commutative' sub_complement') hence "\x. ---x = -x" by (metis (no_types) sub_commutative' sub_complement') thus "-x \ -y = --(-x \ -y)" by (metis sub_sup_closed') qed qed end text \ We introduce a type for the range of complement and show that it is an instance of \boolean_algebra\. \ typedef (overloaded) 'a boolean_subset = "{ x::'a::uminus . \y . x = -y }" by auto lemma simp_boolean_subset[simp]: "\y . Rep_boolean_subset x = -y" using Rep_boolean_subset by simp setup_lifting type_definition_boolean_subset text \Theorem 8.1\ instantiation boolean_subset :: (subset_boolean_algebra) huntington begin lift_definition sup_boolean_subset :: "'a boolean_subset \ 'a boolean_subset \ 'a boolean_subset" is sup using sub_sup_closed by auto lift_definition uminus_boolean_subset :: "'a boolean_subset \ 'a boolean_subset" is uminus by auto instance proof show "\x y z::'a boolean_subset. x \ (y \ z) = x \ y \ z" apply transfer using sub_associative by blast show "\x y::'a boolean_subset. x \ y = y \ x" apply transfer using sub_commutative by blast show "\x y::'a boolean_subset. x = - (- x \ y) \ - (- x \ - y)" apply transfer using sub_complement by blast qed end text \Theorem 8.2\ instantiation boolean_subset :: (subset_boolean_algebra_extended) huntington_extended begin lift_definition inf_boolean_subset :: "'a boolean_subset \ 'a boolean_subset \ 'a boolean_subset" is inf using inf_closed by auto lift_definition minus_boolean_subset :: "'a boolean_subset \ 'a boolean_subset \ 'a boolean_subset" is minus using sub_minus_def by auto lift_definition bot_boolean_subset :: "'a boolean_subset" is bot by (metis complement_top) lift_definition top_boolean_subset :: "'a boolean_subset" is top by (metis complement_bot) lift_definition less_eq_boolean_subset :: "'a boolean_subset \ 'a boolean_subset \ bool" is less_eq . lift_definition less_boolean_subset :: "'a boolean_subset \ 'a boolean_subset \ bool" is less . instance proof show 1: "top = (THE x. \y::'a boolean_subset. x = y \ - y)" proof (rule the_equality[symmetric]) show "\y::'a boolean_subset. top = y \ - y" apply transfer by auto show "\x::'a boolean_subset. \y. x = y \ - y \ x = top" apply transfer by force qed have "(bot::'a boolean_subset) = - top" apply transfer by simp thus "bot = - (THE x. \y::'a boolean_subset. x = y \ - y)" using 1 by simp show "\x y::'a boolean_subset. x \ y = - (- x \ - y)" apply transfer using sub_inf_def by blast show "\x y::'a boolean_subset. x - y = - (- x \ y)" apply transfer using sub_minus_def by blast show "\x y::'a boolean_subset. (x \ y) = (x \ y = y)" apply transfer using sub_less_eq_def by blast show "\x y::'a boolean_subset. (x < y) = (x \ y = y \ y \ x \ x)" apply transfer using sub_less_def by blast qed end section \Subset Boolean algebras with Additional Structure\ text \ We now discuss axioms that make the range of a unary operation a Boolean algebra, but add further properties that are common to the intended models. In the intended models, the unary operation can be a complement, a pseudocomplement or the antidomain operation. For simplicity, we mostly call the unary operation `complement'. We first look at structures based only on join and complement, and then add axioms for the remaining operations of Boolean algebras. In the intended models, the operation that is meet on the range of the complement can be a meet in the whole algebra or composition. \ subsection \Axioms Derived from the New Axiomatisation\ text \ The axioms of the first algebra are based on \boolean_algebra_3\. \ text \Definition 9\ class subset_boolean_algebra_1 = sup + uminus + assumes sba1_associative: "x \ (y \ z) = (x \ y) \ z" assumes sba1_commutative: "x \ y = y \ x" assumes sba1_idempotent[simp]: "x \ x = x" assumes sba1_double_complement[simp]: "---x = -x" assumes sba1_bot_unique: "-(x \ -x) = -(y \ -y)" assumes sba1_export: "-x \ -(-x \ y) = -x \ -y" begin text \Theorem 11.1\ subclass subset_boolean_algebra proof show "\x y z. - x \ (- y \ - z) = - x \ - y \ - z" by (simp add: sba1_associative) show "\x y. - x \ - y = - y \ - x" by (simp add: sba1_commutative) show "\x y. - x = - (- - x \ - y) \ - (- - x \ - - y)" by (smt sba1_bot_unique sba1_commutative sba1_double_complement sba1_export sba1_idempotent) thus "\x y. - x \ - y = - - (- x \ - y)" by (metis sba1_double_complement sba1_export) qed definition "sba1_bot \ THE x . \z . x = -(z \ -z)" lemma sba1_bot: "sba1_bot = -(z \ -z)" using sba1_bot_def sba1_bot_unique by auto end text \Boolean algebra operations based on join and complement\ text \Definition 10\ class subset_extended_1 = sup + inf + minus + uminus + bot + top + ord + assumes ba_bot: "bot = (THE x . \z . x = -(z \ -z))" assumes ba_top: "top = -(THE x . \z . x = -(z \ -z))" assumes ba_inf: "-x \ -y = -(--x \ --y)" assumes ba_minus: "-x - -y = -(--x \ -y)" assumes ba_less_eq: "x \ y \ x \ y = y" assumes ba_less: "x < y \ x \ y = y \ \ (y \ x = x)" class subset_extended_2 = subset_extended_1 + assumes ba_bot_unique: "-(x \ -x) = -(y \ -y)" begin lemma ba_bot_def: "bot = -(z \ -z)" using ba_bot ba_bot_unique by auto lemma ba_top_def: "top = --(z \ -z)" using ba_bot_def ba_top by simp end text \Subset forms Boolean Algebra, extended by Boolean algebra operations\ class subset_boolean_algebra_1_extended = subset_boolean_algebra_1 + subset_extended_1 begin subclass subset_extended_2 proof show "\x y. - (x \ - x) = - (y \ - y)" by (simp add: sba1_bot_unique) qed subclass semilattice_sup proof show "\x y. (x < y) = (x \ y \ \ y \ x)" by (simp add: ba_less ba_less_eq) show "\x. x \ x" by (simp add: ba_less_eq) show "\x y z. x \ y \ y \ z \ x \ z" by (metis sba1_associative ba_less_eq) show "\x y. x \ y \ y \ x \ x = y" by (simp add: sba1_commutative ba_less_eq) show "\x y. x \ x \ y" by (simp add: sba1_associative ba_less_eq) thus "\y x. y \ x \ y" by (simp add: sba1_commutative) show "\y x z. y \ x \ z \ x \ y \ z \ x" by (metis sba1_associative ba_less_eq) qed text \Theorem 11.2\ subclass subset_boolean_algebra_extended proof show "top = (THE x. \y. x = - y \ - - y)" by (smt ba_bot ba_bot_def ba_top sub_sup_closed the_equality) thus "bot = - (THE x. \y. x = - y \ - - y)" using ba_bot_def ba_top_def by force show "\x y. - x \ - y = - (- - x \ - - y)" by (simp add: ba_inf) show "\x y. - x - - y = - (- - x \ - y)" by (simp add: ba_minus) show "\x y. (- x \ - y) = (- x \ - y = - y)" using le_iff_sup by auto show "\x y. (- x < - y) = (- x \ - y = - y \ - y \ - x \ - x)" by (simp add: ba_less) qed end subsection \Stronger Assumptions based on Join and Complement\ text \ We add further axioms covering properties common to the antidomain and (pseudo)complement instances. \ text \Definition 12\ class subset_boolean_algebra_2 = sup + uminus + assumes sba2_associative: "x \ (y \ z) = (x \ y) \ z" assumes sba2_commutative: "x \ y = y \ x" assumes sba2_idempotent[simp]: "x \ x = x" assumes sba2_bot_unit: "x \ -(y \ -y) = x" assumes sba2_sub_sup_demorgan: "-(x \ y) = -(--x \ --y)" assumes sba2_export: "-x \ -(-x \ y) = -x \ -y" begin text \Theorem 13.1\ subclass subset_boolean_algebra_1 proof show "\x y z. x \ (y \ z) = x \ y \ z" by (simp add: sba2_associative) show "\x y. x \ y = y \ x" by (simp add: sba2_commutative) show "\x. x \ x = x" by simp show "\x. - - - x = - x" by (metis sba2_idempotent sba2_sub_sup_demorgan) show "\x y. - (x \ - x) = - (y \ - y)" by (metis sba2_bot_unit sba2_commutative) show "\x y. - x \ - (- x \ y) = - x \ - y" by (simp add: sba2_export) qed text \Theorem 13.2\ lemma double_complement_dist_sup: "--(x \ y) = --x \ --y" by (metis sba2_commutative sba2_export sba2_idempotent sba2_sub_sup_demorgan) lemma maddux_3_3[simp]: "-(x \ y) \ -(x \ -y) = -x" by (metis double_complement_dist_sup sba1_double_complement sba2_commutative sub_complement) lemma huntington_3_pp[simp]: "-(-x \ -y) \ -(-x \ y) = --x" using sba2_commutative maddux_3_3 by fastforce end class subset_boolean_algebra_2_extended = subset_boolean_algebra_2 + subset_extended_1 begin subclass subset_boolean_algebra_1_extended .. subclass bounded_semilattice_sup_bot proof show "\x. bot \ x" using sba2_bot_unit ba_bot_def sup_right_divisibility by auto qed text \Theorem 13.3\ lemma complement_antitone: "x \ y \ -y \ -x" by (metis le_iff_sup maddux_3_3 sba2_export sup_monoid.add_commute) lemma double_complement_isotone: "x \ y \ --x \ --y" by (simp add: complement_antitone) lemma sup_demorgan: "-(x \ y) = -x \ -y" using sba2_sub_sup_demorgan ba_inf by auto end subsection \Axioms for Meet\ text \ We add further axioms of \inf\ covering properties common to the antidomain and pseudocomplement instances. We omit the left distributivity rule and the right zero rule as they do not hold in some models. In particular, the operation \inf\ does not have to be commutative. \ text \Definition 14\ class subset_boolean_algebra_3_extended = subset_boolean_algebra_2_extended + assumes sba3_inf_associative: "x \ (y \ z) = (x \ y) \ z" assumes sba3_inf_right_dist_sup: "(x \ y) \ z = (x \ z) \ (y \ z)" assumes sba3_inf_complement_bot: "-x \ x = bot" assumes sba3_inf_left_unit[simp]: "top \ x = x" assumes sba3_complement_inf_double_complement: "-(x \ --y) = -(x \ y)" begin text \Theorem 15\ lemma inf_left_zero: "bot \ x = bot" by (metis inf_right_unit sba3_inf_associative sba3_inf_complement_bot) lemma inf_double_complement_export: "--(--x \ y) = --x \ --y" by (metis inf_closed sba3_complement_inf_double_complement) lemma inf_left_isotone: "x \ y \ x \ z \ y \ z" using sba3_inf_right_dist_sup sup_right_divisibility by auto lemma inf_complement_export: "--(-x \ y) = -x \ --y" by (metis inf_double_complement_export sba1_double_complement) lemma double_complement_above: "--x \ x = x" by (metis sup_monoid.add_0_right complement_bot inf_demorgan sba1_double_complement sba3_inf_complement_bot sba3_inf_right_dist_sup sba3_inf_left_unit) lemma "x \ y \ z \ x \ z \ y" nitpick [expect=genuine] oops lemma "x \ top = x" nitpick [expect=genuine] oops lemma "x \ y = y \ x" nitpick [expect=genuine] oops end subsection \Stronger Assumptions for Meet\ text \ The following axioms also hold in both models, but follow from the axioms of \subset_boolean_algebra_5_operations\. \ text \Definition 16\ class subset_boolean_algebra_4_extended = subset_boolean_algebra_3_extended + assumes sba4_inf_right_unit[simp]: "x \ top = x" assumes inf_right_isotone: "x \ y \ z \ x \ z \ y" begin lemma "x \ top = top" nitpick [expect=genuine] oops lemma "x \ bot = bot" nitpick [expect=genuine] oops lemma "x \ (y \ z) = (x \ y) \ (x \ z)" nitpick [expect=genuine] oops lemma "(x \ y = bot) = (x \ - y)" nitpick [expect=genuine] oops end section \Boolean Algebras in Stone Algebras\ text \ We specialise \inf\ to meet and complement to pseudocomplement. This puts Stone algebras into the picture; for these it is well known that regular elements form a Boolean subalgebra \cite{Graetzer1971}. \ text \Definition 17\ class subset_boolean_algebra_5_extended = subset_boolean_algebra_3_extended + assumes sba5_inf_commutative: "x \ y = y \ x" assumes sba5_inf_absorb: "x \ (x \ y) = x" begin subclass distrib_lattice_bot proof show "\x y. x \ y \ x" by (metis sba5_inf_commutative sba3_inf_right_dist_sup sba5_inf_absorb sup_right_divisibility) show "\x y. x \ y \ y" by (metis inf_left_isotone sba5_inf_absorb sba5_inf_commutative sup_ge2) show "\x y z. x \ y \ x \ z \ x \ y \ z" by (metis inf_left_isotone sba5_inf_absorb sup.orderE sup_monoid.add_commute) show "\x y z. x \ y \ z = (x \ y) \ (x \ z) " by (metis sba3_inf_right_dist_sup sba5_inf_absorb sba5_inf_commutative sup_assoc) qed lemma inf_demorgan_2: "-(x \ y) = -x \ -y" using sba3_complement_inf_double_complement sba5_inf_commutative sub_sup_closed sub_sup_demorgan by auto lemma inf_export: "x \ -(x \ y) = x \ -y" using inf_demorgan_2 sba3_inf_complement_bot sba3_inf_right_dist_sup sba5_inf_commutative by auto lemma complement_inf[simp]: "x \ -x = bot" using sba3_inf_complement_bot sba5_inf_commutative by auto text \Theorem 18.2\ subclass stone_algebra proof show "\x. x \ top" by (simp add: inf.absorb_iff2) show "\x y. (x \ y = bot) = (x \ - y)" by (metis (full_types) complement_bot complement_inf inf.cobounded1 inf.order_iff inf_export sba3_complement_inf_double_complement sba3_inf_left_unit) show "\x. - x \ - - x = top" by simp qed text \Theorem 18.1\ subclass subset_boolean_algebra_4_extended proof show "\x. x \ top = x" by simp show "\x y z. x \ y \ z \ x \ z \ y" using inf.sup_right_isotone by blast qed end context stone_algebra_extended begin text \Theorem 18.3\ subclass subset_boolean_algebra_5_extended proof show "\x y z. x \ (y \ z) = x \ y \ z" using sup_assoc by auto show "\x y. x \ y = y \ x" by (simp add: sup_commute) show "\x. x \ x = x" by simp show "\x y. x \ - (y \ - y) = x" by simp show "\x y. - (x \ y) = - (- - x \ - - y)" by auto show "\x y. - x \ - (- x \ y) = - x \ - y" by (metis maddux_3_21_pp p_dist_sup regular_closed_p) show "bot = (THE x. \z. x = - (z \ - z))" by simp thus "top = - (THE x. \z. x = - (z \ - z))" using p_bot by blast show "\x y. - x \ - y = - (- - x \ - - y)" by simp show "\x y. - x - - y = - (- - x \ - y)" by auto show "\x y. (x \ y) = (x \ y = y)" by (simp add: le_iff_sup) thus "\x y. (x < y) = (x \ y = y \ y \ x \ x)" by (simp add: less_le_not_le) show "\x y z. x \ (y \ z) = x \ y \ z" by (simp add: inf.sup_monoid.add_assoc) show "\x y z. (x \ y) \ z = x \ z \ y \ z" by (simp add: inf_sup_distrib2) show "\x. - x \ x = bot" by simp show "\x. top \ x = x" by simp show "\x y. - (x \ - - y) = - (x \ y)" by simp show "\x y. x \ y = y \ x" by (simp add: inf_commute) show "\x y. x \ (x \ y) = x" by simp qed end section \Domain Semirings\ text \ The following development of tests in IL-semirings, prepredomain semirings, predomain semirings and domain semirings is mostly based on \cite{MoellerDesharnais2019}; see also \cite{DesharnaisMoeller2014}. See \cite{DesharnaisMoellerStruth2006b} for domain axioms in idempotent semirings. See \cite{DesharnaisJipsenStruth2009,JacksonStokes2004} for domain axioms in semigroups and monoids. Some variants have been implemented in \cite{GomesGuttmannHoefnerStruthWeber2016}. \ subsection \Idempotent Left Semirings\ text \Definition 19\ class il_semiring = sup + inf + bot + top + ord + assumes il_associative: "x \ (y \ z) = (x \ y) \ z" assumes il_commutative: "x \ y = y \ x" assumes il_idempotent[simp]: "x \ x = x" assumes il_bot_unit: "x \ bot = x" assumes il_inf_associative: "x \ (y \ z) = (x \ y) \ z" assumes il_inf_right_dist_sup: "(x \ y) \ z = (x \ z) \ (y \ z)" assumes il_inf_left_unit[simp]: "top \ x = x" assumes il_inf_right_unit[simp]: "x \ top = x" assumes il_sub_inf_left_zero[simp]: "bot \ x = bot" assumes il_sub_inf_right_isotone: "x \ y \ z \ x \ z \ y" assumes il_less_eq: "x \ y \ x \ y = y" assumes il_less_def: "x < y \ x \ y \ \(y \ x)" begin lemma il_unit_bot: "bot \ x = x" using il_bot_unit il_commutative by fastforce subclass order proof show "\x y. (x < y) = (x \ y \ \ y \ x)" by (simp add: il_less_def) show "\x. x \ x" by (simp add: il_less_eq) show "\x y z. x \ y \ y \ z \ x \ z" by (metis il_associative il_less_eq) show "\x y. x \ y \ y \ x \ x = y" by (simp add: il_commutative il_less_eq) qed lemma il_sub_inf_right_isotone_var: "(x \ y) \ (x \ z) \ x \ (y \ z)" by (smt il_associative il_commutative il_idempotent il_less_eq il_sub_inf_right_isotone) lemma il_sub_inf_left_isotone: "x \ y \ x \ z \ y \ z" by (metis il_inf_right_dist_sup il_less_eq) lemma il_sub_inf_left_isotone_var: "(y \ x) \ (z \ x) \ (y \ z) \ x" by (simp add: il_inf_right_dist_sup) lemma sup_left_isotone: "x \ y \ x \ z \ y \ z" by (smt il_associative il_commutative il_idempotent il_less_eq) lemma sup_right_isotone: "x \ y \ z \ x \ z \ y" by (simp add: il_commutative sup_left_isotone) lemma bot_least: "bot \ x" by (simp add: il_less_eq il_unit_bot) lemma less_eq_bot: "x \ bot \ x = bot" by (simp add: il_bot_unit il_less_eq) abbreviation are_complementary :: "'a \ 'a \ bool" where "are_complementary x y \ x \ y = top \ x \ y = bot \ y \ x = bot" abbreviation test :: "'a \ bool" where "test x \ \y . are_complementary x y" definition tests :: "'a set" where "tests = { x . test x }" lemma bot_test: "test bot" by (simp add: il_unit_bot) lemma top_test: "test top" by (simp add: il_bot_unit) lemma test_sub_identity: "test x \ x \ top" using il_associative il_less_eq by auto lemma neg_unique: "are_complementary x y \ are_complementary x z \ y = z" by (metis order.antisym il_inf_left_unit il_inf_right_dist_sup il_inf_right_unit il_sub_inf_right_isotone_var) definition neg :: "'a \ 'a" ("!") where "!x \ THE y . are_complementary x y" lemma neg_char: assumes "test x" shows "are_complementary x (!x)" proof (unfold neg_def) from assms obtain y where 1: "are_complementary x y" by auto show "are_complementary x (THE y. are_complementary x y)" proof (rule theI) show "are_complementary x y" using 1 by simp show "\z. are_complementary x z \ z = y" using 1 neg_unique by blast qed qed lemma are_complementary_symmetric: "are_complementary x y \ are_complementary y x" using il_commutative by auto lemma neg_test: "test x \ test (!x)" using are_complementary_symmetric neg_char by blast lemma are_complementary_test: "test x \ are_complementary x y \ test y" using il_commutative by auto lemma neg_involutive: "test x \ !(!x) = x" using are_complementary_symmetric neg_char neg_unique by blast lemma test_inf_left_below: "test x \ x \ y \ y" by (metis il_associative il_idempotent il_inf_left_unit il_inf_right_dist_sup il_less_eq) lemma test_inf_right_below: "test x \ y \ x \ y" by (metis il_inf_right_unit il_sub_inf_right_isotone test_sub_identity) lemma neg_bot: "!bot = top" using il_unit_bot neg_char by fastforce lemma neg_top: "!top = bot" using bot_test neg_bot neg_involutive by fastforce lemma test_inf_idempotent: "test x \ x \ x = x" by (metis il_bot_unit il_inf_left_unit il_inf_right_dist_sup) lemma test_inf_semicommutative: assumes "test x" and "test y" shows "x \ y \ y \ x" proof - have "x \ y = (y \ x \ y) \ (!y \ x \ y)" by (metis assms(2) il_inf_left_unit il_inf_right_dist_sup neg_char) also have "... \ (y \ x \ y) \ (!y \ y)" proof - obtain z where "are_complementary y z" using assms(2) by blast hence "y \ (x \ y) \ !y \ (x \ y) \ y \ (x \ y)" by (metis assms(1) calculation il_sub_inf_left_isotone il_bot_unit il_idempotent il_inf_associative il_less_eq neg_char test_inf_right_below) thus ?thesis by (simp add: il_associative il_inf_associative il_less_eq) qed also have "... \ (y \ x) \ (!y \ y)" by (metis assms(2) il_bot_unit il_inf_right_unit il_sub_inf_right_isotone neg_char test_sub_identity) also have "... = y \ x" by (simp add: assms(2) il_bot_unit neg_char) finally show ?thesis . qed lemma test_inf_commutative: "test x \ test y \ x \ y = y \ x" by (simp add: order.antisym test_inf_semicommutative) lemma test_inf_bot: "test x \ x \ bot = bot" using il_inf_associative test_inf_idempotent by fastforce lemma test_absorb_1: "test x \ test y \ x \ (x \ y) = x" using il_commutative il_less_eq test_inf_right_below by auto lemma test_absorb_2: "test x \ test y \ x \ (y \ x) = x" by (metis test_absorb_1 test_inf_commutative) lemma test_absorb_3: "test x \ test y \ x \ (x \ y) = x" apply (rule order.antisym) apply (metis il_associative il_inf_right_unit il_less_eq il_sub_inf_right_isotone test_sub_identity) by (metis il_sub_inf_right_isotone_var test_absorb_1 test_inf_idempotent) lemma test_absorb_4: "test x \ test y \ (x \ y) \ x = x" by (smt il_inf_right_dist_sup test_inf_idempotent il_commutative il_less_eq test_inf_left_below) lemma test_import_1: assumes "test x" and "test y" shows "x \ (!x \ y) = x \ y" proof - have "x \ (!x \ y) = x \ ((y \ !y) \ x) \ (!x \ y)" by (simp add: assms(2) neg_char) also have "... = x \ (!y \ x) \ (x \ y) \ (!x \ y)" by (smt assms il_associative il_commutative il_inf_right_dist_sup test_inf_commutative) also have "... = x \ ((x \ !x) \ y)" by (smt calculation il_associative il_commutative il_idempotent il_inf_right_dist_sup) also have "... = x \ y" by (simp add: assms(1) neg_char) finally show ?thesis . qed lemma test_import_2: assumes "test x" and "test y" shows "x \ (y \ !x) = x \ y" proof - obtain z where 1: "are_complementary y z" using assms(2) by moura obtain w where 2: "are_complementary x w" using assms(1) by auto hence "x \ !x = bot" using neg_char by blast hence "!x \ y = y \ !x" using 1 2 by (metis il_commutative neg_char test_inf_commutative) thus ?thesis using 1 2 by (metis test_import_1) qed lemma test_import_3: assumes "test x" shows "(!x \ y) \ x = y \ x" by (simp add: assms(1) il_inf_right_dist_sup il_unit_bot neg_char) lemma test_import_4: assumes "test x" and "test y" shows "(!x \ y) \ x = x \ y" by (metis assms test_import_3 test_inf_commutative) lemma test_inf: "test x \ test y \ test z \ z \ x \ y \ z \ x \ z \ y" apply (rule iffI) using dual_order.trans test_inf_left_below test_inf_right_below apply blast by (smt il_less_eq il_sub_inf_right_isotone test_absorb_4) lemma test_shunting: assumes "test x" and "test y" shows "x \ y \ z \ x \ !y \ z" proof assume 1: "x \ y \ z" have "x = (!y \ x) \ (y \ x)" by (metis assms(2) il_commutative il_inf_left_unit il_inf_right_dist_sup neg_char) also have "... \ !y \ (y \ x)" by (simp add: assms(1) sup_left_isotone test_inf_right_below) also have "... \ !y \ z" using 1 by (simp add: assms sup_right_isotone test_inf_commutative) finally show "x \ !y \ z" . next assume "x \ !y \ z" hence "x \ y \ (!y \ z) \ y" using il_sub_inf_left_isotone by blast also have "... = z \ y" by (simp add: assms(2) test_import_3) also have "... \ z" by (simp add: assms(2) test_inf_right_below) finally show "x \ y \ z" . qed lemma test_shunting_bot: assumes "test x" and "test y" shows "x \ y \ x \ !y \ bot" by (simp add: assms il_bot_unit neg_involutive neg_test test_shunting) lemma test_shunting_bot_eq: assumes "test x" and "test y" shows "x \ y \ x \ !y = bot" by (simp add: assms test_shunting_bot less_eq_bot) lemma neg_antitone: assumes "test x" and "test y" and "x \ y" shows "!y \ !x" proof - have 1: "x \ !y = bot" using assms test_shunting_bot_eq by blast have 2: "x \ !x = top" by (simp add: assms(1) neg_char) have "are_complementary y (!y)" by (simp add: assms(2) neg_char) thus ?thesis using 1 2 by (metis il_unit_bot il_commutative il_inf_left_unit il_inf_right_dist_sup il_inf_right_unit il_sub_inf_right_isotone test_sub_identity) qed lemma test_sup_neg_1: assumes "test x" and "test y" shows "(x \ y) \ (!x \ !y) = top" proof - have "x \ !x = top" by (simp add: assms(1) neg_char) hence "x \ (y \ !x) = top" by (metis assms(2) il_associative il_commutative il_idempotent) hence "x \ (y \ !x \ !y) = top" by (simp add: assms neg_test test_import_2) thus ?thesis by (simp add: il_associative) qed lemma test_sup_neg_2: assumes "test x" and "test y" shows "(x \ y) \ (!x \ !y) = bot" proof - have 1: "are_complementary y (!y)" by (simp add: assms(2) neg_char) obtain z where 2: "are_complementary x z" using assms(1) by auto hence "!x = z" using neg_char neg_unique by blast thus ?thesis using 1 2 by (metis are_complementary_symmetric il_inf_associative neg_involutive test_import_3 test_inf_bot test_inf_commutative) qed lemma de_morgan_1: assumes "test x" and "test y" and "test (x \ y)" shows "!(x \ y) = !x \ !y" proof (rule order.antisym) have 1: "test (!(x \ y))" by (simp add: assms neg_test) have "x \ (x \ y) \ !y" by (metis (full_types) assms il_commutative neg_char test_shunting test_shunting_bot_eq) hence "x \ !(x \ y) \ !y" using 1 by (simp add: assms(1,3) neg_involutive test_shunting) hence "!(x \ y) \ x \ !y" using 1 by (metis assms(1) test_inf_commutative) thus "!(x \ y) \ !x \ !y" using 1 assms(1) test_shunting by blast have 2: "!x \ !(x \ y)" by (simp add: assms neg_antitone test_inf_right_below) have "!y \ !(x \ y)" by (simp add: assms neg_antitone test_inf_left_below) thus "!x \ !y \ !(x \ y)" using 2 by (metis il_associative il_less_eq) qed lemma de_morgan_2: assumes "test x" and "test y" and "test (x \ y)" shows "!(x \ y) = !x \ !y" proof (rule order.antisym) have 1: "!(x \ y) \ !x" by (metis assms il_inf_left_unit il_sub_inf_left_isotone neg_antitone test_absorb_3 test_sub_identity) have "!(x \ y) \ !y" by (metis assms il_commutative il_inf_left_unit il_sub_inf_left_isotone neg_antitone test_absorb_3 test_sub_identity) thus "!(x \ y) \ !x \ !y" using 1 by (simp add: assms neg_test test_inf) have "top \ x \ y \ !(x \ y)" by (simp add: assms(3) neg_char) hence "top \ !x \ y \ !(x \ y)" by (smt assms(1) assms(3) il_commutative il_inf_right_dist_sup il_inf_right_unit il_sub_inf_right_isotone il_unit_bot neg_char test_sub_identity) thus "!x \ !y \ !(x \ y)" by (simp add: assms(1) assms(2) neg_involutive neg_test test_shunting) qed lemma test_inf_closed_sup_complement: assumes "test x" and "test y" and "\u v . test u \ test v \ test (u \ v)" shows "!x \ !y \ (x \ y) = bot" proof - have 1: "!(!x \ !y) = x \ y" by (simp add: assms de_morgan_1 neg_involutive neg_test) have "test (!(!x \ !y))" by (metis assms neg_test) thus ?thesis using 1 by (metis assms(1,2) de_morgan_2 neg_char) qed lemma test_sup_complement_sup_closed: assumes "test x" and "test y" and "\u v . test u \ test v \ !u \ !v \ (u \ v) = bot" shows "test (x \ y)" by (meson assms test_sup_neg_1 test_sup_neg_2) lemma test_inf_closed_sup_closed: assumes "test x" and "test y" and "\u v . test u \ test v \ test (u \ v)" shows "test (x \ y)" using assms test_inf_closed_sup_complement test_sup_complement_sup_closed by simp end subsection \Prepredomain Semirings\ class dom = fixes d :: "'a \ 'a" class ppd_semiring = il_semiring + dom + assumes d_closed: "test (d x)" assumes d1: "x \ d x \ x" begin lemma d_sub_identity: "d x \ top" using d_closed test_sub_identity by blast lemma d1_eq: "x = d x \ x" proof - have "x = (d x \ top) \ x" using d_sub_identity il_less_eq by auto thus ?thesis using d1 il_commutative il_inf_right_dist_sup il_less_eq by force qed lemma d_increasing_sub_identity: "x \ top \ x \ d x" by (metis d1_eq il_inf_right_unit il_sub_inf_right_isotone) lemma d_top: "d top = top" by (simp add: d_increasing_sub_identity d_sub_identity dual_order.antisym) lemma d_bot_only: "d x = bot \ x = bot" by (metis d1_eq il_sub_inf_left_zero) lemma d_strict: "d bot \ bot" nitpick [expect=genuine] oops lemma d_isotone_var: "d x \ d (x \ y)" nitpick [expect=genuine] oops lemma d_fully_strict: "d x = bot \ x = bot" nitpick [expect=genuine] oops lemma test_d_fixpoint: "test x \ d x = x" nitpick [expect=genuine] oops end subsection \Predomain Semirings\ class pd_semiring = ppd_semiring + assumes d2: "test p \ d (p \ x) \ p" begin lemma d_strict: "d bot \ bot" using bot_test d2 by fastforce lemma d_strict_eq: "d bot = bot" using d_strict il_bot_unit il_less_eq by auto lemma test_d_fixpoint: "test x \ d x = x" by (metis order.antisym d1_eq d2 test_inf_idempotent test_inf_right_below) lemma d_surjective: "test x \ \y . d y = x" using test_d_fixpoint by blast lemma test_d_fixpoint_iff: "test x \ d x = x" by (metis d_closed test_d_fixpoint) lemma d_surjective_iff: "test x \ (\y . d y = x)" using d_surjective d_closed by blast lemma tests_d_range: "tests = range d" using tests_def image_def d_surjective_iff by auto lemma llp: assumes "test y" shows "d x \ y \ x \ y \ x" by (metis assms d1_eq d2 order.eq_iff il_sub_inf_left_isotone test_inf_left_below) lemma gla: assumes "test y" shows "y \ !(d x) \ y \ x \ bot" proof - obtain ad where 1: "\x. are_complementary (d x) (ad x)" using d_closed by moura hence 2: "\x y. d (d y \ x) \ d y" using d2 by blast have 3: "\x. ad x \ x = bot" using 1 by (metis d1_eq il_inf_associative il_sub_inf_left_zero) have 4: "\x y. d y \ x \ ad y \ x = top \ x" using 1 by (metis il_inf_right_dist_sup) have 5: "\x y z. z \ y \ x \ y \ (z \ x) \ y \ x \ y" by (simp add: il_inf_right_dist_sup il_less_eq) have 6: "\x. !(d x) = ad x" using 1 neg_char neg_unique by blast have 7: "\x. top \ x = x" by auto hence "\x. y \ x \ !y \ x = x" by (metis assms il_inf_right_dist_sup neg_char) thus ?thesis using 1 2 3 4 5 6 7 by (metis assms d1_eq il_commutative il_less_eq test_d_fixpoint) qed lemma gla_var: "test y \ y \ d x \ bot \ y \ x \ bot" using gla d_closed il_bot_unit test_shunting by auto lemma llp_var: assumes "test y" shows "y \ !(d x) \ x \ !y \ x" apply (rule iffI) apply (metis (no_types, opaque_lifting) assms gla Least_equality il_inf_left_unit il_inf_right_dist_sup il_less_eq il_unit_bot order.refl neg_char) by (metis assms gla gla_var llp il_commutative il_sub_inf_right_isotone neg_char) lemma d_idempotent: "d (d x) = d x" using d_closed test_d_fixpoint_iff by auto lemma d_neg: "test x \ d (!x) = !x" using il_commutative neg_char test_d_fixpoint_iff by fastforce lemma d_fully_strict: "d x = bot \ x = bot" using d_strict_eq d_bot_only by blast lemma d_ad_comp: "!(d x) \ x = bot" proof - have "\x. !(d x) \ d x = bot" by (simp add: d_closed neg_char) thus ?thesis by (metis d1_eq il_inf_associative il_sub_inf_left_zero) qed lemma d_isotone: assumes "x \ y" shows "d x \ d y" proof - obtain ad where 1: "\x. are_complementary (d x) (ad x)" using d_closed by moura hence "ad y \ x \ bot" by (metis assms d1_eq il_inf_associative il_sub_inf_left_zero il_sub_inf_right_isotone) thus ?thesis using 1 by (metis d2 il_bot_unit il_inf_left_unit il_inf_right_dist_sup il_less_eq) qed lemma d_isotone_var: "d x \ d (x \ y)" using d_isotone il_associative il_less_eq by auto lemma d3_conv: "d (x \ y) \ d (x \ d y)" by (metis (mono_tags, opaque_lifting) d1_eq d2 d_closed il_inf_associative) lemma d_test_inf_idempotent: "d x \ d x = d x" by (metis d_idempotent d1_eq) lemma d_test_inf_closed: assumes "test x" and "test y" shows "d (x \ y) = x \ y" proof (rule order.antisym) have "d (x \ y) = d (x \ y) \ d (x \ y)" by (simp add: d_test_inf_idempotent) also have "... \ x \ d (x \ y)" by (simp add: assms(1) d2 il_sub_inf_left_isotone) also have "... \ x \ y" by (metis assms d_isotone il_sub_inf_right_isotone test_inf_left_below test_d_fixpoint) finally show "d (x \ y) \ x \ y" . show "x \ y \ d (x \ y)" using assms d_increasing_sub_identity dual_order.trans test_inf_left_below test_sub_identity by blast qed lemma test_inf_closed: "test x \ test y \ test (x \ y)" using d_test_inf_closed test_d_fixpoint_iff by simp lemma test_sup_closed: "test x \ test y \ test (x \ y)" using test_inf_closed test_inf_closed_sup_closed by simp lemma d_export: assumes "test x" shows "d (x \ y) = x \ d y" proof (rule order.antisym) have 1: "d (x \ y) \ x" by (simp add: assms d2) have "d (x \ y) \ d y" by (metis assms d_isotone_var il_inf_left_unit il_inf_right_dist_sup) thus "d (x \ y) \ x \ d y" using 1 by (metis assms d_idempotent llp dual_order.trans il_sub_inf_right_isotone) have "y = (!x \ y) \ (x \ y)" by (metis assms il_commutative il_inf_left_unit il_inf_right_dist_sup neg_char) also have "... = (!x \ y) \ (d (x \ y) \ x \ y)" by (metis d1_eq il_inf_associative) also have "... = (!x \ y) \ (d (x \ y) \ y)" using 1 by (smt calculation d1_eq il_associative il_commutative il_inf_associative il_inf_right_dist_sup il_less_eq il_sub_inf_right_isotone_var) also have "... = (!x \ d (x \ y)) \ y" by (simp add: il_inf_right_dist_sup) finally have "y \ (!x \ d (x \ y)) \ y" by simp hence "d y \ !x \ d (x \ y)" using assms llp test_sup_closed neg_test d_closed by simp hence "d y \ x \ d (x \ y)" by (simp add: assms d_closed test_shunting) thus "x \ d y \ d (x \ y)" by (metis assms d_closed test_inf_commutative) qed lemma test_inf_left_dist_sup: assumes "test x" and "test y" and "test z" shows "x \ (y \ z) = (x \ y) \ (x \ z)" proof - have "x \ (y \ z) = (y \ z) \ x" using assms test_sup_closed test_inf_commutative by smt also have "... = (y \ x) \ (z \ x)" using il_inf_right_dist_sup by simp also have "... = (x \ y) \ (x \ z)" using assms test_sup_closed test_inf_commutative by smt finally show ?thesis . qed lemma "!x \ !y = !(!(!x \ !y))" nitpick [expect=genuine] oops lemma "d x = !(!x)" nitpick [expect=genuine] oops sublocale subset_boolean_algebra where uminus = "\ x . !(d x)" proof show "\x y z. !(d x) \ (!(d y) \ !(d z)) = !(d x) \ !(d y) \ !(d z)" using il_associative by blast show "\x y. !(d x) \ !(d y) = !(d y) \ !(d x)" by (simp add: il_commutative) show "\x y. !(d x) \ !(d y) = !(d (!(d (!(d x) \ !(d y)))))" proof - fix x y have "test (!(d x)) \ test (!(d y))" by (simp add: d_closed neg_test) hence "test (!(d x) \ !(d y))" by (simp add: test_sup_closed) thus "!(d x) \ !(d y) = !(d (!(d (!(d x) \ !(d y)))))" by (simp add: d_neg neg_involutive test_d_fixpoint) qed show "\x y. !(d x) = !(d (!(d (!(d x))) \ !(d y))) \ !(d (!(d (!(d x))) \ !(d (!(d y)))))" proof - fix x y have "!(d (!(d (!(d x))) \ !(d y))) \ !(d (!(d (!(d x))) \ !(d (!(d y))))) = !(d x \ !(d y)) \ !(d x \ d y)" using d_closed neg_test test_sup_closed neg_involutive test_d_fixpoint by auto also have "... = (!(d x) \ d y) \ (!(d x) \ !(d y))" using d_closed neg_test test_sup_closed neg_involutive de_morgan_2 by auto also have "... = !(d x) \ (d y \ !(d y))" using d_closed neg_test test_inf_left_dist_sup by auto also have "... = !(d x) \ top" by (simp add: neg_char d_closed) finally show "!(d x) = !(d (!(d (!(d x))) \ !(d y))) \ !(d (!(d (!(d x))) \ !(d (!(d y)))))" by simp qed qed lemma d_dist_sup: "d (x \ y) = d x \ d y" proof (rule order.antisym) have "x \ d x \ x" by (simp add: d1) also have "... \ (d x \ d y) \ (x \ y)" using il_associative il_inf_right_dist_sup il_less_eq il_sub_inf_right_isotone by auto finally have 1: "x \ (d x \ d y) \ (x \ y)" . have "y \ d y \ y" by (simp add: d1) also have "... \ (d y \ d x) \ (y \ x)" using il_associative il_idempotent il_inf_right_dist_sup il_less_eq il_sub_inf_right_isotone by simp finally have "y \ (d x \ d y) \ (x \ y)" using il_commutative by auto hence "x \ y \ (d x \ d y) \ (x \ y)" using 1 by (metis il_associative il_less_eq) thus "d (x \ y) \ d x \ d y" using llp test_sup_closed neg_test d_closed by simp show "d x \ d y \ d (x \ y)" using d_isotone_var il_associative il_commutative il_less_eq by fastforce qed end class pd_semiring_extended = pd_semiring + uminus + assumes uminus_def: "-x = !(d x)" begin subclass subset_boolean_algebra by (metis subset_boolean_algebra_axioms uminus_def ext) end subsection \Domain Semirings\ class d_semiring = pd_semiring + assumes d3: "d (x \ d y) \ d (x \ y)" begin lemma d3_eq: "d (x \ d y) = d (x \ y)" by (simp add: order.antisym d3 d3_conv) end text \ Axioms (d1), (d2) and (d3) are independent in IL-semirings. \ context il_semiring begin context fixes d :: "'a \ 'a" assumes d_closed: "test (d x)" begin context assumes d1: "x \ d x \ x" assumes d2: "test p \ d (p \ x) \ p" begin lemma d3: "d (x \ d y) \ d (x \ y)" nitpick [expect=genuine] oops end context assumes d1: "x \ d x \ x" assumes d3: "d (x \ d y) \ d (x \ y)" begin lemma d2: "test p \ d (p \ x) \ p" nitpick [expect=genuine] oops end context assumes d2: "test p \ d (p \ x) \ p" assumes d3: "d (x \ d y) \ d (x \ y)" begin lemma d1: "x \ d x \ x" nitpick [expect=genuine] oops end end end class d_semiring_var = ppd_semiring + assumes d3_var: "d (x \ d y) \ d (x \ y)" assumes d_strict_eq_var: "d bot = bot" begin lemma d2_var: assumes "test p" shows "d (p \ x) \ p" proof - have "!p \ p \ x = bot" by (simp add: assms neg_char) hence "d (!p \ p \ x) = bot" by (simp add: d_strict_eq_var) hence "d (!p \ d (p \ x)) = bot" by (metis d3_var il_inf_associative less_eq_bot) hence "!p \ d (p \ x) = bot" using d_bot_only by blast thus ?thesis by (metis (no_types, opaque_lifting) assms d_sub_identity il_bot_unit il_inf_left_unit il_inf_right_dist_sup il_inf_right_unit il_sub_inf_right_isotone neg_char) qed subclass d_semiring proof show "\p x. test p \ d (p \ x) \ p" by (simp add: d2_var) show "\x y. d (x \ d y) \ d (x \ y)" by (simp add: d3_var) qed end section \Antidomain Semirings\ text \ We now develop prepreantidomain semirings, preantidomain semirings and antidomain semirings. See \cite{DesharnaisStruth2008b,DesharnaisStruth2008a,DesharnaisStruth2011} for related work on internal axioms for antidomain. \ subsection \Prepreantidomain Semirings\ text \Definition 20\ class ppa_semiring = il_semiring + uminus + assumes a_inf_complement_bot: "-x \ x = bot" assumes a_stone[simp]: "-x \ --x = top" begin text \Theorem 21\ lemma l1: "-top = bot" by (metis a_inf_complement_bot il_inf_right_unit) lemma l2: "-bot = top" by (metis l1 a_stone il_unit_bot) lemma l3: "-x \ -y \ -x \ y = bot" by (metis a_inf_complement_bot il_bot_unit il_inf_right_dist_sup il_less_eq) lemma l5: "--x \ --y \ -y \ -x" by (metis (mono_tags, opaque_lifting) l3 a_stone bot_least il_bot_unit il_inf_left_unit il_inf_right_dist_sup il_inf_right_unit il_sub_inf_right_isotone sup_right_isotone) lemma l4: "---x = -x" by (metis l5 a_inf_complement_bot a_stone order.antisym bot_least il_inf_left_unit il_inf_right_dist_sup il_inf_right_unit il_sub_inf_right_isotone il_unit_bot) lemma l6: "-x \ --x = bot" by (metis l3 l5 a_inf_complement_bot a_stone il_inf_left_unit il_inf_right_dist_sup il_inf_right_unit il_less_eq il_sub_inf_right_isotone il_unit_bot) lemma l7: "-x \ -y = -y \ -x" using l6 a_inf_complement_bot a_stone test_inf_commutative by blast lemma l8: "x \ --x \ x" by (metis a_inf_complement_bot a_stone il_idempotent il_inf_left_unit il_inf_right_dist_sup il_less_eq il_unit_bot) sublocale ppa_ppd: ppd_semiring where d = "\x . --x" proof show "\x. test (- - x)" using l4 l6 by force show "\x. x \ - - x \ x" by (simp add: l8) qed (* The following statements have counterexamples, but they take a while to find. lemma "- x = - (- - x \ - y) \ - (- - x \ - - y)" nitpick [card=8, expect=genuine] oops lemma "- x \ - y = - - (- x \ - y)" nitpick [card=8, expect=genuine] oops *) end subsection \Preantidomain Semirings\ text \Definition 22\ class pa_semiring = ppa_semiring + assumes pad2: "--x \ -(-x \ y)" begin text \Theorem 23\ lemma l10: "-x \ y = bot \ -x \ -y" by (metis a_stone il_inf_left_unit il_inf_right_dist_sup il_unit_bot l4 pad2) lemma l10_iff: "-x \ y = bot \ -x \ -y" using l10 l3 by blast lemma l13: "--(--x \ y) \ --x" by (metis l4 l5 pad2) lemma l14: "-(x \ --y) \ -(x \ y)" by (metis il_inf_associative l4 pad2 ppa_ppd.d1_eq) lemma l9: "x \ y \ -y \ -x" by (metis l10 a_inf_complement_bot il_commutative il_less_eq il_sub_inf_right_isotone il_unit_bot) lemma l11: "- x \ - y = - (- - x \ - - y)" proof - have 1: "\x y . x \ y \ x \ y = y" by (simp add: il_less_eq) have 4: "\x y . \(x \ y) \ x \ y = y" using 1 by metis have 5: "\x y z . (x \ y) \ (x \ z) \ x \ (y \ z)" by (simp add: il_sub_inf_right_isotone_var) have 6: "\x y . - - x \ - (- x \ y)" by (simp add: pad2) have 7: "\x y z . x \ (y \ z) = (x \ y) \ z" by (simp add: il_associative) have 8: "\x y z . (x \ y) \ z = x \ (y \ z)" using 7 by metis have 9: "\x y . x \ y = y \ x" by (simp add: il_commutative) have 10: "\x . x \ bot = x" by (simp add: il_bot_unit) have 11: "\x . x \ x = x" by simp have 12: "\x y z . x \ (y \ z) = (x \ y) \ z" by (simp add: il_inf_associative) have 13: "\x y z . (x \ y) \ z = x \ (y \ z)" using 12 by metis have 14: "\x . top \ x = x" by simp have 15: "\x . x \ top = x" by simp have 16: "\x y z . (x \ y) \ z = (x \ z) \ (y \ z)" by (simp add: il_inf_right_dist_sup) have 17: "\x y z . (x \ y) \ (z \ y) = (x \ z) \ y" using 16 by metis have 18: "\x . bot \ x = bot" by simp have 19: "\x . - x \ - - x = top" by simp have 20: "\x . - x \ x = bot" by (simp add: a_inf_complement_bot) have 23: "\x y z . ((x \ y) \ (x \ z)) \ (x \ (y \ z)) = x \ (y \ z)" using 4 5 by metis have 24: "\x y z . (x \ (y \ z)) \ ((x \ y) \ (x \ z)) = x \ (y \ z)" using 9 23 by metis have 25: "\x y . - - x \ - (- x \ y) = - (- x \ y)" using 4 6 by metis have 26: "\x y z . x \ (y \ z) = y \ (x \ z)" using 8 9 by metis have 27: "\x y z . (x \ y) \ ((x \ z) \ (x \ (y \ z))) = x \ (y \ z)" using 9 24 26 by metis have 30: "\x . bot \ x = x" using 9 10 by metis have 31: "\x y . x \ (x \ y) = x \ y" using 8 11 by metis have 34: "\u x y z . ((x \ y) \ z) \ u = (x \ z) \ ((y \ z) \ u)" using 8 17 by metis have 35: "\u x y z . (x \ (y \ z)) \ (u \ z) = ((x \ y) \ u) \ z" using 13 17 by metis have 36: "\u x y z . (x \ y) \ (z \ (u \ y)) = (x \ (z \ u)) \ y" using 13 17 by metis have 39: "\x y . - x \ (- - x \ y) = top \ y" using 8 19 by metis have 41: "\x y . - x \ (x \ y) = bot" using 13 18 20 by metis have 42: "- top = bot" using 15 20 by metis have 43: "\x y . (- x \ y) \ x = y \ x" using 17 20 30 by metis have 44: "\x y . (x \ - y) \ y = x \ y" using 9 17 20 30 by metis have 46: "\x . - bot \ - - x = - bot" using 9 20 25 by metis have 50: "- bot = top" using 19 30 42 by metis have 51: "\x . top \ - - x = top" using 46 50 by metis have 63: "\x y . x \ ((x \ - y) \ (x \ - - y)) = x" using 9 15 19 26 27 by metis have 66: "\x y . (- (x \ y) \ x) \ (- (x \ y) \ y) = bot" using 9 20 27 30 by metis have 67: "\x y z . (x \ - - y) \ (x \ - (- y \ z)) = x \ - (- y \ z)" using 11 25 27 by metis have 70: "\x y . x \ (x \ - - y) = x" using 9 15 27 31 51 by metis have 82: "\x . top \ - x = top" using 9 19 31 by metis have 89: "\x y . x \ (- y \ x) = x" using 14 17 82 by metis have 102: "\x y z . x \ (y \ (x \ - - z)) = y \ x" using 26 70 by metis have 104: "\x y . x \ (x \ - y) = x" using 9 63 102 by metis have 112: "\x y z . (- x \ y) \ ((- - x \ y) \ z) = y \ z" using 14 19 34 by metis have 117: "\x y z . x \ ((x \ - y) \ z) = x \ z" using 8 104 by metis have 120: "\x y z . x \ (y \ (x \ - z)) = y \ x" using 26 104 by metis have 124: "\x . - - x \ x = x" using 14 19 43 by metis have 128: "\x y . - - x \ (x \ y) = x \ y" using 13 124 by metis have 131: "\x . - x \ - - - x = - x" using 9 25 124 by metis have 133: "\x . - - - x = - x" using 9 104 124 131 by metis have 135: "\x y . - x \ - (- - x \ y) = - (- - x \ y)" using 25 133 by metis have 137: "\x y . (- x \ y) \ - - x = y \ - - x" using 43 133 by metis have 145: "\x y z . ((- (x \ y) \ x) \ z) \ y = z \ y" using 20 30 35 by metis have 183: "\x y z . (x \ (- - (y \ z) \ y)) \ z = (x \ y) \ z" using 17 36 124 by metis have 289: "\x y . - x \ - (- x \ y) = top" using 25 39 82 by metis have 316: "\x y . - (- x \ y) \ x = x" using 14 43 289 by metis have 317: "\x y z . - (- x \ y) \ (x \ z) = x \ z" using 13 316 by metis have 320: "\x y . - x \ - - (- x \ y) = - x" using 9 25 316 by metis have 321: "\x y . - - (- x \ y) \ x = bot" using 41 316 by metis have 374: "\x y . - x \ - (x \ y) = - (x \ y)" using 25 128 133 by metis have 388: "\x y . - (x \ y) \ - x = - x" using 128 316 by metis have 389: "\x y . - - (x \ y) \ - x = bot" using 128 321 by metis have 405: "\x y z . - (x \ y) \ (- x \ z) = - x \ z" using 13 388 by metis have 406: "\x y z . - (x \ (y \ z)) \ - (x \ y) = - (x \ y)" using 13 388 by metis have 420: "\x y . - x \ - - (- x \ y) = - - (- x \ y)" using 316 388 by metis have 422: "\x y z . - - (x \ y) \ (- x \ z) = bot" using 13 18 389 by metis have 758: "\x y z . x \ (x \ (- y \ - z)) = x" using 13 104 117 by metis have 1092: "\x y . - (x \ y) \ x = bot" using 9 30 31 66 by metis have 1130: "\x y z . (- (x \ y) \ z) \ x = z \ x" using 17 30 1092 by metis have 1156: "\x y . - - x \ - (- x \ y) = - - x" using 67 104 124 133 by metis have 2098: "\x y . - - (x \ y) \ x = x" using 14 19 1130 by metis have 2125: "\x y . - - (x \ y) \ y = y" using 9 2098 by metis have 2138: "\x y . - x \ - - (x \ y) = top" using 9 289 2098 by metis have 2139: "\x y . - x \ - (x \ y) = - (x \ y)" using 316 2098 by metis have 2192: "\x y . - - x \ (- y \ x) = - y \ x" using 89 2125 by metis have 2202: "\x y . - x \ - - (y \ x) = top" using 9 289 2125 by metis have 2344: "\x y . - (- x \ y) \ - - y = top" using 89 2202 by metis have 2547: "\x y z . - x \ ((- - x \ - y) \ z) = - x \ (- y \ z)" using 112 117 by metis have 3023: "\x y . - x \ - (- y \ - x) = top" using 9 133 2344 by metis have 3134: "\x y . - (- x \ - y) \ y = y" using 14 43 3023 by metis have 3135: "\x y . - x \ (- y \ - x) = - y \ - x" using 14 44 3023 by metis have 3962: "\x y . - - (x \ y) \ - - x = - - x" using 14 137 2138 by metis have 5496: "\x y z . - - (x \ y) \ - (x \ z) = bot" using 422 2139 by metis have 9414: "\x y . - - (- x \ y) \ y = - x \ y" using 9 104 183 320 by metis have 9520: "\x y z . - - (- x \ y) \ - - (x \ z) = bot" using 374 5496 by metis have 11070: "\x y z . - (- - x \ y) \ (- x \ - z) = - (- - x \ y)" using 317 758 by metis have 12371: "\x y . - x \ - (- - x \ y) = - x" using 133 1156 by metis have 12377: "\x y . - x \ - (x \ y) = - x" using 128 133 1156 by metis have 12384: "\x y . - (x \ y) \ - y = - (x \ y)" using 133 1156 2125 by metis have 12394: "\x y . - - (- x \ - y) = - x \ - y" using 1156 3134 9414 by metis have 12640: "\x y . - x \ - (- y \ x) = - x" using 89 12384 by metis have 24648: "\x y . (- x \ - y) \ - (- x \ - y) = top" using 19 12394 by metis have 28270: "\x y z . - - (x \ y) \ - (- x \ z) = - (- x \ z)" using 374 405 by metis have 28339: "\x y . - (- - (x \ y) \ x) = - (x \ y)" using 124 406 12371 by metis have 28423: "\x y . - (- x \ - y) = - (- y \ - x)" using 13 3135 12394 28339 by metis have 28487: "\x y . - x \ - y = - y \ - x" using 2098 3962 12394 28423 by metis have 52423: "\x y . - (- x \ - (- x \ y)) \ y = y" using 14 145 24648 28487 by metis have 52522: "\x y . - x \ - (- x \ y) = - x \ - y" using 13 12377 12394 12640 28487 52423 by metis have 61103: "\x y z . - (- - x \ y) \ z = - x \ (- y \ z)" using 112 2547 12371 52522 by metis have 61158: "\x y . - - (- x \ y) = - x \ - - y" using 420 52522 by metis have 61231: "\x y z . - x \ (- - y \ - (x \ z)) = - x \ - - y" using 13 15 50 133 9520 52522 61158 by metis have 61313: "\x y . - x \ - y = - (- - y \ x)" using 120 11070 61103 by metis have 61393: "\x y . - (- x \ - - y) = - (- x \ y)" using 13 28270 61158 61231 61313 by metis have 61422: "\x y . - (- - x \ y) = - (- - y \ x)" using 13 135 2192 61158 61313 by metis show ?thesis using 61313 61393 61422 by metis qed lemma l12: "- x \ - y = - (x \ y)" proof - have 1: "\x y . x \ y \ x \ y = y" by (simp add: il_less_eq) have 4: "\x y . \(x \ y) \ x \ y = y" using 1 by metis have 5: "\x y z . (x \ y) \ (x \ z) \ x \ (y \ z)" by (simp add: il_sub_inf_right_isotone_var) have 6: "\x y . - - x \ - (- x \ y)" by (simp add: pad2) have 7: "\x y z . x \ (y \ z) = (x \ y) \ z" by (simp add: il_associative) have 8: "\x y z . (x \ y) \ z = x \ (y \ z)" using 7 by metis have 9: "\x y . x \ y = y \ x" by (simp add: il_commutative) have 10: "\x . x \ bot = x" by (simp add: il_bot_unit) have 11: "\x . x \ x = x" by simp have 12: "\x y z . x \ (y \ z) = (x \ y) \ z" by (simp add: il_inf_associative) have 13: "\x y z . (x \ y) \ z = x \ (y \ z)" using 12 by metis have 14: "\x . top \ x = x" by simp have 15: "\x . x \ top = x" by simp have 16: "\x y z . (x \ y) \ z = (x \ z) \ (y \ z)" by (simp add: il_inf_right_dist_sup) have 17: "\x y z . (x \ y) \ (z \ y) = (x \ z) \ y" using 16 by metis have 18: "\x . bot \ x = bot" by simp have 19: "\x . - x \ - - x = top" by simp have 20: "\x . - x \ x = bot" by (simp add: a_inf_complement_bot) have 22: "\x y z . ((x \ y) \ (x \ z)) \ (x \ (y \ z)) = x \ (y \ z)" using 4 5 by metis have 23: "\x y z . (x \ (y \ z)) \ ((x \ y) \ (x \ z)) = x \ (y \ z)" using 9 22 by metis have 24: "\x y . - - x \ - (- x \ y) = - (- x \ y)" using 4 6 by metis have 25: "\x y z . x \ (y \ z) = y \ (x \ z)" using 8 9 by metis have 26: "\x y z . (x \ y) \ ((x \ z) \ (x \ (y \ z))) = x \ (y \ z)" using 9 23 25 by metis have 29: "\x . bot \ x = x" using 9 10 by metis have 30: "\x y . x \ (x \ y) = x \ y" using 8 11 by metis have 32: "\x y . x \ (y \ x) = y \ x" using 8 9 11 by metis have 33: "\u x y z . ((x \ y) \ z) \ u = (x \ z) \ ((y \ z) \ u)" using 8 17 by metis have 34: "\u x y z . (x \ (y \ z)) \ (u \ z) = ((x \ y) \ u) \ z" using 13 17 by metis have 35: "\u x y z . (x \ y) \ (z \ (u \ y)) = (x \ (z \ u)) \ y" using 13 17 by metis have 36: "\x y . (top \ x) \ y = y \ (x \ y)" using 14 17 by metis have 37: "\x y . (x \ top) \ y = y \ (x \ y)" using 9 14 17 by metis have 38: "\x y . - x \ (- - x \ y) = top \ y" using 8 19 by metis have 40: "\x y . - x \ (x \ y) = bot" using 13 18 20 by metis have 41: "- top = bot" using 15 20 by metis have 42: "\x y . (- x \ y) \ x = y \ x" using 17 20 29 by metis have 43: "\x y . (x \ - y) \ y = x \ y" using 9 17 20 29 by metis have 45: "\x . - bot \ - - x = - bot" using 9 20 24 by metis have 46: "\u x y z . (x \ y) \ (z \ (u \ y)) = z \ ((x \ u) \ y)" using 17 25 by metis have 47: "\x y . - x \ (y \ - - x) = y \ top" using 19 25 by metis have 49: "- bot = top" using 19 29 41 by metis have 50: "\x . top \ - - x = top" using 45 49 by metis have 54: "\u x y z . (x \ y) \ ((x \ z) \ ((x \ (y \ z)) \ u)) = (x \ (y \ z)) \ u" using 8 26 by metis have 58: "\u x y z . (x \ (y \ z)) \ ((x \ (y \ u)) \ (x \ (y \ (z \ u)))) = x \ (y \ (z \ u))" using 13 26 by metis have 60: "\x y . x \ ((x \ y) \ (x \ (y \ top))) = x \ (y \ top)" using 15 25 26 by metis have 62: "\x y . x \ ((x \ - y) \ (x \ - - y)) = x" using 9 15 19 25 26 by metis have 65: "\x y . (- (x \ y) \ x) \ (- (x \ y) \ y) = bot" using 9 20 26 29 by metis have 66: "\x y z . (x \ - - y) \ (x \ - (- y \ z)) = x \ - (- y \ z)" using 11 24 26 by metis have 69: "\x y . x \ (x \ - - y) = x" using 9 15 26 30 50 by metis have 81: "\x . top \ - x = top" using 9 19 30 by metis have 82: "\x y z . (x \ y) \ (x \ (y \ z)) = x \ (y \ z)" using 11 26 30 by metis have 83: "\x y . x \ (x \ (y \ top)) = x \ (y \ top)" using 60 82 by metis have 88: "\x y . x \ (- y \ x) = x" using 14 17 81 by metis have 89: "\x y . top \ (x \ - y) = x \ top" using 25 81 by metis have 91: "\x y z . x \ (y \ (z \ x)) = y \ (z \ x)" using 8 32 by metis have 94: "\x y z . x \ (y \ (- z \ x)) = y \ x" using 25 88 by metis have 101: "\x y z . x \ (y \ (x \ - - z)) = y \ x" using 25 69 by metis have 102: "\x . x \ (x \ bot) = x" using 41 49 69 by metis have 103: "\x y . x \ (x \ - y) = x" using 9 62 101 by metis have 109: "\x y . x \ (y \ (x \ bot)) = y \ x" using 25 102 by metis have 111: "\x y z . (- x \ y) \ ((- - x \ y) \ z) = y \ z" using 14 19 33 by metis have 116: "\x y z . x \ ((x \ - y) \ z) = x \ z" using 8 103 by metis have 119: "\x y z . x \ (y \ (x \ - z)) = y \ x" using 25 103 by metis have 123: "\x . - - x \ x = x" using 14 19 42 by metis have 127: "\x y . - - x \ (x \ y) = x \ y" using 13 123 by metis have 130: "\x . - x \ - - - x = - x" using 9 24 123 by metis have 132: "\x . - - - x = - x" using 9 103 123 130 by metis have 134: "\x y . - x \ - (- - x \ y) = - (- - x \ y)" using 24 132 by metis have 136: "\x y . (- x \ y) \ - - x = y \ - - x" using 42 132 by metis have 138: "\x . - x \ - x = - x" using 123 132 by metis have 144: "\x y z . ((- (x \ y) \ x) \ z) \ y = z \ y" using 20 29 34 by metis have 157: "\x y . (- x \ y) \ - x = (top \ y) \ - x" using 17 36 138 by metis have 182: "\x y z . (x \ (- - (y \ z) \ y)) \ z = (x \ y) \ z" using 17 35 123 by metis have 288: "\x y . - x \ - (- x \ y) = top" using 24 38 81 by metis have 315: "\x y . - (- x \ y) \ x = x" using 14 42 288 by metis have 316: "\x y z . - (- x \ y) \ (x \ z) = x \ z" using 13 315 by metis have 319: "\x y . - x \ - - (- x \ y) = - x" using 9 24 315 by metis have 320: "\x y . - - (- x \ y) \ x = bot" using 40 315 by metis have 373: "\x y . - x \ - (x \ y) = - (x \ y)" using 24 127 132 by metis have 387: "\x y . - (x \ y) \ - x = - x" using 127 315 by metis have 388: "\x y . - - (x \ y) \ - x = bot" using 127 320 by metis have 404: "\x y z . - (x \ y) \ (- x \ z) = - x \ z" using 13 387 by metis have 405: "\x y z . - (x \ (y \ z)) \ - (x \ y) = - (x \ y)" using 13 387 by metis have 419: "\x y . - x \ - - (- x \ y) = - - (- x \ y)" using 315 387 by metis have 420: "\x y . - - x \ - - (x \ y) = - - (x \ y)" using 387 by metis have 421: "\x y z . - - (x \ y) \ (- x \ z) = bot" using 13 18 388 by metis have 536: "\x y . (x \ - - y) \ y = (x \ top) \ y" using 42 47 by metis have 662: "\u x y z . (x \ y) \ ((x \ (z \ y)) \ u) = (x \ (z \ y)) \ u" using 9 32 54 by metis have 705: "\u x y z . (x \ (y \ z)) \ ((x \ (y \ (z \ bot))) \ u) = (x \ (y \ z)) \ u" using 25 54 109 662 by metis have 755: "\x y z . (x \ - y) \ (z \ x) = z \ x" using 32 91 116 by metis have 757: "\x y z . x \ (x \ (- y \ - z)) = x" using 13 103 116 by metis have 930: "\x y z . (- (x \ (y \ z)) \ (x \ y)) \ (- (x \ (y \ z)) \ (x \ z)) = bot" using 9 20 29 58 by metis have 1091: "\x y . - (x \ y) \ x = bot" using 9 29 30 65 by metis have 1092: "\x y . - (x \ y) \ y = bot" using 29 30 65 1091 by metis have 1113: "\u x y z . - (x \ ((y \ z) \ u)) \ (x \ (z \ u)) = bot" using 29 46 65 1091 by metis have 1117: "\x y z . - (x \ y) \ (x \ (- z \ y)) = bot" using 29 65 94 1092 by metis have 1128: "\x y z . - (x \ (y \ z)) \ (x \ y) = bot" using 8 1091 by metis have 1129: "\x y z . (- (x \ y) \ z) \ x = z \ x" using 17 29 1091 by metis have 1155: "\x y . - - x \ - (- x \ y) = - - x" using 66 103 123 132 by metis have 1578: "\x y z . - (x \ (y \ z)) \ (x \ y) = bot" using 82 1091 by metis have 1594: "\x y z . - (x \ (y \ z)) \ (x \ z) = bot" using 29 930 1578 by metis have 2094: "\x y z . - (x \ (y \ (z \ top))) \ (x \ y) = bot" using 83 1128 by metis have 2097: "\x y . - - (x \ y) \ x = x" using 14 19 1129 by metis have 2124: "\x y . - - (x \ y) \ y = y" using 9 2097 by metis have 2135: "\x y . - - ((top \ x) \ y) \ y = y" using 36 2097 by metis have 2136: "\x y . - - ((x \ top) \ y) \ y = y" using 37 2097 by metis have 2137: "\x y . - x \ - - (x \ y) = top" using 9 288 2097 by metis have 2138: "\x y . - x \ - (x \ y) = - (x \ y)" using 315 2097 by metis have 2151: "\x y . - x \ - (x \ y) = - x" using 9 132 373 2097 by metis have 2191: "\x y . - - x \ (- y \ x) = - y \ x" using 88 2124 by metis have 2201: "\x y . - x \ - - (y \ x) = top" using 9 288 2124 by metis have 2202: "\x y . - x \ - (y \ x) = - (y \ x)" using 315 2124 by metis have 2320: "\x y . - (x \ (y \ top)) = - x" using 83 373 2151 by metis have 2343: "\x y . - (- x \ y) \ - - y = top" using 88 2201 by metis have 2546: "\x y z . - x \ ((- - x \ - y) \ z) = - x \ (- y \ z)" using 111 116 by metis have 2706: "\x y z . - x \ (y \ - - ((top \ z) \ - x)) = y \ - - ((top \ z) \ - x)" using 755 2135 by metis have 2810: "\x y . - x \ - ((y \ top) \ x) = - ((y \ top) \ x)" using 315 2136 by metis have 3022: "\x y . - x \ - (- y \ - x) = top" using 9 132 2343 by metis have 3133: "\x y . - (- x \ - y) \ y = y" using 14 42 3022 by metis have 3134: "\x y . - x \ (- y \ - x) = - y \ - x" using 14 43 3022 by metis have 3961: "\x y . - - (x \ y) \ - - x = - - x" using 14 136 2137 by metis have 4644: "\x y z . - (x \ - y) \ (x \ - (y \ z)) = bot" using 1594 2151 by metis have 5495: "\x y z . - - (x \ y) \ - (x \ z) = bot" using 421 2138 by metis have 9413: "\x y . - - (- x \ y) \ y = - x \ y" using 9 103 182 319 by metis have 9519: "\x y z . - - (- x \ y) \ - - (x \ z) = bot" using 373 5495 by metis have 11069: "\x y z . - (- - x \ y) \ (- x \ - z) = - (- - x \ y)" using 316 757 by metis have 12370: "\x y . - x \ - (- - x \ y) = - x" using 132 1155 by metis have 12376: "\x y . - x \ - (x \ y) = - x" using 127 132 1155 by metis have 12383: "\x y . - (x \ y) \ - y = - (x \ y)" using 132 1155 2124 by metis have 12393: "\x y . - - (- x \ - y) = - x \ - y" using 1155 3133 9413 by metis have 12407: "\x y . - - x \ - - (x \ y) = - - x" using 1155 2138 by metis have 12639: "\x y . - x \ - (- y \ x) = - x" using 88 12383 by metis have 24647: "\x y . (- x \ - y) \ - (- x \ - y) = top" using 19 12393 by metis have 28269: "\x y z . - - (x \ y) \ - (- x \ z) = - (- x \ z)" using 373 404 by metis have 28338: "\x y . - (- - (x \ y) \ x) = - (x \ y)" using 123 405 12370 by metis have 28422: "\x y . - (- x \ - y) = - (- y \ - x)" using 13 3134 12393 28338 by metis have 28485: "\x y . - x \ - y = - y \ - x" using 2097 3961 12393 28422 by metis have 30411: "\x y . - x \ (x \ (x \ y)) = bot" using 9 82 2094 2320 by metis have 30469: "\x . - x \ (x \ - - x) = bot" using 9 123 132 30411 by metis have 37513: "\x y . - (- x \ - y) \ - (y \ x) = bot" using 2202 4644 by metis have 52421: "\x y . - (- x \ - (- x \ y)) \ y = y" using 14 144 24647 28485 by metis have 52520: "\x y . - x \ - (- x \ y) = - x \ - y" using 13 12376 12393 12639 28485 52421 by metis have 52533: "\x y z . - - (x \ (y \ (z \ top))) \ (x \ y) = x \ y" using 15 49 2094 52421 by metis have 61101: "\x y z . - (- - x \ y) \ z = - x \ (- y \ z)" using 111 2546 12370 52520 by metis have 61156: "\x y . - - (- x \ y) = - x \ - - y" using 419 52520 by metis have 61162: "\x y . - (x \ (x \ y)) = - x" using 15 49 2138 30411 52520 by metis have 61163: "\x . - (x \ - - x) = - x" using 15 49 2138 30469 52520 by metis have 61229: "\x y z . - x \ (- - y \ - (x \ z)) = - x \ - - y" using 13 15 49 132 9519 52520 61156 by metis have 61311: "\x y . - x \ - y = - (- - y \ x)" using 119 11069 61101 by metis have 61391: "\x y . - (- x \ - - y) = - (- x \ y)" using 13 28269 61156 61229 61311 by metis have 61420: "\x y . - (- - x \ y) = - (- - y \ x)" using 13 134 2191 61156 61311 by metis have 61454: "\x y . - (x \ - (- y \ - x)) = - y \ - x" using 9 132 3133 61156 61162 by metis have 61648: "\x y . - x \ (x \ (- y \ - - x)) = bot" using 1117 61163 by metis have 62434: "\x y . - (- - x \ y) \ x = - y \ x" using 43 61311 by metis have 63947: "\x y . - (- x \ y) \ - (- y \ x) = bot" using 37513 61391 by metis have 64227: "\x y . - (x \ (- y \ - - x)) = - x" using 15 49 2138 52520 61648 by metis have 64239: "\x y . - (x \ (- - x \ y)) = - (x \ y)" using 9 25 12407 64227 by metis have 64241: "\x y . - (x \ (- - x \ - y)) = - x" using 28485 64227 by metis have 64260: "\x y . - (x \ - - (x \ y)) = - x" using 420 64241 by metis have 64271: "\x y . - (- x \ (y \ - - (y \ x))) = - (- x \ y)" using 9 25 42 64260 by metis have 64281: "\x y . - (- x \ y) = - (y \ - - ((top \ y) \ - x))" using 9 25 157 2706 64260 by metis have 64282: "\x y . - (x \ - - ((x \ top) \ y)) = - (x \ - - y)" using 9 25 132 536 2810 28485 61311 64260 by metis have 65110: "\x y . - ((- x \ y) \ (- y \ x)) = bot" using 9 14 49 37513 63947 by metis have 65231: "\x y . - (x \ ((- x \ y) \ - y)) = bot" using 9 25 65110 by metis have 65585: "\x y . - (x \ - y) = - - y \ - x" using 61311 61454 64239 by metis have 65615: "\x y . - x \ - ((x \ top) \ y) = - y \ - x" using 132 28485 64282 65585 by metis have 65616: "\x y . - (- x \ y) = - y \ - ((top \ y) \ - x)" using 132 28485 64281 65585 by metis have 65791: "\x y . - x \ - ((top \ x) \ - y) = - - y \ - x" using 89 132 12376 28485 64271 65585 65615 65616 by metis have 65933: "\x y . - (- x \ y) = - - x \ - y" using 65616 65791 by metis have 66082: "\x y z . - (x \ (y \ - z)) = - - z \ - (x \ y)" using 8 65585 by metis have 66204: "\x y . - - x \ - (y \ (- y \ x)) = bot" using 65231 66082 by metis have 66281: "\x y z . - (x \ (- y \ z)) = - - y \ - (x \ z)" using 25 65933 by metis have 67527: "\x y . - - (x \ (- x \ y)) \ y = y" using 14 49 62434 66204 by metis have 67762: "\x y . - (- - x \ (y \ (- y \ x))) = - x" using 61420 67527 by metis have 68018: "\x y z . - (x \ y) \ (x \ (y \ (z \ top))) = bot" using 8 83 1113 2320 by metis have 71989: "\x y z . - (x \ (y \ (z \ top))) = - (x \ y)" using 9 29 52533 67762 68018 by metis have 71997: "\x y z . - ((x \ (y \ top)) \ z) = - (x \ z)" using 17 2320 71989 by metis have 72090: "\x y z . - (x \ ((x \ y) \ z)) = - (x \ z)" using 10 14 705 71997 by metis have 72139: "\x y . - (x \ y) = - x \ - y" using 25 123 132 2138 65933 66281 72090 by metis show ?thesis using 72139 by metis qed lemma l15: "--(x \ y) = --x \ --y" by (simp add: l11 l12 l4) lemma l13_var: "- - (- x \ y) = - x \ - - y" proof - have 1: "\x y . x \ y \ x \ y = y" by (simp add: il_less_eq) have 4: "\x y . \(x \ y) \ x \ y = y" using 1 by metis have 5: "\x y z . (x \ y) \ (x \ z) \ x \ (y \ z)" by (simp add: il_sub_inf_right_isotone_var) have 6: "\x y . - - x \ - (- x \ y)" by (simp add: pad2) have 7: "\x y z . x \ (y \ z) = (x \ y) \ z" by (simp add: il_associative) have 8: "\x y z . (x \ y) \ z = x \ (y \ z)" using 7 by metis have 9: "\x y . x \ y = y \ x" by (simp add: il_commutative) have 10: "\x . x \ bot = x" by (simp add: il_bot_unit) have 11: "\x . x \ x = x" by simp have 12: "\x y z . x \ (y \ z) = (x \ y) \ z" by (simp add: il_inf_associative) have 13: "\x y z . (x \ y) \ z = x \ (y \ z)" using 12 by metis have 14: "\x . top \ x = x" by simp have 15: "\x . x \ top = x" by simp have 16: "\x y z . (x \ y) \ z = (x \ z) \ (y \ z)" by (simp add: il_inf_right_dist_sup) have 17: "\x y z . (x \ y) \ (z \ y) = (x \ z) \ y" using 16 by metis have 19: "\x . - x \ - - x = top" by simp have 20: "\x . - x \ x = bot" by (simp add: a_inf_complement_bot) have 22: "\x y z . ((x \ y) \ (x \ z)) \ (x \ (y \ z)) = x \ (y \ z)" using 4 5 by metis have 23: "\x y z . (x \ (y \ z)) \ ((x \ y) \ (x \ z)) = x \ (y \ z)" using 9 22 by metis have 24: "\x y . - - x \ - (- x \ y) = - (- x \ y)" using 4 6 by metis have 25: "\x y z . x \ (y \ z) = y \ (x \ z)" using 8 9 by metis have 26: "\x y z . (x \ y) \ ((x \ z) \ (x \ (y \ z))) = x \ (y \ z)" using 9 23 25 by metis have 29: "\x . bot \ x = x" using 9 10 by metis have 30: "\x y . x \ (x \ y) = x \ y" using 8 11 by metis have 34: "\u x y z . (x \ (y \ z)) \ (u \ z) = ((x \ y) \ u) \ z" using 13 17 by metis have 35: "\u x y z . (x \ y) \ (z \ (u \ y)) = (x \ (z \ u)) \ y" using 13 17 by metis have 38: "\x y . - x \ (- - x \ y) = top \ y" using 8 19 by metis have 41: "- top = bot" using 15 20 by metis have 42: "\x y . (- x \ y) \ x = y \ x" using 17 20 29 by metis have 43: "\x y . (x \ - y) \ y = x \ y" using 9 17 20 29 by metis have 45: "\x . - bot \ - - x = - bot" using 9 20 24 by metis have 49: "- bot = top" using 19 29 41 by metis have 50: "\x . top \ - - x = top" using 45 49 by metis have 62: "\x y . x \ ((x \ - y) \ (x \ - - y)) = x" using 9 15 19 25 26 by metis have 65: "\x y . (- (x \ y) \ x) \ (- (x \ y) \ y) = bot" using 9 20 26 29 by metis have 66: "\x y z . (x \ - - y) \ (x \ - (- y \ z)) = x \ - (- y \ z)" using 11 24 26 by metis have 69: "\x y . x \ (x \ - - y) = x" using 9 15 26 30 50 by metis have 81: "\x . top \ - x = top" using 9 19 30 by metis have 88: "\x y . x \ (- y \ x) = x" using 14 17 81 by metis have 101: "\x y z . x \ (y \ (x \ - - z)) = y \ x" using 25 69 by metis have 103: "\x y . x \ (x \ - y) = x" using 9 62 101 by metis have 123: "\x . - - x \ x = x" using 14 19 42 by metis have 127: "\x y . - - x \ (x \ y) = x \ y" using 13 123 by metis have 130: "\x . - x \ - - - x = - x" using 9 24 123 by metis have 132: "\x . - - - x = - x" using 9 103 123 130 by metis have 136: "\x y . (- x \ y) \ - - x = y \ - - x" using 42 132 by metis have 144: "\x y z . ((- (x \ y) \ x) \ z) \ y = z \ y" using 20 29 34 by metis have 182: "\x y z . (x \ (- - (y \ z) \ y)) \ z = (x \ y) \ z" using 17 35 123 by metis have 288: "\x y . - x \ - (- x \ y) = top" using 24 38 81 by metis have 315: "\x y . - (- x \ y) \ x = x" using 14 42 288 by metis have 319: "\x y . - x \ - - (- x \ y) = - x" using 9 24 315 by metis have 387: "\x y . - (x \ y) \ - x = - x" using 127 315 by metis have 405: "\x y z . - (x \ (y \ z)) \ - (x \ y) = - (x \ y)" using 13 387 by metis have 419: "\x y . - x \ - - (- x \ y) = - - (- x \ y)" using 315 387 by metis have 1091: "\x y . - (x \ y) \ x = bot" using 9 29 30 65 by metis have 1129: "\x y z . (- (x \ y) \ z) \ x = z \ x" using 17 29 1091 by metis have 1155: "\x y . - - x \ - (- x \ y) = - - x" using 66 103 123 132 by metis have 2097: "\x y . - - (x \ y) \ x = x" using 14 19 1129 by metis have 2124: "\x y . - - (x \ y) \ y = y" using 9 2097 by metis have 2137: "\x y . - x \ - - (x \ y) = top" using 9 288 2097 by metis have 2201: "\x y . - x \ - - (y \ x) = top" using 9 288 2124 by metis have 2343: "\x y . - (- x \ y) \ - - y = top" using 88 2201 by metis have 3022: "\x y . - x \ - (- y \ - x) = top" using 9 132 2343 by metis have 3133: "\x y . - (- x \ - y) \ y = y" using 14 42 3022 by metis have 3134: "\x y . - x \ (- y \ - x) = - y \ - x" using 14 43 3022 by metis have 3961: "\x y . - - (x \ y) \ - - x = - - x" using 14 136 2137 by metis have 9413: "\x y . - - (- x \ y) \ y = - x \ y" using 9 103 182 319 by metis have 12370: "\x y . - x \ - (- - x \ y) = - x" using 132 1155 by metis have 12376: "\x y . - x \ - (x \ y) = - x" using 127 132 1155 by metis have 12383: "\x y . - (x \ y) \ - y = - (x \ y)" using 132 1155 2124 by metis have 12393: "\x y . - - (- x \ - y) = - x \ - y" using 1155 3133 9413 by metis have 12639: "\x y . - x \ - (- y \ x) = - x" using 88 12383 by metis have 24647: "\x y . (- x \ - y) \ - (- x \ - y) = top" using 19 12393 by metis have 28338: "\x y . - (- - (x \ y) \ x) = - (x \ y)" using 123 405 12370 by metis have 28422: "\x y . - (- x \ - y) = - (- y \ - x)" using 13 3134 12393 28338 by metis have 28485: "\x y . - x \ - y = - y \ - x" using 2097 3961 12393 28422 by metis have 52421: "\x y . - (- x \ - (- x \ y)) \ y = y" using 14 144 24647 28485 by metis have 52520: "\x y . - x \ - (- x \ y) = - x \ - y" using 13 12376 12393 12639 28485 52421 by metis have 61156: "\x y . - - (- x \ y) = - x \ - - y" using 419 52520 by metis show ?thesis using 61156 by metis qed text \Theorem 25.1\ subclass subset_boolean_algebra_2 proof show "\x y z. x \ (y \ z) = x \ y \ z" by (simp add: il_associative) show "\x y. x \ y = y \ x" by (simp add: il_commutative) show "\x. x \ x = x" by simp show "\x y. x \ - (y \ - y) = x" using il_bot_unit l12 l6 by auto show "\x y. - (x \ y) = - (- - x \ - - y)" by (metis l15 l4) show "\x y. - x \ - (- x \ y) = - x \ - y" by (smt l11 l15 il_inf_right_dist_sup il_unit_bot l6 l7) qed lemma aa_test: "p = --p \ test p" by (metis ppa_ppd.d_closed) lemma test_aa_increasing: "test p \ p \ --p" by (simp add: ppa_ppd.d_increasing_sub_identity test_sub_identity) lemma "test p \ - - (p \ x) \ p" nitpick [expect=genuine] oops lemma "test p \ --p \ p" nitpick [expect=genuine] oops end class pa_algebra = pa_semiring + minus + assumes pa_minus_def: "-x - -y = -(--x \ -y)" begin subclass subset_boolean_algebra_2_extended proof show "bot = (THE x. \z. x = - (z \ - z))" using l12 l6 by auto thus "top = - (THE x. \z. x = - (z \ - z))" using l2 by blast show "\x y. - x \ - y = - (- - x \ - - y)" by (metis l12 l4) show "\x y. - x - - y = - (- - x \ - y)" by (simp add: pa_minus_def) show "\x y. (x \ y) = (x \ y = y)" by (simp add: il_less_eq) show "\x y. (x < y) = (x \ y = y \ y \ x \ x)" by (simp add: il_less_eq less_le_not_le) qed lemma "\x y. - (x \ - - y) = - (x \ y)" nitpick [expect=genuine] oops end subsection \Antidomain Semirings\ text \Definition 24\ class a_semiring = ppa_semiring + assumes ad3: "-(x \ y) \ -(x \ --y)" begin lemma l16: "- - x \ - (- x \ y)" proof - have 1: "\x y . x \ y \ x \ y = y" by (simp add: il_less_eq) have 3: "\x y z . x \ (y \ z) = (x \ y) \ z" by (simp add: il_associative) have 4: "\x y z . (x \ y) \ z = x \ (y \ z)" using 3 by metis have 5: "\x y . x \ y = y \ x" by (simp add: il_commutative) have 6: "\x . x \ bot = x" by (simp add: il_bot_unit) have 7: "\x . x \ x = x" by simp have 8: "\x y . \(x \ y) \ x \ y = y" using 1 by metis have 9: "\x y . x \ y \ x \ y \ y" using 1 by metis have 10: "\x y z . x \ (y \ z) = (x \ y) \ z" by (simp add: il_inf_associative) have 11: "\x y z . (x \ y) \ z = x \ (y \ z)" using 10 by metis have 12: "\x . top \ x = x" by simp have 13: "\x . x \ top = x" by simp have 14: "\x y z . (x \ y) \ (x \ z) \ x \ (y \ z)" by (simp add: il_sub_inf_right_isotone_var) have 15: "\x y z . (x \ y) \ z = (x \ z) \ (y \ z)" by (simp add: il_inf_right_dist_sup) have 16: "\x y z . (x \ y) \ (z \ y) = (x \ z) \ y" using 15 by metis have 17: "\x . bot \ x = bot" by simp have 18: "\x . - x \ - - x = top" by simp have 19: "\x . - x \ x = bot" by (simp add: a_inf_complement_bot) have 20: "\x y . - (x \ y) \ - (x \ - - y)" by (simp add: ad3) have 22: "\x y z . x \ (y \ z) = y \ (x \ z)" using 4 5 by metis have 25: "\x . bot \ x = x" using 5 6 by metis have 26: "\x y . x \ (x \ y) = x \ y" using 4 7 by metis have 33: "\x y z . (x \ y) \ ((x \ z) \ (x \ (y \ z))) = x \ (y \ z)" using 5 8 14 22 by metis have 47: "\x y . - x \ (- - x \ y) = top \ y" using 4 18 by metis have 48: "\x y . - - x \ (y \ - x) = y \ top" using 4 5 18 by metis have 51: "\x y . - x \ (x \ y) = bot" using 11 17 19 by metis have 52: "- top = bot" using 13 19 by metis have 56: "\x y . (- x \ y) \ x = y \ x" using 16 19 25 by metis have 57: "\x y . (x \ - y) \ y = x \ y" using 5 16 19 25 by metis have 58: "\x y . - (x \ y) \ - (x \ - - y) = - (x \ - - y)" using 8 20 by metis have 60: "\x . - x \ - - - x" using 12 20 by metis have 69: "- bot = top" using 18 25 52 by metis have 74: "\x y . x \ x \ y" using 9 26 by metis have 78: "\x . top \ - x = top" using 5 18 26 by metis have 80: "\x y . x \ y \ x" using 5 74 by metis have 86: "\x y z . x \ y \ x \ (z \ y)" using 22 80 by metis have 95: "\x . - x \ - - - x = - - - x" using 8 60 by metis have 143: "\x y . x \ (x \ - y) = x" using 5 13 26 33 78 by metis have 370: "\x y z . x \ (y \ - z) \ x \ y" using 86 143 by metis have 907: "\x . - x \ - x = - x" using 12 18 57 by metis have 928: "\x y . - x \ (- x \ y) = - x \ y" using 11 907 by metis have 966: "\x y . - (- x \ - - (x \ y)) = top" using 51 58 69 78 by metis have 1535: "\x . - x \ - - - - x = top" using 47 78 95 by metis have 1630: "\x y z . (x \ y) \ - z \ (x \ - z) \ y" using 16 370 by metis have 2422: "\x . - x \ - - - x = - - - x" using 12 57 1535 by metis have 6567: "\x y . - x \ - - (x \ y) = bot" using 12 19 966 by metis have 18123: "\x . - - - x = - x" using 95 143 2422 by metis have 26264: "\x y . - x \ (- y \ - x) \ - - y" using 12 18 1630 by metis have 26279: "\x y . - - (x \ y) \ - - x" using 25 6567 26264 by metis have 26307: "\x y . - - (- x \ y) \ - x" using 928 18123 26279 by metis have 26339: "\x y . - x \ - - (- x \ y) = - x" using 5 8 26307 by metis have 26564: "\x y . - x \ - (- x \ y) = top" using 5 48 78 18123 26339 by metis have 26682: "\x y . - (- x \ y) \ x = x" using 12 56 26564 by metis have 26864: "\x y . - - x \ - (- x \ y)" using 18123 26279 26682 by metis show ?thesis using 26864 by metis qed text \Theorem 25.2\ subclass pa_semiring proof show "\x y. - - x \ - (- x \ y)" by (rule l16) qed lemma l17: "-(x \ y) = -(x \ --y)" by (simp add: ad3 order.antisym l14) lemma a_complement_inf_double_complement: "-(x \ --y) = -(x \ y)" using l17 by auto sublocale a_d: d_semiring_var where d = "\x . --x" proof show "\x y. - - (x \ - - y) \ - - (x \ y)" using l17 by auto show "- - bot = bot" by (simp add: l1 l2) qed lemma "test p \ - - (p \ x) \ p" by (fact a_d.d2) end class a_algebra = a_semiring + minus + assumes a_minus_def: "-x - -y = -(--x \ -y)" begin subclass pa_algebra proof show "\x y. - x - - y = - (- - x \ - y)" by (simp add: a_minus_def) qed text \Theorem 25.4\ subclass subset_boolean_algebra_4_extended proof show "\x y z. x \ (y \ z) = x \ y \ z" by (simp add: il_inf_associative) show "\x y z. (x \ y) \ z = x \ z \ y \ z" by (simp add: il_inf_right_dist_sup) show "\x. - x \ x = bot" by (simp add: a_inf_complement_bot) show "\x. top \ x = x" by simp show "\x y. - (x \ - - y) = - (x \ y)" using l17 by auto show "\x. x \ top = x" by simp show "\x y z. x \ y \ z \ x \ z \ y" by (simp add: il_sub_inf_right_isotone) qed end context subset_boolean_algebra_4_extended begin subclass il_semiring proof show "\x y z. x \ (y \ z) = x \ y \ z" by (simp add: sup_assoc) show "\x y. x \ y = y \ x" by (simp add: sup_commute) show "\x. x \ x = x" by simp show "\x. x \ bot = x" by simp show "\x y z. x \ (y \ z) = x \ y \ z" by (simp add: sba3_inf_associative) show "\x y z. (x \ y) \ z = x \ z \ y \ z" by (simp add: sba3_inf_right_dist_sup) show "\x. top \ x = x" by simp show "\x. x \ top = x" by simp show "\x. bot \ x = bot" by (simp add: inf_left_zero) show "\x y z. x \ y \ z \ x \ z \ y" by (simp add: inf_right_isotone) show "\x y. (x \ y) = (x \ y = y)" by (simp add: le_iff_sup) show "\x y. (x < y) = (x \ y \ \ y \ x)" by (simp add: less_le_not_le) qed subclass a_semiring proof show "\x. - x \ x = bot" by (simp add: sba3_inf_complement_bot) show "\x. - x \ - - x = top" by simp show "\x y. - (x \ y) \ - (x \ - - y)" by (simp add: sba3_complement_inf_double_complement) qed sublocale sba4_a: a_algebra proof show "\x y. - x - - y = - (- - x \ - y)" by (simp add: sub_minus_def) qed end context stone_algebra begin text \Theorem 25.3\ subclass il_semiring proof show "\x y z. x \ (y \ z) = x \ y \ z" by (simp add: sup_assoc) show "\x y. x \ y = y \ x" by (simp add: sup_commute) show "\x. x \ x = x" by simp show "\x. x \ bot = x" by simp show "\x y z. x \ (y \ z) = x \ y \ z" by (simp add: inf.sup_monoid.add_assoc) show "\x y z. (x \ y) \ z = x \ z \ y \ z" by (simp add: inf_sup_distrib2) show "\x. top \ x = x" by simp show "\x. x \ top = x" by simp show "\x. bot \ x = bot" by simp show "\x y z. x \ y \ z \ x \ z \ y" using inf.sup_right_isotone by blast show "\x y. (x \ y) = (x \ y = y)" by (simp add: le_iff_sup) show "\x y. (x < y) = (x \ y \ \ y \ x)" by (simp add: less_le_not_le) qed subclass a_semiring proof show "\x. - x \ x = bot" by simp show "\x. - x \ - - x = top" by simp show "\x y. - (x \ y) \ - (x \ - - y)" by simp qed end end diff --git a/thys/Transformer_Semantics/Kleisli_Transformers.thy b/thys/Transformer_Semantics/Kleisli_Transformers.thy --- a/thys/Transformer_Semantics/Kleisli_Transformers.thy +++ b/thys/Transformer_Semantics/Kleisli_Transformers.thy @@ -1,859 +1,859 @@ (* Title: State Transformers and Predicate Transformers Based on the Powerset Monad Author: Georg Struth Maintainer: Georg Struth *) section \State Transformers and Predicate Transformers Based on the Powerset Monad\ theory Kleisli_Transformers imports Powerset_Monad Sup_Inf_Preserving_Transformers begin subsection \Backward Diamonds from Kleisli Arrows\ text \First I verify the embedding of the Kleisli category of the powerset functor into its Eilenberg-Moore category. This functor maps sets to their mus and functions to their Kleisli liftings. But this is just functoriality of dagger!. I model it as a backward diamond operator in the sense of dynamic logic. It corresponds to a strongest postcondition operator. In the parlance of program semantics, this is an embedding of state into prediate transformers.\ notation klift ("bd\<^sub>\") text \bd stands for backward diamond, the index indicates the setting of Kleisli arrows or nondeterministic functions. ifbd is its inverse.\ abbreviation ifbd :: "('a set \ 'b set) \ 'a \ 'b set" ("bd\<^sup>-\<^sub>\") where "bd\<^sup>-\<^sub>\ \ (\\. \ \ \)" lemma fbd_set: "bd\<^sub>\ f X = {y. \x. y \ f x \ x \ X}" by (force simp: klift_prop) lemma ifbd_set: "bd\<^sup>-\<^sub>\ \ x = {y. y \ \ {x}}" by simp text \The two functors form a bijective pair.\ lemma fbd_ifbd_inv2: "Sup_pres \ \ (bd\<^sub>\ \ bd\<^sup>-\<^sub>\) \ = \" proof - assume h: "Sup_pres \" have "(bd\<^sub>\ \ bd\<^sup>-\<^sub>\) \ = Sup \ \

(\ \ \)" unfolding klift_def by simp also have "... = Sup \ \

\ \ \

\" by (simp add: comp_assoc P_func1) also have "... = \ \ Sup \ \

\" by (simp add: h) also have "... = \ \ id" by force finally show ?thesis by simp qed lemma fbd_ifbd_inv2_inv: "(bd\<^sub>\ \ bd\<^sup>-\<^sub>\) \ = \ \ Sup_pres \" unfolding fun_eq_iff comp_def by (metis (no_types, lifting) Inf.INF_cong UN_extend_simps(8) klift_prop) lemma fbd_ifbd_inv2_iff: "((bd\<^sub>\ \ bd\<^sup>-\<^sub>\) \ = \) = (Sup_pres \)" using fbd_ifbd_inv2 fbd_ifbd_inv2_inv by force lemma fbd_inj: "inj bd\<^sub>\" by (meson inj_on_inverseI klift_eta_inv1) lemma fbd_inj_iff: "(bd\<^sub>\ f = bd\<^sub>\ g) = (f = g)" by (meson injD fbd_inj) lemma ifbd_inj: "Sup_pres \ \ Sup_pres \ \ bd\<^sup>-\<^sub>\ \ = bd\<^sup>-\<^sub>\ \ \ \ = \" proof- assume h1: "Sup_pres \" and h2: "Sup_pres \" and "bd\<^sup>-\<^sub>\ \ = bd\<^sup>-\<^sub>\ \" hence "(bd\<^sub>\ \ bd\<^sup>-\<^sub>\) \ = (bd\<^sub>\ \ bd\<^sup>-\<^sub>\) \" by simp thus ?thesis by (metis h1 h2 fbd_ifbd_inv2) qed lemma ifbd_inj_iff: "Sup_pres \ \ Sup_pres \ \ (bd\<^sup>-\<^sub>\ \ = bd\<^sup>-\<^sub>\ \) = (\ = \)" using ifbd_inj by force lemma fbd_ifbd_galois: "Sup_pres \ \ (bd\<^sup>-\<^sub>\ \ = f) = (bd\<^sub>\ f = \)" using fbd_ifbd_inv2 by force lemma fbd_surj: "Sup_pres \ \ (\f. bd\<^sub>\ f = \)" using fbd_ifbd_inv2 by auto lemma ifbd_surj: "surj bd\<^sup>-\<^sub>\" unfolding surj_def by (metis klift_eta_inv1) text \In addition they preserve the Sup-quantale structure of the powerset algebra. This means that morphisms preserve compositions, units and Sups, but not Infs, hence also bottom but not top.\ lemma fbd_comp_pres: "bd\<^sub>\ (f \\<^sub>K g) = bd\<^sub>\ g \ bd\<^sub>\ f" unfolding kcomp_klift klift_prop1 by simp lemma fbd_Sup_pres: "Sup_pres bd\<^sub>\" by (force simp: fun_eq_iff klift_def) lemma fbd_sup_pres: "sup_pres bd\<^sub>\" using Sup_sup_pres fbd_Sup_pres by blast lemma fbd_Disj: "Sup_pres (bd\<^sub>\ f)" by (simp add: fbd_ifbd_inv2_inv) lemma fbd_disj: "sup_pres (bd\<^sub>\ f)" by (simp add: klift_prop) lemma fbd_bot_pres: "bot_pres bd\<^sub>\" unfolding klift_def by fastforce lemma fbd_zero_pres2 [simp]: "bd\<^sub>\ f {} = {}" by (simp add: klift_prop) lemma fbd_iso: "X \ Y \ bd\<^sub>\ f X \ bd\<^sub>\ f Y" by (metis fbd_disj le_iff_sup) text \The following counterexamples show that Infs are not preserved.\ lemma "top_pres bd\<^sub>\" (*nitpick*) oops lemma "inf_pres bd\<^sub>\" (*nitpick*) oops text \Dual preservation statements hold for ifbd ... and even Inf-preservation.\ lemma ifbd_comp_pres: "Sup_pres \ \ bd\<^sup>-\<^sub>\ (\ \ \) = bd\<^sup>-\<^sub>\ \ \\<^sub>K bd\<^sup>-\<^sub>\ \" by (smt fbd_ifbd_galois fun.map_comp kcomp_def klift_def) lemma ifbd_Sup_pres: "Sup_pres bd\<^sup>-\<^sub>\" by (simp add: fun_eq_iff) lemma ifbd_sup_pres: "sup_pres bd\<^sup>-\<^sub>\" by force lemma ifbd_Inf_pres: "Inf_pres bd\<^sup>-\<^sub>\" by (simp add: fun_eq_iff) lemma ifbd_inf_pres: "inf_pres bd\<^sup>-\<^sub>\" by force lemma ifbd_bot_pres: "bot_pres bd\<^sup>-\<^sub>\" by auto lemma ifbd_top_pres: "top_pres bd\<^sup>-\<^sub>\" by auto text \Preservation of units by the Kleisli lifting has been proved in klift-prop3.\ text \These results estabilish the isomorphism between state and predicate transformers given by backward diamonds. The isomorphism preserves the Sup-quantale structure, but not Infs.\ subsection \Backward Diamonds from Relations\ text \Using the isomorphism between binary relations and Kleisli arrows (or state transformers), it is straightforward to define backward diamonds from relations, by composing isomorphisms. It follows that Sup-quantales of binary relations (under relational composition, the identity relation and Sups) are isomorphic to the Sup-quantales of predicate transformers. Once again, Infs are not preserved.\ definition rbd :: "('a \ 'b) set \ 'a set \ 'b set" ("bd\<^sub>\") where "bd\<^sub>\ = bd\<^sub>\ \ \" definition irbd :: "('a set \ 'b set) \ ('a \ 'b) set" ("bd\<^sup>-\<^sub>\") where "bd\<^sup>-\<^sub>\ = \ \ bd\<^sup>-\<^sub>\" lemma rbd_Im: "bd\<^sub>\ = (``)" unfolding rbd_def klift_def r2f_def fun_eq_iff by force lemma rbd_set: "bd\<^sub>\ R X = {y. \x \ X. (x,y) \ R}" by (force simp: rbd_Im Image_def) lemma irbd_set: "bd\<^sup>-\<^sub>\ \ = {(x,y). y \ (\ \ \) x}" by (simp add: irbd_def f2r_def o_def) lemma irbd_set_var: "bd\<^sup>-\<^sub>\ \ = {(x,y). y \ \ {x}}" by (simp add: irbd_def f2r_def o_def) lemma rbd_irbd_inv1 [simp]: "bd\<^sup>-\<^sub>\ \ bd\<^sub>\ = id" by (metis (no_types, lifting) comp_eq_dest_lhs eq_id_iff fbd_Disj fbd_ifbd_galois irbd_def r2f_f2r_galois rbd_def) lemma irbd_rbd_inv2: "Sup_pres \ \ (bd\<^sub>\ \ bd\<^sup>-\<^sub>\) \ = \" by (metis comp_apply fbd_ifbd_galois irbd_def r2f_f2r_galois rbd_def) lemma irbd_rbd_inv2_inv: "(bd\<^sub>\ \ bd\<^sup>-\<^sub>\) \ = \ \ Sup_pres \" by (simp add: rbd_def irbd_def, metis fbd_Disj) lemma irbd_rbd_inv2_iff: "((bd\<^sub>\ \ bd\<^sup>-\<^sub>\) \ = \) = (Sup_pres \)" using irbd_rbd_inv2 irbd_rbd_inv2_inv by blast lemma rbd_inj: "inj bd\<^sub>\" by (simp add: fbd_inj inj_compose r2f_inj rbd_def) lemma rbd_translate: "(bd\<^sub>\ R = bd\<^sub>\ S) = (R = S)" by (simp add: rbd_inj inj_eq) lemma irbd_inj: "Sup_pres \ \ Sup_pres \ \ bd\<^sup>-\<^sub>\ \ = bd\<^sup>-\<^sub>\ \ \ \ = \" by (metis rbd_Im comp_eq_dest_lhs irbd_rbd_inv2) lemma irbd_inj_iff: "Sup_pres \ \ Sup_pres \ \ (bd\<^sup>-\<^sub>\ \ = bd\<^sup>-\<^sub>\ \) = (\ = \)" using irbd_inj by force lemma rbd_surj: "Sup_pres \ \ (\R. bd\<^sub>\ R = \)" using irbd_rbd_inv2 by force lemma irbd_surj: "surj bd\<^sup>-\<^sub>\" by (metis UNIV_I fun.set_map imageE rbd_irbd_inv1 surj_def surj_id) lemma rbd_irbd_galois: "Sup_pres \ \ (\ = bd\<^sub>\ R) = (R = bd\<^sup>-\<^sub>\ \)" by (smt comp_apply fbd_ifbd_galois irbd_def r2f_f2r_galois rbd_def) lemma rbd_comp_pres: "bd\<^sub>\ (R ; S) = bd\<^sub>\ S \ bd\<^sub>\ R" by (simp add: rbd_def r2f_comp_pres fbd_comp_pres) lemma rbd_Id_pres: "bd\<^sub>\ Id = id" unfolding rbd_def by simp lemma rbd_Un_pres: "Sup_pres bd\<^sub>\" by (simp add: rbd_def Sup_pres_comp fbd_Sup_pres r2f_Sup_pres) lemma rbd_un_pres: "sup_pres bd\<^sub>\" by (simp add: rbd_def fbd_sup_pres r2f_sup_pres) lemma "inf_pres bd\<^sub>\" (*nitpick*) oops lemma rbd_disj: "Sup_pres (bd\<^sub>\ R)" by (simp add: rbd_def fbd_Disj) lemma rbd_disj2: "sup_pres (bd\<^sub>\ R)" by (simp add: Image_Un rbd_Im) lemma rbd_bot_pres: "bot_pres bd\<^sub>\" by (simp add: fbd_bot_pres r2f_bot_pres rbd_def) lemma rbd_zero_pres2 [simp]: "bd\<^sub>\ R {} = {}" by (simp add: rbd_Im) lemma rbd_univ: "bd\<^sub>\ R UNIV = Range R" unfolding rbd_def fun_eq_iff klift_def r2f_def by force lemma rbd_iso: "X \ Y \ bd\<^sub>\ R X \ bd\<^sub>\ R Y" by (metis le_iff_sup rbd_disj2) lemma irbd_comp_pres: "Sup_pres \ \ bd\<^sup>-\<^sub>\ (\ \ \) = bd\<^sup>-\<^sub>\ \ ; bd\<^sup>-\<^sub>\ \" by (simp add: ifbd_comp_pres f2r_kcomp_pres irbd_def) lemma irbd_id_pres [simp]: "bd\<^sup>-\<^sub>\ id = Id" unfolding irbd_def by simp lemma irbd_Sup_pres: "Sup_pres bd\<^sup>-\<^sub>\" by (simp add: irbd_def Sup_pres_comp ifbd_Sup_pres f2r_Sup_pres) lemma irbd_sup_pres: "sup_pres bd\<^sup>-\<^sub>\" by (simp add: irbd_def ifbd_sup_pres f2r_sup_pres) lemma irbd_Inf_pres: "Inf_pres bd\<^sup>-\<^sub>\" by (auto simp: fun_eq_iff irbd_def f2r_def) lemma irbd_inf_pres: "inf_pres bd\<^sup>-\<^sub>\" by (auto simp: fun_eq_iff irbd_def f2r_def) lemma irbd_bot_pres: "bot_pres bd\<^sup>-\<^sub>\" by (metis comp_def ifbd_bot_pres f2r_bot_pres irbd_def) text \This shows that relations are isomorphic to disjunctive forward predicate transformers. In many cases Isabelle picks up the composition of morphisms in proofs.\ subsection \Forward Boxes on Kleisli Arrows\ text \Forward box operators correspond to weakest liberal preconditions in program semantics. Here, Kleisli arrows are mapped to the opposite of the Eilenberg-Moore category, that is, Inf-lattices. It follows that the Inf-quantale structure is preserved. Modelling opposition is based on the fact that Kleisli arrows can be swapped by going through relations.\ definition ffb :: "('a \ 'b set) \ 'b set \ 'a set" ("fb\<^sub>\") where "fb\<^sub>\ = \\<^sub>F \ bd\<^sub>\ \ op\<^sub>K" text \Here, $\partial_F$ is map-dual, which amounts to De Morgan duality. Hence the forward box operator is obtained from the backward diamond by taking the opposite Kleisli arrow, applying the backward diamond, and then De Morgan duality.\ lemma ffb_prop: "fb\<^sub>\ f = \ \ bd\<^sub>\ (op\<^sub>K f) \ \" by (simp add: ffb_def map_dual_def) lemma ffb_prop_var: "fb\<^sub>\ f = uminus \ bd\<^sub>\ (op\<^sub>K f) \ uminus" by (simp add: dual_set_def ffb_prop) lemma ffb_fbd_dual: "\ \ fb\<^sub>\ f = bd\<^sub>\ (op\<^sub>K f) \ \" by (simp add: ffb_prop o_assoc) text \I give a set-theoretic definition of iffb, because the algebraic one below depends on Inf-preservation.\ definition iffb :: "('b set \ 'a set) \ 'a \ 'b set" ("fb\<^sup>-\<^sub>\") where "fb\<^sup>-\<^sub>\ \ = (\x. \{X. x \ \ X})" lemma ffb_set: "fb\<^sub>\ f = (\Y. {x. f x \ Y})" by (force simp: fun_eq_iff ffb_prop_var kop_def klift_def f2r_def r2f_def) text \Forward boxes and backward diamonds are adjoints.\ lemma ffb_fbd_galois: "(bd\<^sub>\ f) \ (fb\<^sub>\ f)" unfolding adj_def ffb_set klift_prop by blast lemma iffb_inv1: "fb\<^sup>-\<^sub>\ \ fb\<^sub>\ = id" unfolding fun_eq_iff ffb_set iffb_def by force lemma iffb_inv2_aux: "Inf_pres \ \ \{X. x \ \ X} \ Y \ x \ \ Y" proof- assume "Inf_pres \" and h1: "\{X. x \ \ X} \ Y" hence h2: "\X. \ (\X) = (\x \ X. \ x)" by (metis comp_eq_dest) hence "\ (\{X. x \ \ X}) \ \ Y" by (metis h1 INF_lower2 cInf_eq_minimum mem_Collect_eq order_refl) hence "(\{\ X |X. x \ \ X}) \ \ Y" by (metis h2 setcompr_eq_image) thus ?thesis by (force simp add: subset_iff) qed lemma iffb_inv2: "Inf_pres \ \ (fb\<^sub>\ \ fb\<^sup>-\<^sub>\) \ = \" proof- assume h: "Inf_pres \" {fix Y have "(fb\<^sub>\ \ fb\<^sup>-\<^sub>\) \ Y = {x. \{X. x \ \ X} \ Y}" by (simp add: ffb_set iffb_def) hence "\x. x \ (fb\<^sub>\ \ fb\<^sup>-\<^sub>\) \ Y \ \{X. x \ \ X} \ Y" by auto hence "\x. x \ (fb\<^sub>\ \ fb\<^sup>-\<^sub>\) \ Y \ x \ \ Y" by (auto simp: h iffb_inv2_aux) hence "(fb\<^sub>\ \ fb\<^sup>-\<^sub>\) \ Y = \ Y" by (simp add: fun_eq_iff set_eq_iff)} thus ?thesis unfolding fun_eq_iff by simp qed lemma iffb_inv2_inv: "(fb\<^sub>\ \ fb\<^sup>-\<^sub>\) \ = \ \ Inf_pres \" by (auto simp: fun_eq_iff ffb_set iffb_def) lemma iffb_inv2_iff: "((fb\<^sub>\ \ fb\<^sup>-\<^sub>\) \ = \) = (Inf_pres \)" using iffb_inv2 iffb_inv2_inv by blast lemma ffb_inj: "inj fb\<^sub>\" unfolding inj_def by (metis iffb_inv1 pointfree_idE) lemma ffb_inj_iff: "(fb\<^sub>\ f = fb\<^sub>\ g) = (f = g)" by (simp add: ffb_inj inj_eq) lemma ffb_iffb_galois: "Inf_pres \ \ (fb\<^sup>-\<^sub>\ \ = f) = (fb\<^sub>\ f = \)" using ffb_inj_iff iffb_inv2 by force lemma iffb_inj: "Inf_pres \ \ Inf_pres \ \ fb\<^sup>-\<^sub>\ \ = fb\<^sup>-\<^sub>\ \ \ \ = \" by (metis ffb_iffb_galois) lemma iffb_inj_iff: "Inf_pres \ \ Inf_pres \ \ (fb\<^sup>-\<^sub>\ \ = fb\<^sup>-\<^sub>\ \) = (\ = \)" using iffb_inj by blast lemma ffb_surj: "Inf_pres \ \ (\f. fb\<^sub>\ f = \)" using iffb_inv2 by auto lemma iffb_surj: "surj fb\<^sup>-\<^sub>\" using surj_def by (metis comp_apply iffb_inv1 surj_id) text \This is now the explicit "definition" of iffb, for Inf-preserving transformers.\ lemma iffb_ifbd_dual: "Inf_pres \ \ fb\<^sup>-\<^sub>\ \ = (op\<^sub>K \ bd\<^sup>-\<^sub>\ \ \\<^sub>F) \" proof- assume h: "Inf_pres \" {fix f have "(fb\<^sup>-\<^sub>\ \ = f) = ((\\<^sub>F \ bd\<^sub>\ \ op\<^sub>K) f = \)" by (simp add: ffb_def ffb_iffb_galois h) also have "... = (op\<^sub>K f = (bd\<^sup>-\<^sub>\ \ \\<^sub>F) \)" by (metis (mono_tags, lifting) comp_apply map_dual_dual ffb_def ffb_surj h klift_eta_inv1 map_dual_dual) finally have "(fb\<^sup>-\<^sub>\ \ = f) = (f = (op\<^sub>K \ bd\<^sup>-\<^sub>\ \ \\<^sub>F) \)" using kop_galois by auto} thus ?thesis by blast qed lemma fbd_ffb_dual: "\\<^sub>F \ fb\<^sub>\ \ op\<^sub>K = bd\<^sub>\" proof- have "\\<^sub>F \ fb\<^sub>\ \ op\<^sub>K = \\<^sub>F \ \\<^sub>F \ bd\<^sub>\ \ (op\<^sub>K \ op\<^sub>K)" by (simp add: comp_def ffb_def) thus ?thesis by simp qed lemma ffbd_ffb_dual_var: "\ \ bd\<^sub>\ f = fb\<^sub>\ (op\<^sub>K f) \ \" by (metis ffb_prop fun_dual1 kop_galois) lemma ifbd_iffb_dual: "Sup_pres \ \ bd\<^sup>-\<^sub>\ \ = (op\<^sub>K \ fb\<^sup>-\<^sub>\ \ \\<^sub>F) \" proof- assume h: "Sup_pres \" hence "Inf_pres (\\<^sub>F \)" using Sup_pres_Inf_pres by blast hence "(op\<^sub>K \ fb\<^sup>-\<^sub>\ \ \\<^sub>F) \ = (op\<^sub>K \ (op\<^sub>K \ bd\<^sup>-\<^sub>\ \ \\<^sub>F) \ \\<^sub>F) \" by (simp add: iffb_ifbd_dual) thus ?thesis by (metis comp_def kop_galois map_dual_dual) qed lemma ffb_kcomp_pres: "fb\<^sub>\ (f \\<^sub>K g) = fb\<^sub>\ f \ fb\<^sub>\ g" proof- have "fb\<^sub>\ (f \\<^sub>K g) = \\<^sub>F (bd\<^sub>\ (op\<^sub>K (f \\<^sub>K g)))" by (simp add: ffb_def) also have "... = \\<^sub>F (bd\<^sub>\ (op\<^sub>K g \\<^sub>K op\<^sub>K f))" by (simp add: kop_contrav) also have "... = \\<^sub>F (bd\<^sub>\ (op\<^sub>K f) \ bd\<^sub>\ (op\<^sub>K g))" by (simp add: fbd_comp_pres) also have "... = \\<^sub>F (bd\<^sub>\ (op\<^sub>K f)) \ \\<^sub>F (bd\<^sub>\ (op\<^sub>K g))" by (simp add: map_dual_func1) finally show ?thesis by (simp add: ffb_def) qed lemma ffb_eta_pres: "fb\<^sub>\ \ = id" unfolding ffb_def by simp lemma ffb_Sup_dual: "Sup_dual fb\<^sub>\" unfolding ffb_prop_var comp_def fun_eq_iff klift_prop kop_def f2r_def r2f_def converse_def by fastforce lemma ffb_Sup_dual_var: "fb\<^sub>\ (\F) = (\f \ F. fb\<^sub>\ f)" unfolding ffb_prop_var comp_def fun_eq_iff klift_prop kop_def f2r_def r2f_def converse_def by fastforce lemma ffb_sup_dual: "sup_dual fb\<^sub>\" using ffb_Sup_dual Sup_sup_dual by force lemma ffb_zero_dual: "fb\<^sub>\ \ = (\X. UNIV)" unfolding ffb_prop_var kop_def klift_prop fun_eq_iff f2r_def r2f_def by simp lemma "inf_dual ffb" (*nitpick*) oops text \Once again, only the Sup-quantale structure is preserved.\ lemma iffb_comp_pres: assumes "Inf_pres \" assumes "Inf_pres \" shows "fb\<^sup>-\<^sub>\ (\ \ \) = fb\<^sup>-\<^sub>\ \ \\<^sub>K fb\<^sup>-\<^sub>\ \" by (metis assms Inf_pres_comp ffb_iffb_galois ffb_kcomp_pres) lemma iffb_id_pres: "fb\<^sup>-\<^sub>\ id = \" unfolding iffb_def by force lemma iffb_Inf_dual: assumes "\\ \ \. Inf_pres \" shows "(fb\<^sup>-\<^sub>\ \ Inf) \ = (Sup \ \

fb\<^sup>-\<^sub>\) \" proof- have "Inf_pres (\\)" using Inf_pres_Inf assms by blast hence "(fb\<^sub>\ \ fb\<^sup>-\<^sub>\) (\\) = \(\

(fb\<^sub>\ \ fb\<^sup>-\<^sub>\) \)" by (metis (mono_tags, lifting) INF_cong INF_identity_eq assms iffb_inv2) hence "(fb\<^sub>\ \ fb\<^sup>-\<^sub>\) (\\) = fb\<^sub>\ (\(\

fb\<^sup>-\<^sub>\ \))" by (simp add: Setcompr_eq_image ffb_Sup_dual_var image_comp) thus ?thesis by (simp add: ffb_inj_iff) qed lemma iffb_Sup_dual: "Sup_dual fb\<^sup>-\<^sub>\" by (auto simp: iffb_def fun_eq_iff) lemma iffb_inf_dual: assumes "Inf_pres \" and "Inf_pres \" shows "fb\<^sup>-\<^sub>\ (\ \ \) = fb\<^sup>-\<^sub>\ \ \ fb\<^sup>-\<^sub>\ \" proof - have f1: "\ \ \ = fb\<^sub>\ (fb\<^sup>-\<^sub>\ \) \ fb\<^sub>\ (fb\<^sup>-\<^sub>\ \)" using assms iffb_inv2 by fastforce have "\ \ \ \ Inter = Inter \ \

(\ \ \)" using assms Inf_pres_inf by blast thus ?thesis by (simp add: f1 ffb_iffb_galois ffb_sup_dual) qed lemma iffb_sup_dual: "fb\<^sup>-\<^sub>\ (\ \ \) = fb\<^sup>-\<^sub>\ \ \ fb\<^sup>-\<^sub>\ \" unfolding iffb_def by fastforce lemma iffb_top_pres [simp]: "fb\<^sup>-\<^sub>\ \ = \" unfolding iffb_def by simp text \This establishes the duality between state transformers and weakest liberal preconditions.\ subsection \Forward Box Operators from Relations\ text \Once again one can compose isomorphisms, linking weakest liberal preconditions with relational semantics. The isomorphism obtained should by now be obvious.\ definition rfb :: "('a \ 'b) set \ 'b set \ 'a set" ("fb\<^sub>\") where "fb\<^sub>\ = fb\<^sub>\ \ \" definition irfb :: "('b set \ 'a set) \ ('a \ 'b) set" ("fb\<^sup>-\<^sub>\") where "fb\<^sup>-\<^sub>\ = \ \ fb\<^sup>-\<^sub>\" lemma rfb_rbd_dual: "fb\<^sub>\ R = \\<^sub>F (bd\<^sub>\ (R\))" by (simp add: rfb_def rbd_def kop_def ffb_def, metis r2f_f2r_galois) lemma rbd_rfb_dual: "bd\<^sub>\ R = \\<^sub>F (fb\<^sub>\ (R\))" by (simp add: rfb_def rbd_def kop_def ffb_def, metis converse_converse map_dual_dual r2f_f2r_galois) lemma irfb_irbd_dual: "Inf_pres \ \ fb\<^sup>-\<^sub>\ \ = ((\) \ bd\<^sup>-\<^sub>\ \ \\<^sub>F) \" by (simp add: irfb_def irbd_def iffb_ifbd_dual kop_def r2f_f2r_galois) lemma irbd_irfb_dual: "Sup_pres \ \ bd\<^sup>-\<^sub>\ \ = ((\) \ fb\<^sup>-\<^sub>\ \ \\<^sub>F) \" by (simp add: irfb_def irbd_def ifbd_iffb_dual kop_def r2f_f2r_galois) lemma rfb_set: "fb\<^sub>\ R Y = {x. \y. (x,y) \ R \ y \ Y}" unfolding rfb_def ffb_prop_var comp_def klift_def f2r_def r2f_def kop_def by force lemma rfb_rbd_galois: "(bd\<^sub>\ R) \ (fb\<^sub>\ R)" by (simp add: ffb_fbd_galois rbd_def rfb_def) lemma irfb_set: "fb\<^sup>-\<^sub>\ \ = {(x, y). \Y. x \ \ Y \ y \ Y}" by (simp add: irfb_def iffb_def f2r_def) lemma irfb_inv1 [simp]: "fb\<^sup>-\<^sub>\ \ fb\<^sub>\ = id" by (simp add: fun_eq_iff rfb_def irfb_def iffb_inv1 pointfree_idE) lemma irfb_inv2: "Inf_pres \ \ (fb\<^sub>\ \ fb\<^sup>-\<^sub>\) \ = \" by (simp add: rfb_def irfb_def, metis ffb_iffb_galois r2f_f2r_galois) lemma rfb_inj: "inj fb\<^sub>\" by (simp add: rfb_def ffb_inj inj_compose r2f_inj) lemma rfb_inj_iff: "(fb\<^sub>\ R = fb\<^sub>\ S) = (R = S)" by (simp add: rfb_inj inj_eq) lemma irfb_inj: "Inf_pres \ \ Inf_pres \ \ fb\<^sup>-\<^sub>\ \ = fb\<^sup>-\<^sub>\ \ \ \ = \" unfolding irfb_def using iffb_inj r2f_inj_iff by fastforce lemma irfb_inf_iff: "Inf_pres \ \ Inf_pres \ \ (fb\<^sup>-\<^sub>\ \ = fb\<^sup>-\<^sub>\ \) = (\ = \)" using irfb_inj by auto lemma rfb_surj: "Inf_pres \ \ (\R. fb\<^sub>\ R = \)" using irfb_inv2 by fastforce lemma irfb_surj: "surj fb\<^sup>-\<^sub>\" by (simp add: irfb_def comp_surj f2r_surj iffb_surj cong del: image_cong_simp) lemma rfb_irfb_galois: "Inf_pres \ \ (fb\<^sup>-\<^sub>\ \ = R) = (fb\<^sub>\ R = \)" by (simp add: irfb_def rfb_def, metis ffb_iffb_galois r2f_f2r_galois) lemma rfb_comp_pres: "fb\<^sub>\ (R ; S) = fb\<^sub>\ R \ fb\<^sub>\ S" by (simp add: ffb_kcomp_pres r2f_comp_pres rfb_def) lemma rfb_Id_pres [simp]: "fb\<^sub>\ Id = id" unfolding rfb_def ffb_prop by force lemma rfb_Sup_dual: "Sup_dual fb\<^sub>\" proof- have "fb\<^sub>\ \ \ = fb\<^sub>\ \ \ \ Sup" by (simp add: rfb_def) also have "... = fb\<^sub>\ \ Sup \ \

\" by (metis fun.map_comp r2f_Sup_pres) also have "... = Inf \ \

fb\<^sub>\ \ \

\" by (simp add: ffb_Sup_dual) also have "... = Inf \ \

(fb\<^sub>\ \ \)" by (simp add: P_func1 comp_assoc) finally show ?thesis by (simp add: rfb_def) qed lemma rfb_Sup_dual_var: "fb\<^sub>\ (\\) = \(\

fb\<^sub>\) \" by (meson comp_eq_dest rfb_Sup_dual) lemma rfb_sup_dual: "sup_dual fb\<^sub>\" by (simp add: rfb_def ffb_sup_dual r2f_sup_pres) lemma "inf_dual fb\<^sub>\" (*nitpick*) oops lemma rfb_Inf_pres: "Inf_pres (fb\<^sub>\ R)" unfolding rfb_def ffb_prop_var comp_def fun_eq_iff klift_def kop_def f2r_def r2f_def converse_def by auto lemma rfb_inf_pres: "inf_pres (fb\<^sub>\ R)" unfolding rfb_def ffb_prop_var comp_def fun_eq_iff klift_def kop_def f2r_def r2f_def converse_def by auto lemma rfb_zero_pres [simp]: "fb\<^sub>\ {} X = UNIV" unfolding rfb_def ffb_prop_var comp_def fun_eq_iff klift_def kop_def f2r_def r2f_def converse_def by auto lemma rfb_zero_pres2 [simp]: "fb\<^sub>\ R {} = - Domain R" unfolding rfb_def ffb_prop_var comp_def fun_eq_iff klift_def kop_def f2r_def r2f_def converse_def by auto lemma rfb_univ [simp]: "fb\<^sub>\ R UNIV = UNIV" unfolding rfb_def ffb_prop_var comp_def fun_eq_iff klift_def kop_def f2r_def r2f_def converse_def by auto lemma rfb_iso: "X \ Y \ fb\<^sub>\ R X \ fb\<^sub>\ R Y" unfolding rfb_def ffb_prop_var comp_def fun_eq_iff klift_def kop_def f2r_def r2f_def converse_def by auto lemma irfb_comp_pres: assumes "Inf_pres \" assumes "Inf_pres \" shows "fb\<^sup>-\<^sub>\ (\ \ \) = fb\<^sup>-\<^sub>\ \ ; fb\<^sup>-\<^sub>\ \" by (metis assms rfb_Inf_pres rfb_comp_pres rfb_irfb_galois) lemma irfb_id_pres [simp]: "fb\<^sup>-\<^sub>\ id = Id" by (simp add: rfb_irfb_galois) lemma irfb_Sup_dual: "Sup_dual fb\<^sup>-\<^sub>\" by (auto simp: fun_eq_iff irfb_def iffb_def f2r_def) lemma irfb_Inf_dual: assumes "\\ \ \. Inf_pres \" shows "(fb\<^sup>-\<^sub>\ \ Inf) \ = (Sup \ \

fb\<^sup>-\<^sub>\) \" proof- have "Inf_pres (\\)" using Inf_pres_Inf assms by blast hence "(fb\<^sub>\ \ fb\<^sup>-\<^sub>\) (\\) = \(\

(fb\<^sub>\ \ fb\<^sup>-\<^sub>\) \)" by (smt INF_identity_eq Sup.SUP_cong assms irfb_inv2) also have "... = \(\

fb\<^sub>\ (\

fb\<^sup>-\<^sub>\ \))" by (simp add: image_comp) also have "... = fb\<^sub>\ (\(\

fb\<^sup>-\<^sub>\ \))" by (simp add: rfb_Sup_dual_var) finally have "(fb\<^sub>\ \ fb\<^sup>-\<^sub>\) (\\) = fb\<^sub>\ (\(\

fb\<^sup>-\<^sub>\ \))". thus ?thesis by (simp add: rfb_inj_iff) qed lemma irfb_sup_dual: "sup_dual fb\<^sup>-\<^sub>\" by (force simp: fun_eq_iff irfb_def iffb_def f2r_def) lemma irfb_inf_dual: assumes "Inf_pres \" and "Inf_pres \" shows "fb\<^sup>-\<^sub>\ (\ \ \) = fb\<^sup>-\<^sub>\ \ \ fb\<^sup>-\<^sub>\ \" by (metis assms rfb_Inf_pres rfb_irfb_galois rfb_sup_dual) lemma irfb_top_pres [simp]: "bd\<^sup>-\<^sub>\ \ = UNIV" unfolding irbd_def f2r_def by auto text \Finally, the adjunctions between the predicate transformers considered so far are revisited.\ lemma ffb_fbd_galois_var: "(bd\<^sub>\ f X \ Y) = (X \ fb\<^sub>\ f Y)" by (meson adj_def ffb_fbd_galois) lemma rfb_rbd_galois_var: "(bd\<^sub>\ R X \ Y) = (X \ fb\<^sub>\ R Y)" by (meson adj_def rfb_rbd_galois) lemma ffb_fbd: "fb\<^sub>\ f Y = \{X. bd\<^sub>\ f X \ Y}" using ffb_fbd_galois_var by fastforce lemma rfb_rbd: "fb\<^sub>\ R Y = \{X. bd\<^sub>\ R X \ Y}" using rfb_rbd_galois_var by fastforce lemma fbd_ffb: "bd\<^sub>\ f X = \{Y. X \ fb\<^sub>\ f Y}" using ffb_fbd_galois_var by fastforce lemma rbd_rfb: "bd\<^sub>\ R X = \{Y. X \ fb\<^sub>\ R Y}" using rfb_rbd_galois_var by fastforce subsection \The Remaining Modalities\ text \Finally I set up the remaining dual transformers: forward diamonds and backward boxes. Most properties are not repeated, only some symmetries and dualities are spelled out.\ text \First, forward diamond operators are introduced, from state transformers and relations; together with their inverses.\ definition ffd :: "('a \ 'b set) \ 'b set \ 'a set" ("fd\<^sub>\") where "fd\<^sub>\ = bd\<^sub>\ \ op\<^sub>K" definition iffd :: "('b set \ 'a set) \ 'a \ 'b set" ("fd\<^sup>-\<^sub>\") where "fd\<^sup>-\<^sub>\ = op\<^sub>K \ bd\<^sup>-\<^sub>\" definition rfd :: "('a \ 'b) set \ 'b set \ 'a set" ("fd\<^sub>\") where "fd\<^sub>\ = fd\<^sub>\ \ \" definition irfd :: "('b set \ 'a set) \ ('a \ 'b) set" ("fd\<^sup>-\<^sub>\") where "fd\<^sup>-\<^sub>\ = \ \ fd\<^sup>-\<^sub>\" text \Second, I introduce forward boxes and their inverses.\ definition fbb :: "('a \ 'b set) \ 'a set \ 'b set" ("bb\<^sub>\") where "bb\<^sub>\ = fb\<^sub>\ \ op\<^sub>K" definition ifbb :: "('a set \ 'b set) \ 'a \ 'b set" ("bb\<^sup>-\<^sub>\") where "bb\<^sup>-\<^sub>\ = op\<^sub>K \ fb\<^sup>-\<^sub>\" definition rbb :: "('a \ 'b) set \ 'a set \ 'b set" ("bb\<^sub>\") where "bb\<^sub>\ = bb\<^sub>\ \ \" definition irbb :: "('a set \ 'b set) \ ('a \ 'b) set" ("bb\<^sup>-\<^sub>\") where "bb\<^sup>-\<^sub>\ = \ \ bb\<^sup>-\<^sub>\" text \Forward and backward operators of the same type (box or diamond) are related by opposition.\ lemma rfd_rbd: "fd\<^sub>\ = bd\<^sub>\ \ (\)" by (simp add: rfd_def rbd_def ffd_def kop_def comp_assoc) lemma irfd_irbd: "fd\<^sup>-\<^sub>\ = (\) \ bd\<^sup>-\<^sub>\" by (simp add: irfd_def iffd_def kop_def irbd_def comp_assoc[symmetric]) lemma fbd_ffd: "bd\<^sub>\ = fd\<^sub>\ \ op\<^sub>K" by (simp add: ffd_def kop_def converse_def f2r_def r2f_def klift_def fun_eq_iff) lemma rbb_rfb: "bb\<^sub>\ = fb\<^sub>\ \ (\)" by (simp add: rfb_def rbb_def, metis fbb_def kop_def r2f_f2r_galois_var2 rewriteR_comp_comp2) lemma irbb_irfb: "bb\<^sup>-\<^sub>\ = (\) \ fb\<^sup>-\<^sub>\" proof- have "bb\<^sup>-\<^sub>\ = \ \ op\<^sub>K \ fb\<^sup>-\<^sub>\" by (simp add: irbb_def ifbb_def o_assoc) also have "... = \ \ \ \ (\) \ \ \ fb\<^sup>-\<^sub>\" by (simp add: kop_def o_assoc) also have "... = (\) \ fb\<^sup>-\<^sub>\" by (simp add: comp_assoc irfb_def) finally show ?thesis. qed text \Complementation is a natural isomorphism between forwards and backward operators of different type.\ lemma ffd_ffb_demorgan: "\ \ fd\<^sub>\ f = fb\<^sub>\ f \ \" by (simp add: comp_assoc ffb_prop ffd_def) lemma iffd_iffb_demorgan: "Sup_pres \ \ fd\<^sup>-\<^sub>\ \ = (fb\<^sup>-\<^sub>\ \ \\<^sub>F) \" by (smt Sup_pres_Inf_pres comp_apply iffb_ifbd_dual iffd_def map_dual_dual) lemma ffb_ffd_demorgan: "\ \ fb\<^sub>\ f = fd\<^sub>\ f \ \" by (simp add: ffb_prop ffd_def rewriteL_comp_comp) lemma iffb_iffd_demorgan: "Inf_pres \ \ fb\<^sup>-\<^sub>\ \ = (fd\<^sup>-\<^sub>\ \ \\<^sub>F) \" by (simp add: iffb_ifbd_dual iffd_def) lemma rfd_rfb_demorgan: "\ \ fd\<^sub>\ R = fb\<^sub>\ R \ \" by (simp add: rfb_def rfd_def ffd_ffb_demorgan) lemma irfd_irfb_demorgan: "Sup_pres \ \ fd\<^sup>-\<^sub>\ \ = (fb\<^sup>-\<^sub>\ \ \\<^sub>F) \" by (simp add: irfb_def irfd_def iffd_iffb_demorgan) lemma rfb_rfd_demorgan: "\ \ fb\<^sub>\ R = fd\<^sub>\ R \ \" by (simp add: ffb_ffd_demorgan rfb_def rfd_def) lemma irfb_irfd_demorgan: "Inf_pres \ \ fb\<^sup>-\<^sub>\ \ = (fd\<^sup>-\<^sub>\ \ \\<^sub>F) \" by (simp add: irfb_irbd_dual irfd_irbd) lemma fbd_fbb_demorgan: "\ \ bd\<^sub>\ f = bb\<^sub>\ f \ \" by (simp add: fbb_def fbd_ffd ffd_ffb_demorgan) lemma ifbd_ifbb_demorgan: "Sup_pres \ \ bd\<^sup>-\<^sub>\ \ = (bb\<^sup>-\<^sub>\ \ \\<^sub>F) \" by (simp add: ifbd_iffb_dual ifbb_def) lemma fbb_fbd_demorgan: "\ \ bb\<^sub>\ R = bd\<^sub>\ R \ \" by (simp add: fbb_def fbd_ffd ffb_ffd_demorgan) lemma ifbb_ifbd_demorgan: "Inf_pres \ \ bb\<^sup>-\<^sub>\ \ = (bd\<^sup>-\<^sub>\ \ \\<^sub>F) \" proof- assume h: "Inf_pres \" have "bb\<^sup>-\<^sub>\ \ = (op\<^sub>K \ fb\<^sup>-\<^sub>\) \" by (simp add: ifbb_def) also have "... = (op\<^sub>K \ op\<^sub>K \ bd\<^sup>-\<^sub>\) (\\<^sub>F \)" by (metis comp_apply h iffb_ifbd_dual) also have "... = (bd\<^sup>-\<^sub>\ \ \\<^sub>F) \" by auto finally show ?thesis. qed lemma rbd_rbb_demorgan: "\ \ bd\<^sub>\ R = bb\<^sub>\ R \ \" by (simp add: rbb_def rbd_def fbd_fbb_demorgan) lemma irbd_irbb_demorgan: "Sup_pres \ \ bd\<^sup>-\<^sub>\ \ = (bb\<^sup>-\<^sub>\ \ \\<^sub>F) \" by (simp add: irbb_irfb irbd_irfb_dual) lemma rbb_rbd_demorgan: "\ \ bb\<^sub>\ R = bd\<^sub>\ R \ \" by (simp add: rbb_def rbd_def fbb_fbd_demorgan) lemma irbb_irbd_demorgan: "Inf_pres \ \ bb\<^sup>-\<^sub>\ \ = (bd\<^sup>-\<^sub>\ \ \\<^sub>F) \" by (simp add: irbb_def irbd_def ifbb_ifbd_demorgan) text \Further symmetries arise by combination.\ lemma ffd_fbb_dual: "\ \ fd\<^sub>\ f = bb\<^sub>\ (op\<^sub>K f) \ \" by (simp add: fbd_fbb_demorgan ffd_def) lemma iffd_ifbb_dual: "Sup_pres \ \ fd\<^sup>-\<^sub>\ \ = (op\<^sub>K \ bb\<^sup>-\<^sub>\ \ \\<^sub>F) \" by (simp add: ifbd_ifbb_demorgan iffd_def) lemma fbb_ffd_dual: "\ \ bb\<^sub>\ f = fd\<^sub>\ (op\<^sub>K f) \ \" by (simp add: fbd_ffd fbb_fbd_demorgan) lemma ifbb_iffd_dual: "Inf_pres \ \ bb\<^sup>-\<^sub>\ \ = (op\<^sub>K \ fd\<^sup>-\<^sub>\ \ \\<^sub>F) \" by (simp add: ifbb_def iffb_iffd_demorgan) lemma rfd_rbb_dual: "\ \ fd\<^sub>\ R = bb\<^sub>\ (R\) \ \" by (metis fun_dual1 map_dual_def rbd_rbb_demorgan rfb_rbd_dual rfd_rfb_demorgan) lemma ifd_ibb_dual: "Sup_pres \ \ fd\<^sup>-\<^sub>\ \ = ((\) \ bb\<^sup>-\<^sub>\ \ \\<^sub>F) \" by (simp add: irbb_irfb irbd_irfb_dual irfd_irbd) lemma rbb_rfd_dual: "\ \ bb\<^sub>\ R = fd\<^sub>\ (R\) \ \" by (simp add: rbb_rfb rfb_rfd_demorgan) lemma irbb_irfd_dual: "Inf_pres \ \ bb\<^sup>-\<^sub>\ \ = ((\) \ fd\<^sup>-\<^sub>\ \ \\<^sub>F) \" by (simp add: irbb_irfb irfb_irbd_dual irfd_irbd) lemma ffd_iffd_galois: "Sup_pres \ \ (\ = fd\<^sub>\ f) = (f = fd\<^sup>-\<^sub>\ \)" unfolding ffd_def iffd_def by (metis comp_apply fbd_surj klift_eta_inv1 kop_galois) lemma rfd_irfd_galois: "Sup_pres \ \ (\ = fd\<^sub>\ R) = (R = fd\<^sup>-\<^sub>\ \)" unfolding irfd_def rfd_def by (metis comp_apply ffd_iffd_galois r2f_f2r_galois) lemma fbb_ifbb_galois: "Inf_pres \ \ (\ = bb\<^sub>\ f) = (f = bb\<^sup>-\<^sub>\ \)" unfolding fbb_def iffb_def by (metis (no_types, lifting) comp_apply ffb_iffb_galois ifbb_ifbd_demorgan iffb_ifbd_dual kop_galois) lemma rbb_irbb_galois: "Inf_pres \ \ (\ = bb\<^sub>\ R) = (R = bb\<^sup>-\<^sub>\ \)" apply (simp add: rbb_def irbb_def) using fbb_ifbb_galois r2f_f2r_galois by blast text \Next I spell out the missing adjunctions.\ lemma ffd_ffb_adj: "fd\<^sub>\ f \ bb\<^sub>\ f" by (simp add: fbb_def ffb_fbd_galois ffd_def) lemma ffd_fbb_galois: "(fd\<^sub>\ f X \ Y) = (X \ bb\<^sub>\ f Y)" by (simp add: fbb_def ffb_fbd_galois_var ffd_def) lemma rfd_rfb_adj: "fd\<^sub>\ f \ bb\<^sub>\ f" by (simp add: ffd_ffb_adj rbb_def rfd_def) lemma rfd_rbb_galois: "(fd\<^sub>\ R X \ Y) = (X \ bb\<^sub>\ R Y)" by (simp add: ffd_fbb_galois rbb_def rfd_def) text \Finally, forward and backward operators of the same type are linked by conjugation.\ lemma ffd_fbd_conjugation: "(fd\<^sub>\ f X \ Y = {}) = (X \ bd\<^sub>\ f Y = {})" proof- have "(fd\<^sub>\ f X \ Y = {}) = (fd\<^sub>\ f X \ -Y)" by (simp add: disjoint_eq_subset_Compl) also have "... = (X \ bb\<^sub>\ f (-Y))" by (simp add: ffd_fbb_galois) also have "... = (X \ - bb\<^sub>\ f (-Y) = {})" by (simp add: disjoint_eq_subset_Compl) also have "... = (X \ \ (bb\<^sub>\ f (\ Y)) = {})" by (simp add: dual_set_def) finally show ?thesis by (metis (no_types, opaque_lifting) comp_apply fbb_fbd_demorgan invol_dual_var) qed lemma rfd_rbd_conjugation: "((fd\<^sub>\ R X) \ Y = {}) = (X \ (bd\<^sub>\ R Y) = {})" by (simp add: rbd_def rfd_def ffd_fbd_conjugation) lemma ffb_fbb_conjugation: "((fb\<^sub>\ f X) \ Y = UNIV) = (X \ (bb\<^sub>\ f Y) = UNIV)" proof- have "((fb\<^sub>\ f X) \ Y = UNIV) = (-Y \ fb\<^sub>\ f X)" by blast also have "... = (bd\<^sub>\ f (\ Y) \ X)" by (simp add: ffb_fbd_galois_var dual_set_def) also have "... = (\ (bb\<^sub>\ f Y) \ X)" by (metis comp_def fbb_fbd_demorgan) also have "... = (X \ (bb\<^sub>\ f Y) = UNIV)" - by (metis compl_le_swap2 dual_set_def join_shunt) + by (metis compl_le_swap2 dual_set_def sup_shunt) finally show ?thesis. qed lemma rfb_rbb_conjugation: "((fb\<^sub>\ R X) \ Y = UNIV) = (X \ (bb\<^sub>\ R Y) = UNIV)" by (simp add: rfb_def rbb_def ffb_fbb_conjugation) end \ No newline at end of file diff --git a/thys/UTP/utp/utp_pred_laws.thy b/thys/UTP/utp/utp_pred_laws.thy --- a/thys/UTP/utp/utp_pred_laws.thy +++ b/thys/UTP/utp/utp_pred_laws.thy @@ -1,963 +1,963 @@ section \ Predicate Calculus Laws \ theory utp_pred_laws imports utp_pred begin subsection \ Propositional Logic \ text \ Showing that predicates form a Boolean Algebra (under the predicate operators as opposed to the lattice operators) gives us many useful laws. \ -interpretation Lattices.boolean_algebra diff_upred not_upred conj_upred "(\)" "(<)" +interpretation boolean_algebra diff_upred not_upred conj_upred "(\)" "(<)" disj_upred false_upred true_upred by (unfold_locales; pred_auto) lemma taut_true [simp]: "`true`" by (pred_auto) lemma taut_false [simp]: "`false` = False" by (pred_auto) lemma taut_conj: "`A \ B` = (`A` \ `B`)" by (rel_auto) lemma taut_conj_elim [elim!]: "\ `A \ B`; \ `A`; `B` \ \ P \ \ P" by (rel_auto) lemma taut_refine_impl: "\ Q \ P; `P` \ \ `Q`" by (rel_auto) lemma taut_shEx_elim: "\ `(\<^bold>\ x \ P x)`; \ x. \ \ P x; \ x. `P x` \ Q \ \ Q" by (rel_blast) text \ Linking refinement and HOL implication \ lemma refine_prop_intro: assumes "\ \ P" "\ \ Q" "`Q` \ `P`" shows "P \ Q" using assms by (pred_auto) lemma taut_not: "\ \ P \ (\ `P`) = `\ P`" by (rel_auto) lemma taut_shAll_intro: "\ x. `P x` \ `\<^bold>\ x \ P x`" by (rel_auto) lemma taut_shAll_intro_2: "\ x y. `P x y` \ `\<^bold>\ (x, y) \ P x y`" by (rel_auto) lemma taut_impl_intro: "\ \ \ P; `P` \ `Q` \ \ `P \ Q`" by (rel_auto) lemma upred_eval_taut: "`P\\b\/&\<^bold>v\` = \P\\<^sub>eb" by (pred_auto) lemma refBy_order: "P \ Q = `Q \ P`" by (pred_auto) lemma conj_idem [simp]: "((P::'\ upred) \ P) = P" by (pred_auto) lemma disj_idem [simp]: "((P::'\ upred) \ P) = P" by (pred_auto) lemma conj_comm: "((P::'\ upred) \ Q) = (Q \ P)" by (pred_auto) lemma disj_comm: "((P::'\ upred) \ Q) = (Q \ P)" by (pred_auto) lemma conj_subst: "P = R \ ((P::'\ upred) \ Q) = (R \ Q)" by (pred_auto) lemma disj_subst: "P = R \ ((P::'\ upred) \ Q) = (R \ Q)" by (pred_auto) lemma conj_assoc:"(((P::'\ upred) \ Q) \ S) = (P \ (Q \ S))" by (pred_auto) lemma disj_assoc:"(((P::'\ upred) \ Q) \ S) = (P \ (Q \ S))" by (pred_auto) lemma conj_disj_abs:"((P::'\ upred) \ (P \ Q)) = P" by (pred_auto) lemma disj_conj_abs:"((P::'\ upred) \ (P \ Q)) = P" by (pred_auto) lemma conj_disj_distr:"((P::'\ upred) \ (Q \ R)) = ((P \ Q) \ (P \ R))" by (pred_auto) lemma disj_conj_distr:"((P::'\ upred) \ (Q \ R)) = ((P \ Q) \ (P \ R))" by (pred_auto) lemma true_disj_zero [simp]: "(P \ true) = true" "(true \ P) = true" by (pred_auto)+ lemma true_conj_zero [simp]: "(P \ false) = false" "(false \ P) = false" by (pred_auto)+ lemma false_sup [simp]: "false \ P = P" "P \ false = P" by (pred_auto)+ lemma true_inf [simp]: "true \ P = P" "P \ true = P" by (pred_auto)+ lemma imp_vacuous [simp]: "(false \ u) = true" by (pred_auto) lemma imp_true [simp]: "(p \ true) = true" by (pred_auto) lemma true_imp [simp]: "(true \ p) = p" by (pred_auto) lemma impl_mp1 [simp]: "(P \ (P \ Q)) = (P \ Q)" by (pred_auto) lemma impl_mp2 [simp]: "((P \ Q) \ P) = (Q \ P)" by (pred_auto) lemma impl_adjoin: "((P \ Q) \ R) = ((P \ R \ Q \ R) \ R)" by (pred_auto) lemma impl_refine_intro: "\ Q\<^sub>1 \ P\<^sub>1; P\<^sub>2 \ (P\<^sub>1 \ Q\<^sub>2) \ \ (P\<^sub>1 \ P\<^sub>2) \ (Q\<^sub>1 \ Q\<^sub>2)" by (pred_auto) lemma spec_refine: "Q \ (P \ R) \ (P \ Q) \ R" by (rel_auto) lemma impl_disjI: "\ `P \ R`; `Q \ R` \ \ `(P \ Q) \ R`" by (rel_auto) lemma conditional_iff: "(P \ Q) = (P \ R) \ `P \ (Q \ R)`" by (pred_auto) lemma p_and_not_p [simp]: "(P \ \ P) = false" by (pred_auto) lemma p_or_not_p [simp]: "(P \ \ P) = true" by (pred_auto) lemma p_imp_p [simp]: "(P \ P) = true" by (pred_auto) lemma p_iff_p [simp]: "(P \ P) = true" by (pred_auto) lemma p_imp_false [simp]: "(P \ false) = (\ P)" by (pred_auto) lemma not_conj_deMorgans [simp]: "(\ ((P::'\ upred) \ Q)) = ((\ P) \ (\ Q))" by (pred_auto) lemma not_disj_deMorgans [simp]: "(\ ((P::'\ upred) \ Q)) = ((\ P) \ (\ Q))" by (pred_auto) lemma conj_disj_not_abs [simp]: "((P::'\ upred) \ ((\P) \ Q)) = (P \ Q)" by (pred_auto) lemma subsumption1: "`P \ Q` \ (P \ Q) = Q" by (pred_auto) lemma subsumption2: "`Q \ P` \ (P \ Q) = P" by (pred_auto) lemma neg_conj_cancel1: "(\ P \ (P \ Q)) = (\ P \ Q :: '\ upred)" by (pred_auto) lemma neg_conj_cancel2: "(\ Q \ (P \ Q)) = (\ Q \ P :: '\ upred)" by (pred_auto) lemma double_negation [simp]: "(\ \ (P::'\ upred)) = P" by (pred_auto) lemma true_not_false [simp]: "true \ false" "false \ true" by (pred_auto)+ lemma closure_conj_distr: "([P]\<^sub>u \ [Q]\<^sub>u) = [P \ Q]\<^sub>u" by (pred_auto) lemma closure_imp_distr: "`[P \ Q]\<^sub>u \ [P]\<^sub>u \ [Q]\<^sub>u`" by (pred_auto) lemma true_iff [simp]: "(P \ true) = P" by (pred_auto) lemma taut_iff_eq: "`P \ Q` \ (P = Q)" by (pred_auto) lemma impl_alt_def: "(P \ Q) = (\ P \ Q)" by (pred_auto) subsection \ Lattice laws \ lemma uinf_or: fixes P Q :: "'\ upred" shows "(P \ Q) = (P \ Q)" by (pred_auto) lemma usup_and: fixes P Q :: "'\ upred" shows "(P \ Q) = (P \ Q)" by (pred_auto) lemma UINF_alt_def: "(\ i | A(i) \ P(i)) = (\ i \ A(i) \ P(i))" by (rel_auto) lemma USUP_true [simp]: "(\ P | F(P) \ true) = true" by (pred_auto) lemma UINF_mem_UNIV [simp]: "(\ x\UNIV \ P(x)) = (\ x \ P(x))" by (pred_auto) lemma USUP_mem_UNIV [simp]: "(\ x\UNIV \ P(x)) = (\ x \ P(x))" by (pred_auto) lemma USUP_false [simp]: "(\ i \ false) = false" by (pred_simp) lemma USUP_mem_false [simp]: "I \ {} \ (\ i\I \ false) = false" by (rel_simp) lemma USUP_where_false [simp]: "(\ i | false \ P(i)) = true" by (rel_auto) lemma UINF_true [simp]: "(\ i \ true) = true" by (pred_simp) lemma UINF_ind_const [simp]: "(\ i \ P) = P" by (rel_auto) lemma UINF_mem_true [simp]: "A \ {} \ (\ i\A \ true) = true" by (pred_auto) lemma UINF_false [simp]: "(\ i | P(i) \ false) = false" by (pred_auto) lemma UINF_where_false [simp]: "(\ i | false \ P(i)) = false" by (rel_auto) lemma UINF_cong_eq: "\ \ x. P\<^sub>1(x) = P\<^sub>2(x); \ x. `P\<^sub>1(x) \ Q\<^sub>1(x) =\<^sub>u Q\<^sub>2(x)` \ \ (\ x | P\<^sub>1(x) \ Q\<^sub>1(x)) = (\ x | P\<^sub>2(x) \ Q\<^sub>2(x))" by (unfold UINF_def, pred_simp, metis) lemma UINF_as_Sup: "(\ P \ \

\ P) = \ \

" apply (simp add: upred_defs bop.rep_eq lit.rep_eq Sup_uexpr_def) apply (pred_simp) apply (rule cong[of "Sup"]) apply (auto) done lemma UINF_as_Sup_collect: "(\P\A \ f(P)) = (\P\A. f(P))" apply (simp add: upred_defs bop.rep_eq lit.rep_eq Sup_uexpr_def) apply (pred_simp) apply (simp add: Setcompr_eq_image) done lemma UINF_as_Sup_collect': "(\P \ f(P)) = (\P. f(P))" apply (simp add: upred_defs bop.rep_eq lit.rep_eq Sup_uexpr_def) apply (pred_simp) apply (simp add: full_SetCompr_eq) done lemma UINF_as_Sup_image: "(\ P | \P\ \\<^sub>u \A\ \ f(P)) = \ (f ` A)" apply (simp add: upred_defs bop.rep_eq lit.rep_eq Sup_uexpr_def) apply (pred_simp) apply (rule cong[of "Sup"]) apply (auto) done lemma USUP_as_Inf: "(\ P \ \

\ P) = \ \

" apply (simp add: upred_defs bop.rep_eq lit.rep_eq Inf_uexpr_def) apply (pred_simp) apply (rule cong[of "Inf"]) apply (auto) done lemma USUP_as_Inf_collect: "(\P\A \ f(P)) = (\P\A. f(P))" apply (pred_simp) apply (simp add: Setcompr_eq_image) done lemma USUP_as_Inf_collect': "(\P \ f(P)) = (\P. f(P))" apply (simp add: upred_defs bop.rep_eq lit.rep_eq Sup_uexpr_def) apply (pred_simp) apply (simp add: full_SetCompr_eq) done lemma USUP_as_Inf_image: "(\ P \ \

\ f(P)) = \ (f ` \

)" apply (simp add: upred_defs bop.rep_eq lit.rep_eq Inf_uexpr_def) apply (pred_simp) apply (rule cong[of "Inf"]) apply (auto) done lemma USUP_image_eq [simp]: "USUP (\i. \i\ \\<^sub>u \f ` A\) g = (\ i\A \ g(f(i)))" by (pred_simp, rule_tac cong[of Inf Inf], auto) lemma UINF_image_eq [simp]: "UINF (\i. \i\ \\<^sub>u \f ` A\) g = (\ i\A \ g(f(i)))" by (pred_simp, rule_tac cong[of Sup Sup], auto) lemma subst_continuous [usubst]: "\ \ (\ A) = (\ {\ \ P | P. P \ A})" by (simp add: UINF_as_Sup[THEN sym] usubst setcompr_eq_image) lemma not_UINF: "(\ (\ i\A\ P(i))) = (\ i\A\ \ P(i))" by (pred_auto) lemma not_USUP: "(\ (\ i\A\ P(i))) = (\ i\A\ \ P(i))" by (pred_auto) lemma not_UINF_ind: "(\ (\ i \ P(i))) = (\ i \ \ P(i))" by (pred_auto) lemma not_USUP_ind: "(\ (\ i \ P(i))) = (\ i \ \ P(i))" by (pred_auto) lemma UINF_empty [simp]: "(\ i \ {} \ P(i)) = false" by (pred_auto) lemma UINF_insert [simp]: "(\ i\insert x xs \ P(i)) = (P(x) \ (\ i\xs \ P(i)))" apply (pred_simp) apply (subst Sup_insert[THEN sym]) apply (rule_tac cong[of Sup Sup]) apply (auto) done lemma UINF_atLeast_first: "P(n) \ (\ i \ {Suc n..} \ P(i)) = (\ i \ {n..} \ P(i))" proof - have "insert n {Suc n..} = {n..}" by (auto) thus ?thesis by (metis UINF_insert) qed lemma UINF_atLeast_Suc: "(\ i \ {Suc m..} \ P(i)) = (\ i \ {m..} \ P(Suc i))" by (rel_simp, metis (full_types) Suc_le_D not_less_eq_eq) lemma USUP_empty [simp]: "(\ i \ {} \ P(i)) = true" by (pred_auto) lemma USUP_insert [simp]: "(\ i\insert x xs \ P(i)) = (P(x) \ (\ i\xs \ P(i)))" apply (pred_simp) apply (subst Inf_insert[THEN sym]) apply (rule_tac cong[of Inf Inf]) apply (auto) done lemma USUP_atLeast_first: "(P(n) \ (\ i \ {Suc n..} \ P(i))) = (\ i \ {n..} \ P(i))" proof - have "insert n {Suc n..} = {n..}" by (auto) thus ?thesis by (metis USUP_insert conj_upred_def) qed lemma USUP_atLeast_Suc: "(\ i \ {Suc m..} \ P(i)) = (\ i \ {m..} \ P(Suc i))" by (rel_simp, metis (full_types) Suc_le_D not_less_eq_eq) lemma conj_UINF_dist: "(P \ (\ Q\S \ F(Q))) = (\ Q\S \ P \ F(Q))" by (simp add: upred_defs bop.rep_eq lit.rep_eq, pred_auto) lemma conj_UINF_ind_dist: "(P \ (\ Q \ F(Q))) = (\ Q \ P \ F(Q))" by pred_auto lemma disj_UINF_dist: "S \ {} \ (P \ (\ Q\S \ F(Q))) = (\ Q\S \ P \ F(Q))" by (simp add: upred_defs bop.rep_eq lit.rep_eq, pred_auto) lemma UINF_conj_UINF [simp]: "((\ i\I \ P(i)) \ (\ i\I \ Q(i))) = (\ i\I \ P(i) \ Q(i))" by (rel_auto) lemma conj_USUP_dist: "S \ {} \ (P \ (\ Q\S \ F(Q))) = (\ Q\S \ P \ F(Q))" by (subst uexpr_eq_iff, auto simp add: conj_upred_def USUP.rep_eq inf_uexpr.rep_eq bop.rep_eq lit.rep_eq) lemma USUP_conj_USUP [simp]: "((\ P \ A \ F(P)) \ (\ P \ A \ G(P))) = (\ P \ A \ F(P) \ G(P))" by (simp add: upred_defs bop.rep_eq lit.rep_eq, pred_auto) lemma UINF_all_cong [cong]: assumes "\ P. F(P) = G(P)" shows "(\ P \ F(P)) = (\ P \ G(P))" by (simp add: UINF_as_Sup_collect assms) lemma UINF_cong: assumes "\ P. P \ A \ F(P) = G(P)" shows "(\ P\A \ F(P)) = (\ P\A \ G(P))" by (simp add: UINF_as_Sup_collect assms) lemma USUP_all_cong: assumes "\ P. F(P) = G(P)" shows "(\ P \ F(P)) = (\ P \ G(P))" by (simp add: assms) lemma USUP_cong: assumes "\ P. P \ A \ F(P) = G(P)" shows "(\ P\A \ F(P)) = (\ P\A \ G(P))" by (simp add: USUP_as_Inf_collect assms) lemma UINF_subset_mono: "A \ B \ (\ P\B \ F(P)) \ (\ P\A \ F(P))" by (simp add: SUP_subset_mono UINF_as_Sup_collect) lemma USUP_subset_mono: "A \ B \ (\ P\A \ F(P)) \ (\ P\B \ F(P))" by (simp add: INF_superset_mono USUP_as_Inf_collect) lemma UINF_impl: "(\ P\A \ F(P) \ G(P)) = ((\ P\A \ F(P)) \ (\ P\A \ G(P)))" by (pred_auto) lemma USUP_is_forall: "(\ x \ P(x)) = (\<^bold>\ x \ P(x))" by (pred_simp) lemma USUP_ind_is_forall: "(\ x\A \ P(x)) = (\<^bold>\ x\\A\ \ P(x))" by (pred_auto) lemma UINF_is_exists: "(\ x \ P(x)) = (\<^bold>\ x \ P(x))" by (pred_simp) lemma UINF_all_nats [simp]: fixes P :: "nat \ '\ upred" shows "(\ n \ \ i\{0..n} \ P(i)) = (\ n \ P(n))" by (pred_auto) lemma USUP_all_nats [simp]: fixes P :: "nat \ '\ upred" shows "(\ n \ \ i\{0..n} \ P(i)) = (\ n \ P(n))" by (pred_auto) lemma UINF_upto_expand_first: "m < n \ (\ i \ {m.. P(i)) = ((P(m) :: '\ upred) \ (\ i \ {Suc m.. P(i)))" apply (rel_auto) using Suc_leI le_eq_less_or_eq by auto lemma UINF_upto_expand_last: "(\ i \ {0.. P(i)) = ((\ i \ {0.. P(i)) \ P(n))" apply (rel_auto) using less_SucE by blast lemma UINF_Suc_shift: "(\ i \ {Suc 0.. P(i)) = (\ i \ {0.. P(Suc i))" apply (rel_simp) apply (rule cong[of Sup], auto) using less_Suc_eq_0_disj by auto lemma USUP_upto_expand_first: "(\ i \ {0.. P(i)) = (P(0) \ (\ i \ {1.. P(i)))" apply (rel_auto) using not_less by auto lemma USUP_Suc_shift: "(\ i \ {Suc 0.. P(i)) = (\ i \ {0.. P(Suc i))" apply (rel_simp) apply (rule cong[of Inf], auto) using less_Suc_eq_0_disj by auto lemma UINF_list_conv: "(\ i \ {0.. f (xs ! i)) = foldr (\) (map f xs) false" apply (induct xs) apply (rel_auto) apply (simp add: UINF_upto_expand_first UINF_Suc_shift) done lemma USUP_list_conv: "(\ i \ {0.. f (xs ! i)) = foldr (\) (map f xs) true" apply (induct xs) apply (rel_auto) apply (simp_all add: USUP_upto_expand_first USUP_Suc_shift) done lemma UINF_refines: "\ \ i. i\I \ P \ Q i \ \ P \ (\ i\I \ Q i)" by (simp add: UINF_as_Sup_collect, metis SUP_least) lemma UINF_refines': assumes "\ i. P \ Q(i)" shows "P \ (\ i \ Q(i))" using assms apply (rel_auto) using Sup_le_iff by fastforce lemma UINF_pred_ueq [simp]: "(\ x | \x\ =\<^sub>u v \ P(x)) = (P x)\x\v\" by (pred_auto) lemma UINF_pred_lit_eq [simp]: "(\ x | \x = v\ \ P(x)) = (P v)" by (pred_auto) subsection \ Equality laws \ lemma eq_upred_refl [simp]: "(x =\<^sub>u x) = true" by (pred_auto) lemma eq_upred_sym: "(x =\<^sub>u y) = (y =\<^sub>u x)" by (pred_auto) lemma eq_cong_left: assumes "vwb_lens x" "$x \ Q" "$x\ \ Q" "$x \ R" "$x\ \ R" shows "(($x\ =\<^sub>u $x \ Q) = ($x\ =\<^sub>u $x \ R)) \ (Q = R)" using assms by (pred_simp, (meson mwb_lens_def vwb_lens_mwb weak_lens_def)+) lemma conj_eq_in_var_subst: fixes x :: "('a \ '\)" assumes "vwb_lens x" shows "(P \ $x =\<^sub>u v) = (P\v/$x\ \ $x =\<^sub>u v)" using assms by (pred_simp, (metis vwb_lens_wb wb_lens.get_put)+) lemma conj_eq_out_var_subst: fixes x :: "('a \ '\)" assumes "vwb_lens x" shows "(P \ $x\ =\<^sub>u v) = (P\v/$x\\ \ $x\ =\<^sub>u v)" using assms by (pred_simp, (metis vwb_lens_wb wb_lens.get_put)+) lemma conj_pos_var_subst: assumes "vwb_lens x" shows "($x \ Q) = ($x \ Q\true/$x\)" using assms by (pred_auto, metis (full_types) vwb_lens_wb wb_lens.get_put, metis (full_types) vwb_lens_wb wb_lens.get_put) lemma conj_neg_var_subst: assumes "vwb_lens x" shows "(\ $x \ Q) = (\ $x \ Q\false/$x\)" using assms by (pred_auto, metis (full_types) vwb_lens_wb wb_lens.get_put, metis (full_types) vwb_lens_wb wb_lens.get_put) lemma upred_eq_true [simp]: "(p =\<^sub>u true) = p" by (pred_auto) lemma upred_eq_false [simp]: "(p =\<^sub>u false) = (\ p)" by (pred_auto) lemma upred_true_eq [simp]: "(true =\<^sub>u p) = p" by (pred_auto) lemma upred_false_eq [simp]: "(false =\<^sub>u p) = (\ p)" by (pred_auto) lemma conj_var_subst: assumes "vwb_lens x" shows "(P \ var x =\<^sub>u v) = (P\v/x\ \ var x =\<^sub>u v)" using assms by (pred_simp, (metis (full_types) vwb_lens_def wb_lens.get_put)+) subsection \ HOL Variable Quantifiers \ lemma shEx_unbound [simp]: "(\<^bold>\ x \ P) = P" by (pred_auto) lemma shEx_bool [simp]: "shEx P = (P True \ P False)" by (pred_simp, metis (full_types)) lemma shEx_commute: "(\<^bold>\ x \ \<^bold>\ y \ P x y) = (\<^bold>\ y \ \<^bold>\ x \ P x y)" by (pred_auto) lemma shEx_cong: "\ \ x. P x = Q x \ \ shEx P = shEx Q" by (pred_auto) lemma shEx_insert: "(\<^bold>\ x \ insert\<^sub>u y A \ P(x)) = (P(x)\x\y\ \ (\<^bold>\ x \ A \ P(x)))" by (pred_auto) lemma shEx_one_point: "(\<^bold>\ x \ \x\ =\<^sub>u v \ P(x)) = P(x)\x\v\" by (rel_auto) lemma shAll_unbound [simp]: "(\<^bold>\ x \ P) = P" by (pred_auto) lemma shAll_bool [simp]: "shAll P = (P True \ P False)" by (pred_simp, metis (full_types)) lemma shAll_cong: "\ \ x. P x = Q x \ \ shAll P = shAll Q" by (pred_auto) text \ Quantifier lifting \ named_theorems uquant_lift lemma shEx_lift_conj_1 [uquant_lift]: "((\<^bold>\ x \ P(x)) \ Q) = (\<^bold>\ x \ P(x) \ Q)" by (pred_auto) lemma shEx_lift_conj_2 [uquant_lift]: "(P \ (\<^bold>\ x \ Q(x))) = (\<^bold>\ x \ P \ Q(x))" by (pred_auto) subsection \ Case Splitting \ lemma eq_split_subst: assumes "vwb_lens x" shows "(P = Q) \ (\ v. P\\v\/x\ = Q\\v\/x\)" using assms by (pred_auto, metis vwb_lens_wb wb_lens.source_stability) lemma eq_split_substI: assumes "vwb_lens x" "\ v. P\\v\/x\ = Q\\v\/x\" shows "P = Q" using assms(1) assms(2) eq_split_subst by blast lemma taut_split_subst: assumes "vwb_lens x" shows "`P` \ (\ v. `P\\v\/x\`)" using assms by (pred_auto, metis vwb_lens_wb wb_lens.source_stability) lemma eq_split: assumes "`P \ Q`" "`Q \ P`" shows "P = Q" using assms by (pred_auto) lemma bool_eq_splitI: assumes "vwb_lens x" "P\true/x\ = Q\true/x\" "P\false/x\ = Q\false/x\" shows "P = Q" by (metis (full_types) assms eq_split_subst false_alt_def true_alt_def) lemma subst_bool_split: assumes "vwb_lens x" shows "`P` = `(P\false/x\ \ P\true/x\)`" proof - from assms have "`P` = (\ v. `P\\v\/x\`)" by (subst taut_split_subst[of x], auto) also have "... = (`P\\True\/x\` \ `P\\False\/x\`)" by (metis (mono_tags, lifting)) also have "... = `(P\false/x\ \ P\true/x\)`" by (pred_auto) finally show ?thesis . qed lemma subst_eq_replace: fixes x :: "('a \ '\)" shows "(p\u/x\ \ u =\<^sub>u v) = (p\v/x\ \ u =\<^sub>u v)" by (pred_auto) subsection \ UTP Quantifiers \ lemma one_point: assumes "mwb_lens x" "x \ v" shows "(\ x \ P \ var x =\<^sub>u v) = P\v/x\" using assms by (pred_auto) lemma exists_twice: "mwb_lens x \ (\ x \ \ x \ P) = (\ x \ P)" by (pred_auto) lemma all_twice: "mwb_lens x \ (\ x \ \ x \ P) = (\ x \ P)" by (pred_auto) lemma exists_sub: "\ mwb_lens y; x \\<^sub>L y \ \ (\ x \ \ y \ P) = (\ y \ P)" by (pred_auto) lemma all_sub: "\ mwb_lens y; x \\<^sub>L y \ \ (\ x \ \ y \ P) = (\ y \ P)" by (pred_auto) lemma ex_commute: assumes "x \ y" shows "(\ x \ \ y \ P) = (\ y \ \ x \ P)" using assms apply (pred_auto) using lens_indep_comm apply fastforce+ done lemma all_commute: assumes "x \ y" shows "(\ x \ \ y \ P) = (\ y \ \ x \ P)" using assms apply (pred_auto) using lens_indep_comm apply fastforce+ done lemma ex_equiv: assumes "x \\<^sub>L y" shows "(\ x \ P) = (\ y \ P)" using assms by (pred_simp, metis (no_types, lifting) lens.select_convs(2)) lemma all_equiv: assumes "x \\<^sub>L y" shows "(\ x \ P) = (\ y \ P)" using assms by (pred_simp, metis (no_types, lifting) lens.select_convs(2)) lemma ex_zero: "(\ \ \ P) = P" by (pred_auto) lemma all_zero: "(\ \ \ P) = P" by (pred_auto) lemma ex_plus: "(\ y;x \ P) = (\ x \ \ y \ P)" by (pred_auto) lemma all_plus: "(\ y;x \ P) = (\ x \ \ y \ P)" by (pred_auto) lemma closure_all: "[P]\<^sub>u = (\ \ \ P)" by (pred_auto) lemma unrest_as_exists: "vwb_lens x \ (x \ P) \ ((\ x \ P) = P)" by (pred_simp, metis vwb_lens.put_eq) lemma ex_mono: "P \ Q \ (\ x \ P) \ (\ x \ Q)" by (pred_auto) lemma ex_weakens: "wb_lens x \ (\ x \ P) \ P" by (pred_simp, metis wb_lens.get_put) lemma all_mono: "P \ Q \ (\ x \ P) \ (\ x \ Q)" by (pred_auto) lemma all_strengthens: "wb_lens x \ P \ (\ x \ P)" by (pred_simp, metis wb_lens.get_put) lemma ex_unrest: "x \ P \ (\ x \ P) = P" by (pred_auto) lemma all_unrest: "x \ P \ (\ x \ P) = P" by (pred_auto) lemma not_ex_not: "\ (\ x \ \ P) = (\ x \ P)" by (pred_auto) lemma not_all_not: "\ (\ x \ \ P) = (\ x \ P)" by (pred_auto) lemma ex_conj_contr_left: "x \ P \ (\ x \ P \ Q) = (P \ (\ x \ Q))" by (pred_auto) lemma ex_conj_contr_right: "x \ Q \ (\ x \ P \ Q) = ((\ x \ P) \ Q)" by (pred_auto) subsection \ Variable Restriction \ lemma var_res_all: "P \\<^sub>v \ = P" by (rel_auto) lemma var_res_twice: "mwb_lens x \ P \\<^sub>v x \\<^sub>v x = P \\<^sub>v x" by (pred_auto) subsection \ Conditional laws \ lemma cond_def: "(P \ b \ Q) = ((b \ P) \ ((\ b) \ Q))" by (pred_auto) lemma cond_idem [simp]:"(P \ b \ P) = P" by (pred_auto) lemma cond_true_false [simp]: "true \ b \ false = b" by (pred_auto) lemma cond_symm:"(P \ b \ Q) = (Q \ \ b \ P)" by (pred_auto) lemma cond_assoc: "((P \ b \ Q) \ c \ R) = (P \ b \ c \ (Q \ c \ R))" by (pred_auto) lemma cond_distr: "(P \ b \ (Q \ c \ R)) = ((P \ b \ Q) \ c \ (P \ b \ R))" by (pred_auto) lemma cond_unit_T [simp]:"(P \ true \ Q) = P" by (pred_auto) lemma cond_unit_F [simp]:"(P \ false \ Q) = Q" by (pred_auto) lemma cond_conj_not: "((P \ b \ Q) \ (\ b)) = (Q \ (\ b))" by (rel_auto) lemma cond_and_T_integrate: "((P \ b) \ (Q \ b \ R)) = ((P \ Q) \ b \ R)" by (pred_auto) lemma cond_L6: "(P \ b \ (Q \ b \ R)) = (P \ b \ R)" by (pred_auto) lemma cond_L7: "(P \ b \ (P \ c \ Q)) = (P \ b \ c \ Q)" by (pred_auto) lemma cond_and_distr: "((P \ Q) \ b \ (R \ S)) = ((P \ b \ R) \ (Q \ b \ S))" by (pred_auto) lemma cond_or_distr: "((P \ Q) \ b \ (R \ S)) = ((P \ b \ R) \ (Q \ b \ S))" by (pred_auto) lemma cond_imp_distr: "((P \ Q) \ b \ (R \ S)) = ((P \ b \ R) \ (Q \ b \ S))" by (pred_auto) lemma cond_eq_distr: "((P \ Q) \ b \ (R \ S)) = ((P \ b \ R) \ (Q \ b \ S))" by (pred_auto) lemma cond_conj_distr:"(P \ (Q \ b \ S)) = ((P \ Q) \ b \ (P \ S))" by (pred_auto) lemma cond_disj_distr:"(P \ (Q \ b \ S)) = ((P \ Q) \ b \ (P \ S))" by (pred_auto) lemma cond_neg: "\ (P \ b \ Q) = ((\ P) \ b \ (\ Q))" by (pred_auto) lemma cond_conj: "P \ b \ c \ Q = (P \ c \ Q) \ b \ Q" by (pred_auto) lemma spec_cond_dist: "(P \ (Q \ b \ R)) = ((P \ Q) \ b \ (P \ R))" by (pred_auto) lemma cond_USUP_dist: "(\ P\S \ F(P)) \ b \ (\ P\S \ G(P)) = (\ P\S \ F(P) \ b \ G(P))" by (pred_auto) lemma cond_UINF_dist: "(\ P\S \ F(P)) \ b \ (\ P\S \ G(P)) = (\ P\S \ F(P) \ b \ G(P))" by (pred_auto) lemma cond_var_subst_left: assumes "vwb_lens x" shows "(P\true/x\ \ var x \ Q) = (P \ var x \ Q)" using assms by (pred_auto, metis (full_types) vwb_lens_wb wb_lens.get_put) lemma cond_var_subst_right: assumes "vwb_lens x" shows "(P \ var x \ Q\false/x\) = (P \ var x \ Q)" using assms by (pred_auto, metis (full_types) vwb_lens.put_eq) lemma cond_var_split: "vwb_lens x \ (P\true/x\ \ var x \ P\false/x\) = P" by (rel_simp, (metis (full_types) vwb_lens.put_eq)+) lemma cond_assign_subst: "vwb_lens x \ (P \ utp_expr.var x =\<^sub>u v \ Q) = (P\v/x\ \ utp_expr.var x =\<^sub>u v \ Q)" apply (rel_simp) using vwb_lens.put_eq by force lemma conj_conds: "(P1 \ b \ Q1 \ P2 \ b \ Q2) = (P1 \ P2) \ b \ (Q1 \ Q2)" by pred_auto lemma disj_conds: "(P1 \ b \ Q1 \ P2 \ b \ Q2) = (P1 \ P2) \ b \ (Q1 \ Q2)" by pred_auto lemma cond_mono: "\ P\<^sub>1 \ P\<^sub>2; Q\<^sub>1 \ Q\<^sub>2 \ \ (P\<^sub>1 \ b \ Q\<^sub>1) \ (P\<^sub>2 \ b \ Q\<^sub>2)" by (rel_auto) lemma cond_monotonic: "\ mono P; mono Q \ \ mono (\ X. P X \ b \ Q X)" by (simp add: mono_def, rel_blast) subsection \ Additional Expression Laws \ lemma le_pred_refl [simp]: fixes x :: "('a::preorder, '\) uexpr" shows "(x \\<^sub>u x) = true" by (pred_auto) lemma uzero_le_laws [simp]: "(0 :: ('a::{linordered_semidom}, '\) uexpr) \\<^sub>u numeral x = true" "(1 :: ('a::{linordered_semidom}, '\) uexpr) \\<^sub>u numeral x = true" "(0 :: ('a::{linordered_semidom}, '\) uexpr) \\<^sub>u 1 = true" by (pred_simp)+ lemma unumeral_le_1 [simp]: assumes "(numeral i :: 'a::{numeral,ord}) \ numeral j" shows "(numeral i :: ('a, '\) uexpr) \\<^sub>u numeral j = true" using assms by (pred_auto) lemma unumeral_le_2 [simp]: assumes "(numeral i :: 'a::{numeral,linorder}) > numeral j" shows "(numeral i :: ('a, '\) uexpr) \\<^sub>u numeral j = false" using assms by (pred_auto) lemma uset_laws [simp]: "x \\<^sub>u {}\<^sub>u = false" "x \\<^sub>u {m..n}\<^sub>u = (m \\<^sub>u x \ x \\<^sub>u n)" by (pred_auto)+ lemma ulit_eq [simp]: "x = y \ (\x\ =\<^sub>u \y\) = true" by (rel_auto) lemma ulit_neq [simp]: "x \ y \ (\x\ =\<^sub>u \y\) = false" by (rel_auto) lemma uset_mems [simp]: "x \\<^sub>u {y}\<^sub>u = (x =\<^sub>u y)" "x \\<^sub>u A \\<^sub>u B = (x \\<^sub>u A \ x \\<^sub>u B)" "x \\<^sub>u A \\<^sub>u B = (x \\<^sub>u A \ x \\<^sub>u B)" by (rel_auto)+ subsection \ Refinement By Observation \ text \ Function to obtain the set of observations of a predicate \ definition obs_upred :: "'\ upred \ '\ set" ("\_\\<^sub>o") where [upred_defs]: "\P\\<^sub>o = {b. \P\\<^sub>eb}" lemma obs_upred_refine_iff: "P \ Q \ \Q\\<^sub>o \ \P\\<^sub>o" by (pred_auto) text \ A refinement can be demonstrated by considering only the observations of the predicates which are relevant, i.e. not unrestricted, for them. In other words, if the alphabet can be split into two disjoint segments, $x$ and $y$, and neither predicate refers to $y$ then only $x$ need be considered when checking for observations. \ lemma refine_by_obs: assumes "x \ y" "bij_lens (x +\<^sub>L y)" "y \ P" "y \ Q" "{v. `P\\v\/x\`} \ {v. `Q\\v\/x\`}" shows "Q \ P" using assms(3-5) apply (simp add: obs_upred_refine_iff subset_eq) apply (pred_simp) apply (rename_tac b) apply (drule_tac x="get\<^bsub>x\<^esub>b" in spec) apply (auto simp add: assms) apply (metis assms(1) assms(2) bij_lens.axioms(2) bij_lens_axioms_def lens_override_def lens_override_plus)+ done subsection \ Cylindric Algebra \ lemma C1: "(\ x \ false) = false" by (pred_auto) lemma C2: "wb_lens x \ `P \ (\ x \ P)`" by (pred_simp, metis wb_lens.get_put) lemma C3: "mwb_lens x \ (\ x \ (P \ (\ x \ Q))) = ((\ x \ P) \ (\ x \ Q))" by (pred_auto) lemma C4a: "x \\<^sub>L y \ (\ x \ \ y \ P) = (\ y \ \ x \ P)" by (pred_simp, metis (no_types, lifting) lens.select_convs(2))+ lemma C4b: "x \ y \ (\ x \ \ y \ P) = (\ y \ \ x \ P)" using ex_commute by blast lemma C5: fixes x :: "('a \ '\)" shows "(&x =\<^sub>u &x) = true" by (pred_auto) lemma C6: assumes "wb_lens x" "x \ y" "x \ z" shows "(&y =\<^sub>u &z) = (\ x \ &y =\<^sub>u &x \ &x =\<^sub>u &z)" using assms by (pred_simp, (metis lens_indep_def)+) lemma C7: assumes "weak_lens x" "x \ y" shows "((\ x \ &x =\<^sub>u &y \ P) \ (\ x \ &x =\<^sub>u &y \ \ P)) = false" using assms by (pred_simp, simp add: lens_indep_sym) end \ No newline at end of file diff --git a/thys/Word_Lib/Bits_Int.thy b/thys/Word_Lib/Bits_Int.thy --- a/thys/Word_Lib/Bits_Int.thy +++ b/thys/Word_Lib/Bits_Int.thy @@ -1,1573 +1,1573 @@ (* * Copyright Brian Huffman, PSU; Jeremy Dawson and Gerwin Klein, NICTA * * SPDX-License-Identifier: BSD-2-Clause *) section \Bitwise Operations on integers\ theory Bits_Int imports "Word_Lib.Most_significant_bit" "Word_Lib.Least_significant_bit" "Word_Lib.Generic_set_bit" "Word_Lib.Bit_Comprehension" begin subsection \Implicit bit representation of \<^typ>\int\\ lemma bin_last_def: "(odd :: int \ bool) w \ w mod 2 = 1" by (fact odd_iff_mod_2_eq_one) lemma bin_last_numeral_simps [simp]: "\ odd (0 :: int)" "odd (1 :: int)" "odd (- 1 :: int)" "odd (Numeral1 :: int)" "\ odd (numeral (Num.Bit0 w) :: int)" "odd (numeral (Num.Bit1 w) :: int)" "\ odd (- numeral (Num.Bit0 w) :: int)" "odd (- numeral (Num.Bit1 w) :: int)" by simp_all lemma bin_rest_numeral_simps [simp]: "(\k::int. k div 2) 0 = 0" "(\k::int. k div 2) 1 = 0" "(\k::int. k div 2) (- 1) = - 1" "(\k::int. k div 2) Numeral1 = 0" "(\k::int. k div 2) (numeral (Num.Bit0 w)) = numeral w" "(\k::int. k div 2) (numeral (Num.Bit1 w)) = numeral w" "(\k::int. k div 2) (- numeral (Num.Bit0 w)) = - numeral w" "(\k::int. k div 2) (- numeral (Num.Bit1 w)) = - numeral (w + Num.One)" by simp_all lemma bin_rl_eqI: "\(\k::int. k div 2) x = (\k::int. k div 2) y; odd x = odd y\ \ x = y" by (auto elim: oddE) lemma [simp]: shows bin_rest_lt0: "(\k::int. k div 2) i < 0 \ i < 0" and bin_rest_ge_0: "(\k::int. k div 2) i \ 0 \ i \ 0" by auto lemma bin_rest_gt_0 [simp]: "(\k::int. k div 2) x > 0 \ x > 1" by auto subsection \Bit projection\ lemma bin_nth_eq_iff: "(bit :: int \ nat \ bool) x = (bit :: int \ nat \ bool) y \ x = y" by (simp add: bit_eq_iff fun_eq_iff) lemma bin_eqI: "x = y" if "\n. (bit :: int \ nat \ bool) x n \ (bit :: int \ nat \ bool) y n" using that by (rule bit_eqI) lemma bin_eq_iff: "x = y \ (\n. (bit :: int \ nat \ bool) x n = (bit :: int \ nat \ bool) y n)" by (fact bit_eq_iff) lemma bin_nth_zero [simp]: "\ (bit :: int \ nat \ bool) 0 n" by simp lemma bin_nth_1 [simp]: "(bit :: int \ nat \ bool) 1 n \ n = 0" by (cases n) (simp_all add: bit_Suc) lemma bin_nth_minus1 [simp]: "(bit :: int \ nat \ bool) (- 1) n" by simp lemma bin_nth_numeral: "(\k::int. k div 2) x = y \ (bit :: int \ nat \ bool) x (numeral n) = (bit :: int \ nat \ bool) y (pred_numeral n)" by (simp add: numeral_eq_Suc bit_Suc) lemmas bin_nth_numeral_simps [simp] = bin_nth_numeral [OF bin_rest_numeral_simps(8)] lemmas bin_nth_simps = bit_0 bit_Suc bin_nth_zero bin_nth_minus1 bin_nth_numeral_simps lemma nth_2p_bin: "(bit :: int \ nat \ bool) (2 ^ n) m = (m = n)" \ \for use when simplifying with \bin_nth_Bit\\ by (auto simp add: bit_exp_iff) lemma nth_rest_power_bin: "(bit :: int \ nat \ bool) (((\k::int. k div 2) ^^ k) w) n = (bit :: int \ nat \ bool) w (n + k)" apply (induct k arbitrary: n) apply clarsimp apply clarsimp apply (simp only: bit_Suc [symmetric] add_Suc) done lemma bin_nth_numeral_unfold: "(bit :: int \ nat \ bool) (numeral (num.Bit0 x)) n \ n > 0 \ (bit :: int \ nat \ bool) (numeral x) (n - 1)" "(bit :: int \ nat \ bool) (numeral (num.Bit1 x)) n \ (n > 0 \ (bit :: int \ nat \ bool) (numeral x) (n - 1))" by (cases n; simp)+ subsection \Truncating\ definition bin_sign :: "int \ int" where "bin_sign k = (if k \ 0 then 0 else - 1)" lemma bin_sign_simps [simp]: "bin_sign 0 = 0" "bin_sign 1 = 0" "bin_sign (- 1) = - 1" "bin_sign (numeral k) = 0" "bin_sign (- numeral k) = -1" by (simp_all add: bin_sign_def) lemma bin_sign_rest [simp]: "bin_sign ((\k::int. k div 2) w) = bin_sign w" by (simp add: bin_sign_def) lemma bintrunc_mod2p: "(take_bit :: nat \ int \ int) n w = w mod 2 ^ n" by (fact take_bit_eq_mod) lemma sbintrunc_mod2p: "(signed_take_bit :: nat \ int \ int) n w = (w + 2 ^ n) mod 2 ^ Suc n - 2 ^ n" by (simp add: bintrunc_mod2p signed_take_bit_eq_take_bit_shift) lemma sbintrunc_eq_take_bit: \(signed_take_bit :: nat \ int \ int) n k = take_bit (Suc n) (k + 2 ^ n) - 2 ^ n\ by (fact signed_take_bit_eq_take_bit_shift) lemma sign_bintr: "bin_sign ((take_bit :: nat \ int \ int) n w) = 0" by (simp add: bin_sign_def) lemma bintrunc_n_0: "(take_bit :: nat \ int \ int) n 0 = 0" by (fact take_bit_of_0) lemma sbintrunc_n_0: "(signed_take_bit :: nat \ int \ int) n 0 = 0" by (fact signed_take_bit_of_0) lemma sbintrunc_n_minus1: "(signed_take_bit :: nat \ int \ int) n (- 1) = -1" by (fact signed_take_bit_of_minus_1) lemma bintrunc_Suc_numeral: "(take_bit :: nat \ int \ int) (Suc n) 1 = 1" "(take_bit :: nat \ int \ int) (Suc n) (- 1) = 1 + 2 * (take_bit :: nat \ int \ int) n (- 1)" "(take_bit :: nat \ int \ int) (Suc n) (numeral (Num.Bit0 w)) = 2 * (take_bit :: nat \ int \ int) n (numeral w)" "(take_bit :: nat \ int \ int) (Suc n) (numeral (Num.Bit1 w)) = 1 + 2 * (take_bit :: nat \ int \ int) n (numeral w)" "(take_bit :: nat \ int \ int) (Suc n) (- numeral (Num.Bit0 w)) = 2 * (take_bit :: nat \ int \ int) n (- numeral w)" "(take_bit :: nat \ int \ int) (Suc n) (- numeral (Num.Bit1 w)) = 1 + 2 * (take_bit :: nat \ int \ int) n (- numeral (w + Num.One))" by (simp_all add: take_bit_Suc) lemma sbintrunc_0_numeral [simp]: "(signed_take_bit :: nat \ int \ int) 0 1 = -1" "(signed_take_bit :: nat \ int \ int) 0 (numeral (Num.Bit0 w)) = 0" "(signed_take_bit :: nat \ int \ int) 0 (numeral (Num.Bit1 w)) = -1" "(signed_take_bit :: nat \ int \ int) 0 (- numeral (Num.Bit0 w)) = 0" "(signed_take_bit :: nat \ int \ int) 0 (- numeral (Num.Bit1 w)) = -1" by simp_all lemma sbintrunc_Suc_numeral: "(signed_take_bit :: nat \ int \ int) (Suc n) 1 = 1" "(signed_take_bit :: nat \ int \ int) (Suc n) (numeral (Num.Bit0 w)) = 2 * (signed_take_bit :: nat \ int \ int) n (numeral w)" "(signed_take_bit :: nat \ int \ int) (Suc n) (numeral (Num.Bit1 w)) = 1 + 2 * (signed_take_bit :: nat \ int \ int) n (numeral w)" "(signed_take_bit :: nat \ int \ int) (Suc n) (- numeral (Num.Bit0 w)) = 2 * (signed_take_bit :: nat \ int \ int) n (- numeral w)" "(signed_take_bit :: nat \ int \ int) (Suc n) (- numeral (Num.Bit1 w)) = 1 + 2 * (signed_take_bit :: nat \ int \ int) n (- numeral (w + Num.One))" by (simp_all add: signed_take_bit_Suc) lemma bin_sign_lem: "(bin_sign ((signed_take_bit :: nat \ int \ int) n bin) = -1) = bit bin n" by (simp add: bin_sign_def) lemma nth_bintr: "(bit :: int \ nat \ bool) ((take_bit :: nat \ int \ int) m w) n \ n < m \ (bit :: int \ nat \ bool) w n" by (fact bit_take_bit_iff) lemma nth_sbintr: "(bit :: int \ nat \ bool) ((signed_take_bit :: nat \ int \ int) m w) n = (if n < m then (bit :: int \ nat \ bool) w n else (bit :: int \ nat \ bool) w m)" by (simp add: bit_signed_take_bit_iff min_def) lemma bin_nth_Bit0: "(bit :: int \ nat \ bool) (numeral (Num.Bit0 w)) n \ (\m. n = Suc m \ (bit :: int \ nat \ bool) (numeral w) m)" using bit_double_iff [of \numeral w :: int\ n] by (auto intro: exI [of _ \n - 1\]) lemma bin_nth_Bit1: "(bit :: int \ nat \ bool) (numeral (Num.Bit1 w)) n \ n = 0 \ (\m. n = Suc m \ (bit :: int \ nat \ bool) (numeral w) m)" using even_bit_succ_iff [of \2 * numeral w :: int\ n] bit_double_iff [of \numeral w :: int\ n] by auto lemma bintrunc_bintrunc_l: "n \ m \ (take_bit :: nat \ int \ int) m ((take_bit :: nat \ int \ int) n w) = (take_bit :: nat \ int \ int) n w" by (simp add: min.absorb2) lemma sbintrunc_sbintrunc_l: "n \ m \ (signed_take_bit :: nat \ int \ int) m ((signed_take_bit :: nat \ int \ int) n w) = (signed_take_bit :: nat \ int \ int) n w" by (simp add: min.absorb2) lemma bintrunc_bintrunc_ge: "n \ m \ (take_bit :: nat \ int \ int) n ((take_bit :: nat \ int \ int) m w) = (take_bit :: nat \ int \ int) n w" by (rule bin_eqI) (auto simp: nth_bintr) lemma bintrunc_bintrunc_min [simp]: "(take_bit :: nat \ int \ int) m ((take_bit :: nat \ int \ int) n w) = (take_bit :: nat \ int \ int) (min m n) w" by (rule take_bit_take_bit) lemma sbintrunc_sbintrunc_min [simp]: "(signed_take_bit :: nat \ int \ int) m ((signed_take_bit :: nat \ int \ int) n w) = (signed_take_bit :: nat \ int \ int) (min m n) w" by (rule signed_take_bit_signed_take_bit) lemmas sbintrunc_Suc_Pls = signed_take_bit_Suc [where a="0::int", simplified bin_last_numeral_simps bin_rest_numeral_simps] lemmas sbintrunc_Suc_Min = signed_take_bit_Suc [where a="-1::int", simplified bin_last_numeral_simps bin_rest_numeral_simps] lemmas sbintrunc_Sucs = sbintrunc_Suc_Pls sbintrunc_Suc_Min sbintrunc_Suc_numeral lemmas sbintrunc_Pls = signed_take_bit_0 [where a="0::int", simplified bin_last_numeral_simps bin_rest_numeral_simps] lemmas sbintrunc_Min = signed_take_bit_0 [where a="-1::int", simplified bin_last_numeral_simps bin_rest_numeral_simps] lemmas sbintrunc_0_simps = sbintrunc_Pls sbintrunc_Min lemmas sbintrunc_simps = sbintrunc_0_simps sbintrunc_Sucs lemma bintrunc_minus: "0 < n \ (take_bit :: nat \ int \ int) (Suc (n - 1)) w = (take_bit :: nat \ int \ int) n w" by auto lemma sbintrunc_minus: "0 < n \ (signed_take_bit :: nat \ int \ int) (Suc (n - 1)) w = (signed_take_bit :: nat \ int \ int) n w" by auto lemmas sbintrunc_minus_simps = sbintrunc_Sucs [THEN [2] sbintrunc_minus [symmetric, THEN trans]] lemma sbintrunc_BIT_I: \0 < n \ (signed_take_bit :: nat \ int \ int) (n - 1) 0 = y \ (signed_take_bit :: nat \ int \ int) n 0 = 2 * y\ by simp lemma sbintrunc_Suc_Is: \(signed_take_bit :: nat \ int \ int) n (- 1) = y \ (signed_take_bit :: nat \ int \ int) (Suc n) (- 1) = 1 + 2 * y\ by auto lemma sbintrunc_Suc_lem: "(signed_take_bit :: nat \ int \ int) (Suc n) x = y \ m = Suc n \ (signed_take_bit :: nat \ int \ int) m x = y" by (rule ssubst) lemmas sbintrunc_Suc_Ialts = sbintrunc_Suc_Is [THEN sbintrunc_Suc_lem] lemma sbintrunc_bintrunc_lt: "m > n \ (signed_take_bit :: nat \ int \ int) n ((take_bit :: nat \ int \ int) m w) = (signed_take_bit :: nat \ int \ int) n w" by (rule bin_eqI) (auto simp: nth_sbintr nth_bintr) lemma bintrunc_sbintrunc_le: "m \ Suc n \ (take_bit :: nat \ int \ int) m ((signed_take_bit :: nat \ int \ int) n w) = (take_bit :: nat \ int \ int) m w" by (rule take_bit_signed_take_bit) lemmas bintrunc_sbintrunc [simp] = order_refl [THEN bintrunc_sbintrunc_le] lemmas sbintrunc_bintrunc [simp] = lessI [THEN sbintrunc_bintrunc_lt] lemmas bintrunc_bintrunc [simp] = order_refl [THEN bintrunc_bintrunc_l] lemmas sbintrunc_sbintrunc [simp] = order_refl [THEN sbintrunc_sbintrunc_l] lemma bintrunc_sbintrunc' [simp]: "0 < n \ (take_bit :: nat \ int \ int) n ((signed_take_bit :: nat \ int \ int) (n - 1) w) = (take_bit :: nat \ int \ int) n w" by (cases n) simp_all lemma sbintrunc_bintrunc' [simp]: "0 < n \ (signed_take_bit :: nat \ int \ int) (n - 1) ((take_bit :: nat \ int \ int) n w) = (signed_take_bit :: nat \ int \ int) (n - 1) w" by (cases n) simp_all lemma bin_sbin_eq_iff: "(take_bit :: nat \ int \ int) (Suc n) x = (take_bit :: nat \ int \ int) (Suc n) y \ (signed_take_bit :: nat \ int \ int) n x = (signed_take_bit :: nat \ int \ int) n y" apply (rule iffI) apply (rule box_equals [OF _ sbintrunc_bintrunc sbintrunc_bintrunc]) apply simp apply (rule box_equals [OF _ bintrunc_sbintrunc bintrunc_sbintrunc]) apply simp done lemma bin_sbin_eq_iff': "0 < n \ (take_bit :: nat \ int \ int) n x = (take_bit :: nat \ int \ int) n y \ (signed_take_bit :: nat \ int \ int) (n - 1) x = (signed_take_bit :: nat \ int \ int) (n - 1) y" by (cases n) (simp_all add: bin_sbin_eq_iff) lemmas bintrunc_sbintruncS0 [simp] = bintrunc_sbintrunc' [unfolded One_nat_def] lemmas sbintrunc_bintruncS0 [simp] = sbintrunc_bintrunc' [unfolded One_nat_def] lemmas bintrunc_bintrunc_l' = le_add1 [THEN bintrunc_bintrunc_l] lemmas sbintrunc_sbintrunc_l' = le_add1 [THEN sbintrunc_sbintrunc_l] (* although bintrunc_minus_simps, if added to default simpset, tends to get applied where it's not wanted in developing the theories, we get a version for when the word length is given literally *) lemmas nat_non0_gr = trans [OF iszero_def [THEN Not_eq_iff [THEN iffD2]] refl] lemma bintrunc_numeral: "(take_bit :: nat \ int \ int) (numeral k) x = of_bool (odd x) + 2 * (take_bit :: nat \ int \ int) (pred_numeral k) (x div 2)" by (simp add: numeral_eq_Suc take_bit_Suc mod_2_eq_odd) lemma sbintrunc_numeral: "(signed_take_bit :: nat \ int \ int) (numeral k) x = of_bool (odd x) + 2 * (signed_take_bit :: nat \ int \ int) (pred_numeral k) (x div 2)" by (simp add: numeral_eq_Suc signed_take_bit_Suc mod2_eq_if) lemma bintrunc_numeral_simps [simp]: "(take_bit :: nat \ int \ int) (numeral k) (numeral (Num.Bit0 w)) = 2 * (take_bit :: nat \ int \ int) (pred_numeral k) (numeral w)" "(take_bit :: nat \ int \ int) (numeral k) (numeral (Num.Bit1 w)) = 1 + 2 * (take_bit :: nat \ int \ int) (pred_numeral k) (numeral w)" "(take_bit :: nat \ int \ int) (numeral k) (- numeral (Num.Bit0 w)) = 2 * (take_bit :: nat \ int \ int) (pred_numeral k) (- numeral w)" "(take_bit :: nat \ int \ int) (numeral k) (- numeral (Num.Bit1 w)) = 1 + 2 * (take_bit :: nat \ int \ int) (pred_numeral k) (- numeral (w + Num.One))" "(take_bit :: nat \ int \ int) (numeral k) 1 = 1" by (simp_all add: bintrunc_numeral) lemma sbintrunc_numeral_simps [simp]: "(signed_take_bit :: nat \ int \ int) (numeral k) (numeral (Num.Bit0 w)) = 2 * (signed_take_bit :: nat \ int \ int) (pred_numeral k) (numeral w)" "(signed_take_bit :: nat \ int \ int) (numeral k) (numeral (Num.Bit1 w)) = 1 + 2 * (signed_take_bit :: nat \ int \ int) (pred_numeral k) (numeral w)" "(signed_take_bit :: nat \ int \ int) (numeral k) (- numeral (Num.Bit0 w)) = 2 * (signed_take_bit :: nat \ int \ int) (pred_numeral k) (- numeral w)" "(signed_take_bit :: nat \ int \ int) (numeral k) (- numeral (Num.Bit1 w)) = 1 + 2 * (signed_take_bit :: nat \ int \ int) (pred_numeral k) (- numeral (w + Num.One))" "(signed_take_bit :: nat \ int \ int) (numeral k) 1 = 1" by (simp_all add: sbintrunc_numeral) lemma no_bintr_alt1: "(take_bit :: nat \ int \ int) n = (\w. w mod 2 ^ n :: int)" by (rule ext) (rule bintrunc_mod2p) lemma range_bintrunc: "range ((take_bit :: nat \ int \ int) n) = {i. 0 \ i \ i < 2 ^ n}" by (auto simp add: take_bit_eq_mod image_iff) (metis mod_pos_pos_trivial) lemma no_sbintr_alt2: "(signed_take_bit :: nat \ int \ int) n = (\w. (w + 2 ^ n) mod 2 ^ Suc n - 2 ^ n :: int)" by (rule ext) (simp add : sbintrunc_mod2p) lemma range_sbintrunc: "range ((signed_take_bit :: nat \ int \ int) n) = {i. - (2 ^ n) \ i \ i < 2 ^ n}" proof - have \surj (\k::int. k + 2 ^ n)\ by (rule surjI [of _ \(\k. k - 2 ^ n)\]) simp moreover have \(signed_take_bit :: nat \ int \ int) n = ((\k. k - 2 ^ n) \ take_bit (Suc n) \ (\k. k + 2 ^ n))\ by (simp add: sbintrunc_eq_take_bit fun_eq_iff) ultimately show ?thesis apply (simp only: fun.set_map range_bintrunc) apply (auto simp add: image_iff) apply presburger done qed lemma sbintrunc_inc: \k + 2 ^ Suc n \ (signed_take_bit :: nat \ int \ int) n k\ if \k < - (2 ^ n)\ using that by (fact signed_take_bit_int_greater_eq) lemma sbintrunc_dec: \(signed_take_bit :: nat \ int \ int) n k \ k - 2 ^ (Suc n)\ if \k \ 2 ^ n\ using that by (fact signed_take_bit_int_less_eq) lemma bintr_ge0: "0 \ (take_bit :: nat \ int \ int) n w" by (simp add: bintrunc_mod2p) lemma bintr_lt2p: "(take_bit :: nat \ int \ int) n w < 2 ^ n" by (simp add: bintrunc_mod2p) lemma bintr_Min: "(take_bit :: nat \ int \ int) n (- 1) = 2 ^ n - 1" by (simp add: stable_imp_take_bit_eq) lemma sbintr_ge: "- (2 ^ n) \ (signed_take_bit :: nat \ int \ int) n w" by (fact signed_take_bit_int_greater_eq_minus_exp) lemma sbintr_lt: "(signed_take_bit :: nat \ int \ int) n w < 2 ^ n" by (fact signed_take_bit_int_less_exp) lemma sign_Pls_ge_0: "bin_sign bin = 0 \ bin \ 0" for bin :: int by (simp add: bin_sign_def) lemma sign_Min_lt_0: "bin_sign bin = -1 \ bin < 0" for bin :: int by (simp add: bin_sign_def) lemma bin_rest_trunc: "(\k::int. k div 2) ((take_bit :: nat \ int \ int) n bin) = (take_bit :: nat \ int \ int) (n - 1) ((\k::int. k div 2) bin)" by (simp add: take_bit_rec [of n bin]) lemma bin_rest_power_trunc: "((\k::int. k div 2) ^^ k) ((take_bit :: nat \ int \ int) n bin) = (take_bit :: nat \ int \ int) (n - k) (((\k::int. k div 2) ^^ k) bin)" by (induct k) (auto simp: bin_rest_trunc) lemma bin_rest_trunc_i: "(take_bit :: nat \ int \ int) n ((\k::int. k div 2) bin) = (\k::int. k div 2) ((take_bit :: nat \ int \ int) (Suc n) bin)" by (auto simp add: take_bit_Suc) lemma bin_rest_strunc: "(\k::int. k div 2) ((signed_take_bit :: nat \ int \ int) (Suc n) bin) = (signed_take_bit :: nat \ int \ int) n ((\k::int. k div 2) bin)" by (simp add: signed_take_bit_Suc) lemma bintrunc_rest [simp]: "(take_bit :: nat \ int \ int) n ((\k::int. k div 2) ((take_bit :: nat \ int \ int) n bin)) = (\k::int. k div 2) ((take_bit :: nat \ int \ int) n bin)" by (induct n arbitrary: bin) (simp_all add: take_bit_Suc) lemma sbintrunc_rest [simp]: "(signed_take_bit :: nat \ int \ int) n ((\k::int. k div 2) ((signed_take_bit :: nat \ int \ int) n bin)) = (\k::int. k div 2) ((signed_take_bit :: nat \ int \ int) n bin)" by (induct n arbitrary: bin) (simp_all add: signed_take_bit_Suc mod2_eq_if) lemma bintrunc_rest': "(take_bit :: nat \ int \ int) n \ (\k::int. k div 2) \ (take_bit :: nat \ int \ int) n = (\k::int. k div 2) \ (take_bit :: nat \ int \ int) n" by (rule ext) auto lemma sbintrunc_rest': "(signed_take_bit :: nat \ int \ int) n \ (\k::int. k div 2) \ (signed_take_bit :: nat \ int \ int) n = (\k::int. k div 2) \ (signed_take_bit :: nat \ int \ int) n" by (rule ext) auto lemma rco_lem: "f \ g \ f = g \ f \ f \ (g \ f) ^^ n = g ^^ n \ f" apply (rule ext) apply (induct_tac n) apply (simp_all (no_asm)) apply (drule fun_cong) apply (unfold o_def) apply (erule trans) apply simp done lemmas rco_bintr = bintrunc_rest' [THEN rco_lem [THEN fun_cong], unfolded o_def] lemmas rco_sbintr = sbintrunc_rest' [THEN rco_lem [THEN fun_cong], unfolded o_def] subsection \Splitting and concatenation\ definition bin_split :: \nat \ int \ int \ int\ where [simp]: \bin_split n k = (drop_bit n k, take_bit n k)\ lemma [code]: "bin_split (Suc n) w = (let (w1, w2) = bin_split n (w div 2) in (w1, of_bool (odd w) + 2 * w2))" "bin_split 0 w = (w, 0)" by (simp_all add: drop_bit_Suc take_bit_Suc mod_2_eq_odd) lemma bin_cat_eq_push_bit_add_take_bit: \concat_bit n l k = push_bit n k + take_bit n l\ by (simp add: concat_bit_eq) lemma bin_sign_cat: "bin_sign ((\k n l. concat_bit n l k) x n y) = bin_sign x" proof - have \0 \ x\ if \0 \ x * 2 ^ n + y mod 2 ^ n\ proof - have \y mod 2 ^ n < 2 ^ n\ using pos_mod_bound [of \2 ^ n\ y] by simp then have \\ y mod 2 ^ n \ 2 ^ n\ by (simp add: less_le) with that have \x \ - 1\ by auto have *: \- 1 \ (- (y mod 2 ^ n)) div 2 ^ n\ by (simp add: zdiv_zminus1_eq_if) from that have \- (y mod 2 ^ n) \ x * 2 ^ n\ by simp then have \(- (y mod 2 ^ n)) div 2 ^ n \ (x * 2 ^ n) div 2 ^ n\ using zdiv_mono1 zero_less_numeral zero_less_power by blast with * have \- 1 \ x * 2 ^ n div 2 ^ n\ by simp with \x \ - 1\ show ?thesis by simp qed then show ?thesis by (simp add: bin_sign_def not_le not_less bin_cat_eq_push_bit_add_take_bit push_bit_eq_mult take_bit_eq_mod) qed lemma bin_cat_assoc: "(\k n l. concat_bit n l k) ((\k n l. concat_bit n l k) x m y) n z = (\k n l. concat_bit n l k) x (m + n) ((\k n l. concat_bit n l k) y n z)" by (fact concat_bit_assoc) lemma bin_cat_assoc_sym: "(\k n l. concat_bit n l k) x m ((\k n l. concat_bit n l k) y n z) = (\k n l. concat_bit n l k) ((\k n l. concat_bit n l k) x (m - n) y) (min m n) z" by (fact concat_bit_assoc_sym) definition bin_rcat :: \nat \ int list \ int\ where \bin_rcat n = horner_sum (take_bit n) (2 ^ n) \ rev\ lemma bin_rcat_eq_foldl: \bin_rcat n = foldl (\u v. (\k n l. concat_bit n l k) u n v) 0\ proof fix ks :: \int list\ show \bin_rcat n ks = foldl (\u v. (\k n l. concat_bit n l k) u n v) 0 ks\ by (induction ks rule: rev_induct) (simp_all add: bin_rcat_def concat_bit_eq push_bit_eq_mult) qed fun bin_rsplit_aux :: "nat \ nat \ int \ int list \ int list" where "bin_rsplit_aux n m c bs = (if m = 0 \ n = 0 then bs else let (a, b) = bin_split n c in bin_rsplit_aux n (m - n) a (b # bs))" definition bin_rsplit :: "nat \ nat \ int \ int list" where "bin_rsplit n w = bin_rsplit_aux n (fst w) (snd w) []" fun bin_rsplitl_aux :: "nat \ nat \ int \ int list \ int list" where "bin_rsplitl_aux n m c bs = (if m = 0 \ n = 0 then bs else let (a, b) = bin_split (min m n) c in bin_rsplitl_aux n (m - n) a (b # bs))" definition bin_rsplitl :: "nat \ nat \ int \ int list" where "bin_rsplitl n w = bin_rsplitl_aux n (fst w) (snd w) []" declare bin_rsplit_aux.simps [simp del] declare bin_rsplitl_aux.simps [simp del] lemma bin_nth_cat: "(bit :: int \ nat \ bool) ((\k n l. concat_bit n l k) x k y) n = (if n < k then (bit :: int \ nat \ bool) y n else (bit :: int \ nat \ bool) x (n - k))" by (simp add: bit_concat_bit_iff) lemma bin_nth_drop_bit_iff: \(bit :: int \ nat \ bool) (drop_bit n c) k \ (bit :: int \ nat \ bool) c (n + k)\ by (simp add: bit_drop_bit_eq) lemma bin_nth_take_bit_iff: \(bit :: int \ nat \ bool) (take_bit n c) k \ k < n \ (bit :: int \ nat \ bool) c k\ by (fact bit_take_bit_iff) lemma bin_nth_split: "bin_split n c = (a, b) \ (\k. (bit :: int \ nat \ bool) a k = (bit :: int \ nat \ bool) c (n + k)) \ (\k. (bit :: int \ nat \ bool) b k = (k < n \ (bit :: int \ nat \ bool) c k))" by (auto simp add: bin_nth_drop_bit_iff bin_nth_take_bit_iff) lemma bin_cat_zero [simp]: "(\k n l. concat_bit n l k) 0 n w = (take_bit :: nat \ int \ int) n w" by (simp add: bin_cat_eq_push_bit_add_take_bit) lemma bintr_cat1: "(take_bit :: nat \ int \ int) (k + n) ((\k n l. concat_bit n l k) a n b) = (\k n l. concat_bit n l k) ((take_bit :: nat \ int \ int) k a) n b" by (metis bin_cat_assoc bin_cat_zero) lemma bintr_cat: "(take_bit :: nat \ int \ int) m ((\k n l. concat_bit n l k) a n b) = (\k n l. concat_bit n l k) ((take_bit :: nat \ int \ int) (m - n) a) n ((take_bit :: nat \ int \ int) (min m n) b)" by (rule bin_eqI) (auto simp: bin_nth_cat nth_bintr) lemma bintr_cat_same [simp]: "(take_bit :: nat \ int \ int) n ((\k n l. concat_bit n l k) a n b) = (take_bit :: nat \ int \ int) n b" by (auto simp add : bintr_cat) lemma cat_bintr [simp]: "(\k n l. concat_bit n l k) a n ((take_bit :: nat \ int \ int) n b) = (\k n l. concat_bit n l k) a n b" by (simp add: bin_cat_eq_push_bit_add_take_bit) lemma split_bintrunc: "bin_split n c = (a, b) \ b = (take_bit :: nat \ int \ int) n c" by simp lemma bin_cat_split: "bin_split n w = (u, v) \ w = (\k n l. concat_bit n l k) u n v" by (auto simp add: bin_cat_eq_push_bit_add_take_bit bits_ident) lemma drop_bit_bin_cat_eq: \drop_bit n ((\k n l. concat_bit n l k) v n w) = v\ by (rule bit_eqI) (simp add: bit_drop_bit_eq bit_concat_bit_iff) lemma take_bit_bin_cat_eq: \take_bit n ((\k n l. concat_bit n l k) v n w) = take_bit n w\ by (rule bit_eqI) (simp add: bit_concat_bit_iff) lemma bin_split_cat: "bin_split n ((\k n l. concat_bit n l k) v n w) = (v, (take_bit :: nat \ int \ int) n w)" by (simp add: drop_bit_bin_cat_eq take_bit_bin_cat_eq) lemma bin_split_zero [simp]: "bin_split n 0 = (0, 0)" by simp lemma bin_split_minus1 [simp]: "bin_split n (- 1) = (- 1, (take_bit :: nat \ int \ int) n (- 1))" by simp lemma bin_split_trunc: "bin_split (min m n) c = (a, b) \ bin_split n ((take_bit :: nat \ int \ int) m c) = ((take_bit :: nat \ int \ int) (m - n) a, b)" apply (induct n arbitrary: m b c, clarsimp) apply (simp add: bin_rest_trunc Let_def split: prod.split_asm) apply (case_tac m) apply (auto simp: Let_def drop_bit_Suc take_bit_Suc mod_2_eq_odd split: prod.split_asm) done lemma bin_split_trunc1: "bin_split n c = (a, b) \ bin_split n ((take_bit :: nat \ int \ int) m c) = ((take_bit :: nat \ int \ int) (m - n) a, (take_bit :: nat \ int \ int) m b)" apply (induct n arbitrary: m b c, clarsimp) apply (simp add: bin_rest_trunc Let_def split: prod.split_asm) apply (case_tac m) apply (auto simp: Let_def drop_bit_Suc take_bit_Suc mod_2_eq_odd split: prod.split_asm) done lemma bin_cat_num: "(\k n l. concat_bit n l k) a n b = a * 2 ^ n + (take_bit :: nat \ int \ int) n b" by (simp add: bin_cat_eq_push_bit_add_take_bit push_bit_eq_mult) lemma bin_split_num: "bin_split n b = (b div 2 ^ n, b mod 2 ^ n)" by (simp add: drop_bit_eq_div take_bit_eq_mod) lemmas bin_rsplit_aux_simps = bin_rsplit_aux.simps bin_rsplitl_aux.simps lemmas rsplit_aux_simps = bin_rsplit_aux_simps lemmas th_if_simp1 = if_split [where P = "(=) l", THEN iffD1, THEN conjunct1, THEN mp] for l lemmas th_if_simp2 = if_split [where P = "(=) l", THEN iffD1, THEN conjunct2, THEN mp] for l lemmas rsplit_aux_simp1s = rsplit_aux_simps [THEN th_if_simp1] lemmas rsplit_aux_simp2ls = rsplit_aux_simps [THEN th_if_simp2] \ \these safe to \[simp add]\ as require calculating \m - n\\ lemmas bin_rsplit_aux_simp2s [simp] = rsplit_aux_simp2ls [unfolded Let_def] lemmas rbscl = bin_rsplit_aux_simp2s (2) lemmas rsplit_aux_0_simps [simp] = rsplit_aux_simp1s [OF disjI1] rsplit_aux_simp1s [OF disjI2] lemma bin_rsplit_aux_append: "bin_rsplit_aux n m c (bs @ cs) = bin_rsplit_aux n m c bs @ cs" apply (induct n m c bs rule: bin_rsplit_aux.induct) apply (subst bin_rsplit_aux.simps) apply (subst bin_rsplit_aux.simps) apply (clarsimp split: prod.split) done lemma bin_rsplitl_aux_append: "bin_rsplitl_aux n m c (bs @ cs) = bin_rsplitl_aux n m c bs @ cs" apply (induct n m c bs rule: bin_rsplitl_aux.induct) apply (subst bin_rsplitl_aux.simps) apply (subst bin_rsplitl_aux.simps) apply (clarsimp split: prod.split) done lemmas rsplit_aux_apps [where bs = "[]"] = bin_rsplit_aux_append bin_rsplitl_aux_append lemmas rsplit_def_auxs = bin_rsplit_def bin_rsplitl_def lemmas rsplit_aux_alts = rsplit_aux_apps [unfolded append_Nil rsplit_def_auxs [symmetric]] lemma bin_split_minus: "0 < n \ bin_split (Suc (n - 1)) w = bin_split n w" by auto lemma bin_split_pred_simp [simp]: "(0::nat) < numeral bin \ bin_split (numeral bin) w = (let (w1, w2) = bin_split (numeral bin - 1) ((\k::int. k div 2) w) in (w1, of_bool (odd w) + 2 * w2))" by (simp add: take_bit_rec drop_bit_rec mod_2_eq_odd) lemma bin_rsplit_aux_simp_alt: "bin_rsplit_aux n m c bs = (if m = 0 \ n = 0 then bs else let (a, b) = bin_split n c in bin_rsplit n (m - n, a) @ b # bs)" apply (simp add: bin_rsplit_aux.simps [of n m c bs]) apply (subst rsplit_aux_alts) apply (simp add: bin_rsplit_def) done lemmas bin_rsplit_simp_alt = trans [OF bin_rsplit_def bin_rsplit_aux_simp_alt] lemmas bthrs = bin_rsplit_simp_alt [THEN [2] trans] lemma bin_rsplit_size_sign' [rule_format]: "n > 0 \ rev sw = bin_rsplit n (nw, w) \ \v\set sw. (take_bit :: nat \ int \ int) n v = v" apply (induct sw arbitrary: nw w) apply clarsimp apply clarsimp apply (drule bthrs) apply (simp (no_asm_use) add: Let_def split: prod.split_asm if_split_asm) apply clarify apply simp done lemmas bin_rsplit_size_sign = bin_rsplit_size_sign' [OF asm_rl rev_rev_ident [THEN trans] set_rev [THEN equalityD2 [THEN subsetD]]] lemma bin_nth_rsplit [rule_format] : "n > 0 \ m < n \ \w k nw. rev sw = bin_rsplit n (nw, w) \ k < size sw \ (bit :: int \ nat \ bool) (sw ! k) m = (bit :: int \ nat \ bool) w (k * n + m)" apply (induct sw) apply clarsimp apply clarsimp apply (drule bthrs) apply (simp (no_asm_use) add: Let_def split: prod.split_asm if_split_asm) apply (erule allE, erule impE, erule exI) apply (case_tac k) apply clarsimp prefer 2 apply clarsimp apply (erule allE) apply (erule (1) impE) apply (simp add: bit_drop_bit_eq ac_simps) apply (simp add: bit_take_bit_iff ac_simps) done lemma bin_rsplit_all: "0 < nw \ nw \ n \ bin_rsplit n (nw, w) = [(take_bit :: nat \ int \ int) n w]" by (auto simp: bin_rsplit_def rsplit_aux_simp2ls split: prod.split dest!: split_bintrunc) lemma bin_rsplit_l [rule_format]: "\bin. bin_rsplitl n (m, bin) = bin_rsplit n (m, (take_bit :: nat \ int \ int) m bin)" apply (rule_tac a = "m" in wf_less_than [THEN wf_induct]) apply (simp (no_asm) add: bin_rsplitl_def bin_rsplit_def) apply (rule allI) apply (subst bin_rsplitl_aux.simps) apply (subst bin_rsplit_aux.simps) apply (clarsimp simp: Let_def split: prod.split) apply (simp add: ac_simps) apply (subst rsplit_aux_alts(1)) apply (subst rsplit_aux_alts(2)) apply clarsimp unfolding bin_rsplit_def bin_rsplitl_def apply (simp add: drop_bit_take_bit) apply (case_tac \x < n\) apply (simp_all add: not_less min_def) done lemma bin_rsplit_rcat [rule_format]: "n > 0 \ bin_rsplit n (n * size ws, bin_rcat n ws) = map ((take_bit :: nat \ int \ int) n) ws" apply (unfold bin_rsplit_def bin_rcat_eq_foldl) apply (rule_tac xs = ws in rev_induct) apply clarsimp apply clarsimp apply (subst rsplit_aux_alts) apply (simp add: drop_bit_bin_cat_eq take_bit_bin_cat_eq) done lemma bin_rsplit_aux_len_le [rule_format] : "\ws m. n \ 0 \ ws = bin_rsplit_aux n nw w bs \ length ws \ m \ nw + length bs * n \ m * n" proof - have *: R if d: "i \ j \ m < j'" and R1: "i * k \ j * k \ R" and R2: "Suc m * k' \ j' * k' \ R" for i j j' k k' m :: nat and R using d apply safe apply (rule R1, erule mult_le_mono1) apply (rule R2, erule Suc_le_eq [THEN iffD2 [THEN mult_le_mono1]]) done have **: "0 < sc \ sc - n + (n + lb * n) \ m * n \ sc + lb * n \ m * n" for sc m n lb :: nat apply safe apply arith apply (case_tac "sc \ n") apply arith apply (insert linorder_le_less_linear [of m lb]) apply (erule_tac k=n and k'=n in *) apply arith apply simp done show ?thesis apply (induct n nw w bs rule: bin_rsplit_aux.induct) apply (subst bin_rsplit_aux.simps) apply (simp add: ** Let_def split: prod.split) done qed lemma bin_rsplit_len_le: "n \ 0 \ ws = bin_rsplit n (nw, w) \ length ws \ m \ nw \ m * n" by (auto simp: bin_rsplit_def bin_rsplit_aux_len_le) lemma bin_rsplit_aux_len: "n \ 0 \ length (bin_rsplit_aux n nw w cs) = (nw + n - 1) div n + length cs" apply (induct n nw w cs rule: bin_rsplit_aux.induct) apply (subst bin_rsplit_aux.simps) apply (clarsimp simp: Let_def split: prod.split) apply (erule thin_rl) apply (case_tac m) apply simp apply (case_tac "m \ n") apply (auto simp add: div_add_self2) done lemma bin_rsplit_len: "n \ 0 \ length (bin_rsplit n (nw, w)) = (nw + n - 1) div n" by (auto simp: bin_rsplit_def bin_rsplit_aux_len) lemma bin_rsplit_aux_len_indep: "n \ 0 \ length bs = length cs \ length (bin_rsplit_aux n nw v bs) = length (bin_rsplit_aux n nw w cs)" proof (induct n nw w cs arbitrary: v bs rule: bin_rsplit_aux.induct) case (1 n m w cs v bs) show ?case proof (cases "m = 0") case True with \length bs = length cs\ show ?thesis by simp next case False from "1.hyps" [of \bin_split n w\ \drop_bit n w\ \take_bit n w\] \m \ 0\ \n \ 0\ have hyp: "\v bs. length bs = Suc (length cs) \ length (bin_rsplit_aux n (m - n) v bs) = length (bin_rsplit_aux n (m - n) (drop_bit n w) (take_bit n w # cs))" using bin_rsplit_aux_len by fastforce from \length bs = length cs\ \n \ 0\ show ?thesis by (auto simp add: bin_rsplit_aux_simp_alt Let_def bin_rsplit_len split: prod.split) qed qed lemma bin_rsplit_len_indep: "n \ 0 \ length (bin_rsplit n (nw, v)) = length (bin_rsplit n (nw, w))" apply (unfold bin_rsplit_def) apply (simp (no_asm)) apply (erule bin_rsplit_aux_len_indep) apply (rule refl) done subsection \Logical operations\ abbreviation (input) bin_sc :: \nat \ bool \ int \ int\ where \bin_sc n b i \ set_bit i n b\ lemma bin_sc_0 [simp]: "bin_sc 0 b w = of_bool b + 2 * (\k::int. k div 2) w" by (simp add: set_bit_int_def) lemma bin_sc_Suc [simp]: "bin_sc (Suc n) b w = of_bool (odd w) + 2 * bin_sc n b (w div 2)" by (simp add: set_bit_int_def set_bit_Suc unset_bit_Suc bin_last_def) lemma bin_nth_sc [bit_simps]: "bit (bin_sc n b w) n \ b" by (simp add: bit_simps) lemma bin_sc_sc_same [simp]: "bin_sc n c (bin_sc n b w) = bin_sc n c w" by (induction n arbitrary: w) (simp_all add: bit_Suc) lemma bin_sc_sc_diff: "m \ n \ bin_sc m c (bin_sc n b w) = bin_sc n b (bin_sc m c w)" apply (induct n arbitrary: w m) apply (case_tac [!] m) apply auto done lemma bin_nth_sc_gen: "(bit :: int \ nat \ bool) (bin_sc n b w) m = (if m = n then b else (bit :: int \ nat \ bool) w m)" apply (induct n arbitrary: w m) apply (case_tac m; simp add: bit_Suc) apply (case_tac m; simp add: bit_Suc) done lemma bin_sc_eq: \bin_sc n False = unset_bit n\ \bin_sc n True = Bit_Operations.set_bit n\ apply (simp_all add: fun_eq_iff bit_eq_iff) apply (simp_all add: bit_simps bin_nth_sc_gen) done lemma bin_sc_nth [simp]: "bin_sc n ((bit :: int \ nat \ bool) w n) w = w" by (rule bit_eqI) (simp add: bin_nth_sc_gen) lemma bin_sign_sc [simp]: "bin_sign (bin_sc n b w) = bin_sign w" proof (induction n arbitrary: w) case 0 then show ?case by (auto simp add: bin_sign_def) (use bin_rest_ge_0 in fastforce) next case (Suc n) from Suc [of \w div 2\] show ?case by (auto simp add: bin_sign_def split: if_splits) qed lemma bin_sc_bintr [simp]: "(take_bit :: nat \ int \ int) m (bin_sc n x ((take_bit :: nat \ int \ int) m w)) = (take_bit :: nat \ int \ int) m (bin_sc n x w)" apply (rule bit_eqI) apply (cases x) apply (auto simp add: bit_simps bin_sc_eq) done lemma bin_clr_le: "bin_sc n False w \ w" by (simp add: set_bit_int_def unset_bit_less_eq) lemma bin_set_ge: "bin_sc n True w \ w" by (simp add: set_bit_int_def set_bit_greater_eq) lemma bintr_bin_clr_le: "(take_bit :: nat \ int \ int) n (bin_sc m False w) \ (take_bit :: nat \ int \ int) n w" by (simp add: set_bit_int_def take_bit_unset_bit_eq unset_bit_less_eq) lemma bintr_bin_set_ge: "(take_bit :: nat \ int \ int) n (bin_sc m True w) \ (take_bit :: nat \ int \ int) n w" by (simp add: set_bit_int_def take_bit_set_bit_eq set_bit_greater_eq) lemma bin_sc_FP [simp]: "bin_sc n False 0 = 0" by (induct n) auto lemma bin_sc_TM [simp]: "bin_sc n True (- 1) = - 1" by (induct n) auto lemmas bin_sc_simps = bin_sc_0 bin_sc_Suc bin_sc_TM bin_sc_FP lemma bin_sc_minus: "0 < n \ bin_sc (Suc (n - 1)) b w = bin_sc n b w" by auto lemmas bin_sc_Suc_minus = trans [OF bin_sc_minus [symmetric] bin_sc_Suc] lemma bin_sc_numeral [simp]: "bin_sc (numeral k) b w = of_bool (odd w) + 2 * bin_sc (pred_numeral k) b (w div 2)" by (simp add: numeral_eq_Suc) lemmas bin_sc_minus_simps = bin_sc_simps (2,3,4) [THEN [2] trans, OF bin_sc_minus [THEN sym]] lemma int_set_bit_0 [simp]: fixes x :: int shows "set_bit x 0 b = of_bool b + 2 * (x div 2)" by (fact bin_sc_0) lemma int_set_bit_Suc: fixes x :: int shows "set_bit x (Suc n) b = of_bool (odd x) + 2 * set_bit (x div 2) n b" by (fact bin_sc_Suc) lemma bin_last_set_bit: "odd (set_bit x n b :: int) = (if n > 0 then odd x else b)" by (cases n) (simp_all add: int_set_bit_Suc) lemma bin_rest_set_bit: "(set_bit x n b :: int) div 2 = (if n > 0 then set_bit (x div 2) (n - 1) b else x div 2)" by (cases n) (simp_all add: int_set_bit_Suc) lemma int_set_bit_numeral: fixes x :: int shows "set_bit x (numeral w) b = of_bool (odd x) + 2 * set_bit (x div 2) (pred_numeral w) b" by (fact bin_sc_numeral) lemmas int_set_bit_numerals [simp] = int_set_bit_numeral[where x="numeral w'"] int_set_bit_numeral[where x="- numeral w'"] int_set_bit_numeral[where x="Numeral1"] int_set_bit_numeral[where x="1"] int_set_bit_numeral[where x="0"] int_set_bit_Suc[where x="numeral w'"] int_set_bit_Suc[where x="- numeral w'"] int_set_bit_Suc[where x="Numeral1"] int_set_bit_Suc[where x="1"] int_set_bit_Suc[where x="0"] for w' lemma msb_set_bit [simp]: "msb (set_bit (x :: int) n b) \ msb x" by (simp add: msb_int_def set_bit_int_def) lemma word_set_bit_def: \set_bit a n x = word_of_int (bin_sc n x (uint a))\ apply (rule bit_word_eqI) apply (cases x) apply (simp_all add: bit_simps bin_sc_eq) done lemma set_bit_word_of_int: "set_bit (word_of_int x) n b = word_of_int (bin_sc n b x)" unfolding word_set_bit_def by (rule word_eqI) (simp add: word_size bin_nth_sc_gen nth_bintr bit_simps) lemma word_set_numeral [simp]: "set_bit (numeral bin::'a::len word) n b = word_of_int (bin_sc n b (numeral bin))" unfolding word_numeral_alt by (rule set_bit_word_of_int) lemma word_set_neg_numeral [simp]: "set_bit (- numeral bin::'a::len word) n b = word_of_int (bin_sc n b (- numeral bin))" unfolding word_neg_numeral_alt by (rule set_bit_word_of_int) lemma word_set_bit_0 [simp]: "set_bit 0 n b = word_of_int (bin_sc n b 0)" unfolding word_0_wi by (rule set_bit_word_of_int) lemma word_set_bit_1 [simp]: "set_bit 1 n b = word_of_int (bin_sc n b 1)" unfolding word_1_wi by (rule set_bit_word_of_int) lemma shiftl_int_def: "push_bit n x = x * 2 ^ n" for x :: int by (fact push_bit_eq_mult) lemma shiftr_int_def: "drop_bit n x = x div 2 ^ n" for x :: int by (fact drop_bit_eq_div) subsubsection \Basic simplification rules\ context includes bit_operations_syntax begin lemmas int_not_def = not_int_def lemma int_not_simps: "NOT (0::int) = -1" "NOT (1::int) = -2" "NOT (- 1::int) = 0" "NOT (numeral w::int) = - numeral (w + Num.One)" "NOT (- numeral (Num.Bit0 w)::int) = numeral (Num.BitM w)" "NOT (- numeral (Num.Bit1 w)::int) = numeral (Num.Bit0 w)" by (simp_all add: not_int_def) lemma int_not_not: "NOT (NOT x) = x" for x :: int by (fact bit.double_compl) lemma int_and_0 [simp]: "0 AND x = 0" for x :: int by (fact bit.conj_zero_left) lemma int_and_m1 [simp]: "-1 AND x = x" for x :: int - by (fact bit.conj_one_left) + by (fact and.left_neutral) lemma int_or_zero [simp]: "0 OR x = x" for x :: int - by (fact bit.disj_zero_left) + by (fact or.left_neutral) lemma int_or_minus1 [simp]: "-1 OR x = -1" for x :: int by (fact bit.disj_one_left) lemma int_xor_zero [simp]: "0 XOR x = x" for x :: int - by (fact bit.xor_zero_left) + by (fact xor.left_neutral) subsubsection \Binary destructors\ lemma bin_rest_NOT [simp]: "(\k::int. k div 2) (NOT x) = NOT ((\k::int. k div 2) x)" by (fact not_int_div_2) lemma bin_last_NOT [simp]: "(odd :: int \ bool) (NOT x) \ \ (odd :: int \ bool) x" by simp lemma bin_rest_AND [simp]: "(\k::int. k div 2) (x AND y) = (\k::int. k div 2) x AND (\k::int. k div 2) y" by (subst and_int_rec) auto lemma bin_last_AND [simp]: "(odd :: int \ bool) (x AND y) \ (odd :: int \ bool) x \ (odd :: int \ bool) y" by (subst and_int_rec) auto lemma bin_rest_OR [simp]: "(\k::int. k div 2) (x OR y) = (\k::int. k div 2) x OR (\k::int. k div 2) y" by (subst or_int_rec) auto lemma bin_last_OR [simp]: "(odd :: int \ bool) (x OR y) \ (odd :: int \ bool) x \ (odd :: int \ bool) y" by (subst or_int_rec) auto lemma bin_rest_XOR [simp]: "(\k::int. k div 2) (x XOR y) = (\k::int. k div 2) x XOR (\k::int. k div 2) y" by (subst xor_int_rec) auto lemma bin_last_XOR [simp]: "(odd :: int \ bool) (x XOR y) \ ((odd :: int \ bool) x \ (odd :: int \ bool) y) \ \ ((odd :: int \ bool) x \ (odd :: int \ bool) y)" by (subst xor_int_rec) auto lemma bin_nth_ops: "\x y. (bit :: int \ nat \ bool) (x AND y) n \ (bit :: int \ nat \ bool) x n \ (bit :: int \ nat \ bool) y n" "\x y. (bit :: int \ nat \ bool) (x OR y) n \ (bit :: int \ nat \ bool) x n \ (bit :: int \ nat \ bool) y n" "\x y. (bit :: int \ nat \ bool) (x XOR y) n \ (bit :: int \ nat \ bool) x n \ (bit :: int \ nat \ bool) y n" "\x. (bit :: int \ nat \ bool) (NOT x) n \ \ (bit :: int \ nat \ bool) x n" by (simp_all add: bit_and_iff bit_or_iff bit_xor_iff bit_not_iff) subsubsection \Derived properties\ lemma int_xor_minus1 [simp]: "-1 XOR x = NOT x" for x :: int by (fact bit.xor_one_left) lemma int_xor_extra_simps [simp]: "w XOR 0 = w" "w XOR -1 = NOT w" for w :: int by simp_all lemma int_or_extra_simps [simp]: "w OR 0 = w" "w OR -1 = -1" for w :: int by simp_all lemma int_and_extra_simps [simp]: "w AND 0 = 0" "w AND -1 = w" for w :: int by simp_all text \Commutativity of the above.\ lemma bin_ops_comm: fixes x y :: int shows int_and_comm: "x AND y = y AND x" and int_or_comm: "x OR y = y OR x" and int_xor_comm: "x XOR y = y XOR x" by (simp_all add: ac_simps) lemma bin_ops_same [simp]: "x AND x = x" "x OR x = x" "x XOR x = 0" for x :: int by simp_all lemmas bin_log_esimps = int_and_extra_simps int_or_extra_simps int_xor_extra_simps int_and_0 int_and_m1 int_or_zero int_or_minus1 int_xor_zero int_xor_minus1 subsubsection \Basic properties of logical (bit-wise) operations\ lemma bbw_ao_absorb: "x AND (y OR x) = x \ x OR (y AND x) = x" for x y :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma bbw_ao_absorbs_other: "x AND (x OR y) = x \ (y AND x) OR x = x" "(y OR x) AND x = x \ x OR (x AND y) = x" "(x OR y) AND x = x \ (x AND y) OR x = x" for x y :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemmas bbw_ao_absorbs [simp] = bbw_ao_absorb bbw_ao_absorbs_other lemma int_xor_not: "(NOT x) XOR y = NOT (x XOR y) \ x XOR (NOT y) = NOT (x XOR y)" for x y :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma int_and_assoc: "(x AND y) AND z = x AND (y AND z)" for x y z :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma int_or_assoc: "(x OR y) OR z = x OR (y OR z)" for x y z :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma int_xor_assoc: "(x XOR y) XOR z = x XOR (y XOR z)" for x y z :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemmas bbw_assocs = int_and_assoc int_or_assoc int_xor_assoc (* BH: Why are these declared as simp rules??? *) lemma bbw_lcs [simp]: "y AND (x AND z) = x AND (y AND z)" "y OR (x OR z) = x OR (y OR z)" "y XOR (x XOR z) = x XOR (y XOR z)" for x y :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma bbw_not_dist: "NOT (x OR y) = (NOT x) AND (NOT y)" "NOT (x AND y) = (NOT x) OR (NOT y)" for x y :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma bbw_oa_dist: "(x AND y) OR z = (x OR z) AND (y OR z)" for x y z :: int by (auto simp add: bin_eq_iff bin_nth_ops) lemma bbw_ao_dist: "(x OR y) AND z = (x AND z) OR (y AND z)" for x y z :: int by (auto simp add: bin_eq_iff bin_nth_ops) subsubsection \Simplification with numerals\ text \Cases for \0\ and \-1\ are already covered by other simp rules.\ lemma bin_rest_neg_numeral_BitM [simp]: "(\k::int. k div 2) (- numeral (Num.BitM w)) = - numeral w" by simp lemma bin_last_neg_numeral_BitM [simp]: "(odd :: int \ bool) (- numeral (Num.BitM w))" by simp subsubsection \Interactions with arithmetic\ lemma le_int_or: "bin_sign y = 0 \ x \ x OR y" for x y :: int by (simp add: bin_sign_def or_greater_eq split: if_splits) lemmas int_and_le = xtrans(3) [OF bbw_ao_absorbs (2) [THEN conjunct2, symmetric] le_int_or] text \Interaction between bit-wise and arithmetic: good example of \bin_induction\.\ lemma bin_add_not: "x + NOT x = (-1::int)" by (simp add: not_int_def) lemma AND_mod: "x AND (2 ^ n - 1) = x mod 2 ^ n" for x :: int by (simp flip: take_bit_eq_mod add: take_bit_eq_mask mask_eq_exp_minus_1) subsubsection \Truncating results of bit-wise operations\ lemma bin_trunc_ao: "(take_bit :: nat \ int \ int) n x AND (take_bit :: nat \ int \ int) n y = (take_bit :: nat \ int \ int) n (x AND y)" "(take_bit :: nat \ int \ int) n x OR (take_bit :: nat \ int \ int) n y = (take_bit :: nat \ int \ int) n (x OR y)" by simp_all lemma bin_trunc_xor: "(take_bit :: nat \ int \ int) n ((take_bit :: nat \ int \ int) n x XOR (take_bit :: nat \ int \ int) n y) = (take_bit :: nat \ int \ int) n (x XOR y)" by simp lemma bin_trunc_not: "(take_bit :: nat \ int \ int) n (NOT ((take_bit :: nat \ int \ int) n x)) = (take_bit :: nat \ int \ int) n (NOT x)" by (fact take_bit_not_take_bit) text \Want theorems of the form of \bin_trunc_xor\.\ lemma bintr_bintr_i: "x = (take_bit :: nat \ int \ int) n y \ (take_bit :: nat \ int \ int) n x = (take_bit :: nat \ int \ int) n y" by auto lemmas bin_trunc_and = bin_trunc_ao(1) [THEN bintr_bintr_i] lemmas bin_trunc_or = bin_trunc_ao(2) [THEN bintr_bintr_i] subsubsection \More lemmas\ lemma not_int_cmp_0 [simp]: fixes i :: int shows "0 < NOT i \ i < -1" "0 \ NOT i \ i < 0" "NOT i < 0 \ i \ 0" "NOT i \ 0 \ i \ -1" by(simp_all add: int_not_def) arith+ lemma bbw_ao_dist2: "(x :: int) AND (y OR z) = x AND y OR x AND z" by (fact bit.conj_disj_distrib) lemmas int_and_ac = bbw_lcs(1) int_and_comm int_and_assoc lemma int_nand_same [simp]: fixes x :: int shows "x AND NOT x = 0" by simp lemma int_nand_same_middle: fixes x :: int shows "x AND y AND NOT x = 0" by (simp add: bit_eq_iff bit_and_iff bit_not_iff) lemma and_xor_dist: fixes x :: int shows "x AND (y XOR z) = (x AND y) XOR (x AND z)" by (fact bit.conj_xor_distrib) lemma int_and_lt0 [simp]: \x AND y < 0 \ x < 0 \ y < 0\ for x y :: int by (fact and_negative_int_iff) lemma int_and_ge0 [simp]: \x AND y \ 0 \ x \ 0 \ y \ 0\ for x y :: int by (fact and_nonnegative_int_iff) lemma int_and_1: fixes x :: int shows "x AND 1 = x mod 2" by (fact and_one_eq) lemma int_1_and: fixes x :: int shows "1 AND x = x mod 2" by (fact one_and_eq) lemma int_or_lt0 [simp]: \x OR y < 0 \ x < 0 \ y < 0\ for x y :: int by (fact or_negative_int_iff) lemma int_or_ge0 [simp]: \x OR y \ 0 \ x \ 0 \ y \ 0\ for x y :: int by (fact or_nonnegative_int_iff) lemma int_xor_lt0 [simp]: \x XOR y < 0 \ (x < 0) \ (y < 0)\ for x y :: int by (fact xor_negative_int_iff) lemma int_xor_ge0 [simp]: \x XOR y \ 0 \ (x \ 0 \ y \ 0)\ for x y :: int by (fact xor_nonnegative_int_iff) lemma even_conv_AND: \even i \ i AND 1 = 0\ for i :: int by (simp add: and_one_eq mod2_eq_if) lemma bin_last_conv_AND: "(odd :: int \ bool) i \ i AND 1 \ 0" by (simp add: and_one_eq mod2_eq_if) lemma bitval_bin_last: "of_bool ((odd :: int \ bool) i) = i AND 1" by (simp add: and_one_eq mod2_eq_if) lemma bin_sign_and: "bin_sign (i AND j) = - (bin_sign i * bin_sign j)" by(simp add: bin_sign_def) lemma int_not_neg_numeral: "NOT (- numeral n) = (Num.sub n num.One :: int)" by(simp add: int_not_def) lemma int_neg_numeral_pOne_conv_not: "- numeral (n + num.One) = (NOT (numeral n) :: int)" by(simp add: int_not_def) subsection \Setting and clearing bits\ lemma int_shiftl_BIT: fixes x :: int shows int_shiftl0: "push_bit 0 x = x" and int_shiftl_Suc: "push_bit (Suc n) x = 2 * push_bit n x" by (auto simp add: shiftl_int_def) lemma int_0_shiftl: "push_bit n 0 = (0 :: int)" by (fact push_bit_of_0) lemma bin_last_shiftl: "odd (push_bit n x) \ n = 0 \ (odd :: int \ bool) x" by simp lemma bin_rest_shiftl: "(\k::int. k div 2) (push_bit n x) = (if n > 0 then push_bit (n - 1) x else (\k::int. k div 2) x)" by (cases n) (simp_all add: push_bit_eq_mult) lemma bin_nth_shiftl: "(bit :: int \ nat \ bool) (push_bit n x) m \ n \ m \ (bit :: int \ nat \ bool) x (m - n)" by (fact bit_push_bit_iff_int) lemma bin_last_shiftr: "odd (drop_bit n x) \ bit x n" for x :: int by (simp add: bit_iff_odd_drop_bit) lemma bin_rest_shiftr: "(\k::int. k div 2) (drop_bit n x) = drop_bit (Suc n) x" by (simp add: drop_bit_Suc drop_bit_half) lemma bin_nth_shiftr: "(bit :: int \ nat \ bool) (drop_bit n x) m = (bit :: int \ nat \ bool) x (n + m)" by (simp add: bit_simps) lemma bin_nth_conv_AND: fixes x :: int shows "(bit :: int \ nat \ bool) x n \ x AND (push_bit n 1) \ 0" by (fact bit_iff_and_push_bit_not_eq_0) lemma int_shiftl_numeral [simp]: "push_bit (numeral w') (numeral w :: int) = push_bit (pred_numeral w') (numeral (num.Bit0 w))" "push_bit (numeral w') (- numeral w :: int) = push_bit (pred_numeral w') (- numeral (num.Bit0 w))" by(simp_all add: numeral_eq_Suc shiftl_int_def) (metis add_One mult_inc semiring_norm(11) semiring_norm(13) semiring_norm(2) semiring_norm(6) semiring_norm(87))+ lemma int_shiftl_One_numeral [simp]: "push_bit (numeral w) (1::int) = push_bit (pred_numeral w) 2" using int_shiftl_numeral [of Num.One w] by (simp add: numeral_eq_Suc) lemma shiftl_ge_0: fixes i :: int shows "push_bit n i \ 0 \ i \ 0" by (fact push_bit_nonnegative_int_iff) lemma shiftl_lt_0: fixes i :: int shows "push_bit n i < 0 \ i < 0" by (fact push_bit_negative_int_iff) lemma int_shiftl_test_bit: "bit (push_bit i n :: int) m \ m \ i \ bit n (m - i)" by (fact bit_push_bit_iff_int) lemma int_0shiftr: "drop_bit x (0 :: int) = 0" by (fact drop_bit_of_0) lemma int_minus1_shiftr: "drop_bit x (-1 :: int) = -1" by (fact drop_bit_minus_one) lemma int_shiftr_ge_0: fixes i :: int shows "drop_bit n i \ 0 \ i \ 0" by (fact drop_bit_nonnegative_int_iff) lemma int_shiftr_lt_0 [simp]: fixes i :: int shows "drop_bit n i < 0 \ i < 0" by (fact drop_bit_negative_int_iff) lemma int_shiftr_numeral [simp]: "drop_bit (numeral w') (1 :: int) = 0" "drop_bit (numeral w') (numeral num.One :: int) = 0" "drop_bit (numeral w') (numeral (num.Bit0 w) :: int) = drop_bit (pred_numeral w') (numeral w)" "drop_bit (numeral w') (numeral (num.Bit1 w) :: int) = drop_bit (pred_numeral w') (numeral w)" "drop_bit (numeral w') (- numeral (num.Bit0 w) :: int) = drop_bit (pred_numeral w') (- numeral w)" "drop_bit (numeral w') (- numeral (num.Bit1 w) :: int) = drop_bit (pred_numeral w') (- numeral (Num.inc w))" by (simp_all add: numeral_eq_Suc add_One drop_bit_Suc) lemma int_shiftr_numeral_Suc0 [simp]: "drop_bit (Suc 0) (1 :: int) = 0" "drop_bit (Suc 0) (numeral num.One :: int) = 0" "drop_bit (Suc 0) (numeral (num.Bit0 w) :: int) = numeral w" "drop_bit (Suc 0) (numeral (num.Bit1 w) :: int) = numeral w" "drop_bit (Suc 0) (- numeral (num.Bit0 w) :: int) = - numeral w" "drop_bit (Suc 0) (- numeral (num.Bit1 w) :: int) = - numeral (Num.inc w)" by (simp_all add: drop_bit_Suc add_One) lemma bin_nth_minus_p2: assumes sign: "bin_sign x = 0" and y: "y = push_bit n 1" and m: "m < n" and x: "x < y" shows "bit (x - y) m = bit x m" proof - from sign y x have \x \ 0\ and \y = 2 ^ n\ and \x < 2 ^ n\ by (simp_all add: bin_sign_def push_bit_eq_mult split: if_splits) from \0 \ x\ \x < 2 ^ n\ \m < n\ have \bit x m \ bit (x - 2 ^ n) m\ proof (induction m arbitrary: x n) case 0 then show ?case by simp next case (Suc m) moreover define q where \q = n - 1\ ultimately have n: \n = Suc q\ by simp have \(x - 2 ^ Suc q) div 2 = x div 2 - 2 ^ q\ by simp moreover from Suc.IH [of \x div 2\ q] Suc.prems have \bit (x div 2) m \ bit (x div 2 - 2 ^ q) m\ by (simp add: n) ultimately show ?case by (simp add: bit_Suc n) qed with \y = 2 ^ n\ show ?thesis by simp qed lemma bin_clr_conv_NAND: "bin_sc n False i = i AND NOT (push_bit n 1)" by (rule bit_eqI) (auto simp add: bin_sc_eq bit_simps) lemma bin_set_conv_OR: "bin_sc n True i = i OR (push_bit n 1)" by (rule bit_eqI) (auto simp add: bin_sc_eq bit_simps) end subsection \More lemmas on words\ lemma msb_conv_bin_sign: "msb x \ bin_sign x = -1" by (simp add: bin_sign_def not_le msb_int_def) lemma msb_bin_sc [simp]: "msb (bin_sc n b x) \ msb x" by (simp add: msb_conv_bin_sign) lemma msb_word_def: \msb a \ bin_sign (signed_take_bit (LENGTH('a) - 1) (uint a)) = - 1\ for a :: \'a::len word\ by (simp add: bin_sign_def bit_simps msb_word_iff_bit) lemma word_msb_def: "msb a \ bin_sign (sint a) = - 1" by (simp add: msb_word_def sint_uint) lemma word_rcat_eq: \word_rcat ws = word_of_int (bin_rcat (LENGTH('a::len)) (map uint ws))\ for ws :: \'a::len word list\ apply (simp add: word_rcat_def bin_rcat_def rev_map) apply transfer apply (simp add: horner_sum_foldr foldr_map comp_def) done lemma sign_uint_Pls [simp]: "bin_sign (uint x) = 0" by (simp add: sign_Pls_ge_0) lemmas bin_log_bintrs = bin_trunc_not bin_trunc_xor bin_trunc_and bin_trunc_or \ \following definitions require both arithmetic and bit-wise word operations\ \ \to get \word_no_log_defs\ from \word_log_defs\, using \bin_log_bintrs\\ lemmas wils1 = bin_log_bintrs [THEN word_of_int_eq_iff [THEN iffD2], folded uint_word_of_int_eq, THEN eq_reflection] \ \the binary operations only\ (* BH: why is this needed? *) lemmas word_log_binary_defs = word_and_def word_or_def word_xor_def lemma setBit_no: "Bit_Operations.set_bit n (numeral bin) = word_of_int (bin_sc n True (numeral bin))" by (rule bit_word_eqI) (simp add: bit_simps) lemma clearBit_no: "unset_bit n (numeral bin) = word_of_int (bin_sc n False (numeral bin))" by (rule bit_word_eqI) (simp add: bit_simps) lemma eq_mod_iff: "0 < n \ b = b mod n \ 0 \ b \ b < n" for b n :: int by auto (metis pos_mod_conj)+ lemma split_uint_lem: "bin_split n (uint w) = (a, b) \ a = take_bit (LENGTH('a) - n) a \ b = take_bit (LENGTH('a)) b" for w :: "'a::len word" by transfer (simp add: drop_bit_take_bit ac_simps) \ \limited hom result\ lemma word_cat_hom: "LENGTH('a::len) \ LENGTH('b::len) + LENGTH('c::len) \ (word_cat (word_of_int w :: 'b word) (b :: 'c word) :: 'a word) = word_of_int ((\k n l. concat_bit n l k) w (size b) (uint b))" by transfer (simp add: take_bit_concat_bit_eq) lemma bintrunc_shiftl: "take_bit n (push_bit i m) = push_bit i (take_bit (n - i) m)" for m :: int by (fact take_bit_push_bit) lemma uint_shiftl: "uint (push_bit i n) = take_bit (size n) (push_bit i (uint n))" by (simp add: unsigned_push_bit_eq word_size) lemma bin_mask_conv_pow2: "mask n = 2 ^ n - (1 :: int)" by (fact mask_eq_exp_minus_1) lemma bin_mask_ge0: "mask n \ (0 :: int)" by (fact mask_nonnegative_int) context includes bit_operations_syntax begin lemma and_bin_mask_conv_mod: "x AND mask n = x mod 2 ^ n" for x :: int by (simp flip: take_bit_eq_mod add: take_bit_eq_mask) end lemma bin_mask_numeral: "mask (numeral n) = (1 :: int) + 2 * mask (pred_numeral n)" by (fact mask_numeral) lemma bin_nth_mask: "bit (mask n :: int) i \ i < n" by (simp add: bit_mask_iff) lemma bin_sign_mask [simp]: "bin_sign (mask n) = 0" by (simp add: bin_sign_def bin_mask_conv_pow2) lemma bin_mask_p1_conv_shift: "mask n + 1 = push_bit n (1 :: int)" by (simp add: bin_mask_conv_pow2 shiftl_int_def) lemma sbintrunc_eq_in_range: "((signed_take_bit :: nat \ int \ int) n x = x) = (x \ range ((signed_take_bit :: nat \ int \ int) n))" "(x = (signed_take_bit :: nat \ int \ int) n x) = (x \ range ((signed_take_bit :: nat \ int \ int) n))" apply (simp_all add: image_def) apply (metis sbintrunc_sbintrunc)+ done lemma sbintrunc_If: "- 3 * (2 ^ n) \ x \ x < 3 * (2 ^ n) \ (signed_take_bit :: nat \ int \ int) n x = (if x < - (2 ^ n) then x + 2 * (2 ^ n) else if x \ 2 ^ n then x - 2 * (2 ^ n) else x)" apply (simp add: no_sbintr_alt2, safe) apply (simp add: mod_pos_geq) apply (subst mod_add_self1[symmetric], simp) done lemma sint_range': \- (2 ^ (LENGTH('a) - Suc 0)) \ sint x \ sint x < 2 ^ (LENGTH('a) - Suc 0)\ for x :: \'a::len word\ apply transfer using sbintr_ge sbintr_lt apply auto done lemma signed_arith_eq_checks_to_ord: "(sint a + sint b = sint (a + b )) = ((a <=s a + b) = (0 <=s b))" "(sint a - sint b = sint (a - b )) = ((0 <=s a - b) = (b <=s a))" "(- sint a = sint (- a)) = (0 <=s (- a) = (a <=s 0))" using sint_range'[where x=a] sint_range'[where x=b] by (simp_all add: sint_word_ariths word_sle_eq word_sless_alt sbintrunc_If) lemma signed_mult_eq_checks_double_size: assumes mult_le: "(2 ^ (len_of TYPE ('a) - 1) + 1) ^ 2 \ (2 :: int) ^ (len_of TYPE ('b) - 1)" and le: "2 ^ (LENGTH('a) - 1) \ (2 :: int) ^ (len_of TYPE ('b) - 1)" shows "(sint (a :: 'a :: len word) * sint b = sint (a * b)) = (scast a * scast b = (scast (a * b) :: 'b :: len word))" proof - have P: "(signed_take_bit :: nat \ int \ int) (size a - 1) (sint a * sint b) \ range ((signed_take_bit :: nat \ int \ int) (size a - 1))" by simp have abs: "!! x :: 'a word. abs (sint x) < 2 ^ (size a - 1) + 1" apply (cut_tac x=x in sint_range') apply (simp add: abs_le_iff word_size) done have abs_ab: "abs (sint a * sint b) < 2 ^ (LENGTH('b) - 1)" using abs_mult_less[OF abs[where x=a] abs[where x=b]] mult_le by (simp add: abs_mult power2_eq_square word_size) define r s where \r = LENGTH('a) - 1\ \s = LENGTH('b) - 1\ then have \LENGTH('a) = Suc r\ \LENGTH('b) = Suc s\ \size a = Suc r\ \size b = Suc r\ by (simp_all add: word_size) then show ?thesis using P[unfolded range_sbintrunc] abs_ab le apply clarsimp apply (transfer fixing: r s) apply (auto simp add: signed_take_bit_int_eq_self min.absorb2 simp flip: signed_take_bit_eq_iff_take_bit_eq) done qed lemma bintrunc_id: "\m \ int n; 0 < m\ \ take_bit n m = m" by (simp add: take_bit_int_eq_self_iff le_less_trans less_exp) lemma bin_cat_cong: "concat_bit n b a = concat_bit m d c" if "n = m" "a = c" "take_bit m b = take_bit m d" using that(3) unfolding that(1,2) by (simp add: bin_cat_eq_push_bit_add_take_bit) lemma bin_cat_eqD1: "concat_bit n b a = concat_bit n d c \ a = c" by (metis drop_bit_bin_cat_eq) lemma bin_cat_eqD2: "concat_bit n b a = concat_bit n d c \ take_bit n b = take_bit n d" by (metis take_bit_bin_cat_eq) lemma bin_cat_inj: "(concat_bit n b a) = concat_bit n d c \ a = c \ take_bit n b = take_bit n d" by (auto intro: bin_cat_cong bin_cat_eqD1 bin_cat_eqD2) code_identifier code_module Bits_Int \ (SML) Bit_Operations and (OCaml) Bit_Operations and (Haskell) Bit_Operations and (Scala) Bit_Operations end diff --git a/thys/Word_Lib/Legacy_Aliases.thy b/thys/Word_Lib/Legacy_Aliases.thy --- a/thys/Word_Lib/Legacy_Aliases.thy +++ b/thys/Word_Lib/Legacy_Aliases.thy @@ -1,212 +1,271 @@ (* * Copyright Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section \Legacy aliases\ theory Legacy_Aliases imports "HOL-Library.Word" begin +context abstract_boolean_algebra +begin + +lemma conj_assoc: "(x \<^bold>\ y) \<^bold>\ z = x \<^bold>\ (y \<^bold>\ z)" + by (fact conj.assoc) + +lemma conj_commute: "x \<^bold>\ y = y \<^bold>\ x" + by (fact conj.commute) + +lemmas conj_left_commute = conj.left_commute +lemmas conj_ac = conj.assoc conj.commute conj.left_commute + +lemma conj_one_left: "\<^bold>1 \<^bold>\ x = x" + by (fact conj.left_neutral) + +lemma conj_left_absorb: "x \<^bold>\ (x \<^bold>\ y) = x \<^bold>\ y" + by (fact conj.left_idem) + +lemma conj_absorb: "x \<^bold>\ x = x" + by (fact conj.idem) + +lemma disj_assoc: "(x \<^bold>\ y) \<^bold>\ z = x \<^bold>\ (y \<^bold>\ z)" + by (fact disj.assoc) + +lemma disj_commute: "x \<^bold>\ y = y \<^bold>\ x" + by (fact disj.commute) + +lemmas disj_left_commute = disj.left_commute + +lemmas disj_ac = disj.assoc disj.commute disj.left_commute + +lemma disj_zero_left: "\<^bold>0 \<^bold>\ x = x" + by (fact disj.left_neutral) + +lemma disj_left_absorb: "x \<^bold>\ (x \<^bold>\ y) = x \<^bold>\ y" + by (fact disj.left_idem) + +lemma disj_absorb: "x \<^bold>\ x = x" + by (fact disj.idem) + +end + +context abstract_boolean_algebra_sym_diff +begin + +lemmas xor_assoc = xor.assoc +lemmas xor_commute = xor.commute +lemmas xor_left_commute = xor.left_commute + +lemmas xor_ac = xor.assoc xor.commute xor.left_commute + +lemma xor_zero_right: "x \<^bold>\ \<^bold>0 = x" + by (fact xor.comm_neutral) + +lemma xor_zero_left: "\<^bold>0 \<^bold>\ x = x" + by (fact xor.left_neutral) + +end + abbreviation (input) test_bit :: \'a::semiring_bits \ nat \ bool\ where \test_bit \ bit\ abbreviation (input) bin_nth :: \int \ nat \ bool\ where \bin_nth \ bit\ abbreviation (input) bin_last :: \int \ bool\ where \bin_last \ odd\ abbreviation (input) bin_rest :: \int \ int\ where \bin_rest w \ w div 2\ abbreviation (input) bintrunc :: \nat \ int \ int\ where \bintrunc \ take_bit\ abbreviation (input) sbintrunc :: \nat \ int \ int\ where \sbintrunc \ signed_take_bit\ abbreviation (input) bin_cat :: \int \ nat \ int \ int\ where \bin_cat k n l \ concat_bit n l k\ abbreviation (input) norm_sint :: \nat \ int \ int\ where \norm_sint n \ signed_take_bit (n - 1)\ abbreviation (input) max_word :: \'a::len word\ where \max_word \ - 1\ abbreviation (input) complement :: \'a::len word \ 'a word\ where \complement \ not\ lemma complement_mask: "complement (2 ^ n - 1) = NOT (mask n)" unfolding mask_eq_decr_exp by simp abbreviation (input) shiftl1 :: \'a::len word \ 'a word\ where \shiftl1 \ (*) 2\ abbreviation (input) shiftr1 :: \'a::len word \ 'a word\ where \shiftr1 w \ w div 2\ abbreviation (input) sshiftr1 :: \'a::len word \ 'a word\ where \sshiftr1 \ signed_drop_bit (Suc 0)\ context includes bit_operations_syntax begin abbreviation (input) bshiftr1 :: \bool \ 'a::len word \ 'a word\ where \bshiftr1 b w \ w div 2 OR push_bit (LENGTH('a) - Suc 0) (of_bool b) \ end lemma shiftr1_1: "shiftr1 (1::'a::len word) = 0" by (fact bits_1_div_2) lemma sshiftr1_eq: \sshiftr1 w = word_of_int (sint w div 2)\ by (rule bit_word_eqI) (auto simp add: bit_simps min_def simp flip: bit_Suc elim: le_SucE) lemma shiftl1_eq: \shiftl1 w = word_of_int (2 * uint w)\ by (rule bit_word_eqI) (simp add: bit_simps) lemma bit_shiftl1_iff: \bit (shiftl1 w) n \ 0 < n \ n < LENGTH('a) \ bit w (n - 1)\ for w :: \'a::len word\ by (auto simp add: bit_simps) lemma bit_shiftr1_iff: \bit (shiftr1 w) n \ bit w (Suc n)\ by (simp add: bit_Suc) lemma shiftr1_eq: \shiftr1 w = word_of_int (uint w div 2)\ by (rule bit_word_eqI) (simp add: bit_simps flip: bit_Suc) lemma shiftl1_rev: "shiftl1 w = word_reverse (shiftr1 (word_reverse w))" by (rule bit_word_eqI) (auto simp add: bit_simps Suc_diff_Suc simp flip: bit_Suc) lemma shiftl1_p: "shiftl1 w = w + w" for w :: "'a::len word" by (fact mult_2) lemma shiftr1_bintr: "(shiftr1 (numeral w) :: 'a::len word) = word_of_int (take_bit LENGTH('a) (numeral w) div 2)" by (rule bit_word_eqI) (simp add: bit_simps flip: bit_Suc) lemma sshiftr1_sbintr: "(sshiftr1 (numeral w) :: 'a::len word) = word_of_int (signed_take_bit (LENGTH('a) - 1) (numeral w) div 2)" apply (cases \LENGTH('a)\) apply simp_all apply (rule bit_word_eqI) apply (auto simp add: bit_simps min_def simp flip: bit_Suc elim: le_SucE) done lemma shiftl1_wi: "shiftl1 (word_of_int w) = word_of_int (2 * w)" by transfer simp lemma shiftl1_numeral: "shiftl1 (numeral w) = numeral (Num.Bit0 w)" unfolding word_numeral_alt shiftl1_wi by simp lemma shiftl1_neg_numeral: "shiftl1 (- numeral w) = - numeral (Num.Bit0 w)" unfolding word_neg_numeral_alt shiftl1_wi by simp lemma shiftl1_0: "shiftl1 0 = 0" by (fact mult_zero_right) lemma shiftl1_def_u: "shiftl1 w = word_of_int (2 * uint w)" by (fact shiftl1_eq) lemma shiftl1_def_s: "shiftl1 w = word_of_int (2 * sint w)" by simp lemma shiftr1_0: "shiftr1 0 = 0" by (fact bits_div_0) lemma sshiftr1_0: "sshiftr1 0 = 0" by (fact signed_drop_bit_of_0) lemma sshiftr1_n1: "sshiftr1 (- 1) = - 1" by (fact signed_drop_bit_of_minus_1) lemma uint_shiftr1: "uint (shiftr1 w) = uint w div 2" by (rule bit_eqI) (simp add: bit_simps flip: bit_Suc) lemma shiftr1_div_2: "uint (shiftr1 w) = uint w div 2" by (fact uint_shiftr1) lemma sshiftr1_div_2: "sint (sshiftr1 w) = sint w div 2" by (rule bit_eqI) (auto simp add: bit_simps ac_simps min_def simp flip: bit_Suc elim: le_SucE) lemma bit_sshiftr1_iff: \bit (sshiftr1 w) n \ bit w (if n = LENGTH('a) - 1 then LENGTH('a) - 1 else Suc n)\ for w :: \'a::len word\ by (auto simp add: bit_simps) lemma bit_bshiftr1_iff: \bit (bshiftr1 b w) n \ b \ n = LENGTH('a) - 1 \ bit w (Suc n)\ for w :: \'a::len word\ by (auto simp add: bit_simps simp flip: bit_Suc) lemma nth_shiftl1: "bit (shiftl1 w) n \ n < size w \ n > 0 \ bit w (n - 1)" by (auto simp add: word_size bit_simps) lemma nth_shiftr1: "bit (shiftr1 w) n = bit w (Suc n)" by (simp add: bit_Suc) lemma nth_sshiftr1: "bit (sshiftr1 w) n = (if n = size w - 1 then bit w n else bit w (Suc n))" by (auto simp add: word_size bit_simps) lemma shiftl_power: "(shiftl1 ^^ x) (y::'a::len word) = 2 ^ x * y" by (induction x) simp_all lemma le_shiftr1: \shiftr1 u \ shiftr1 v\ if \u \ v\ using that by (simp add: word_le_nat_alt unat_div div_le_mono) lemma le_shiftr1': "\ shiftr1 u \ shiftr1 v ; shiftr1 u \ shiftr1 v \ \ u \ v" by (meson dual_order.antisym le_cases le_shiftr1) abbreviation (input) setBit :: \'a::len word \ nat \ 'a word\ where \setBit w n \ set_bit n w\ abbreviation (input) clearBit :: \'a::len word \ nat \ 'a word\ where \clearBit w n \ unset_bit n w\ lemma bit_setBit_iff: \bit (setBit w m) n \ (m = n \ n < LENGTH('a) \ bit w n)\ for w :: \'a::len word\ by (auto simp add: bit_simps) lemma bit_clearBit_iff: \bit (clearBit w m) n \ m \ n \ bit w n\ for w :: \'a::len word\ by (auto simp add: bit_simps) lemmas less_def = less_eq [symmetric] lemmas le_def = not_less [symmetric, where ?'a = nat] end diff --git a/thys/Word_Lib/Word_Lemmas.thy b/thys/Word_Lib/Word_Lemmas.thy --- a/thys/Word_Lib/Word_Lemmas.thy +++ b/thys/Word_Lib/Word_Lemmas.thy @@ -1,1940 +1,1944 @@ (* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) section "Lemmas with Generic Word Length" theory Word_Lemmas imports Type_Syntax Signed_Division_Word Signed_Words More_Word Most_significant_bit Enumeration_Word Aligned Bit_Shifts_Infix_Syntax begin context includes bit_operations_syntax begin lemma ucast_le_ucast_eq: fixes x y :: "'a::len word" assumes x: "x < 2 ^ n" assumes y: "y < 2 ^ n" assumes n: "n = LENGTH('b::len)" shows "(UCAST('a \ 'b) x \ UCAST('a \ 'b) y) = (x \ y)" apply (rule iffI) apply (cases "LENGTH('b) < LENGTH('a)") apply (subst less_mask_eq[OF x, symmetric]) apply (subst less_mask_eq[OF y, symmetric]) apply (unfold n) apply (subst ucast_ucast_mask[symmetric])+ apply (simp add: ucast_le_ucast)+ apply (erule ucast_mono_le[OF _ y[unfolded n]]) done lemma ucast_zero_is_aligned: \is_aligned w n\ if \UCAST('a::len \ 'b::len) w = 0\ \n \ LENGTH('b)\ proof (rule is_aligned_bitI) fix q assume \q < n\ moreover have \bit (UCAST('a::len \ 'b::len) w) q = bit 0 q\ using that by simp with \q < n\ \n \ LENGTH('b)\ show \\ bit w q\ by (simp add: bit_simps) qed lemma unat_ucast_eq_unat_and_mask: "unat (UCAST('b::len \ 'a::len) w) = unat (w AND mask LENGTH('a))" apply (simp flip: take_bit_eq_mask) apply transfer apply (simp add: ac_simps) done lemma le_max_word_ucast_id: \UCAST('b \ 'a) (UCAST('a \ 'b) x) = x\ if \x \ UCAST('b::len \ 'a) (- 1)\ for x :: \'a::len word\ proof - from that have a1: \x \ word_of_int (uint (word_of_int (2 ^ LENGTH('b) - 1) :: 'b word))\ by simp have f2: "((\i ia. (0::int) \ i \ \ 0 \ i + - 1 * ia \ i mod ia \ i) \ \ (0::int) \ - 1 + 2 ^ LENGTH('b) \ (0::int) \ - 1 + 2 ^ LENGTH('b) + - 1 * 2 ^ LENGTH('b) \ (- (1::int) + 2 ^ LENGTH('b)) mod 2 ^ LENGTH('b) = - 1 + 2 ^ LENGTH('b)) = ((\i ia. (0::int) \ i \ \ 0 \ i + - 1 * ia \ i mod ia \ i) \ \ (1::int) \ 2 ^ LENGTH('b) \ 2 ^ LENGTH('b) + - (1::int) * ((- 1 + 2 ^ LENGTH('b)) mod 2 ^ LENGTH('b)) = 1)" by force have f3: "\i ia. \ (0::int) \ i \ 0 \ i + - 1 * ia \ i mod ia = i" using mod_pos_pos_trivial by force have "(1::int) \ 2 ^ LENGTH('b)" by simp then have "2 ^ LENGTH('b) + - (1::int) * ((- 1 + 2 ^ LENGTH('b)) mod 2 ^ len_of TYPE ('b)) = 1" using f3 f2 by blast then have f4: "- (1::int) + 2 ^ LENGTH('b) = (- 1 + 2 ^ LENGTH('b)) mod 2 ^ LENGTH('b)" by linarith have f5: "x \ word_of_int (uint (word_of_int (- 1 + 2 ^ LENGTH('b))::'b word))" using a1 by force have f6: "2 ^ LENGTH('b) + - (1::int) = - 1 + 2 ^ LENGTH('b)" by force have f7: "- (1::int) * 1 = - 1" by auto have "\x0 x1. (x1::int) - x0 = x1 + - 1 * x0" by force then have "x \ 2 ^ LENGTH('b) - 1" using f7 f6 f5 f4 by (metis uint_word_of_int wi_homs(2) word_arith_wis(8) word_of_int_2p) then have \uint x \ uint (2 ^ LENGTH('b) - (1 :: 'a word))\ by (simp add: word_le_def) then have \uint x \ 2 ^ LENGTH('b) - 1\ by (simp add: uint_word_ariths) (metis \1 \ 2 ^ LENGTH('b)\ \uint x \ uint (2 ^ LENGTH('b) - 1)\ linorder_not_less lt2p_lem uint_1 uint_minus_simple_alt uint_power_lower word_le_def zle_diff1_eq) then show ?thesis apply (simp add: unsigned_ucast_eq take_bit_word_eq_self_iff) apply (meson \x \ 2 ^ LENGTH('b) - 1\ not_le word_less_sub_le) done qed lemma uint_shiftr_eq: \uint (w >> n) = uint w div 2 ^ n\ by transfer (simp flip: drop_bit_eq_div add: drop_bit_take_bit min_def le_less less_diff_conv) lemma bit_shiftl_word_iff [bit_simps]: \bit (w << m) n \ m \ n \ n < LENGTH('a) \ bit w (n - m)\ for w :: \'a::len word\ by (simp add: bit_push_bit_iff not_le) lemma shiftl_def: \w << n = ((*) 2 ^^ n) w\ for w :: \'a::len word\ proof - have \push_bit n = (((*) 2 ^^ n) :: int \ int)\ for n by (induction n) (simp_all add: fun_eq_iff funpow_swap1, simp add: ac_simps) then show ?thesis by transfer simp qed lemma shiftr_def: \w >> n = ((\w. w div 2) ^^ n) w\ for w :: \'a::len word\ proof - have \(\w. w div 2) ^^ n = (drop_bit n :: 'a word \ 'a word)\ by (induction n) (simp_all add: drop_bit_half drop_bit_Suc) then show ?thesis by simp qed lemma bit_shiftr_word_iff: \bit (w >> m) n \ bit w (m + n)\ for w :: \'a::len word\ by (simp add: bit_simps) lemma sshiftr_eq_funpow_sshiftr1: \w >>> n = (signed_drop_bit (Suc 0) ^^ n) w\ apply (rule sym) apply (induction n) apply simp_all done lemma uint_sshiftr_eq: \uint (w >>> n) = take_bit LENGTH('a) (sint w div 2 ^ n)\ for w :: \'a::len word\ by transfer (simp flip: drop_bit_eq_div) lemma sshiftr_0: "0 >>> n = 0" by (fact signed_drop_bit_of_0) lemma sshiftr_n1: "-1 >>> n = -1" by (fact signed_drop_bit_of_minus_1) lemma bit_sshiftr_word_iff: \bit (w >>> m) n \ bit w (if LENGTH('a) - m \ n \ n < LENGTH('a) then LENGTH('a) - 1 else (m + n))\ for w :: \'a::len word\ by (fact bit_signed_drop_bit_iff) lemma nth_sshiftr : "bit (w >>> m) n = (n < size w \ (if n + m \ size w then bit w (size w - 1) else bit w (n + m)))" apply (auto simp add: bit_simps word_size ac_simps not_less) apply (meson bit_imp_le_length bit_shiftr_word_iff leD) done lemma sshiftr_numeral: \(numeral k >>> numeral n :: 'a::len word) = word_of_int (drop_bit (numeral n) (signed_take_bit (LENGTH('a) - 1) (numeral k)))\ by (fact signed_drop_bit_word_numeral) lemma sshiftr_div_2n: "sint (w >>> n) = sint w div 2 ^ n" using sint_signed_drop_bit_eq [of n w] by (simp add: drop_bit_eq_div) lemma mask_eq: \mask n = (1 << n) - (1 :: 'a::len word)\ by (simp add: mask_eq_exp_minus_1 push_bit_of_1) lemma shiftl_0: "(0::'a::len word) << n = 0" by (fact push_bit_of_0) lemma shiftr_0: "(0::'a::len word) >> n = 0" by (fact drop_bit_of_0) lemma nth_shiftl': "bit (w << m) n \ n < size w \ n >= m \ bit w (n - m)" for w :: "'a::len word" by transfer (auto simp add: bit_push_bit_iff) lemmas nth_shiftl = nth_shiftl' [unfolded word_size] lemma nth_shiftr: "bit (w >> m) n = bit w (n + m)" for w :: "'a::len word" by (simp add: bit_simps ac_simps) lemma shiftr_div_2n: "uint (shiftr w n) = uint w div 2 ^ n" by (fact uint_shiftr_eq) lemma shiftl_rev: "shiftl w n = word_reverse (shiftr (word_reverse w) n)" by (rule bit_word_eqI) (auto simp add: bit_simps) lemma rev_shiftl: "word_reverse w << n = word_reverse (w >> n)" by (simp add: shiftl_rev) lemma shiftr_rev: "w >> n = word_reverse (word_reverse w << n)" by (simp add: rev_shiftl) lemma rev_shiftr: "word_reverse w >> n = word_reverse (w << n)" by (simp add: shiftr_rev) lemmas ucast_up = rc1 [simplified rev_shiftr [symmetric] revcast_ucast [symmetric]] lemmas ucast_down = rc2 [simplified rev_shiftr revcast_ucast [symmetric]] lemma shiftl_zero_size: "size x \ n \ x << n = 0" for x :: "'a::len word" apply transfer apply (simp add: take_bit_push_bit) done lemma shiftl_t2n: "shiftl w n = 2 ^ n * w" for w :: "'a::len word" by (simp add: push_bit_eq_mult) lemma word_shift_by_2: "x * 4 = (x::'a::len word) << 2" by (simp add: shiftl_t2n) lemma slice_shiftr: "slice n w = ucast (w >> n)" apply (rule bit_word_eqI) apply (cases \n \ LENGTH('b)\) apply (auto simp add: bit_slice_iff bit_ucast_iff bit_shiftr_word_iff ac_simps dest: bit_imp_le_length) done lemma shiftr_zero_size: "size x \ n \ x >> n = 0" for x :: "'a :: len word" by (rule word_eqI) (auto simp add: nth_shiftr dest: test_bit_size) lemma shiftr_x_0: "x >> 0 = x" for x :: "'a::len word" by simp lemma shiftl_x_0: "x << 0 = x" for x :: "'a::len word" by simp lemma shiftl_1: "(1::'a::len word) << n = 2^n" by (fact push_bit_of_1) lemma shiftr_1: "(1::'a::len word) >> n = (if n = 0 then 1 else 0)" by simp lemma shiftl0: "x << 0 = (x :: 'a :: len word)" by (fact shiftl_x_0) lemma and_not_mask: "w AND NOT (mask n) = (w >> n) << n" for w :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_simps) lemma and_mask: "w AND mask n = (w << (size w - n)) >> (size w - n)" for w :: \'a::len word\ by (rule bit_word_eqI) (auto simp add: bit_simps word_size) lemma shiftr_div_2n_w: "n < size w \ w >> n = w div (2^n :: 'a :: len word)" apply (unfold word_div_def) apply (simp add: uint_2p_alt word_size) apply (metis uint_shiftr_eq word_of_int_uint) done lemma le_shiftr: "u \ v \ u >> (n :: nat) \ (v :: 'a :: len word) >> n" apply transfer apply (simp add: take_bit_drop_bit) apply (simp add: drop_bit_eq_div zdiv_mono1) done lemma le_shiftr': "\ u >> n \ v >> n ; u >> n \ v >> n \ \ (u::'a::len word) \ v" apply (metis le_cases le_shiftr verit_la_disequality) done lemma shiftr_mask_le: "n <= m \ mask n >> m = (0 :: 'a::len word)" by (rule bit_word_eqI) (auto simp add: bit_simps) lemma shiftr_mask [simp]: \mask m >> m = (0::'a::len word)\ by (rule shiftr_mask_le) simp lemma le_mask_iff: "(w \ mask n) = (w >> n = 0)" for w :: \'a::len word\ apply safe apply (rule word_le_0_iff [THEN iffD1]) apply (rule xtrans(3)) apply (erule_tac [2] le_shiftr) apply simp apply (rule word_leI) apply (rename_tac n') apply (drule_tac x = "n' - n" in word_eqD) apply (simp add : nth_shiftr word_size bit_simps) apply (case_tac "n <= n'") by auto lemma and_mask_eq_iff_shiftr_0: "(w AND mask n = w) = (w >> n = 0)" for w :: \'a::len word\ apply (unfold test_bit_eq_iff [THEN sym]) apply (rule iffI) apply (rule ext) apply (rule_tac [2] ext) apply (auto simp add : word_ao_nth nth_shiftr) apply (drule arg_cong) apply (drule iffD2) apply assumption apply (simp add : word_ao_nth) prefer 2 apply (simp add : word_size test_bit_bin) apply transfer apply (auto simp add: fun_eq_iff bit_simps) apply (metis add_diff_inverse_nat) done lemma mask_shiftl_decompose: "mask m << n = mask (m + n) AND NOT (mask n :: 'a::len word)" by (rule bit_word_eqI) (auto simp add: bit_simps) lemma shiftl_over_and_dist: fixes a::"'a::len word" shows "(a AND b) << c = (a << c) AND (b << c)" by (fact push_bit_and) lemma shiftr_over_and_dist: fixes a::"'a::len word" shows "a AND b >> c = (a >> c) AND (b >> c)" by (fact drop_bit_and) lemma sshiftr_over_and_dist: fixes a::"'a::len word" shows "a AND b >>> c = (a >>> c) AND (b >>> c)" apply(rule word_eqI) apply(simp add:nth_sshiftr word_ao_nth word_size) done lemma shiftl_over_or_dist: fixes a::"'a::len word" shows "a OR b << c = (a << c) OR (b << c)" by (fact push_bit_or) lemma shiftr_over_or_dist: fixes a::"'a::len word" shows "a OR b >> c = (a >> c) OR (b >> c)" by (fact drop_bit_or) lemma sshiftr_over_or_dist: fixes a::"'a::len word" shows "a OR b >>> c = (a >>> c) OR (b >>> c)" apply(rule word_eqI) apply(simp add:nth_sshiftr word_ao_nth word_size) done lemmas shift_over_ao_dists = shiftl_over_or_dist shiftr_over_or_dist sshiftr_over_or_dist shiftl_over_and_dist shiftr_over_and_dist sshiftr_over_and_dist lemma shiftl_shiftl: fixes a::"'a::len word" shows "a << b << c = a << (b + c)" apply(rule word_eqI) apply(auto simp:word_size nth_shiftl add.commute add.left_commute) done lemma shiftr_shiftr: fixes a::"'a::len word" shows "a >> b >> c = a >> (b + c)" apply(rule word_eqI) apply(simp add:word_size nth_shiftr add.left_commute add.commute) done lemma shiftl_shiftr1: fixes a::"'a::len word" shows "c \ b \ a << b >> c = a AND (mask (size a - b)) << (b - c)" apply (rule word_eqI) apply (auto simp add: bit_simps not_le word_size ac_simps) done lemma shiftl_shiftr2: fixes a::"'a::len word" shows "b < c \ a << b >> c = (a >> (c - b)) AND (mask (size a - c))" apply(rule word_eqI) apply(auto simp:nth_shiftr nth_shiftl word_size word_ao_nth bit_simps) done lemma shiftr_shiftl1: fixes a::"'a::len word" shows "c \ b \ a >> b << c = (a >> (b - c)) AND (NOT (mask c))" by (rule bit_word_eqI) (auto simp add: bit_simps) lemma shiftr_shiftl2: fixes a::"'a::len word" shows "b < c \ a >> b << c = (a << (c - b)) AND (NOT (mask c))" apply (rule word_eqI) apply (auto simp add: bit_simps not_le word_size ac_simps) done lemmas multi_shift_simps = shiftl_shiftl shiftr_shiftr shiftl_shiftr1 shiftl_shiftr2 shiftr_shiftl1 shiftr_shiftl2 lemma shiftr_mask2: "n \ LENGTH('a) \ (mask n >> m :: ('a :: len) word) = mask (n - m)" by (rule bit_word_eqI) (auto simp add: bit_simps) lemma word_shiftl_add_distrib: fixes x :: "'a :: len word" shows "(x + y) << n = (x << n) + (y << n)" by (simp add: shiftl_t2n ring_distribs) lemma mask_shift: "(x AND NOT (mask y)) >> y = x >> y" for x :: \'a::len word\ apply (rule bit_eqI) apply (simp add: bit_and_iff bit_not_iff bit_shiftr_word_iff bit_mask_iff not_le) using bit_imp_le_length apply auto done lemma shiftr_div_2n': "unat (w >> n) = unat w div 2 ^ n" apply (unfold unat_eq_nat_uint) apply (subst shiftr_div_2n) apply (subst nat_div_distrib) apply simp apply (simp add: nat_power_eq) done lemma shiftl_shiftr_id: assumes nv: "n < LENGTH('a)" and xv: "x < 2 ^ (LENGTH('a) - n)" shows "x << n >> n = (x::'a::len word)" apply (simp add: shiftl_t2n) apply (rule word_eq_unatI) apply (subst shiftr_div_2n') apply (cases n) apply simp apply (subst iffD1 [OF unat_mult_lem])+ apply (subst unat_power_lower[OF nv]) apply (rule nat_less_power_trans [OF _ order_less_imp_le [OF nv]]) apply (rule order_less_le_trans [OF unat_mono [OF xv] order_eq_refl]) apply (rule unat_power_lower) apply simp apply (subst unat_power_lower[OF nv]) apply simp done lemma ucast_shiftl_eq_0: fixes w :: "'a :: len word" shows "\ n \ LENGTH('b) \ \ ucast (w << n) = (0 :: 'b :: len word)" by transfer (simp add: take_bit_push_bit) lemma word_shift_nonzero: "\ (x::'a::len word) \ 2 ^ m; m + n < LENGTH('a::len); x \ 0\ \ x << n \ 0" apply (simp only: word_neq_0_conv word_less_nat_alt shiftl_t2n mod_0 unat_word_ariths unat_power_lower word_le_nat_alt) apply (subst mod_less) apply (rule order_le_less_trans) apply (erule mult_le_mono2) apply (subst power_add[symmetric]) apply (rule power_strict_increasing) apply simp apply simp apply simp done lemma word_shiftr_lt: fixes w :: "'a::len word" shows "unat (w >> n) < (2 ^ (LENGTH('a) - n))" apply (subst shiftr_div_2n') apply transfer apply (simp flip: drop_bit_eq_div add: drop_bit_nat_eq drop_bit_take_bit) done lemma shiftr_less_t2n': "\ x AND mask (n + m) = x; m < LENGTH('a) \ \ x >> n < 2 ^ m" for x :: "'a :: len word" apply (simp add: word_size mask_eq_iff_w2p [symmetric] flip: take_bit_eq_mask) apply transfer apply (simp add: take_bit_drop_bit ac_simps) done lemma shiftr_less_t2n: "x < 2 ^ (n + m) \ x >> n < 2 ^ m" for x :: "'a :: len word" apply (rule shiftr_less_t2n') apply (erule less_mask_eq) apply (rule ccontr) apply (simp add: not_less) apply (subst (asm) p2_eq_0[symmetric]) apply (simp add: power_add) done lemma shiftr_eq_0: "n \ LENGTH('a) \ ((w::'a::len word) >> n) = 0" apply (cut_tac shiftr_less_t2n'[of w n 0], simp) apply (simp add: mask_eq_iff) apply (simp add: lt2p_lem) apply simp done lemma shiftl_less_t2n: fixes x :: "'a :: len word" shows "\ x < (2 ^ (m - n)); m < LENGTH('a) \ \ (x << n) < 2 ^ m" apply (simp add: word_size mask_eq_iff_w2p [symmetric] flip: take_bit_eq_mask) apply transfer apply (simp add: take_bit_push_bit) done lemma shiftl_less_t2n': "(x::'a::len word) < 2 ^ m \ m+n < LENGTH('a) \ x << n < 2 ^ (m + n)" by (rule shiftl_less_t2n) simp_all lemma scast_bit_test [simp]: "scast ((1 :: 'a::len signed word) << n) = (1 :: 'a word) << n" by (rule bit_word_eqI) (simp add: bit_simps) lemma signed_shift_guard_to_word: \unat x * 2 ^ y < 2 ^ n \ x = 0 \ x < 1 << n >> y\ if \n < LENGTH('a)\ \0 < n\ for x :: \'a::len word\ proof (cases \x = 0\) case True then show ?thesis by simp next case False then have \unat x \ 0\ by (simp add: unat_eq_0) then have \unat x \ 1\ by simp show ?thesis proof (cases \y < n\) case False then have \n \ y\ by simp then obtain q where \y = n + q\ using le_Suc_ex by blast moreover have \(2 :: nat) ^ n >> n + q \ 1\ by (simp add: drop_bit_eq_div power_add) ultimately show ?thesis using \x \ 0\ \unat x \ 1\ \n < LENGTH('a)\ by (simp add: power_add not_less word_le_nat_alt unat_drop_bit_eq push_bit_of_1) next case True with that have \y < LENGTH('a)\ by simp show ?thesis proof (cases \2 ^ n = unat x * 2 ^ y\) case True moreover have \unat x * 2 ^ y < 2 ^ LENGTH('a)\ using \n < LENGTH('a)\ by (simp flip: True) moreover have \(word_of_nat (2 ^ n) :: 'a word) = word_of_nat (unat x * 2 ^ y)\ using True by simp then have \2 ^ n = x * 2 ^ y\ by simp ultimately show ?thesis using \y < LENGTH('a)\ by (auto simp add: push_bit_of_1 drop_bit_eq_div word_less_nat_alt unat_div unat_word_ariths) next case False with \y < n\ have *: \unat x \ 2 ^ n div 2 ^ y\ by (auto simp flip: power_sub power_add) have \unat x * 2 ^ y < 2 ^ n \ unat x * 2 ^ y \ 2 ^ n\ using False by (simp add: less_le) also have \\ \ unat x \ 2 ^ n div 2 ^ y\ by (simp add: less_eq_div_iff_mult_less_eq) also have \\ \ unat x < 2 ^ n div 2 ^ y\ using * by (simp add: less_le) finally show ?thesis using that \x \ 0\ by (simp flip: push_bit_eq_mult drop_bit_eq_div add: push_bit_of_1 unat_drop_bit_eq word_less_iff_unsigned [where ?'a = nat]) qed qed qed lemma shiftr_not_mask_0: "n+m \ LENGTH('a :: len) \ ((w::'a::len word) >> n) AND NOT (mask m) = 0" by (rule bit_word_eqI) (auto simp add: bit_simps word_size dest: bit_imp_le_length) lemma shiftl_mask_is_0[simp]: "(x << n) AND mask n = 0" for x :: \'a::len word\ by (simp flip: take_bit_eq_mask add: take_bit_push_bit) lemma rshift_sub_mask_eq: "(a >> (size a - b)) AND mask b = a >> (size a - b)" for a :: \'a::len word\ using shiftl_shiftr2[where a=a and b=0 and c="size a - b"] apply (cases "b < size a") apply simp apply (simp add: linorder_not_less mask_eq_decr_exp word_size p2_eq_0[THEN iffD2]) done lemma shiftl_shiftr3: "b \ c \ a << b >> c = (a >> c - b) AND mask (size a - c)" for a :: \'a::len word\ apply (cases "b = c") apply (simp add: shiftl_shiftr1) apply (simp add: shiftl_shiftr2) done lemma and_mask_shiftr_comm: "m \ size w \ (w AND mask m) >> n = (w >> n) AND mask (m-n)" for w :: \'a::len word\ by (simp add: and_mask shiftr_shiftr) (simp add: word_size shiftl_shiftr3) lemma and_mask_shiftl_comm: "m+n \ size w \ (w AND mask m) << n = (w << n) AND mask (m+n)" for w :: \'a::len word\ by (simp add: and_mask word_size shiftl_shiftl) (simp add: shiftl_shiftr1) lemma le_mask_shiftl_le_mask: "s = m + n \ x \ mask n \ x << m \ mask s" for x :: \'a::len word\ by (simp add: le_mask_iff shiftl_shiftr3) lemma word_and_1_shiftl: "x AND (1 << n) = (if bit x n then (1 << n) else 0)" for x :: "'a :: len word" apply (rule bit_word_eqI; transfer) apply (auto simp add: bit_simps not_le ac_simps) done lemmas word_and_1_shiftls' = word_and_1_shiftl[where n=0] word_and_1_shiftl[where n=1] word_and_1_shiftl[where n=2] lemmas word_and_1_shiftls = word_and_1_shiftls' [simplified] lemma word_and_mask_shiftl: "x AND (mask n << m) = ((x >> m) AND mask n) << m" for x :: \'a::len word\ apply (rule bit_word_eqI; transfer) apply (auto simp add: bit_simps not_le ac_simps) done lemma shift_times_fold: "(x :: 'a :: len word) * (2 ^ n) << m = x << (m + n)" by (simp add: shiftl_t2n ac_simps power_add) lemma of_bool_nth: "of_bool (bit x v) = (x >> v) AND 1" for x :: \'a::len word\ by (simp add: bit_iff_odd_drop_bit word_and_1) lemma shiftr_mask_eq: "(x >> n) AND mask (size x - n) = x >> n" for x :: "'a :: len word" apply (simp flip: take_bit_eq_mask) apply transfer apply (simp add: take_bit_drop_bit) done lemma shiftr_mask_eq': "m = (size x - n) \ (x >> n) AND mask m = x >> n" for x :: "'a :: len word" by (simp add: shiftr_mask_eq) lemma and_eq_0_is_nth: fixes x :: "'a :: len word" shows "y = 1 << n \ ((x AND y) = 0) = (\ (bit x n))" by (simp add: and_exp_eq_0_iff_not_bit push_bit_of_1) lemma word_shift_zero: "\ x << n = 0; x \ 2^m; m + n < LENGTH('a)\ \ (x::'a::len word) = 0" apply (rule ccontr) apply (drule (2) word_shift_nonzero) apply simp done lemma mask_shift_and_negate[simp]:"(w AND mask n << m) AND NOT (mask n << m) = 0" for w :: \'a::len word\ by (rule bit_word_eqI) (simp add: bit_simps) (* The seL4 bitfield generator produces functions containing mask and shift operations, such that * invoking two of them consecutively can produce something like the following. *) lemma bitfield_op_twice: "(x AND NOT (mask n << m) OR ((y AND mask n) << m)) AND NOT (mask n << m) = x AND NOT (mask n << m)" for x :: \'a::len word\ by (induct n arbitrary: m) (auto simp: word_ao_dist) lemma bitfield_op_twice'': "\NOT a = b << c; \x. b = mask x\ \ (x AND a OR (y AND b << c)) AND a = x AND a" for a b :: \'a::len word\ apply clarsimp apply (cut_tac n=xa and m=c and x=x and y=y in bitfield_op_twice) apply (clarsimp simp:mask_eq_decr_exp) apply (drule not_switch) apply clarsimp done lemma shiftr1_unfold: "x div 2 = x >> 1" by (simp add: drop_bit_eq_div) lemma shiftr1_is_div_2: "(x::('a::len) word) >> 1 = x div 2" by (simp add: drop_bit_eq_div) lemma shiftl1_is_mult: "(x << 1) = (x :: 'a::len word) * 2" by (metis One_nat_def mult_2 mult_2_right one_add_one power_0 power_Suc shiftl_t2n) lemma shiftr1_lt:"x \ 0 \ (x::('a::len) word) >> 1 < x" apply (subst shiftr1_is_div_2) apply (rule div_less_dividend_word) apply simp+ done lemma shiftr1_0_or_1:"(x::('a::len) word) >> 1 = 0 \ x = 0 \ x = 1" apply (subst (asm) shiftr1_is_div_2) apply (drule word_less_div) apply (case_tac "LENGTH('a) = 1") apply (simp add:degenerate_word) apply (erule disjE) apply (subgoal_tac "(2::'a word) \ 0") apply simp apply (rule not_degenerate_imp_2_neq_0) apply (subgoal_tac "LENGTH('a) \ 0") apply arith apply simp apply (rule x_less_2_0_1', simp+) done lemma shiftr1_irrelevant_lsb: "bit (x::('a::len) word) 0 \ x >> 1 = (x + 1) >> 1" apply (cases \LENGTH('a)\; transfer) apply (simp_all add: take_bit_drop_bit) apply (simp add: drop_bit_take_bit drop_bit_Suc) done lemma shiftr1_0_imp_only_lsb:"((x::('a::len) word) + 1) >> 1 = 0 \ x = 0 \ x + 1 = 0" by (metis One_nat_def shiftr1_0_or_1 word_less_1 word_overflow) lemma shiftr1_irrelevant_lsb': "\ (bit (x::('a::len) word) 0) \ x >> 1 = (x + 1) >> 1" by (metis shiftr1_irrelevant_lsb) (* Perhaps this one should be a simp lemma, but it seems a little dangerous. *) lemma cast_chunk_assemble_id: "\n = LENGTH('a::len); m = LENGTH('b::len); n * 2 = m\ \ (((ucast ((ucast (x::'b word))::'a word))::'b word) OR (((ucast ((ucast (x >> n))::'a word))::'b word) << n)) = x" apply (subgoal_tac "((ucast ((ucast (x >> n))::'a word))::'b word) = x >> n") apply clarsimp apply (subst and_not_mask[symmetric]) apply (subst ucast_ucast_mask) apply (subst word_ao_dist2[symmetric]) apply clarsimp apply (rule ucast_ucast_len) apply (rule shiftr_less_t2n') apply (subst and_mask_eq_iff_le_mask) apply (simp_all add: mask_eq_decr_exp flip: mult_2_right) apply (metis add_diff_cancel_left' len_gt_0 mult_2_right zero_less_diff) done lemma cast_chunk_scast_assemble_id: "\n = LENGTH('a::len); m = LENGTH('b::len); n * 2 = m\ \ (((ucast ((scast (x::'b word))::'a word))::'b word) OR (((ucast ((scast (x >> n))::'a word))::'b word) << n)) = x" apply (subgoal_tac "((scast x)::'a word) = ((ucast x)::'a word)") apply (subgoal_tac "((scast (x >> n))::'a word) = ((ucast (x >> n))::'a word)") apply (simp add:cast_chunk_assemble_id) apply (subst down_cast_same[symmetric], subst is_down, arith, simp)+ done lemma unat_shiftr_less_t2n: fixes x :: "'a :: len word" shows "unat x < 2 ^ (n + m) \ unat (x >> n) < 2 ^ m" by (simp add: shiftr_div_2n' power_add mult.commute less_mult_imp_div_less) lemma ucast_less_shiftl_helper: "\ LENGTH('b) + 2 < LENGTH('a); 2 ^ (LENGTH('b) + 2) \ n\ \ (ucast (x :: 'b::len word) << 2) < (n :: 'a::len word)" apply (erule order_less_le_trans[rotated]) using ucast_less[where x=x and 'a='a] apply (simp only: shiftl_t2n field_simps) apply (rule word_less_power_trans2; simp) done (* negating a mask which has been shifted to the very left *) lemma NOT_mask_shifted_lenword: "NOT (mask len << (LENGTH('a) - len) ::'a::len word) = mask (LENGTH('a) - len)" by (rule bit_word_eqI) (auto simp add: word_size bit_not_iff bit_push_bit_iff bit_mask_iff) (* Comparisons between different word sizes. *) lemma shiftr_less: "(w::'a::len word) < k \ w >> n < k" by (metis div_le_dividend le_less_trans shiftr_div_2n' unat_arith_simps(2)) lemma word_and_notzeroD: "w AND w' \ 0 \ w \ 0 \ w' \ 0" by auto lemma shiftr_le_0: "unat (w::'a::len word) < 2 ^ n \ w >> n = (0::'a::len word)" apply (auto simp add: take_bit_word_eq_self_iff word_less_nat_alt simp flip: take_bit_eq_self_iff_drop_bit_eq_0) apply (rule ccontr) apply (simp add: not_le) done lemma of_nat_shiftl: "(of_nat x << n) = (of_nat (x * 2 ^ n) :: ('a::len) word)" proof - have "(of_nat x::'a word) << n = of_nat (2 ^ n) * of_nat x" using shiftl_t2n by (metis word_unat_power) thus ?thesis by simp qed lemma shiftl_1_not_0: "n < LENGTH('a) \ (1::'a::len word) << n \ 0" by (simp add: shiftl_t2n) (* continue sorting out from here *) (* usually: x,y = (len_of TYPE ('a)) *) lemma bitmagic_zeroLast_leq_or1Last: "(a::('a::len) word) AND (mask len << x - len) \ a OR mask (y - len)" by (meson le_word_or2 order_trans word_and_le2) lemma zero_base_lsb_imp_set_eq_as_bit_operation: fixes base ::"'a::len word" assumes valid_prefix: "mask (LENGTH('a) - len) AND base = 0" shows "(base = NOT (mask (LENGTH('a) - len)) AND a) \ (a \ {base .. base OR mask (LENGTH('a) - len)})" proof have helper3: "x OR y = x OR y AND NOT x" for x y ::"'a::len word" by (simp add: word_oa_dist2) from assms show "base = NOT (mask (LENGTH('a) - len)) AND a \ a \ {base..base OR mask (LENGTH('a) - len)}" apply(simp add: word_and_le1) apply(metis helper3 le_word_or2 word_bw_comms(1) word_bw_comms(2)) done next assume "a \ {base..base OR mask (LENGTH('a) - len)}" hence a: "base \ a \ a \ base OR mask (LENGTH('a) - len)" by simp show "base = NOT (mask (LENGTH('a) - len)) AND a" proof - have f2: "\x\<^sub>0. base AND NOT (mask x\<^sub>0) \ a AND NOT (mask x\<^sub>0)" using a neg_mask_mono_le by blast have f3: "\x\<^sub>0. a AND NOT (mask x\<^sub>0) \ (base OR mask (LENGTH('a) - len)) AND NOT (mask x\<^sub>0)" using a neg_mask_mono_le by blast have f4: "base = base AND NOT (mask (LENGTH('a) - len))" using valid_prefix by (metis mask_eq_0_eq_x word_bw_comms(1)) hence f5: "\x\<^sub>6. (base OR x\<^sub>6) AND NOT (mask (LENGTH('a) - len)) = base OR x\<^sub>6 AND NOT (mask (LENGTH('a) - len))" using word_ao_dist by (metis) have f6: "\x\<^sub>2 x\<^sub>3. a AND NOT (mask x\<^sub>2) \ x\<^sub>3 \ \ (base OR mask (LENGTH('a) - len)) AND NOT (mask x\<^sub>2) \ x\<^sub>3" using f3 dual_order.trans by auto have "base = (base OR mask (LENGTH('a) - len)) AND NOT (mask (LENGTH('a) - len))" using f5 by auto hence "base = a AND NOT (mask (LENGTH('a) - len))" using f2 f4 f6 by (metis eq_iff) thus "base = NOT (mask (LENGTH('a) - len)) AND a" by (metis word_bw_comms(1)) qed qed lemma of_nat_eq_signed_scast: "(of_nat x = (y :: ('a::len) signed word)) = (of_nat x = (scast y :: 'a word))" by (metis scast_of_nat scast_scast_id(2)) lemma word_aligned_add_no_wrap_bounded: "\ w + 2^n \ x; w + 2^n \ 0; is_aligned w n \ \ (w::'a::len word) < x" by (blast dest: is_aligned_no_overflow le_less_trans word_leq_le_minus_one) lemma mask_Suc: "mask (Suc n) = (2 :: 'a::len word) ^ n + mask n" by (simp add: mask_eq_decr_exp) lemma mask_mono: "sz' \ sz \ mask sz' \ (mask sz :: 'a::len word)" by (simp add: le_mask_iff shiftr_mask_le) lemma aligned_mask_disjoint: "\ is_aligned (a :: 'a :: len word) n; b \ mask n \ \ a AND b = 0" by (metis and_zero_eq is_aligned_mask le_mask_imp_and_mask word_bw_lcs(1)) lemma word_and_or_mask_aligned: "\ is_aligned a n; b \ mask n \ \ a + b = a OR b" by (simp add: aligned_mask_disjoint word_plus_and_or_coroll) lemma word_and_or_mask_aligned2: \is_aligned b n \ a \ mask n \ a + b = a OR b\ using word_and_or_mask_aligned [of b n a] by (simp add: ac_simps) lemma is_aligned_ucastI: "is_aligned w n \ is_aligned (ucast w) n" by (simp add: bit_ucast_iff is_aligned_nth) lemma ucast_le_maskI: "a \ mask n \ UCAST('a::len \ 'b::len) a \ mask n" by (metis and_mask_eq_iff_le_mask ucast_and_mask) lemma ucast_add_mask_aligned: "\ a \ mask n; is_aligned b n \ \ UCAST ('a::len \ 'b::len) (a + b) = ucast a + ucast b" by (metis add.commute is_aligned_ucastI ucast_le_maskI ucast_or_distrib word_and_or_mask_aligned) lemma ucast_shiftl: "LENGTH('b) \ LENGTH ('a) \ UCAST ('a::len \ 'b::len) x << n = ucast (x << n)" by word_eqI_solve lemma ucast_leq_mask: "LENGTH('a) \ n \ ucast (x::'a::len word) \ mask n" apply (simp add: less_eq_mask_iff_take_bit_eq_self) apply transfer apply (simp add: ac_simps) done lemma shiftl_inj: "\ x << n = y << n; x \ mask (LENGTH('a)-n); y \ mask (LENGTH('a)-n) \ \ x = (y :: 'a :: len word)" apply word_eqI apply (rename_tac n') apply (case_tac "LENGTH('a) - n \ n'", simp) by (metis add.commute add.right_neutral diff_add_inverse le_diff_conv linorder_not_less zero_order(1)) lemma distinct_word_add_ucast_shift_inj: \p' = p \ off' = off\ if *: \p + (UCAST('a::len \ 'b::len) off << n) = p' + (ucast off' << n)\ and \is_aligned p n'\ \is_aligned p' n'\ \n' = n + LENGTH('a)\ \n' < LENGTH('b)\ proof - from \n' = n + LENGTH('a)\ have [simp]: \n' - n = LENGTH('a)\ \n + LENGTH('a) = n'\ by simp_all from \is_aligned p n'\ obtain q where p: \p = push_bit n' (word_of_nat q)\ \q < 2 ^ (LENGTH('b) - n')\ by (rule is_alignedE') from \is_aligned p' n'\ obtain q' where p': \p' = push_bit n' (word_of_nat q')\ \q' < 2 ^ (LENGTH('b) - n')\ by (rule is_alignedE') define m :: nat where \m = unat off\ then have off: \off = word_of_nat m\ by simp define m' :: nat where \m' = unat off'\ then have off': \off' = word_of_nat m'\ by simp have \push_bit n' q + take_bit n' (push_bit n m) < 2 ^ LENGTH('b)\ by (metis id_apply is_aligned_no_wrap''' of_nat_eq_id of_nat_push_bit p(1) p(2) take_bit_nat_eq_self_iff take_bit_nat_less_exp take_bit_push_bit that(2) that(5) unsigned_of_nat) moreover have \push_bit n' q' + take_bit n' (push_bit n m') < 2 ^ LENGTH('b)\ by (metis \n' - n = LENGTH('a)\ id_apply is_aligned_no_wrap''' m'_def of_nat_eq_id of_nat_push_bit off' p'(1) p'(2) take_bit_nat_eq_self_iff take_bit_push_bit that(3) that(5) unsigned_of_nat) ultimately have \push_bit n' q + take_bit n' (push_bit n m) = push_bit n' q' + take_bit n' (push_bit n m')\ using * by (simp add: p p' off off' push_bit_of_nat push_bit_take_bit word_of_nat_inj flip: of_nat_add) then have \int (push_bit n' q + take_bit n' (push_bit n m)) = int (push_bit n' q' + take_bit n' (push_bit n m'))\ by simp then have \concat_bit n' (int (push_bit n m)) (int q) = concat_bit n' (int (push_bit n m')) (int q')\ by (simp add: of_nat_push_bit of_nat_take_bit concat_bit_eq) then show ?thesis by (simp add: p p' off off' take_bit_of_nat take_bit_push_bit word_of_nat_eq_iff concat_bit_eq_iff) (simp add: push_bit_eq_mult) qed lemma word_upto_Nil: "y < x \ [x .e. y ::'a::len word] = []" by (simp add: upto_enum_red not_le word_less_nat_alt) lemma word_enum_decomp_elem: assumes "[x .e. (y ::'a::len word)] = as @ a # bs" shows "x \ a \ a \ y" proof - have "set as \ set [x .e. y] \ a \ set [x .e. y]" using assms by (auto dest: arg_cong[where f=set]) then show ?thesis by auto qed lemma word_enum_prefix: "[x .e. (y ::'a::len word)] = as @ a # bs \ as = (if x < a then [x .e. a - 1] else [])" apply (induct as arbitrary: x; clarsimp) apply (case_tac "x < y") prefer 2 apply (case_tac "x = y", simp) apply (simp add: not_less) apply (drule (1) dual_order.not_eq_order_implies_strict) apply (simp add: word_upto_Nil) apply (simp add: word_upto_Cons_eq) apply (case_tac "x < y") prefer 2 apply (case_tac "x = y", simp) apply (simp add: not_less) apply (drule (1) dual_order.not_eq_order_implies_strict) apply (simp add: word_upto_Nil) apply (clarsimp simp: word_upto_Cons_eq) apply (frule word_enum_decomp_elem) apply clarsimp apply (rule conjI) prefer 2 apply (subst word_Suc_le[symmetric]; clarsimp) apply (drule meta_spec) apply (drule (1) meta_mp) apply clarsimp apply (rule conjI; clarsimp) apply (subst (2) word_upto_Cons_eq) apply unat_arith apply simp done lemma word_enum_decomp_set: "[x .e. (y ::'a::len word)] = as @ a # bs \ a \ set as" by (metis distinct_append distinct_enum_upto' not_distinct_conv_prefix) lemma word_enum_decomp: assumes "[x .e. (y ::'a::len word)] = as @ a # bs" shows "x \ a \ a \ y \ a \ set as \ (\z \ set as. x \ z \ z \ y)" proof - from assms have "set as \ set [x .e. y] \ a \ set [x .e. y]" by (auto dest: arg_cong[where f=set]) with word_enum_decomp_set[OF assms] show ?thesis by auto qed lemma of_nat_unat_le_mask_ucast: "\of_nat (unat t) = w; t \ mask LENGTH('a)\ \ t = UCAST('a::len \ 'b::len) w" by (clarsimp simp: ucast_nat_def ucast_ucast_mask simp flip: and_mask_eq_iff_le_mask) lemma less_diff_gt0: "a < b \ (0 :: 'a :: len word) < b - a" by unat_arith lemma unat_plus_gt: "unat ((a :: 'a :: len word) + b) \ unat a + unat b" by (clarsimp simp: unat_plus_if_size) lemma const_less: "\ (a :: 'a :: len word) - 1 < b; a \ b \ \ a < b" by (metis less_1_simp word_le_less_eq) lemma add_mult_aligned_neg_mask: \(x + y * m) AND NOT(mask n) = (x AND NOT(mask n)) + y * m\ if \m AND (2 ^ n - 1) = 0\ for x y m :: \'a::len word\ by (metis (no_types, opaque_lifting) add.assoc add.commute add.right_neutral add_uminus_conv_diff mask_eq_decr_exp mask_eqs(2) mask_eqs(6) mult.commute mult_zero_left subtract_mask(1) that) lemma unat_of_nat_minus_1: "\ n < 2 ^ LENGTH('a); n \ 0 \ \ unat ((of_nat n:: 'a :: len word) - 1) = n - 1" by (simp add: of_nat_diff unat_eq_of_nat) lemma word_eq_zeroI: "a \ a - 1 \ a = 0" for a :: "'a :: len word" by (simp add: word_must_wrap) lemma word_add_format: "(-1 :: 'a :: len word) + b + c = b + (c - 1)" by simp lemma upto_enum_word_nth: "\ i \ j; k \ unat (j - i) \ \ [i .e. j] ! k = i + of_nat k" apply (clarsimp simp: upto_enum_def nth_append) apply (clarsimp simp: word_le_nat_alt[symmetric]) apply (rule conjI, clarsimp) apply (subst toEnum_of_nat, unat_arith) apply unat_arith apply (clarsimp simp: not_less unat_sub[symmetric]) apply unat_arith done lemma upto_enum_step_nth: "\ a \ c; n \ unat ((c - a) div (b - a)) \ \ [a, b .e. c] ! n = a + of_nat n * (b - a)" by (clarsimp simp: upto_enum_step_def not_less[symmetric] upto_enum_word_nth) lemma upto_enum_inc_1_len: "a < - 1 \ [(0 :: 'a :: len word) .e. 1 + a] = [0 .e. a] @ [1 + a]" apply (simp add: upto_enum_word) apply (subgoal_tac "unat (1+a) = 1 + unat a") apply simp apply (subst unat_plus_simple[THEN iffD1]) apply (metis add.commute no_plus_overflow_neg olen_add_eqv) apply unat_arith done lemma neg_mask_add: "y AND mask n = 0 \ x + y AND NOT(mask n) = (x AND NOT(mask n)) + y" for x y :: \'a::len word\ by (clarsimp simp: mask_out_sub_mask mask_eqs(7)[symmetric] mask_twice) lemma shiftr_shiftl_shiftr[simp]: "(x :: 'a :: len word) >> a << a >> a = x >> a" by word_eqI_solve lemma add_right_shift: "\ x AND mask n = 0; y AND mask n = 0; x \ x + y \ \ (x + y :: ('a :: len) word) >> n = (x >> n) + (y >> n)" apply (simp add: no_olen_add_nat is_aligned_mask[symmetric]) apply (simp add: unat_arith_simps shiftr_div_2n' split del: if_split) apply (subst if_P) apply (erule order_le_less_trans[rotated]) apply (simp add: add_mono) apply (simp add: shiftr_div_2n' is_aligned_iff_dvd_nat) done lemma sub_right_shift: "\ x AND mask n = 0; y AND mask n = 0; y \ x \ \ (x - y) >> n = (x >> n :: 'a :: len word) - (y >> n)" using add_right_shift[where x="x - y" and y=y and n=n] by (simp add: aligned_sub_aligned is_aligned_mask[symmetric] word_sub_le) lemma and_and_mask_simple: "y AND mask n = mask n \ (x AND y) AND mask n = x AND mask n" by (simp add: ac_simps) lemma and_and_mask_simple_not: "y AND mask n = 0 \ (x AND y) AND mask n = 0" by (simp add: ac_simps) lemma word_and_le': "b \ c \ (a :: 'a :: len word) AND b \ c" by (metis word_and_le1 order_trans) lemma word_and_less': "b < c \ (a :: 'a :: len word) AND b < c" by transfer simp lemma shiftr_w2p: "x < LENGTH('a) \ 2 ^ x = (2 ^ (LENGTH('a) - 1) >> (LENGTH('a) - 1 - x) :: 'a :: len word)" by word_eqI_solve lemma t2p_shiftr: "\ b \ a; a < LENGTH('a) \ \ (2 :: 'a :: len word) ^ a >> b = 2 ^ (a - b)" by word_eqI_solve lemma scast_1[simp]: "scast (1 :: 'a :: len signed word) = (1 :: 'a word)" by simp lemma unsigned_uminus1 [simp]: \(unsigned (-1::'b::len word)::'c::len word) = mask LENGTH('b)\ by (rule bit_word_eqI) (auto simp add: bit_simps) lemma ucast_ucast_mask_eq: "\ UCAST('a::len \ 'b::len) x = y; x AND mask LENGTH('b) = x \ \ x = ucast y" by (drule sym) (simp flip: take_bit_eq_mask add: unsigned_ucast_eq) lemma ucast_up_eq: "\ ucast x = (ucast y::'b::len word); LENGTH('a) \ LENGTH ('b) \ \ ucast x = (ucast y::'a::len word)" by word_eqI_solve lemma ucast_up_neq: "\ ucast x \ (ucast y::'b::len word); LENGTH('b) \ LENGTH ('a) \ \ ucast x \ (ucast y::'a::len word)" by (fastforce dest: ucast_up_eq) lemma mask_AND_less_0: "\ x AND mask n = 0; m \ n \ \ x AND mask m = 0" for x :: \'a::len word\ by (metis mask_twice2 word_and_notzeroD) lemma mask_len_id [simp]: "(x :: 'a :: len word) AND mask LENGTH('a) = x" using uint_lt2p [of x] by (simp add: mask_eq_iff) lemma scast_ucast_down_same: "LENGTH('b) \ LENGTH('a) \ SCAST('a \ 'b) = UCAST('a::len \ 'b::len)" by (simp add: down_cast_same is_down) lemma word_aligned_0_sum: "\ a + b = 0; is_aligned (a :: 'a :: len word) n; b \ mask n; n < LENGTH('a) \ \ a = 0 \ b = 0" by (simp add: word_plus_and_or_coroll aligned_mask_disjoint word_or_zero) lemma mask_eq1_nochoice: "\ LENGTH('a) > 1; (x :: 'a :: len word) AND 1 = x \ \ x = 0 \ x = 1" by (metis word_and_1) lemma shiftr_and_eq_shiftl: "(w >> n) AND x = y \ w AND (x << n) = (y << n)" for y :: "'a:: len word" - by (metis (no_types, lifting) and_not_mask bit.conj_ac(1) bit.conj_ac(2) mask_eq_0_eq_x shiftl_mask_is_0 shiftl_over_and_dist) + apply (drule sym) + apply simp + apply (rule bit_word_eqI) + apply (auto simp add: bit_simps) + done lemma add_mask_lower_bits': "\ len = LENGTH('a); is_aligned (x :: 'a :: len word) n; \n' \ n. n' < len \ \ bit p n' \ \ x + p AND NOT(mask n) = x" using add_mask_lower_bits by auto lemma leq_mask_shift: "(x :: 'a :: len word) \ mask (low_bits + high_bits) \ (x >> low_bits) \ mask high_bits" by (simp add: le_mask_iff shiftr_shiftr ac_simps) lemma ucast_ucast_eq_mask_shift: "(x :: 'a :: len word) \ mask (low_bits + LENGTH('b)) \ ucast((ucast (x >> low_bits)) :: 'b :: len word) = x >> low_bits" by (meson and_mask_eq_iff_le_mask eq_ucast_ucast_eq not_le_imp_less shiftr_less_t2n' ucast_ucast_len) lemma const_le_unat: "\ b < 2 ^ LENGTH('a); of_nat b \ a \ \ b \ unat (a :: 'a :: len word)" apply (simp add: word_le_def) apply (simp only: uint_nat zle_int) apply transfer apply (simp add: take_bit_nat_eq_self) done lemma upt_enum_offset_trivial: "\ x < 2 ^ LENGTH('a) - 1 ; n \ unat x \ \ ([(0 :: 'a :: len word) .e. x] ! n) = of_nat n" apply (induct x arbitrary: n) apply simp by (simp add: upto_enum_word_nth) lemma word_le_mask_out_plus_2sz: "x \ (x AND NOT(mask sz)) + 2 ^ sz - 1" for x :: \'a::len word\ by (metis add_diff_eq word_neg_and_le) lemma ucast_add: "ucast (a + (b :: 'a :: len word)) = ucast a + (ucast b :: ('a signed word))" by transfer (simp add: take_bit_add) lemma ucast_minus: "ucast (a - (b :: 'a :: len word)) = ucast a - (ucast b :: ('a signed word))" apply (insert ucast_add[where a=a and b="-b"]) apply (metis (no_types, opaque_lifting) add_diff_eq diff_add_cancel ucast_add) done lemma scast_ucast_add_one [simp]: "scast (ucast (x :: 'a::len word) + (1 :: 'a signed word)) = x + 1" apply (subst ucast_1[symmetric]) apply (subst ucast_add[symmetric]) apply clarsimp done lemma word_and_le_plus_one: "a > 0 \ (x :: 'a :: len word) AND (a - 1) < a" by (simp add: gt0_iff_gem1 word_and_less') lemma unat_of_ucast_then_shift_eq_unat_of_shift[simp]: "LENGTH('b) \ LENGTH('a) \ unat ((ucast (x :: 'a :: len word) :: 'b :: len word) >> n) = unat (x >> n)" by (simp add: shiftr_div_2n' unat_ucast_up_simp) lemma unat_of_ucast_then_mask_eq_unat_of_mask[simp]: "LENGTH('b) \ LENGTH('a) \ unat ((ucast (x :: 'a :: len word) :: 'b :: len word) AND mask m) = unat (x AND mask m)" by (metis ucast_and_mask unat_ucast_up_simp) lemma shiftr_less_t2n3: "\ (2 :: 'a word) ^ (n + m) = 0; m < LENGTH('a) \ \ (x :: 'a :: len word) >> n < 2 ^ m" by (fastforce intro: shiftr_less_t2n' simp: mask_eq_decr_exp power_overflow) lemma unat_shiftr_le_bound: "\ 2 ^ (LENGTH('a :: len) - n) - 1 \ bnd; 0 < n \ \ unat ((x :: 'a word) >> n) \ bnd" apply transfer apply (simp add: take_bit_drop_bit) apply (simp add: drop_bit_take_bit) apply (rule order_trans) defer apply assumption apply (simp add: nat_le_iff of_nat_diff) done lemma shiftr_eqD: "\ x >> n = y >> n; is_aligned x n; is_aligned y n \ \ x = y" by (metis is_aligned_shiftr_shiftl) lemma word_shiftr_shiftl_shiftr_eq_shiftr: "a \ b \ (x :: 'a :: len word) >> a << b >> b = x >> a" apply (rule bit_word_eqI) apply (auto simp add: bit_simps dest: bit_imp_le_length) done lemma of_int_uint_ucast: "of_int (uint (x :: 'a::len word)) = (ucast x :: 'b::len word)" by (fact Word.of_int_uint) lemma mod_mask_drop: "\ m = 2 ^ n; 0 < m; mask n AND msk = mask n \ \ (x mod m) AND msk = x mod m" for x :: \'a::len word\ by (simp add: word_mod_2p_is_mask word_bw_assocs) lemma mask_eq_ucast_eq: "\ x AND mask LENGTH('a) = (x :: ('c :: len word)); LENGTH('a) \ LENGTH('b)\ \ ucast (ucast x :: ('a :: len word)) = (ucast x :: ('b :: len word))" by (metis ucast_and_mask ucast_id ucast_ucast_mask ucast_up_eq) lemma of_nat_less_t2n: "of_nat i < (2 :: ('a :: len) word) ^ n \ n < LENGTH('a) \ unat (of_nat i :: 'a word) < 2 ^ n" by (metis order_less_trans p2_gt_0 unat_less_power word_neq_0_conv) lemma two_power_increasing_less_1: "\ n \ m; m \ LENGTH('a) \ \ (2 :: 'a :: len word) ^ n - 1 \ 2 ^ m - 1" by (metis diff_diff_cancel le_m1_iff_lt less_imp_diff_less p2_gt_0 two_power_increasing word_1_le_power word_le_minus_mono_left word_less_sub_1) lemma word_sub_mono4: "\ y + x \ z + x; y \ y + x; z \ z + x \ \ y \ z" for y :: "'a :: len word" by (simp add: word_add_le_iff2) lemma eq_or_less_helperD: "\ n = unat (2 ^ m - 1 :: 'a :: len word) \ n < unat (2 ^ m - 1 :: 'a word); m < LENGTH('a) \ \ n < 2 ^ m" by (meson le_less_trans nat_less_le unat_less_power word_power_less_1) lemma mask_sub: "n \ m \ mask m - mask n = mask m AND NOT(mask n :: 'a::len word)" by (metis (full_types) and_mask_eq_iff_shiftr_0 mask_out_sub_mask shiftr_mask_le word_bw_comms(1)) lemma neg_mask_diff_bound: "sz'\ sz \ (ptr AND NOT(mask sz')) - (ptr AND NOT(mask sz)) \ 2 ^ sz - 2 ^ sz'" (is "_ \ ?lhs \ ?rhs") for ptr :: \'a::len word\ proof - assume lt: "sz' \ sz" hence "?lhs = ptr AND (mask sz AND NOT(mask sz'))" by (metis add_diff_cancel_left' multiple_mask_trivia) also have "\ \ ?rhs" using lt by (metis (mono_tags) add_diff_eq diff_eq_eq eq_iff mask_2pm1 mask_sub word_and_le') finally show ?thesis by simp qed lemma mask_out_eq_0: "\ idx < 2 ^ sz; sz < LENGTH('a) \ \ (of_nat idx :: 'a :: len word) AND NOT(mask sz) = 0" by (simp add: of_nat_power less_mask_eq mask_eq_0_eq_x) lemma is_aligned_neg_mask_eq': "is_aligned ptr sz = (ptr AND NOT(mask sz) = ptr)" using is_aligned_mask mask_eq_0_eq_x by blast lemma neg_mask_mask_unat: "sz < LENGTH('a) \ unat ((ptr :: 'a :: len word) AND NOT(mask sz)) + unat (ptr AND mask sz) = unat ptr" by (metis AND_NOT_mask_plus_AND_mask_eq unat_plus_simple word_and_le2) lemma unat_pow_le_intro: "LENGTH('a) \ n \ unat (x :: 'a :: len word) < 2 ^ n" by (metis lt2p_lem not_le of_nat_le_iff of_nat_numeral semiring_1_class.of_nat_power uint_nat) lemma unat_shiftl_less_t2n: "\ unat (x :: 'a :: len word) < 2 ^ (m - n); m < LENGTH('a) \ \ unat (x << n) < 2 ^ m" by (metis More_Word.of_nat_power nat_mult_power_less_eq numeral_2_eq_2 of_nat_push_bit push_bit_eq_mult unat_less_power unat_of_nat_len unsigned_less word_eq_unatI zero_less_Suc) lemma unat_is_aligned_add: "\ is_aligned p n; unat d < 2 ^ n \ \ unat (p + d AND mask n) = unat d \ unat (p + d AND NOT(mask n)) = unat p" by (metis add.right_neutral and_mask_eq_iff_le_mask and_not_mask le_mask_iff mask_add_aligned mask_out_add_aligned mult_zero_right shiftl_t2n shiftr_le_0) lemma unat_shiftr_shiftl_mask_zero: "\ c + a \ LENGTH('a) + b ; c < LENGTH('a) \ \ unat (((q :: 'a :: len word) >> a << b) AND NOT(mask c)) = 0" by (fastforce intro: unat_is_aligned_add[where p=0 and n=c, simplified, THEN conjunct2] unat_shiftl_less_t2n unat_shiftr_less_t2n unat_pow_le_intro) lemmas of_nat_ucast = ucast_of_nat[symmetric] lemma shift_then_mask_eq_shift_low_bits: "x \ mask (low_bits + high_bits) \ (x >> low_bits) AND mask high_bits = x >> low_bits" for x :: \'a::len word\ by (simp add: leq_mask_shift le_mask_imp_and_mask) lemma leq_low_bits_iff_zero: "\ x \ mask (low bits + high bits); x >> low_bits = 0 \ \ (x AND mask low_bits = 0) = (x = 0)" for x :: \'a::len word\ using and_mask_eq_iff_shiftr_0 by force lemma unat_less_iff: "\ unat (a :: 'a :: len word) = b; c < 2 ^ LENGTH('a) \ \ (a < of_nat c) = (b < c)" using unat_ucast_less_no_overflow_simp by blast lemma is_aligned_no_overflow3: "\ is_aligned (a :: 'a :: len word) n; n < LENGTH('a); b < 2 ^ n; c \ 2 ^ n; b < c \ \ a + b \ a + (c - 1)" by (meson is_aligned_no_wrap' le_m1_iff_lt not_le word_less_sub_1 word_plus_mono_right) lemma mask_add_aligned_right: "is_aligned p n \ (q + p) AND mask n = q AND mask n" by (simp add: mask_add_aligned add.commute) lemma leq_high_bits_shiftr_low_bits_leq_bits_mask: "x \ mask high_bits \ (x :: 'a :: len word) << low_bits \ mask (low_bits + high_bits)" by (metis le_mask_shiftl_le_mask) lemma word_two_power_neg_ineq: "2 ^ m \ (0 :: 'a word) \ 2 ^ n \ - (2 ^ m :: 'a :: len word)" apply (cases "n < LENGTH('a)"; simp add: power_overflow) apply (cases "m < LENGTH('a)"; simp add: power_overflow) apply (simp add: word_le_nat_alt unat_minus word_size) apply (cases "LENGTH('a)"; simp) apply (simp add: less_Suc_eq_le) apply (drule power_increasing[where a=2 and n=n] power_increasing[where a=2 and n=m], simp)+ apply (drule(1) add_le_mono) apply simp done lemma unat_shiftl_absorb: "\ x \ 2 ^ p; p + k < LENGTH('a) \ \ unat (x :: 'a :: len word) * 2 ^ k = unat (x * 2 ^ k)" by (smt add_diff_cancel_right' add_lessD1 le_add2 le_less_trans mult.commute nat_le_power_trans unat_lt2p unat_mult_lem unat_power_lower word_le_nat_alt) lemma word_plus_mono_right_split: "\ unat ((x :: 'a :: len word) AND mask sz) + unat z < 2 ^ sz; sz < LENGTH('a) \ \ x \ x + z" apply (subgoal_tac "(x AND NOT(mask sz)) + (x AND mask sz) \ (x AND NOT(mask sz)) + ((x AND mask sz) + z)") apply (simp add:word_plus_and_or_coroll2 field_simps) apply (rule word_plus_mono_right) apply (simp add: less_le_trans no_olen_add_nat) using of_nat_power is_aligned_no_wrap' by force lemma mul_not_mask_eq_neg_shiftl: "NOT(mask n :: 'a::len word) = -1 << n" by (simp add: NOT_mask shiftl_t2n) lemma shiftr_mul_not_mask_eq_and_not_mask: "(x >> n) * NOT(mask n) = - (x AND NOT(mask n))" for x :: \'a::len word\ by (metis NOT_mask and_not_mask mult_minus_left semiring_normalization_rules(7) shiftl_t2n) lemma mask_eq_n1_shiftr: "n \ LENGTH('a) \ (mask n :: 'a :: len word) = -1 >> (LENGTH('a) - n)" by (metis diff_diff_cancel eq_refl mask_full shiftr_mask2) lemma is_aligned_mask_out_add_eq: "is_aligned p n \ (p + x) AND NOT(mask n) = p + (x AND NOT(mask n))" by (simp add: mask_out_sub_mask mask_add_aligned) lemmas is_aligned_mask_out_add_eq_sub = is_aligned_mask_out_add_eq[where x="a - b" for a b, simplified field_simps] lemma aligned_bump_down: "is_aligned x n \ (x - 1) AND NOT(mask n) = x - 2 ^ n" by (drule is_aligned_mask_out_add_eq[where x="-1"]) (simp add: NOT_mask) lemma unat_2tp_if: "unat (2 ^ n :: ('a :: len) word) = (if n < LENGTH ('a) then 2 ^ n else 0)" by (split if_split, simp_all add: power_overflow) lemma mask_of_mask: "mask (n::nat) AND mask (m::nat) = (mask (min m n) :: 'a::len word)" by word_eqI_solve lemma unat_signed_ucast_less_ucast: "LENGTH('a) \ LENGTH('b) \ unat (ucast (x :: 'a :: len word) :: 'b :: len signed word) = unat x" by (simp add: unat_ucast_up_simp) lemma toEnum_of_ucast: "LENGTH('b) \ LENGTH('a) \ (toEnum (unat (b::'b :: len word))::'a :: len word) = of_nat (unat b)" by (simp add: unat_pow_le_intro) lemmas unat_ucast_mask = unat_ucast_eq_unat_and_mask[where w=a for a] lemma t2n_mask_eq_if: "2 ^ n AND mask m = (if n < m then 2 ^ n else (0 :: 'a::len word))" by (rule word_eqI) (auto simp add: bit_simps) lemma unat_ucast_le: "unat (ucast (x :: 'a :: len word) :: 'b :: len word) \ unat x" by (simp add: ucast_nat_def word_unat_less_le) lemma ucast_le_up_down_iff: "\ LENGTH('a) \ LENGTH('b); (x :: 'b :: len word) \ ucast (- 1 :: 'a :: len word) \ \ (ucast x \ (y :: 'a word)) = (x \ ucast y)" using le_max_word_ucast_id ucast_le_ucast by metis lemma ucast_ucast_mask_shift: "a \ LENGTH('a) + b \ ucast (ucast (p AND mask a >> b) :: 'a :: len word) = p AND mask a >> b" by (metis add.commute le_mask_iff shiftr_mask_le ucast_ucast_eq_mask_shift word_and_le') lemma unat_ucast_mask_shift: "a \ LENGTH('a) + b \ unat (ucast (p AND mask a >> b) :: 'a :: len word) = unat (p AND mask a >> b)" by (metis linear ucast_ucast_mask_shift unat_ucast_up_simp) lemma mask_overlap_zero: "a \ b \ (p AND mask a) AND NOT(mask b) = 0" for p :: \'a::len word\ by (metis NOT_mask_AND_mask mask_lower_twice2 max_def) lemma mask_shifl_overlap_zero: "a + c \ b \ (p AND mask a << c) AND NOT(mask b) = 0" for p :: \'a::len word\ by (metis and_mask_0_iff_le_mask mask_mono mask_shiftl_decompose order_trans shiftl_over_and_dist word_and_le' word_and_le2) lemma mask_overlap_zero': "a \ b \ (p AND NOT(mask a)) AND mask b = 0" for p :: \'a::len word\ using mask_AND_NOT_mask mask_AND_less_0 by blast lemma mask_rshift_mult_eq_rshift_lshift: "((a :: 'a :: len word) >> b) * (1 << c) = (a >> b << c)" by (simp add: shiftl_t2n) lemma shift_alignment: "a \ b \ is_aligned (p >> a << a) b" using is_aligned_shift is_aligned_weaken by blast lemma mask_split_sum_twice: "a \ b \ (p AND NOT(mask a)) + ((p AND mask a) AND NOT(mask b)) + (p AND mask b) = p" for p :: \'a::len word\ by (simp add: add.commute multiple_mask_trivia word_bw_comms(1) word_bw_lcs(1) word_plus_and_or_coroll2) lemma mask_shift_eq_mask_mask: "(p AND mask a >> b << b) = (p AND mask a) AND NOT(mask b)" for p :: \'a::len word\ by (simp add: and_not_mask) lemma mask_shift_sum: "\ a \ b; unat n = unat (p AND mask b) \ \ (p AND NOT(mask a)) + (p AND mask a >> b) * (1 << b) + n = (p :: 'a :: len word)" apply (simp add: push_bit_of_1 flip: push_bit_eq_mult) apply (subst disjunctive_add) apply (auto simp add: bit_simps) apply (smt (z3) AND_NOT_mask_plus_AND_mask_eq and.comm_neutral and.right_idem and_not_mask bit.conj_disj_distrib bit.disj_cancel_right mask_out_first_mask_some word_bw_assocs(1) word_bw_comms(1) word_bw_comms(2) word_eq_unatI) done lemma is_up_compose: "\ is_up uc; is_up uc' \ \ is_up (uc' \ uc)" unfolding is_up_def by (simp add: Word.target_size Word.source_size) lemma of_int_sint_scast: "of_int (sint (x :: 'a :: len word)) = (scast x :: 'b :: len word)" by (fact Word.of_int_sint) lemma scast_of_nat_to_signed [simp]: "scast (of_nat x :: 'a :: len word) = (of_nat x :: 'a signed word)" by (rule bit_word_eqI) (simp add: bit_simps) lemma scast_of_nat_signed_to_unsigned_add: "scast (of_nat x + of_nat y :: 'a :: len signed word) = (of_nat x + of_nat y :: 'a :: len word)" by (metis of_nat_add scast_of_nat) lemma scast_of_nat_unsigned_to_signed_add: "(scast (of_nat x + of_nat y :: 'a :: len word)) = (of_nat x + of_nat y :: 'a :: len signed word)" by (metis Abs_fnat_hom_add scast_of_nat_to_signed) lemma and_mask_cases: fixes x :: "'a :: len word" assumes len: "n < LENGTH('a)" shows "x AND mask n \ of_nat ` set [0 ..< 2 ^ n]" apply (simp flip: take_bit_eq_mask) apply (rule image_eqI [of _ _ \unat (take_bit n x)\]) using len apply simp_all apply transfer apply simp done lemma sint_eq_uint_2pl: "\ (a :: 'a :: len word) < 2 ^ (LENGTH('a) - 1) \ \ sint a = uint a" by (simp add: not_msb_from_less sint_eq_uint word_2p_lem word_size) lemma pow_sub_less: "\ a + b \ LENGTH('a); unat (x :: 'a :: len word) = 2 ^ a \ \ unat (x * 2 ^ b - 1) < 2 ^ (a + b)" by (smt (z3) eq_or_less_helperD le_add2 le_eq_less_or_eq le_trans power_add unat_mult_lem unat_pow_le_intro unat_power_lower word_eq_unatI) lemma sle_le_2pl: "\ (b :: 'a :: len word) < 2 ^ (LENGTH('a) - 1); a \ b \ \ a <=s b" by (simp add: not_msb_from_less word_sle_msb_le) lemma sless_less_2pl: "\ (b :: 'a :: len word) < 2 ^ (LENGTH('a) - 1); a < b \ \ a > n = w AND mask (size w - n)" for w :: \'a::len word\ by (cases "n \ size w"; clarsimp simp: word_and_le2 and_mask shiftl_zero_size) lemma aligned_sub_aligned_simple: "\ is_aligned a n; is_aligned b n \ \ is_aligned (a - b) n" by (simp add: aligned_sub_aligned) lemma minus_one_shift: "- (1 << n) = (-1 << n :: 'a::len word)" by (simp flip: mul_not_mask_eq_neg_shiftl minus_exp_eq_not_mask add: push_bit_of_1) lemma ucast_eq_mask: "(UCAST('a::len \ 'b::len) x = UCAST('a \ 'b) y) = (x AND mask LENGTH('b) = y AND mask LENGTH('b))" by transfer (simp flip: take_bit_eq_mask add: ac_simps) context fixes w :: "'a::len word" begin private lemma sbintrunc_uint_ucast: assumes "Suc n = LENGTH('b::len)" shows "signed_take_bit n (uint (ucast w :: 'b word)) = signed_take_bit n (uint w)" by (rule bit_eqI) (use assms in \simp add: bit_simps\) private lemma test_bit_sbintrunc: assumes "i < LENGTH('a)" shows "bit (word_of_int (signed_take_bit n (uint w)) :: 'a word) i = (if n < i then bit w n else bit w i)" using assms by (simp add: bit_simps) private lemma test_bit_sbintrunc_ucast: assumes len_a: "i < LENGTH('a)" shows "bit (word_of_int (signed_take_bit (LENGTH('b) - 1) (uint (ucast w :: 'b word))) :: 'a word) i = (if LENGTH('b::len) \ i then bit w (LENGTH('b) - 1) else bit w i)" using len_a by (auto simp add: sbintrunc_uint_ucast bit_simps) lemma scast_ucast_high_bits: \scast (ucast w :: 'b::len word) = w \ (\ i \ {LENGTH('b) ..< size w}. bit w i = bit w (LENGTH('b) - 1))\ proof (cases \LENGTH('a) \ LENGTH('b)\) case True moreover define m where \m = LENGTH('b) - LENGTH('a)\ ultimately have \LENGTH('b) = m + LENGTH('a)\ by simp then show ?thesis apply (simp_all add: signed_ucast_eq word_size) apply (rule bit_word_eqI) apply (simp add: bit_signed_take_bit_iff) done next case False define q where \q = LENGTH('b) - 1\ then have \LENGTH('b) = Suc q\ by simp moreover define m where \m = Suc LENGTH('a) - LENGTH('b)\ with False \LENGTH('b) = Suc q\ have \LENGTH('a) = m + q\ by (simp add: not_le) ultimately show ?thesis apply (simp_all add: signed_ucast_eq word_size) apply (transfer fixing: m q) apply (simp add: signed_take_bit_take_bit) apply rule apply (subst bit_eq_iff) apply (simp add: bit_take_bit_iff bit_signed_take_bit_iff min_def) apply (auto simp add: Suc_le_eq) using less_imp_le_nat apply blast using less_imp_le_nat apply blast done qed lemma scast_ucast_mask_compare: "scast (ucast w :: 'b::len word) = w \ (w \ mask (LENGTH('b) - 1) \ NOT(mask (LENGTH('b) - 1)) \ w)" apply (clarsimp simp: le_mask_high_bits neg_mask_le_high_bits scast_ucast_high_bits word_size) apply (rule iffI; clarsimp) apply (rename_tac i j; case_tac "i = LENGTH('b) - 1"; case_tac "j = LENGTH('b) - 1") by auto lemma ucast_less_shiftl_helper': "\ LENGTH('b) + (a::nat) < LENGTH('a); 2 ^ (LENGTH('b) + a) \ n\ \ (ucast (x :: 'b::len word) << a) < (n :: 'a::len word)" apply (erule order_less_le_trans[rotated]) using ucast_less[where x=x and 'a='a] apply (simp only: shiftl_t2n field_simps) apply (rule word_less_power_trans2; simp) done end lemma ucast_ucast_mask2: "is_down (UCAST ('a \ 'b)) \ UCAST ('b::len \ 'c::len) (UCAST ('a::len \ 'b::len) x) = UCAST ('a \ 'c) (x AND mask LENGTH('b))" apply (simp flip: take_bit_eq_mask) apply transfer apply simp done lemma ucast_NOT: "ucast (NOT x) = NOT(ucast x) AND mask (LENGTH('a))" for x::"'a::len word" by word_eqI lemma ucast_NOT_down: "is_down UCAST('a::len \ 'b::len) \ UCAST('a \ 'b) (NOT x) = NOT(UCAST('a \ 'b) x)" by word_eqI lemma upto_enum_step_shift: "\ is_aligned p n \ \ ([p , p + 2 ^ m .e. p + 2 ^ n - 1]) = map ((+) p) [0, 2 ^ m .e. 2 ^ n - 1]" apply (erule is_aligned_get_word_bits) prefer 2 apply (simp add: map_idI) apply (clarsimp simp: upto_enum_step_def) apply (frule is_aligned_no_overflow) apply (simp add: linorder_not_le [symmetric]) done lemma upto_enum_step_shift_red: "\ is_aligned p sz; sz < LENGTH('a); us \ sz \ \ [p :: 'a :: len word, p + 2 ^ us .e. p + 2 ^ sz - 1] = map (\x. p + of_nat x * 2 ^ us) [0 ..< 2 ^ (sz - us)]" apply (subst upto_enum_step_shift, assumption) apply (simp add: upto_enum_step_red) done lemma upto_enum_step_subset: "set [x, y .e. z] \ {x .. z}" apply (clarsimp simp: upto_enum_step_def linorder_not_less) apply (drule div_to_mult_word_lt) apply (rule conjI) apply (erule word_random[rotated]) apply simp apply (rule order_trans) apply (erule word_plus_mono_right) apply simp apply simp done lemma ucast_distrib: fixes M :: "'a::len word \ 'a::len word \ 'a::len word" fixes M' :: "'b::len word \ 'b::len word \ 'b::len word" fixes L :: "int \ int \ int" assumes lift_M: "\x y. uint (M x y) = L (uint x) (uint y) mod 2 ^ LENGTH('a)" assumes lift_M': "\x y. uint (M' x y) = L (uint x) (uint y) mod 2 ^ LENGTH('b)" assumes distrib: "\x y. (L (x mod (2 ^ LENGTH('b))) (y mod (2 ^ LENGTH('b)))) mod (2 ^ LENGTH('b)) = (L x y) mod (2 ^ LENGTH('b))" assumes is_down: "is_down (ucast :: 'a word \ 'b word)" shows "ucast (M a b) = M' (ucast a) (ucast b)" apply (simp only: ucast_eq) apply (subst lift_M) apply (subst of_int_uint [symmetric], subst lift_M') apply (metis local.distrib local.is_down take_bit_eq_mod ucast_down_wi uint_word_of_int_eq word_of_int_uint) done lemma ucast_down_add: "is_down (ucast:: 'a word \ 'b word) \ ucast ((a :: 'a::len word) + b) = (ucast a + ucast b :: 'b::len word)" by (rule ucast_distrib [where L="(+)"], (clarsimp simp: uint_word_ariths)+, presburger, simp) lemma ucast_down_minus: "is_down (ucast:: 'a word \ 'b word) \ ucast ((a :: 'a::len word) - b) = (ucast a - ucast b :: 'b::len word)" apply (rule ucast_distrib [where L="(-)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_diff_left_eq mod_diff_right_eq) apply simp done lemma ucast_down_mult: "is_down (ucast:: 'a word \ 'b word) \ ucast ((a :: 'a::len word) * b) = (ucast a * ucast b :: 'b::len word)" apply (rule ucast_distrib [where L="(*)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_mult_eq) apply simp done lemma scast_distrib: fixes M :: "'a::len word \ 'a::len word \ 'a::len word" fixes M' :: "'b::len word \ 'b::len word \ 'b::len word" fixes L :: "int \ int \ int" assumes lift_M: "\x y. uint (M x y) = L (uint x) (uint y) mod 2 ^ LENGTH('a)" assumes lift_M': "\x y. uint (M' x y) = L (uint x) (uint y) mod 2 ^ LENGTH('b)" assumes distrib: "\x y. (L (x mod (2 ^ LENGTH('b))) (y mod (2 ^ LENGTH('b)))) mod (2 ^ LENGTH('b)) = (L x y) mod (2 ^ LENGTH('b))" assumes is_down: "is_down (scast :: 'a word \ 'b word)" shows "scast (M a b) = M' (scast a) (scast b)" apply (subst (1 2 3) down_cast_same [symmetric]) apply (insert is_down) apply (clarsimp simp: is_down_def target_size source_size is_down) apply (rule ucast_distrib [where L=L, OF lift_M lift_M' distrib]) apply (insert is_down) apply (clarsimp simp: is_down_def target_size source_size is_down) done lemma scast_down_add: "is_down (scast:: 'a word \ 'b word) \ scast ((a :: 'a::len word) + b) = (scast a + scast b :: 'b::len word)" by (rule scast_distrib [where L="(+)"], (clarsimp simp: uint_word_ariths)+, presburger, simp) lemma scast_down_minus: "is_down (scast:: 'a word \ 'b word) \ scast ((a :: 'a::len word) - b) = (scast a - scast b :: 'b::len word)" apply (rule scast_distrib [where L="(-)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_diff_left_eq mod_diff_right_eq) apply simp done lemma scast_down_mult: "is_down (scast:: 'a word \ 'b word) \ scast ((a :: 'a::len word) * b) = (scast a * scast b :: 'b::len word)" apply (rule scast_distrib [where L="(*)"], (clarsimp simp: uint_word_ariths)+) apply (metis mod_mult_eq) apply simp done lemma scast_ucast_1: "\ is_down (ucast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (scast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_eq ucast_down_wi) lemma scast_ucast_3: "\ is_down (ucast :: 'a word \ 'c word); is_down (ucast :: 'b word \ 'c word) \ \ (scast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_eq ucast_down_wi) lemma scast_ucast_4: "\ is_up (ucast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (scast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_eq ucast_down_wi) lemma scast_scast_b: "\ is_up (scast :: 'a word \ 'b word) \ \ (scast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis scast_eq sint_up_scast) lemma ucast_scast_1: "\ is_down (scast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (ucast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis scast_eq ucast_down_wi) lemma ucast_scast_3: "\ is_down (scast :: 'a word \ 'c word); is_down (ucast :: 'b word \ 'c word) \ \ (ucast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis scast_eq ucast_down_wi) lemma ucast_scast_4: "\ is_up (scast :: 'a word \ 'b word); is_down (ucast :: 'b word \ 'c word) \ \ (ucast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" by (metis down_cast_same scast_eq sint_up_scast) lemma ucast_ucast_a: "\ is_down (ucast :: 'b word \ 'c word) \ \ (ucast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis down_cast_same ucast_eq ucast_down_wi) lemma ucast_ucast_b: "\ is_up (ucast :: 'a word \ 'b word) \ \ (ucast (ucast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = ucast a" by (metis ucast_up_ucast) lemma scast_scast_a: "\ is_down (scast :: 'b word \ 'c word) \ \ (scast (scast (a :: 'a::len word) :: 'b::len word) :: 'c::len word) = scast a" apply (simp only: scast_eq) apply (metis down_cast_same is_up_down scast_eq ucast_down_wi) done lemma scast_down_wi [OF refl]: "uc = scast \ is_down uc \ uc (word_of_int x) = word_of_int x" by (metis down_cast_same is_up_down ucast_down_wi) lemmas cast_simps = is_down is_up scast_down_add scast_down_minus scast_down_mult ucast_down_add ucast_down_minus ucast_down_mult scast_ucast_1 scast_ucast_3 scast_ucast_4 ucast_scast_1 ucast_scast_3 ucast_scast_4 ucast_ucast_a ucast_ucast_b scast_scast_a scast_scast_b ucast_down_wi scast_down_wi ucast_of_nat scast_of_nat uint_up_ucast sint_up_scast up_scast_surj up_ucast_surj lemma sdiv_word_max: "(sint (a :: ('a::len) word) sdiv sint (b :: ('a::len) word) < (2 ^ (size a - 1))) = ((a \ - (2 ^ (size a - 1)) \ (b \ -1)))" (is "?lhs = (\ ?a_int_min \ \ ?b_minus1)") proof (rule classical) assume not_thesis: "\ ?thesis" have not_zero: "b \ 0" using not_thesis by (clarsimp) let ?range = \{- (2 ^ (size a - 1))..<2 ^ (size a - 1)} :: int set\ have result_range: "sint a sdiv sint b \ ?range \ {2 ^ (size a - 1)}" using sdiv_word_min [of a b] sdiv_word_max [of a b] by auto have result_range_overflow: "(sint a sdiv sint b = 2 ^ (size a - 1)) = (?a_int_min \ ?b_minus1)" apply (rule iffI [rotated]) apply (clarsimp simp: signed_divide_int_def sgn_if word_size sint_int_min) apply (rule classical) apply (case_tac "?a_int_min") apply (clarsimp simp: word_size sint_int_min) apply (metis diff_0_right int_sdiv_negated_is_minus1 minus_diff_eq minus_int_code(2) power_eq_0_iff sint_minus1 zero_neq_numeral) apply (subgoal_tac "abs (sint a) < 2 ^ (size a - 1)") apply (insert sdiv_int_range [where a="sint a" and b="sint b"])[1] apply (clarsimp simp: word_size) apply (insert sdiv_int_range [where a="sint a" and b="sint b"])[1] apply auto apply (cases \size a\) apply simp_all apply (smt (z3) One_nat_def diff_Suc_1 signed_word_eqI sint_int_min sint_range_size wsst_TYs(3)) done have result_range_simple: "(sint a sdiv sint b \ ?range) \ ?thesis" apply (insert sdiv_int_range [where a="sint a" and b="sint b"]) apply (clarsimp simp: word_size sint_int_min) done show ?thesis apply (rule UnE [OF result_range result_range_simple]) apply simp apply (clarsimp simp: word_size) using result_range_overflow apply (clarsimp simp: word_size) done qed lemmas sdiv_word_min' = sdiv_word_min [simplified word_size, simplified] lemmas sdiv_word_max' = sdiv_word_max [simplified word_size, simplified] lemma signed_arith_ineq_checks_to_eq: "((- (2 ^ (size a - 1)) \ (sint a + sint b)) \ (sint a + sint b \ (2 ^ (size a - 1) - 1))) = (sint a + sint b = sint (a + b ))" "((- (2 ^ (size a - 1)) \ (sint a - sint b)) \ (sint a - sint b \ (2 ^ (size a - 1) - 1))) = (sint a - sint b = sint (a - b))" "((- (2 ^ (size a - 1)) \ (- sint a)) \ (- sint a) \ (2 ^ (size a - 1) - 1)) = ((- sint a) = sint (- a))" "((- (2 ^ (size a - 1)) \ (sint a * sint b)) \ (sint a * sint b \ (2 ^ (size a - 1) - 1))) = (sint a * sint b = sint (a * b))" "((- (2 ^ (size a - 1)) \ (sint a sdiv sint b)) \ (sint a sdiv sint b \ (2 ^ (size a - 1) - 1))) = (sint a sdiv sint b = sint (a sdiv b))" "((- (2 ^ (size a - 1)) \ (sint a smod sint b)) \ (sint a smod sint b \ (2 ^ (size a - 1) - 1))) = (sint a smod sint b = sint (a smod b))" by (auto simp: sint_word_ariths word_size signed_div_arith signed_mod_arith signed_take_bit_int_eq_self_iff intro: sym dest: sym) lemma signed_arith_sint: "((- (2 ^ (size a - 1)) \ (sint a + sint b)) \ (sint a + sint b \ (2 ^ (size a - 1) - 1))) \ sint (a + b) = (sint a + sint b)" "((- (2 ^ (size a - 1)) \ (sint a - sint b)) \ (sint a - sint b \ (2 ^ (size a - 1) - 1))) \ sint (a - b) = (sint a - sint b)" "((- (2 ^ (size a - 1)) \ (- sint a)) \ (- sint a) \ (2 ^ (size a - 1) - 1)) \ sint (- a) = (- sint a)" "((- (2 ^ (size a - 1)) \ (sint a * sint b)) \ (sint a * sint b \ (2 ^ (size a - 1) - 1))) \ sint (a * b) = (sint a * sint b)" "((- (2 ^ (size a - 1)) \ (sint a sdiv sint b)) \ (sint a sdiv sint b \ (2 ^ (size a - 1) - 1))) \ sint (a sdiv b) = (sint a sdiv sint b)" "((- (2 ^ (size a - 1)) \ (sint a smod sint b)) \ (sint a smod sint b \ (2 ^ (size a - 1) - 1))) \ sint (a smod b) = (sint a smod sint b)" by (subst (asm) signed_arith_ineq_checks_to_eq; simp)+ end end