diff --git a/thys/Stone_Kleene_Relation_Algebras/Kleene_Algebras.thy b/thys/Stone_Kleene_Relation_Algebras/Kleene_Algebras.thy --- a/thys/Stone_Kleene_Relation_Algebras/Kleene_Algebras.thy +++ b/thys/Stone_Kleene_Relation_Algebras/Kleene_Algebras.thy @@ -1,716 +1,729 @@ (* Title: Kleene Algebras Author: Walter Guttmann Maintainer: Walter Guttmann *) section \Kleene Algebras\ text \ Kleene algebras have been axiomatised by Kozen to describe the equational theory of regular languages \cite{Kozen1994}. Binary relations are another important model. This theory implements variants of Kleene algebras based on idempotent left semirings \cite{Moeller2007}. The weakening of some semiring axioms allows the treatment of further computation models. The presented algebras are special cases of iterings, so many results can be inherited. \ theory Kleene_Algebras imports Iterings begin text \ We start with left Kleene algebras, which use the left unfold and left induction axioms of Kleene algebras. \ class star = fixes star :: "'a \ 'a" ("_\<^sup>\" [100] 100) class left_kleene_algebra = idempotent_left_semiring + star + assumes star_left_unfold : "1 \ y * y\<^sup>\ \ y\<^sup>\" assumes star_left_induct : "z \ y * x \ x \ y\<^sup>\ * z \ x" begin no_notation trancl ("(_\<^sup>+)" [1000] 999) abbreviation tc ("_\<^sup>+" [100] 100) where "tc x \ x * x\<^sup>\" lemma star_left_unfold_equal: "1 \ x * x\<^sup>\ = x\<^sup>\" by (metis sup_right_isotone antisym mult_right_isotone mult_1_right star_left_induct star_left_unfold) text \ This means that for some properties of Kleene algebras, only one inequality can be derived, as exemplified by the following sliding rule. \ lemma star_left_slide: "(x * y)\<^sup>\ * x \ x * (y * x)\<^sup>\" by (metis mult_assoc mult_left_sub_dist_sup mult_1_right star_left_induct star_left_unfold_equal) lemma star_isotone: "x \ y \ x\<^sup>\ \ y\<^sup>\" by (metis sup_right_isotone mult_left_isotone order_trans star_left_unfold mult_1_right star_left_induct) lemma star_sup_1: "(x \ y)\<^sup>\ = x\<^sup>\ * (y * x\<^sup>\)\<^sup>\" proof (rule antisym) have "y * x\<^sup>\ * (y * x\<^sup>\)\<^sup>\ \ (y * x\<^sup>\)\<^sup>\" using sup_right_divisibility star_left_unfold_equal by auto also have "... \ x\<^sup>\ * (y * x\<^sup>\)\<^sup>\" using mult_left_isotone sup_left_divisibility star_left_unfold_equal by fastforce finally have "(x \ y) * (x\<^sup>\ * (y * x\<^sup>\)\<^sup>\) \ x\<^sup>\ * (y * x\<^sup>\)\<^sup>\" by (metis le_supI mult_right_dist_sup mult_right_sub_dist_sup_right mult_assoc star_left_unfold_equal) hence "1 \ (x \ y) * (x\<^sup>\ * (y * x\<^sup>\)\<^sup>\) \ x\<^sup>\ * (y * x\<^sup>\)\<^sup>\" using reflexive_mult_closed star_left_unfold by auto thus "(x \ y)\<^sup>\ \ x\<^sup>\ * (y * x\<^sup>\)\<^sup>\" using star_left_induct by force next have "x\<^sup>\ * (y * x\<^sup>\)\<^sup>\ \ x\<^sup>\ * (y * (x \ y)\<^sup>\)\<^sup>\" by (simp add: mult_right_isotone star_isotone) also have "... \ x\<^sup>\ * ((x \ y) * (x \ y)\<^sup>\)\<^sup>\" by (simp add: mult_right_isotone mult_right_sub_dist_sup_right star_isotone) also have "... \ x\<^sup>\ * (x \ y)\<^sup>\\<^sup>\" using mult_right_isotone star_left_unfold star_isotone by auto also have "... \ (x \ y)\<^sup>\ * (x \ y)\<^sup>\\<^sup>\" by (simp add: mult_left_isotone star_isotone) also have "... \ (x \ y)\<^sup>\" by (metis sup.bounded_iff mult_1_right star_left_induct star_left_unfold) finally show "x\<^sup>\ * (y * x\<^sup>\)\<^sup>\ \ (x \ y)\<^sup>\" by simp qed end text \ We now show that left Kleene algebras form iterings. A sublocale is used instead of a subclass, because iterings use a different iteration operation. \ sublocale left_kleene_algebra < star: left_conway_semiring where circ = star apply unfold_locales apply (rule star_left_unfold_equal) apply (rule star_left_slide) by (rule star_sup_1) context left_kleene_algebra begin text \ A number of lemmas in this class are taken from Georg Struth's Kleene algebra theory \cite{ArmstrongGomesStruthWeber2016}. \ lemma star_sub_one: "x \ 1 \ x\<^sup>\ = 1" by (metis sup_right_isotone eq_iff le_iff_sup mult_1_right star.circ_plus_one star_left_induct) lemma star_one: "1\<^sup>\ = 1" by (simp add: star_sub_one) lemma star_left_induct_mult: "x * y \ y \ x\<^sup>\ * y \ y" by (simp add: star_left_induct) lemma star_left_induct_mult_iff: "x * y \ y \ x\<^sup>\ * y \ y" using mult_left_isotone order_trans star.circ_increasing star_left_induct_mult by blast lemma star_involutive: "x\<^sup>\ = x\<^sup>\\<^sup>\" using star.circ_circ_sup star_sup_1 star_one by auto lemma star_sup_one: "(1 \ x)\<^sup>\ = x\<^sup>\" using star.circ_circ_sup star_involutive by auto lemma star_left_induct_equal: "z \ x * y = y \ x\<^sup>\ * z \ y" by (simp add: star_left_induct) lemma star_left_induct_mult_equal: "x * y = y \ x\<^sup>\ * y \ y" by (simp add: star_left_induct_mult) lemma star_star_upper_bound: "x\<^sup>\ \ z\<^sup>\ \ x\<^sup>\\<^sup>\ \ z\<^sup>\" using star_involutive by auto lemma star_simulation_left: assumes "x * z \ z * y" shows "x\<^sup>\ * z \ z * y\<^sup>\" proof - have "x * z * y\<^sup>\ \ z * y * y\<^sup>\" by (simp add: assms mult_left_isotone) also have "... \ z * y\<^sup>\" by (simp add: mult_right_isotone star.left_plus_below_circ mult_assoc) finally have "z \ x * z * y\<^sup>\ \ z * y\<^sup>\" using star.circ_back_loop_prefixpoint by auto thus ?thesis by (simp add: star_left_induct mult_assoc) qed lemma quasicomm_1: "y * x \ x * (x \ y)\<^sup>\ \ y\<^sup>\ * x \ x * (x \ y)\<^sup>\" by (metis mult_isotone order_refl order_trans star.circ_increasing star_involutive star_simulation_left) lemma star_rtc_3: "1 \ x \ y * y = y \ x\<^sup>\ \ y" by (metis sup.bounded_iff le_iff_sup mult_left_sub_dist_sup_left mult_1_right star_left_induct_mult_iff star.circ_sub_dist) lemma star_decompose_1: "(x \ y)\<^sup>\ = (x\<^sup>\ * y\<^sup>\)\<^sup>\" apply (rule antisym) apply (simp add: star.circ_sup_2) using star.circ_sub_dist_3 star_isotone star_involutive by fastforce lemma star_sum: "(x \ y)\<^sup>\ = (x\<^sup>\ \ y\<^sup>\)\<^sup>\" using star_decompose_1 star_involutive by auto lemma star_decompose_3: "(x\<^sup>\ * y\<^sup>\)\<^sup>\ = x\<^sup>\ * (y * x\<^sup>\)\<^sup>\" using star_sup_1 star_decompose_1 by auto text \ In contrast to iterings, we now obtain that the iteration operation results in least fixpoints. \ lemma star_loop_least_fixpoint: "y * x \ z = x \ y\<^sup>\ * z \ x" by (simp add: sup_commute star_left_induct_equal) lemma star_loop_is_least_fixpoint: "is_least_fixpoint (\x . y * x \ z) (y\<^sup>\ * z)" by (simp add: is_least_fixpoint_def star.circ_loop_fixpoint star_loop_least_fixpoint) lemma star_loop_mu: "\ (\x . y * x \ z) = y\<^sup>\ * z" by (metis least_fixpoint_same star_loop_is_least_fixpoint) lemma affine_has_least_fixpoint: "has_least_fixpoint (\x . y * x \ z)" by (metis has_least_fixpoint_def star_loop_is_least_fixpoint) lemma star_outer_increasing: "x \ y\<^sup>\ * x * y\<^sup>\" by (metis star.circ_back_loop_prefixpoint star.circ_loop_fixpoint sup.boundedE) (* lemma circ_sup: "(x\<^sup>\ * y)\<^sup>\ * x\<^sup>\ = (x \ y)\<^sup>\" nitpick [expect=genuine,card=7] oops lemma circ_mult: "1 \ x * (y * x)\<^sup>\ * y = (x * y)\<^sup>\" nitpick [expect=genuine] oops lemma circ_plus_same: "x\<^sup>\ * x = x * x\<^sup>\" nitpick [expect=genuine] oops lemma circ_unfold_sum: "(x \ y)\<^sup>\ = x\<^sup>\ \ x\<^sup>\ * y * (x \ y)\<^sup>\" nitpick [expect=genuine,card=7] oops lemma mult_zero_sup_circ_2: "(x \ y * bot)\<^sup>\ = x\<^sup>\ \ x\<^sup>\ * y * bot" nitpick [expect=genuine,card=7] oops lemma circ_simulate_left: "x * z \ z * y \ w \ x\<^sup>\ * z \ (z \ x\<^sup>\ * w) * y\<^sup>\" nitpick [expect=genuine] oops lemma circ_simulate_1: "y * x \ x * y \ y\<^sup>\ * x\<^sup>\ \ x\<^sup>\ * y\<^sup>\" nitpick [expect=genuine,card=7] oops lemma circ_separate_1: "y * x \ x * y \ (x \ y)\<^sup>\ = x\<^sup>\ * y\<^sup>\" nitpick [expect=genuine,card=7] oops lemma atomicity_refinement: "s = s * q \ x = q * x \ q * b = bot \ r * b \ b * r \ r * l \ l * r \ x * l \ l * x \ b * l \ l * b \ q * l \ l * q \ r\<^sup>\ * q \ q * r\<^sup>\ \ q \ 1 \ s * (x \ b \ r \ l)\<^sup>\ * q \ s * (x * b\<^sup>\ * q \ r \ l)\<^sup>\" nitpick [expect=genuine] oops lemma circ_simulate_left_plus: "x * z \ z * y\<^sup>\ \ w \ x\<^sup>\ * z \ (z \ x\<^sup>\ * w) * y\<^sup>\" nitpick [expect=genuine] oops lemma circ_separate_unfold: "(y * x\<^sup>\)\<^sup>\ = y\<^sup>\ \ y\<^sup>\ * y * x * x\<^sup>\ * (y * x\<^sup>\)\<^sup>\" nitpick [expect=genuine] oops lemma separation: "y * x \ x * y\<^sup>\ \ (x \ y)\<^sup>\ = x\<^sup>\ * y\<^sup>\" nitpick [expect=genuine,card=7] oops lemma circ_simulate_4: "y * x \ x * x\<^sup>\ * (1 \ y) \ y\<^sup>\ * x\<^sup>\ \ x\<^sup>\ * y\<^sup>\" nitpick [expect=genuine,card=7] oops lemma circ_simulate_5: "y * x \ x * x\<^sup>\ * (x \ y) \ y\<^sup>\ * x\<^sup>\ \ x\<^sup>\ * y\<^sup>\" nitpick [expect=genuine,card=7] oops lemma circ_simulate_6: "y * x \ x * (x \ y) \ y\<^sup>\ * x\<^sup>\ \ x\<^sup>\ * y\<^sup>\" nitpick [expect=genuine,card=7] oops lemma circ_separate_4: "y * x \ x * x\<^sup>\ * (1 \ y) \ (x \ y)\<^sup>\ = x\<^sup>\ * y\<^sup>\" nitpick [expect=genuine,card=7] oops lemma circ_separate_5: "y * x \ x * x\<^sup>\ * (x \ y) \ (x \ y)\<^sup>\ = x\<^sup>\ * y\<^sup>\" nitpick [expect=genuine,card=7] oops lemma circ_separate_6: "y * x \ x * (x \ y) \ (x \ y)\<^sup>\ = x\<^sup>\ * y\<^sup>\" nitpick [expect=genuine,card=7] oops *) end text \ We next add the right induction rule, which allows us to strengthen many inequalities of left Kleene algebras to equalities. \ class strong_left_kleene_algebra = left_kleene_algebra + assumes star_right_induct: "z \ x * y \ x \ z * y\<^sup>\ \ x" begin lemma star_plus: "y\<^sup>\ * y = y * y\<^sup>\" proof (rule antisym) show "y\<^sup>\ * y \ y * y\<^sup>\" by (simp add: star.circ_plus_sub) next have "y\<^sup>\ * y * y \ y\<^sup>\ * y" by (simp add: mult_left_isotone star.right_plus_below_circ) hence "y \ y\<^sup>\ * y * y \ y\<^sup>\ * y" by (simp add: star.circ_mult_increasing_2) thus "y * y\<^sup>\ \ y\<^sup>\ * y" using star_right_induct by blast qed lemma star_slide: "(x * y)\<^sup>\ * x = x * (y * x)\<^sup>\" proof (rule antisym) show "(x * y)\<^sup>\ * x \ x * (y * x)\<^sup>\" by (rule star_left_slide) next have "x \ (x * y)\<^sup>\ * x * y * x \ (x * y)\<^sup>\ * x" by (metis (full_types) sup.commute eq_refl star.circ_loop_fixpoint mult.assoc star_plus) thus "x * (y * x)\<^sup>\ \ (x * y)\<^sup>\ * x" by (simp add: mult_assoc star_right_induct) qed lemma star_simulation_right: assumes "z * x \ y * z" shows "z * x\<^sup>\ \ y\<^sup>\ * z" proof - have "y\<^sup>\ * z * x \ y\<^sup>\ * z" by (metis assms dual_order.trans mult_isotone mult_left_sub_dist_sup_right star.circ_loop_fixpoint star.circ_transitive_equal sup.cobounded1 mult_assoc) thus ?thesis by (metis le_supI star.circ_loop_fixpoint star_right_induct sup.cobounded2) qed end text \ Again we inherit results from the itering hierarchy. \ sublocale strong_left_kleene_algebra < star: itering_1 where circ = star apply unfold_locales apply (simp add: star_slide) by (simp add: star_simulation_right) context strong_left_kleene_algebra begin lemma star_right_induct_mult: "y * x \ y \ y * x\<^sup>\ \ y" by (simp add: star_right_induct) lemma star_right_induct_mult_iff: "y * x \ y \ y * x\<^sup>\ \ y" using mult_right_isotone order_trans star.circ_increasing star_right_induct_mult by blast lemma star_simulation_right_equal: "z * x = y * z \ z * x\<^sup>\ = y\<^sup>\ * z" by (metis eq_iff star_simulation_left star_simulation_right) lemma star_simulation_star: "x * y \ y * x \ x\<^sup>\ * y\<^sup>\ \ y\<^sup>\ * x\<^sup>\" by (simp add: star_simulation_left star_simulation_right) lemma star_right_induct_equal: "z \ y * x = y \ z * x\<^sup>\ \ y" by (simp add: star_right_induct) lemma star_right_induct_mult_equal: "y * x = y \ y * x\<^sup>\ \ y" by (simp add: star_right_induct_mult) lemma star_back_loop_least_fixpoint: "x * y \ z = x \ z * y\<^sup>\ \ x" by (simp add: sup_commute star_right_induct_equal) lemma star_back_loop_is_least_fixpoint: "is_least_fixpoint (\x . x * y \ z) (z * y\<^sup>\)" proof (unfold is_least_fixpoint_def, rule conjI) have "(z * y\<^sup>\ * y \ z) * y \ z * y\<^sup>\ * y \ z" using le_supI1 mult_left_isotone star.circ_back_loop_prefixpoint by auto hence "z * y\<^sup>\ \ z * y\<^sup>\ * y \ z" by (simp add: star_right_induct) thus "z * y\<^sup>\ * y \ z = z * y\<^sup>\" using antisym star.circ_back_loop_prefixpoint by auto next show "\x. x * y \ z = x \ z * y\<^sup>\ \ x" by (simp add: star_back_loop_least_fixpoint) qed lemma star_back_loop_mu: "\ (\x . x * y \ z) = z * y\<^sup>\" by (metis least_fixpoint_same star_back_loop_is_least_fixpoint) lemma star_square: "x\<^sup>\ = (1 \ x) * (x * x)\<^sup>\" proof - let ?f = "\y . y * x \ 1" have 1: "isotone ?f" by (metis sup_left_isotone isotone_def mult_left_isotone) have "?f \ ?f = (\y . y * (x * x) \ (1 \ x))" by (simp add: sup_assoc sup_commute mult_assoc mult_right_dist_sup o_def) thus ?thesis using 1 by (metis mu_square mult_left_one star_back_loop_mu has_least_fixpoint_def star_back_loop_is_least_fixpoint) qed lemma star_square_2: "x\<^sup>\ = (x * x)\<^sup>\ * (x \ 1)" proof - have "(1 \ x) * (x * x)\<^sup>\ = (x * x)\<^sup>\ * 1 \ x * (x * x)\<^sup>\" using mult_right_dist_sup by force thus ?thesis by (metis (no_types) antisym mult_left_sub_dist_sup star.circ_square_2 star_slide sup_commute star_square) qed lemma star_circ_simulate_right_plus: assumes "z * x \ y * y\<^sup>\ * z \ w" shows "z * x\<^sup>\ \ y\<^sup>\ * (z \ w * x\<^sup>\)" proof - have "(z \ w * x\<^sup>\) * x \ z * x \ w * x\<^sup>\" using mult_right_dist_sup star.circ_back_loop_prefixpoint sup_right_isotone by auto also have "... \ y * y\<^sup>\ * z \ w \ w * x\<^sup>\" using assms sup_left_isotone by blast also have "... \ y * y\<^sup>\ * z \ w * x\<^sup>\" using le_supI1 star.circ_back_loop_prefixpoint sup_commute by auto also have "... \ y\<^sup>\ * (z \ w * x\<^sup>\)" by (metis sup.bounded_iff mult_isotone mult_left_isotone mult_left_one mult_left_sub_dist_sup_left star.circ_reflexive star.left_plus_below_circ) finally have "y\<^sup>\ * (z \ w * x\<^sup>\) * x \ y\<^sup>\ * (z \ w * x\<^sup>\)" by (metis mult_assoc mult_right_isotone star.circ_transitive_equal) thus ?thesis by (metis sup.bounded_iff star_right_induct mult_left_sub_dist_sup_left star.circ_loop_fixpoint) qed lemma transitive_star: "x * x \ x \ x\<^sup>\ = 1 \ x" by (metis order.antisym star.circ_mult_increasing_2 star.circ_plus_same star_left_induct_mult star_left_unfold_equal) (* lemma star_circ_simulate_left_plus: "x * z \ z * y\<^sup>\ \ w \ x\<^sup>\ * z \ (z \ x\<^sup>\ * w) * y\<^sup>\" nitpick [expect=genuine,card=7] oops *) end text \ The following class contains a generalisation of Kleene algebras, which lacks the right zero axiom. \ class left_zero_kleene_algebra = idempotent_left_zero_semiring + strong_left_kleene_algebra begin lemma star_star_absorb: "y\<^sup>\ * (y\<^sup>\ * x)\<^sup>\ * y\<^sup>\ = (y\<^sup>\ * x)\<^sup>\ * y\<^sup>\" by (metis star.circ_transitive_equal star_slide mult_assoc) lemma star_circ_simulate_left_plus: assumes "x * z \ z * y\<^sup>\ \ w" shows "x\<^sup>\ * z \ (z \ x\<^sup>\ * w) * y\<^sup>\" proof - have "x * (x\<^sup>\ * (w * y\<^sup>\)) \ x\<^sup>\ * (w * y\<^sup>\)" by (metis (no_types) mult_right_sub_dist_sup_left star.circ_loop_fixpoint mult_assoc) hence "x * ((z \ x\<^sup>\ * w) * y\<^sup>\) \ x * z * y\<^sup>\ \ x\<^sup>\ * w * y\<^sup>\" using mult_left_dist_sup mult_right_dist_sup sup_right_isotone mult_assoc by presburger also have "... \ (z * y\<^sup>\ \ w) * y\<^sup>\ \ x\<^sup>\ * w * y\<^sup>\" using assms mult_isotone semiring.add_right_mono by blast also have "... = z * y\<^sup>\ \ w * y\<^sup>\ \ x\<^sup>\ * w * y\<^sup>\" by (simp add: mult_right_dist_sup star.circ_transitive_equal mult_assoc) also have "... = (z \ w \ x\<^sup>\ * w) * y\<^sup>\" by (simp add: mult_right_dist_sup) also have "... = (z \ x\<^sup>\ * w) * y\<^sup>\" by (metis sup_assoc sup_ge2 le_iff_sup star.circ_loop_fixpoint) finally show ?thesis by (metis sup.bounded_iff mult_left_sub_dist_sup_left mult_1_right star.circ_right_unfold_1 star_left_induct) qed lemma star_one_sup_below: "x * y\<^sup>\ * (1 \ z) \ x * (y \ z)\<^sup>\" proof - have "y\<^sup>\ * z \ (y \ z)\<^sup>\" using sup_ge2 order_trans star.circ_increasing star.circ_mult_upper_bound star.circ_sub_dist by blast hence "y\<^sup>\ \ y\<^sup>\ * z \ (y \ z)\<^sup>\" by (simp add: star.circ_sup_upper_bound star.circ_sub_dist) hence "y\<^sup>\ * (1 \ z) \ (y \ z)\<^sup>\" by (simp add: mult_left_dist_sup) thus ?thesis by (metis mult_right_isotone mult_assoc) qed text \ The following theorem is similar to the puzzle where persons insert themselves always in the middle between two groups of people in a line. Here, however, items in the middle annihilate each other, leaving just one group of items behind. \ lemma cancel_separate: assumes "x * y \ 1" shows "x\<^sup>\ * y\<^sup>\ \ x\<^sup>\ \ y\<^sup>\" proof - have "x * y\<^sup>\ = x \ x * y * y\<^sup>\" by (metis mult_assoc mult_left_dist_sup mult_1_right star_left_unfold_equal) also have "... \ x \ y\<^sup>\" by (meson assms dual_order.trans order.refl star.circ_mult_upper_bound star.circ_reflexive sup_right_isotone) also have "... \ x\<^sup>\ \ y\<^sup>\" using star.circ_increasing sup_left_isotone by auto finally have 1: "x * y\<^sup>\ \ x\<^sup>\ \ y\<^sup>\" . have "x * (x\<^sup>\ \ y\<^sup>\) = x * x\<^sup>\ \ x * y\<^sup>\" by (simp add: mult_left_dist_sup) also have "... \ x\<^sup>\ \ y\<^sup>\" using 1 by (metis sup.bounded_iff sup_ge1 order_trans star.left_plus_below_circ) finally have 2: "x * (x\<^sup>\ \ y\<^sup>\) \ x\<^sup>\ \ y\<^sup>\" . have "y\<^sup>\ \ x\<^sup>\ \ y\<^sup>\" by simp hence "y\<^sup>\ \ x * (x\<^sup>\ \ y\<^sup>\) \ x\<^sup>\ \ y\<^sup>\" using 2 sup.bounded_iff by blast thus ?thesis by (metis star_left_induct) qed lemma star_separate: assumes "x * y = bot" and "y * y = bot" shows "(x \ y)\<^sup>\ = x\<^sup>\ \ y * x\<^sup>\" proof - have 1: "y\<^sup>\ = 1 \ y" using assms(2) by (simp add: transitive_star) have "(x \ y)\<^sup>\ = y\<^sup>\ * (x * y\<^sup>\)\<^sup>\" by (simp add: star.circ_decompose_6 star_sup_1) also have "... = y\<^sup>\ * (x * (1 \ y * y\<^sup>\))\<^sup>\" by (simp add: star_left_unfold_equal) also have "... = (1 \ y) * x\<^sup>\" using 1 by (simp add: assms mult_left_dist_sup) also have "... = x\<^sup>\ \ y * x\<^sup>\" by (simp add: mult_right_dist_sup) finally show ?thesis . qed end text \ We can now inherit from the strongest variant of iterings. \ sublocale left_zero_kleene_algebra < star: itering where circ = star apply unfold_locales apply (metis star.circ_sup_9) apply (metis star.circ_mult_1) apply (simp add: star_circ_simulate_right_plus) by (simp add: star_circ_simulate_left_plus) context left_zero_kleene_algebra begin lemma star_absorb: "x * y = bot \ x * y\<^sup>\ = x" by (metis sup.bounded_iff antisym_conv star.circ_back_loop_prefixpoint star.circ_elimination) lemma star_separate_2: assumes "x * z\<^sup>+ * y = bot" and "y * z\<^sup>+ * y = bot" and "z * x = bot" shows "(x\<^sup>\ \ y * x\<^sup>\) * (z * (1 \ y * x\<^sup>\))\<^sup>\ = z\<^sup>\ * (x\<^sup>\ \ y * x\<^sup>\) * z\<^sup>\" proof - have 1: "x\<^sup>\ * z\<^sup>+ * y = z\<^sup>+ * y" by (metis assms mult_assoc mult_1_left mult_left_zero star.circ_zero star_simulation_right_equal) have 2: "z\<^sup>\ * (x\<^sup>\ \ y * x\<^sup>\) * z\<^sup>+ \ z\<^sup>\ * (x\<^sup>\ \ y * x\<^sup>\) * z\<^sup>\" by (simp add: mult_right_isotone star.left_plus_below_circ) have "z\<^sup>\ * z\<^sup>+ * y * x\<^sup>\ \ z\<^sup>\ * y * x\<^sup>\" by (metis mult_left_isotone star.left_plus_below_circ star.right_plus_circ star_plus) also have "... \ z\<^sup>\ * (x\<^sup>\ \ y * x\<^sup>\)" by (simp add: mult_assoc mult_left_sub_dist_sup_right) also have "... \ z\<^sup>\ * (x\<^sup>\ \ y * x\<^sup>\) * z\<^sup>\" using sup_right_divisibility star.circ_back_loop_fixpoint by blast finally have 3: "z\<^sup>\ * z\<^sup>+ * y * x\<^sup>\ \ z\<^sup>\ * (x\<^sup>\ \ y * x\<^sup>\) * z\<^sup>\" . have "z\<^sup>\ * (x\<^sup>\ \ y * x\<^sup>\) * z\<^sup>\ * (z * (1 \ y * x\<^sup>\)) = z\<^sup>\ * (x\<^sup>\ \ y * x\<^sup>\) * z\<^sup>+ \ z\<^sup>\ * (x\<^sup>\ \ y * x\<^sup>\) * z\<^sup>+ * y * x\<^sup>\" by (metis mult_1_right semiring.distrib_left star.circ_plus_same mult_assoc) also have "... = z\<^sup>\ * (x\<^sup>\ \ y * x\<^sup>\) * z\<^sup>+ \ z\<^sup>\ * (1 \ y) * x\<^sup>\ * z\<^sup>+ * y * x\<^sup>\" by (simp add: semiring.distrib_right mult_assoc) also have "... = z\<^sup>\ * (x\<^sup>\ \ y * x\<^sup>\) * z\<^sup>+ \ z\<^sup>\ * (1 \ y) * z\<^sup>+ * y * x\<^sup>\" using 1 by (simp add: mult_assoc) also have "... = z\<^sup>\ * (x\<^sup>\ \ y * x\<^sup>\) * z\<^sup>+ \ z\<^sup>\ * z\<^sup>+ * y * x\<^sup>\ \ z\<^sup>\ * y * z\<^sup>+ * y * x\<^sup>\" using mult_left_dist_sup mult_right_dist_sup sup_assoc by auto also have "... = z\<^sup>\ * (x\<^sup>\ \ y * x\<^sup>\) * z\<^sup>+ \ z\<^sup>\ * z\<^sup>+ * y * x\<^sup>\" by (metis assms(2) mult_left_dist_sup mult_left_zero sup_commute sup_monoid.add_0_left mult_assoc) also have "... \ z\<^sup>\ * (x\<^sup>\ \ y * x\<^sup>\) * z\<^sup>\" using 2 3 by simp finally have "(x\<^sup>\ \ y * x\<^sup>\) \ z\<^sup>\ * (x\<^sup>\ \ y * x\<^sup>\) * z\<^sup>\ * (z * (1 \ y * x\<^sup>\)) \ z\<^sup>\ * (x\<^sup>\ \ y * x\<^sup>\) * z\<^sup>\" by (simp add: star_outer_increasing) hence 4: "(x\<^sup>\ \ y * x\<^sup>\) * (z * (1 \ y * x\<^sup>\))\<^sup>\ \ z\<^sup>\ * (x\<^sup>\ \ y * x\<^sup>\) * z\<^sup>\" by (simp add: star_right_induct) have 5: "(x\<^sup>\ \ y * x\<^sup>\) * z\<^sup>\ \ (x\<^sup>\ \ y * x\<^sup>\) * (z * (1 \ y * x\<^sup>\))\<^sup>\" by (metis sup_ge1 mult_right_isotone mult_1_right star_isotone) have "z * (x\<^sup>\ \ y * x\<^sup>\) = z * x\<^sup>\ \ z * y * x\<^sup>\" by (simp add: mult_assoc mult_left_dist_sup) also have "... = z \ z * y * x\<^sup>\" by (simp add: assms star_absorb) also have "... = z * (1 \ y * x\<^sup>\)" by (simp add: mult_assoc mult_left_dist_sup) also have "... \ (z * (1 \ y * x\<^sup>\))\<^sup>\" by (simp add: star.circ_increasing) also have "... \ (x\<^sup>\ \ y * x\<^sup>\) * (z * (1 \ y * x\<^sup>\))\<^sup>\" by (metis le_supE mult_right_sub_dist_sup_left star.circ_loop_fixpoint) finally have "z * (x\<^sup>\ \ y * x\<^sup>\) \ (x\<^sup>\ \ y * x\<^sup>\) * (z * (1 \ y * x\<^sup>\))\<^sup>\" . hence "z * (x\<^sup>\ \ y * x\<^sup>\) * (z * (1 \ y * x\<^sup>\))\<^sup>\ \ (x\<^sup>\ \ y * x\<^sup>\) * (z * (1 \ y * x\<^sup>\))\<^sup>\" by (metis mult_assoc mult_left_isotone star.circ_transitive_equal) hence "z\<^sup>\ * (x\<^sup>\ \ y * x\<^sup>\) * z\<^sup>\ \ (x\<^sup>\ \ y * x\<^sup>\) * (z * (1 \ y * x\<^sup>\))\<^sup>\" using 5 by (metis star_left_induct sup.bounded_iff mult_assoc) thus ?thesis using 4 by (simp add: antisym) qed lemma cancel_separate_eq: "x * y \ 1 \ x\<^sup>\ * y\<^sup>\ = x\<^sup>\ \ y\<^sup>\" by (metis antisym cancel_separate star.circ_plus_one star.circ_sup_sub_sup_one_1 star_involutive) lemma cancel_separate_1: assumes "x * y \ 1" shows "(x \ y)\<^sup>\ = y\<^sup>\ * x\<^sup>\" proof - have "y\<^sup>\ * x\<^sup>\ * y = y\<^sup>\ * x\<^sup>\ * x * y \ y\<^sup>\ * y" by (metis mult_right_dist_sup star.circ_back_loop_fixpoint) also have "... \ y\<^sup>\ * x\<^sup>\ \ y\<^sup>\ * y" by (metis assms semiring.add_right_mono mult_right_isotone mult_1_right mult_assoc) also have "... \ y\<^sup>\ * x\<^sup>\ \ y\<^sup>\" using semiring.add_left_mono star.right_plus_below_circ by simp also have "... = y\<^sup>\ * x\<^sup>\" by (metis star.circ_back_loop_fixpoint sup.left_idem sup_commute) finally have "y\<^sup>\ * x\<^sup>\ * y \ y\<^sup>\ * x\<^sup>\" by simp hence "y\<^sup>\ * x\<^sup>\ * (x \ y) \ y\<^sup>\ * x\<^sup>\ * x \ y\<^sup>\ * x\<^sup>\" using mult_left_dist_sup order_lesseq_imp by fastforce also have "... = y\<^sup>\ * x\<^sup>\" by (metis star.circ_back_loop_fixpoint sup.left_idem) finally have "(x \ y)\<^sup>\ \ y\<^sup>\ * x\<^sup>\" by (metis star.circ_decompose_7 star_right_induct_mult sup_commute) thus ?thesis using antisym star.circ_sub_dist_3 sup_commute by fastforce qed lemma plus_sup: "(x \ y)\<^sup>+ = (x\<^sup>\ * y)\<^sup>\ * x\<^sup>+ \ (x\<^sup>\ * y)\<^sup>+" by (metis semiring.distrib_left star.circ_sup_9 star_plus mult_assoc) lemma plus_half_bot: "x * y * x = bot \ (x * y)\<^sup>+ = x * y" by (metis star_absorb star_slide mult_assoc) lemma cancel_separate_1_sup: assumes "x * y \ 1" and "y * x \ 1" shows "(x \ y)\<^sup>\ = x\<^sup>\ \ y\<^sup>\" - by (simp add: assms cancel_separate_1 cancel_separate_eq local.sup_commute) + by (simp add: assms cancel_separate_1 cancel_separate_eq sup_commute) + +lemma star_separate_3: + assumes "y * x\<^sup>\ * y \ y" + shows "(x \ y)\<^sup>\ = x\<^sup>\ \ x\<^sup>\ * y * x\<^sup>\" +proof (rule antisym) + have "x\<^sup>\ * y * (x\<^sup>\ * y)\<^sup>\ * x\<^sup>\ \ x\<^sup>\ * y * x\<^sup>\" + by (metis assms mult_left_isotone mult_right_isotone star_right_induct_mult mult_assoc) + thus "(x \ y)\<^sup>\ \ x\<^sup>\ \ x\<^sup>\ * y * x\<^sup>\" + by (metis antisym semiring.add_left_mono star.circ_sup_2 star.circ_sup_sub star.circ_unfold_sum star_decompose_3 star_slide mult_assoc) +next + show "x\<^sup>\ \ x\<^sup>\ * y * x\<^sup>\ \ (x \ y)\<^sup>\" + using mult_isotone star.circ_increasing star.circ_sub_dist star.circ_sup_9 by auto +qed end text \ A Kleene algebra is obtained by requiring an idempotent semiring. \ class kleene_algebra = left_zero_kleene_algebra + idempotent_semiring text \ The following classes are variants of Kleene algebras expanded by an additional iteration operation. This is useful to study the Kleene star in computation models that do not use least fixpoints in the refinement order as the semantics of recursion. \ class left_kleene_conway_semiring = left_kleene_algebra + left_conway_semiring begin lemma star_below_circ: "x\<^sup>\ \ x\<^sup>\" by (metis circ_left_unfold mult_1_right order_refl star_left_induct) lemma star_zero_below_circ_mult: "x\<^sup>\ * bot \ x\<^sup>\ * y" by (simp add: mult_isotone star_below_circ) lemma star_mult_circ: "x\<^sup>\ * x\<^sup>\ = x\<^sup>\" by (metis sup_right_divisibility antisym circ_left_unfold star_left_induct_mult star.circ_loop_fixpoint) lemma circ_mult_star: "x\<^sup>\ * x\<^sup>\ = x\<^sup>\" by (metis sup_assoc sup.bounded_iff circ_left_unfold circ_rtc_2 eq_iff left_plus_circ star.circ_sup_sub star.circ_back_loop_prefixpoint star.circ_increasing star_below_circ star_mult_circ star_sup_one) lemma circ_star: "x\<^sup>\\<^sup>\ = x\<^sup>\" by (metis antisym circ_reflexive circ_transitive_equal star.circ_increasing star.circ_sup_one_right_unfold star_left_induct_mult_equal) lemma star_circ: "x\<^sup>\\<^sup>\ = x\<^sup>\\<^sup>\" by (metis antisym circ_circ_sup circ_sub_dist le_iff_sup star.circ_rtc_2 star_below_circ) lemma circ_sup_3: "(x\<^sup>\ * y\<^sup>\)\<^sup>\ \ (x \ y)\<^sup>\" using circ_star circ_sub_dist_3 star_isotone by fastforce end class left_zero_kleene_conway_semiring = left_zero_kleene_algebra + itering begin subclass left_kleene_conway_semiring .. lemma circ_isolate: "x\<^sup>\ = x\<^sup>\ * bot \ x\<^sup>\" by (metis sup_commute antisym circ_sup_upper_bound circ_mult_star circ_simulate_absorb star.left_plus_below_circ star_below_circ zero_right_mult_decreasing) lemma circ_isolate_mult: "x\<^sup>\ * y = x\<^sup>\ * bot \ x\<^sup>\ * y" by (metis circ_isolate mult_assoc mult_left_zero mult_right_dist_sup) lemma circ_isolate_mult_sub: "x\<^sup>\ * y \ x\<^sup>\ \ x\<^sup>\ * y" by (metis sup_left_isotone circ_isolate_mult zero_right_mult_decreasing) lemma circ_sub_decompose: "(x\<^sup>\ * y)\<^sup>\ \ (x\<^sup>\ * y)\<^sup>\ * x\<^sup>\" proof - have "x\<^sup>\ * y \ x\<^sup>\ * bot = x\<^sup>\ * y" by (metis sup.commute circ_isolate_mult) hence "(x\<^sup>\ * y)\<^sup>\ * x\<^sup>\ = ((x\<^sup>\ * y)\<^sup>\ \ x\<^sup>\)\<^sup>\" by (metis circ_star circ_sup_9 circ_sup_mult_zero star_decompose_1) thus ?thesis by (metis circ_star le_iff_sup star.circ_decompose_7 star.circ_unfold_sum) qed lemma circ_sup_4: "(x \ y)\<^sup>\ = (x\<^sup>\ * y)\<^sup>\ * x\<^sup>\" apply (rule antisym) apply (metis circ_sup circ_sub_decompose circ_transitive_equal mult_assoc mult_left_isotone) by (metis circ_sup circ_isotone mult_left_isotone star_below_circ) lemma circ_sup_5: "(x\<^sup>\ * y)\<^sup>\ * x\<^sup>\ = (x\<^sup>\ * y)\<^sup>\ * x\<^sup>\" using circ_sup_4 circ_sup_9 by auto lemma plus_circ: "(x\<^sup>\ * x)\<^sup>\ = x\<^sup>\" by (metis sup_idem circ_sup_4 circ_decompose_7 circ_star star.circ_decompose_5 star.right_plus_circ) (* lemma "(x\<^sup>\ * y * x\<^sup>\)\<^sup>\ = (x\<^sup>\ * y)\<^sup>\" nitpick [expect=genuine] oops *) end text \ The following classes add a greatest element. \ class bounded_left_kleene_algebra = bounded_idempotent_left_semiring + left_kleene_algebra sublocale bounded_left_kleene_algebra < star: bounded_left_conway_semiring where circ = star .. class bounded_left_zero_kleene_algebra = bounded_idempotent_left_semiring + left_zero_kleene_algebra sublocale bounded_left_zero_kleene_algebra < star: bounded_itering where circ = star .. class bounded_kleene_algebra = bounded_idempotent_semiring + kleene_algebra sublocale bounded_kleene_algebra < star: bounded_itering where circ = star .. text \ We conclude with an alternative axiomatisation of Kleene algebras. \ class kleene_algebra_var = idempotent_semiring + star + assumes star_left_unfold_var : "1 \ y * y\<^sup>\ \ y\<^sup>\" assumes star_left_induct_var : "y * x \ x \ y\<^sup>\ * x \ x" assumes star_right_induct_var : "x * y \ x \ x * y\<^sup>\ \ x" begin subclass kleene_algebra apply unfold_locales apply (rule star_left_unfold_var) apply (meson sup.bounded_iff mult_right_isotone order_trans star_left_induct_var) by (meson sup.bounded_iff mult_left_isotone order_trans star_right_induct_var) end end diff --git a/thys/Stone_Kleene_Relation_Algebras/Kleene_Relation_Algebras.thy b/thys/Stone_Kleene_Relation_Algebras/Kleene_Relation_Algebras.thy --- a/thys/Stone_Kleene_Relation_Algebras/Kleene_Relation_Algebras.thy +++ b/thys/Stone_Kleene_Relation_Algebras/Kleene_Relation_Algebras.thy @@ -1,3173 +1,3227 @@ (* Title: Kleene Relation Algebras Author: Walter Guttmann Maintainer: Walter Guttmann *) section \Kleene Relation Algebras\ text \ This theory combines Kleene algebras with Stone relation algebras. Relation algebras with transitive closure have been studied by \cite{Ng1984}. The weakening to Stone relation algebras allows us to talk about reachability in weighted graphs, for example. Many results in this theory are used in the correctness proof of Prim's minimum spanning tree algorithm. In particular, they are concerned with the exchange property, preservation of parts of the invariant and with establishing parts of the postcondition. \ theory Kleene_Relation_Algebras imports Stone_Relation_Algebras.Relation_Algebras Kleene_Algebras begin text \ We first note that bounded distributive lattices can be expanded to Kleene algebras by reusing some of the operations. \ sublocale bounded_distrib_lattice < comp_inf: bounded_kleene_algebra where star = "\x . top" and one = top and times = inf apply unfold_locales apply (simp add: inf.assoc) apply simp apply simp apply (simp add: le_infI2) apply (simp add: inf_sup_distrib2) apply simp apply simp apply simp apply simp apply simp apply (simp add: inf_sup_distrib1) apply simp apply simp by (simp add: inf_assoc) text \ We add the Kleene star operation to each of bounded distributive allegories, pseudocomplemented distributive allegories and Stone relation algebras. We start with single-object bounded distributive allegories. \ class bounded_distrib_kleene_allegory = bounded_distrib_allegory + kleene_algebra begin subclass bounded_kleene_algebra .. lemma conv_star_conv: "x\<^sup>\ \ x\<^sup>T\<^sup>\\<^sup>T" proof - have "x\<^sup>T\<^sup>\ * x\<^sup>T \ x\<^sup>T\<^sup>\" by (simp add: star.right_plus_below_circ) hence 1: "x * x\<^sup>T\<^sup>\\<^sup>T \ x\<^sup>T\<^sup>\\<^sup>T" using conv_dist_comp conv_isotone by fastforce have "1 \ x\<^sup>T\<^sup>\\<^sup>T" by (simp add: reflexive_conv_closed star.circ_reflexive) hence "1 \ x * x\<^sup>T\<^sup>\\<^sup>T \ x\<^sup>T\<^sup>\\<^sup>T" using 1 by simp thus ?thesis using star_left_induct by fastforce qed text \ It follows that star and converse commute. \ lemma conv_star_commute: "x\<^sup>\\<^sup>T = x\<^sup>T\<^sup>\" proof (rule antisym) show "x\<^sup>\\<^sup>T \ x\<^sup>T\<^sup>\" using conv_star_conv conv_isotone by fastforce next show "x\<^sup>T\<^sup>\ \ x\<^sup>\\<^sup>T" by (metis conv_star_conv conv_involutive) qed lemma conv_plus_commute: "x\<^sup>+\<^sup>T = x\<^sup>T\<^sup>+" by (simp add: conv_dist_comp conv_star_commute star_plus) +lemma reflexive_inf_star: + assumes "reflexive y" + shows "y \ x\<^sup>\ = 1 \ (y \ x\<^sup>+)" + by (simp add: assms star_left_unfold_equal sup.absorb2 sup_inf_distrib1) + text \ The following results are variants of a separation lemma of Kleene algebras. \ lemma cancel_separate_2: assumes "x * y \ 1" shows "((w \ x) \ (z \ y))\<^sup>\ = (z \ y)\<^sup>\ * (w \ x)\<^sup>\" proof - have "(w \ x) * (z \ y) \ 1" by (meson assms comp_isotone order.trans inf.cobounded2) thus ?thesis using cancel_separate_1 sup_commute by simp qed lemma cancel_separate_3: assumes "x * y \ 1" shows "(w \ x)\<^sup>\ * (z \ y)\<^sup>\ = (w \ x)\<^sup>\ \ (z \ y)\<^sup>\" proof - have "(w \ x) * (z \ y) \ 1" by (meson assms comp_isotone order.trans inf.cobounded2) thus ?thesis by (simp add: cancel_separate_eq) qed lemma cancel_separate_4: assumes "z * y \ 1" and "w \ y \ z" and "x \ y \ z" shows "w\<^sup>\ * x\<^sup>\ = (w \ y)\<^sup>\ * ((w \ z)\<^sup>\ \ (x \ y)\<^sup>\) * (x \ z)\<^sup>\" proof - have "w\<^sup>\ * x\<^sup>\ = ((w \ y) \ (w \ z))\<^sup>\ * ((x \ y) \ (x \ z))\<^sup>\" by (metis assms(2,3) inf.orderE inf_sup_distrib1) also have "... = (w \ y)\<^sup>\ * ((w \ z)\<^sup>\ * (x \ y)\<^sup>\) * (x \ z)\<^sup>\" by (metis assms(1) cancel_separate_2 sup_commute mult_assoc) finally show ?thesis by (simp add: assms(1) cancel_separate_3) qed lemma cancel_separate_5: assumes "w * z\<^sup>T \ 1" shows "w \ x * (y \ z) \ y" proof - have "w \ x * (y \ z) \ (x \ w * (y \ z)\<^sup>T) * (y \ z)" by (metis dedekind_2 inf_commute) also have "... \ w * z\<^sup>T * (y \ z)" by (simp add: conv_dist_inf inf.coboundedI2 mult_left_isotone mult_right_isotone) also have "... \ y \ z" by (metis assms mult_1_left mult_left_isotone) finally show ?thesis by simp qed lemma cancel_separate_6: assumes "z * y \ 1" and "w \ y \ z" and "x \ y \ z" and "v * z\<^sup>T \ 1" and "v \ y\<^sup>\ = bot" shows "v \ w\<^sup>\ * x\<^sup>\ \ x \ w" proof - have "v \ (w \ y)\<^sup>\ * (x \ y)\<^sup>\ \ v \ y\<^sup>\ * (x \ y)\<^sup>\" using comp_inf.mult_right_isotone mult_left_isotone star_isotone by simp also have "... \ v \ y\<^sup>\" by (simp add: inf.coboundedI2 star.circ_increasing star.circ_mult_upper_bound star_right_induct_mult) finally have 1: "v \ (w \ y)\<^sup>\ * (x \ y)\<^sup>\ = bot" using assms(5) le_bot by simp have "v \ w\<^sup>\ * x\<^sup>\ = v \ (w \ y)\<^sup>\ * ((w \ z)\<^sup>\ \ (x \ y)\<^sup>\) * (x \ z)\<^sup>\" using assms(1-3) cancel_separate_4 by simp also have "... = (v \ (w \ y)\<^sup>\ * ((w \ z)\<^sup>\ \ (x \ y)\<^sup>\) * (x \ z)\<^sup>\ * (x \ z)) \ (v \ (w \ y)\<^sup>\ * ((w \ z)\<^sup>\ \ (x \ y)\<^sup>\))" by (metis inf_sup_distrib1 star.circ_back_loop_fixpoint) also have "... \ x \ (v \ (w \ y)\<^sup>\ * ((w \ z)\<^sup>\ \ (x \ y)\<^sup>\))" using assms(4) cancel_separate_5 semiring.add_right_mono by simp also have "... = x \ (v \ (w \ y)\<^sup>\ * (w \ z)\<^sup>\)" using 1 by (simp add: inf_sup_distrib1 mult_left_dist_sup sup_monoid.add_assoc) also have "... = x \ (v \ (w \ y)\<^sup>\ * (w \ z)\<^sup>\ * (w \ z)) \ (v \ (w \ y)\<^sup>\)" by (metis comp_inf.semiring.distrib_left star.circ_back_loop_fixpoint sup_assoc) also have "... \ x \ w \ (v \ (w \ y)\<^sup>\)" using assms(4) cancel_separate_5 sup_left_isotone sup_right_isotone by simp also have "... \ x \ w \ (v \ y\<^sup>\)" using comp_inf.mult_right_isotone star_isotone sup_right_isotone by simp finally show ?thesis using assms(5) le_bot by simp qed text \ We show several results about the interaction of vectors and the Kleene star. \ lemma vector_star_1: assumes "vector x" shows "x\<^sup>T * (x * x\<^sup>T)\<^sup>\ \ x\<^sup>T" proof - have "x\<^sup>T * (x * x\<^sup>T)\<^sup>\ = (x\<^sup>T * x)\<^sup>\ * x\<^sup>T" by (simp add: star_slide) also have "... \ top * x\<^sup>T" by (simp add: mult_left_isotone) also have "... = x\<^sup>T" using assms vector_conv_covector by auto finally show ?thesis . qed lemma vector_star_2: "vector x \ x\<^sup>T * (x * x\<^sup>T)\<^sup>\ \ x\<^sup>T * bot\<^sup>\" by (simp add: star_absorb vector_star_1) lemma vector_vector_star: "vector v \ (v * v\<^sup>T)\<^sup>\ = 1 \ v * v\<^sup>T" by (simp add: transitive_star vv_transitive) text \ The following equivalence relation characterises the component trees of a forest. This is a special case of undirected reachability in a directed graph. \ abbreviation "forest_components f \ f\<^sup>T\<^sup>\ * f\<^sup>\" lemma forest_components_equivalence: "injective x \ equivalence (forest_components x)" apply (intro conjI) apply (simp add: reflexive_mult_closed star.circ_reflexive) apply (metis cancel_separate_1 eq_iff star.circ_transitive_equal) by (simp add: conv_dist_comp conv_star_commute) lemma forest_components_increasing: "x \ forest_components x" by (metis order.trans mult_left_isotone mult_left_one star.circ_increasing star.circ_reflexive) lemma forest_components_isotone: "x \ y \ forest_components x \ forest_components y" by (simp add: comp_isotone conv_isotone star_isotone) lemma forest_components_idempotent: "injective x \ forest_components (forest_components x) = forest_components x" by (metis forest_components_equivalence cancel_separate_1 star.circ_transitive_equal star_involutive) lemma forest_components_star: "injective x \ (forest_components x)\<^sup>\ = forest_components x" using forest_components_equivalence forest_components_idempotent star.circ_transitive_equal by simp text \ The following lemma shows that the nodes reachable in the graph can be reached by only using edges between reachable nodes. \ lemma reachable_restrict: assumes "vector r" shows "r\<^sup>T * g\<^sup>\ = r\<^sup>T * ((r\<^sup>T * g\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\) \ g)\<^sup>\" proof - have 1: "r\<^sup>T \ r\<^sup>T * ((r\<^sup>T * g\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\) \ g)\<^sup>\" using mult_right_isotone mult_1_right star.circ_reflexive by fastforce have 2: "covector (r\<^sup>T * g\<^sup>\)" using assms covector_mult_closed vector_conv_covector by auto have "r\<^sup>T * ((r\<^sup>T * g\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\) \ g)\<^sup>\ * g \ r\<^sup>T * g\<^sup>\ * g" by (simp add: mult_left_isotone mult_right_isotone star_isotone) also have "... \ r\<^sup>T * g\<^sup>\" by (simp add: mult_assoc mult_right_isotone star.left_plus_below_circ star_plus) finally have "r\<^sup>T * ((r\<^sup>T * g\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\) \ g)\<^sup>\ * g = r\<^sup>T * ((r\<^sup>T * g\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\) \ g)\<^sup>\ * g \ r\<^sup>T * g\<^sup>\" by (simp add: le_iff_inf) also have "... = r\<^sup>T * ((r\<^sup>T * g\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\) \ g)\<^sup>\ * (g \ r\<^sup>T * g\<^sup>\)" using assms covector_comp_inf covector_mult_closed vector_conv_covector by auto also have "... = (r\<^sup>T * ((r\<^sup>T * g\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\) \ g)\<^sup>\ \ r\<^sup>T * g\<^sup>\) * (g \ r\<^sup>T * g\<^sup>\)" by (simp add: inf.absorb2 inf_commute mult_right_isotone star_isotone) also have "... = r\<^sup>T * ((r\<^sup>T * g\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\) \ g)\<^sup>\ * (g \ r\<^sup>T * g\<^sup>\ \ (r\<^sup>T * g\<^sup>\)\<^sup>T)" using 2 by (metis comp_inf_vector_1) also have "... = r\<^sup>T * ((r\<^sup>T * g\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\) \ g)\<^sup>\ * ((r\<^sup>T * g\<^sup>\)\<^sup>T \ r\<^sup>T * g\<^sup>\ \ g)" using inf_commute inf_assoc by simp also have "... = r\<^sup>T * ((r\<^sup>T * g\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\) \ g)\<^sup>\ * ((r\<^sup>T * g\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\) \ g)" using 2 by (metis covector_conv_vector inf_top.right_neutral vector_inf_comp) also have "... \ r\<^sup>T * ((r\<^sup>T * g\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\) \ g)\<^sup>\" by (simp add: mult_assoc mult_right_isotone star.left_plus_below_circ star_plus) finally have "r\<^sup>T * g\<^sup>\ \ r\<^sup>T * ((r\<^sup>T * g\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\) \ g)\<^sup>\" using 1 star_right_induct by auto thus ?thesis by (simp add: inf.eq_iff mult_right_isotone star_isotone) qed lemma kruskal_acyclic_inv_1: assumes "injective f" and "e * forest_components f * e = bot" shows "(f \ top * e * f\<^sup>T\<^sup>\)\<^sup>T * f\<^sup>\ * e = bot" proof - let ?q = "top * e * f\<^sup>T\<^sup>\" let ?F = "forest_components f" have "(f \ ?q)\<^sup>T * f\<^sup>\ * e = ?q\<^sup>T \ f\<^sup>T * f\<^sup>\ * e" by (metis (mono_tags) comp_associative conv_dist_inf covector_conv_vector inf_vector_comp vector_top_closed) also have "... \ ?q\<^sup>T \ ?F * e" using comp_inf.mult_right_isotone mult_left_isotone star.circ_increasing by simp also have "... = f\<^sup>\ * e\<^sup>T * top \ ?F * e" by (simp add: conv_dist_comp conv_star_commute mult_assoc) also have "... \ ?F * e\<^sup>T * top \ ?F * e" by (metis conv_dist_comp conv_star_commute conv_top inf.sup_left_isotone star.circ_right_top star_outer_increasing mult_assoc) also have "... = ?F * (e\<^sup>T * top \ ?F * e)" by (metis assms(1) forest_components_equivalence equivalence_comp_dist_inf mult_assoc) also have "... = (?F \ top * e) * ?F * e" by (simp add: comp_associative comp_inf_vector_1 conv_dist_comp inf_vector_comp) also have "... \ top * e * ?F * e" by (simp add: mult_left_isotone) also have "... = bot" using assms(2) mult_assoc by simp finally show ?thesis by (simp add: bot_unique) qed lemma kruskal_forest_components_inf_1: assumes "f \ w \ w\<^sup>T" and "injective w" and "f \ forest_components g" shows "f * forest_components (forest_components g \ w) \ forest_components (forest_components g \ w)" proof - let ?f = "forest_components g" let ?w = "forest_components (?f \ w)" have "f * ?w = (f \ (w \ w\<^sup>T)) * ?w" by (simp add: assms(1) inf.absorb1) also have "... = (f \ w) * ?w \ (f \ w\<^sup>T) * ?w" by (simp add: inf_sup_distrib1 semiring.distrib_right) also have "... \ (?f \ w) * ?w \ (f \ w\<^sup>T) * ?w" using assms(3) inf.sup_left_isotone mult_left_isotone sup_left_isotone by simp also have "... \ (?f \ w) * ?w \ (?f \ w\<^sup>T) * ?w" using assms(3) inf.sup_left_isotone mult_left_isotone sup_right_isotone by simp also have "... = (?f \ w) * ?w \ (?f \ w)\<^sup>T * ?w" by (simp add: conv_dist_comp conv_dist_inf conv_star_commute) also have "... \ (?f \ w) * ?w \ ?w" by (metis star.circ_loop_fixpoint sup_ge1 sup_right_isotone) also have "... = ?w \ (?f \ w) * (?f \ w)\<^sup>\ \ (?f \ w) * (?f \ w)\<^sup>T\<^sup>+ * (?f \ w)\<^sup>\" by (metis comp_associative mult_left_dist_sup star.circ_loop_fixpoint sup_commute sup_assoc) also have "... \ ?w \ (?f \ w)\<^sup>\ \ (?f \ w) * (?f \ w)\<^sup>T\<^sup>+ * (?f \ w)\<^sup>\" using star.left_plus_below_circ sup_left_isotone sup_right_isotone by auto also have "... = ?w \ (?f \ w) * (?f \ w)\<^sup>T\<^sup>+ * (?f \ w)\<^sup>\" by (metis star.circ_loop_fixpoint sup.right_idem) also have "... \ ?w \ w * w\<^sup>T * ?w" using comp_associative conv_dist_inf mult_isotone sup_right_isotone by simp also have "... = ?w" by (metis assms(2) coreflexive_comp_top_inf inf.cobounded2 sup.orderE) finally show ?thesis by simp qed lemma kruskal_forest_components_inf: assumes "f \ w \ w\<^sup>T" and "injective w" shows "forest_components f \ forest_components (forest_components f \ w)" proof - let ?f = "forest_components f" let ?w = "forest_components (?f \ w)" have 1: "1 \ ?w" by (simp add: reflexive_mult_closed star.circ_reflexive) have "f * ?w \ ?w" using assms forest_components_increasing kruskal_forest_components_inf_1 by simp hence 2: "f\<^sup>\ \ ?w" using 1 star_left_induct by fastforce have "f\<^sup>T * ?w \ ?w" apply (rule kruskal_forest_components_inf_1) apply (metis assms(1) conv_dist_sup conv_involutive conv_isotone sup_commute) apply (simp add: assms(2)) by (metis le_supI2 star.circ_back_loop_fixpoint star.circ_increasing) thus "?f \ ?w" using 2 star_left_induct by simp qed end text \ We next add the Kleene star to single-object pseudocomplemented distributive allegories. \ class pd_kleene_allegory = pd_allegory + bounded_distrib_kleene_allegory begin text \ The following definitions and results concern acyclic graphs and forests. \ abbreviation acyclic :: "'a \ bool" where "acyclic x \ x\<^sup>+ \ -1" abbreviation forest :: "'a \ bool" where "forest x \ injective x \ acyclic x" lemma forest_bot: "forest bot" by simp lemma acyclic_star_below_complement: "acyclic w \ w\<^sup>T\<^sup>\ \ -w" by (simp add: conv_star_commute schroeder_4_p) lemma acyclic_star_below_complement_1: "acyclic w \ w\<^sup>\ \ w\<^sup>T = bot" using pseudo_complement schroeder_5_p by force lemma acyclic_star_inf_conv: assumes "acyclic w" shows "w\<^sup>\ \ w\<^sup>T\<^sup>\ = 1" proof - have "w\<^sup>+ \ w\<^sup>T\<^sup>\ \ (w \ w\<^sup>T\<^sup>\) * w\<^sup>\" by (metis conv_star_commute dedekind_2 star.circ_transitive_equal) also have "... = bot" by (metis assms conv_star_commute p_antitone_iff pseudo_complement schroeder_4_p semiring.mult_not_zero star.circ_circ_mult star_involutive star_one) finally have "w\<^sup>\ \ w\<^sup>T\<^sup>\ \ 1" by (metis eq_iff le_bot mult_left_zero star.circ_plus_one star.circ_zero star_left_unfold_equal sup_inf_distrib1) thus ?thesis by (simp add: inf.antisym star.circ_reflexive) qed lemma acyclic_asymmetric: "acyclic w \ asymmetric w" - by (simp add: local.dual_order.trans local.pseudo_complement local.schroeder_5_p local.star.circ_increasing) + by (simp add: dual_order.trans pseudo_complement schroeder_5_p star.circ_increasing) lemma forest_separate: assumes "forest x" shows "x\<^sup>\ * x\<^sup>T\<^sup>\ \ x\<^sup>T * x \ 1" proof - have "x\<^sup>\ * 1 \ -x\<^sup>T" using assms schroeder_5_p by force hence 1: "x\<^sup>\ \ x\<^sup>T = bot" by (simp add: pseudo_complement) have "x\<^sup>\ \ x\<^sup>T * x = (1 \ x\<^sup>\ * x) \ x\<^sup>T * x" using star.circ_right_unfold_1 by simp also have "... = (1 \ x\<^sup>T * x) \ (x\<^sup>\ * x \ x\<^sup>T * x)" by (simp add: inf_sup_distrib2) also have "... \ 1 \ (x\<^sup>\ * x \ x\<^sup>T * x)" using sup_left_isotone by simp also have "... = 1 \ (x\<^sup>\ \ x\<^sup>T) * x" by (simp add: assms injective_comp_right_dist_inf) also have "... = 1" using 1 by simp finally have 2: "x\<^sup>\ \ x\<^sup>T * x \ 1" . hence 3: "x\<^sup>T\<^sup>\ \ x\<^sup>T * x \ 1" by (metis (mono_tags, lifting) conv_star_commute conv_dist_comp conv_dist_inf conv_involutive coreflexive_symmetric) have "x\<^sup>\ * x\<^sup>T\<^sup>\ \ x\<^sup>T * x \ (x\<^sup>\ \ x\<^sup>T\<^sup>\) \ x\<^sup>T * x" using assms cancel_separate inf.sup_left_isotone by simp also have "... \ 1" using 2 3 by (simp add: inf_sup_distrib2) finally show ?thesis . qed text \ The following definition captures the components of undirected weighted graphs. \ abbreviation "components g \ (--g)\<^sup>\" lemma components_equivalence: "symmetric x \ equivalence (components x)" by (simp add: conv_star_commute conv_complement star.circ_reflexive star.circ_transitive_equal) lemma components_increasing: "x \ components x" using order_trans pp_increasing star.circ_increasing by blast lemma components_isotone: "x \ y \ components x \ components y" by (simp add: pp_isotone star_isotone) lemma cut_reachable: assumes "v\<^sup>T = r\<^sup>T * t\<^sup>\" and "t \ g" shows "v * -v\<^sup>T \ g \ (r\<^sup>T * g\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\)" proof - have "v * -v\<^sup>T \ g \ v * top \ g" using inf.sup_left_isotone mult_right_isotone top_greatest by blast also have "... = (r\<^sup>T * t\<^sup>\)\<^sup>T * top \ g" by (metis assms(1) conv_involutive) also have "... \ (r\<^sup>T * g\<^sup>\)\<^sup>T * top \ g" using assms(2) conv_isotone inf.sup_left_isotone mult_left_isotone mult_right_isotone star_isotone by auto also have "... \ (r\<^sup>T * g\<^sup>\)\<^sup>T * ((r\<^sup>T * g\<^sup>\) * g)" by (metis conv_involutive dedekind_1 inf_top.left_neutral) also have "... \ (r\<^sup>T * g\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\)" by (simp add: mult_assoc mult_right_isotone star.left_plus_below_circ star_plus) finally show ?thesis . qed text \ The following lemma shows that the predecessors of visited nodes in the minimum spanning tree extending the current tree have all been visited. \ lemma predecessors_reachable: assumes "vector r" and "injective r" and "v\<^sup>T = r\<^sup>T * t\<^sup>\" and "forest w" and "t \ w" and "w \ (r\<^sup>T * g\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\) \ g" and "r\<^sup>T * g\<^sup>\ \ r\<^sup>T * w\<^sup>\" shows "w * v \ v" proof - have "w * r \ (r\<^sup>T * g\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\) * r" using assms(6) mult_left_isotone by auto also have "... \ (r\<^sup>T * g\<^sup>\)\<^sup>T * top" by (simp add: mult_assoc mult_right_isotone) also have "... = (r\<^sup>T * g\<^sup>\)\<^sup>T" by (simp add: assms(1) comp_associative conv_dist_comp) also have "... \ (r\<^sup>T * w\<^sup>\)\<^sup>T" by (simp add: assms(7) conv_isotone) also have "... = w\<^sup>T\<^sup>\ * r" by (simp add: conv_dist_comp conv_star_commute) also have "... \ -w * r" using assms(4) by (simp add: mult_left_isotone acyclic_star_below_complement) also have "... \ -(w * r)" by (simp add: assms(2) comp_injective_below_complement) finally have 1: "w * r = bot" by (simp add: le_iff_inf) have "v = t\<^sup>T\<^sup>\ * r" by (metis assms(3) conv_dist_comp conv_involutive conv_star_commute) also have "... = t\<^sup>T * v \ r" by (simp add: calculation star.circ_loop_fixpoint) also have "... \ w\<^sup>T * v \ r" using assms(5) comp_isotone conv_isotone semiring.add_right_mono by auto finally have "w * v \ w * w\<^sup>T * v \ w * r" by (simp add: comp_left_dist_sup mult_assoc mult_right_isotone) also have "... = w * w\<^sup>T * v" using 1 by simp also have "... \ v" using assms(4) by (simp add: star_left_induct_mult_iff star_sub_one) finally show ?thesis . qed subsection \Prim's Algorithm\ text \ The following results are used for proving the correctness of Prim's minimum spanning tree algorithm. \ subsubsection \Preservation of Invariant\ text \ We first treat the preservation of the invariant. The following lemma shows that the while-loop preserves that \v\ represents the nodes of the constructed tree. The remaining lemmas in this section show that \t\ is a spanning tree. The exchange property is treated in the following two sections. \ lemma reachable_inv: assumes "vector v" and "e \ v * -v\<^sup>T" and "e * t = bot" and "v\<^sup>T = r\<^sup>T * t\<^sup>\" shows "(v \ e\<^sup>T * top)\<^sup>T = r\<^sup>T * (t \ e)\<^sup>\" proof - have 1: "v\<^sup>T \ r\<^sup>T * (t \ e)\<^sup>\" by (simp add: assms(4) mult_right_isotone star.circ_sub_dist) have 2: "(e\<^sup>T * top)\<^sup>T = top * e" by (simp add: conv_dist_comp) also have "... = top * (v * -v\<^sup>T \ e)" by (simp add: assms(2) inf_absorb2) also have "... \ top * (v * top \ e)" using inf.sup_left_isotone mult_right_isotone top_greatest by blast also have "... = top * v\<^sup>T * e" by (simp add: comp_inf_vector inf.sup_monoid.add_commute) also have "... = v\<^sup>T * e" using assms(1) vector_conv_covector by auto also have "... \ r\<^sup>T * (t \ e)\<^sup>\ * e" using 1 by (simp add: mult_left_isotone) also have "... \ r\<^sup>T * (t \ e)\<^sup>\ * (t \ e)" by (simp add: mult_right_isotone) also have "... \ r\<^sup>T * (t \ e)\<^sup>\" by (simp add: comp_associative mult_right_isotone star.right_plus_below_circ) finally have 3: "(v \ e\<^sup>T * top)\<^sup>T \ r\<^sup>T * (t \ e)\<^sup>\" using 1 by (simp add: conv_dist_sup) have "r\<^sup>T \ r\<^sup>T * t\<^sup>\" using sup.bounded_iff star.circ_back_loop_prefixpoint by blast also have "... \ (v \ e\<^sup>T * top)\<^sup>T" by (metis assms(4) conv_isotone sup_ge1) finally have 4: "r\<^sup>T \ (v \ e\<^sup>T * top)\<^sup>T" . have "(v \ e\<^sup>T * top)\<^sup>T * (t \ e) = (v \ e\<^sup>T * top)\<^sup>T * t \ (v \ e\<^sup>T * top)\<^sup>T * e" by (simp add: mult_left_dist_sup) also have "... \ (v \ e\<^sup>T * top)\<^sup>T * t \ top * e" using comp_isotone semiring.add_left_mono by auto also have "... = v\<^sup>T * t \ top * e * t \ top * e" using 2 by (simp add: conv_dist_sup mult_right_dist_sup) also have "... = v\<^sup>T * t \ top * e" by (simp add: assms(3) comp_associative) also have "... \ r\<^sup>T * t\<^sup>\ \ top * e" by (metis assms(4) star.circ_back_loop_fixpoint sup_ge1 sup_left_isotone) also have "... = v\<^sup>T \ top * e" by (simp add: assms(4)) finally have 5: "(v \ e\<^sup>T * top)\<^sup>T * (t \ e) \ (v \ e\<^sup>T * top)\<^sup>T" using 2 by (simp add: conv_dist_sup) have "r\<^sup>T * (t \ e)\<^sup>\ \ (v \ e\<^sup>T * top)\<^sup>T * (t \ e)\<^sup>\" using 4 by (simp add: mult_left_isotone) also have "... \ (v \ e\<^sup>T * top)\<^sup>T" using 5 by (simp add: star_right_induct_mult) finally show ?thesis using 3 by (simp add: inf.eq_iff) qed text \ The next result is used to show that the while-loop preserves acyclicity of the constructed tree. \ lemma acyclic_inv: assumes "acyclic t" and "vector v" and "e \ v * -v\<^sup>T" and "t \ v * v\<^sup>T" shows "acyclic (t \ e)" proof - have "t\<^sup>+ * e \ t\<^sup>+ * v * -v\<^sup>T" by (simp add: assms(3) comp_associative mult_right_isotone) also have "... \ v * v\<^sup>T * t\<^sup>\ * v * -v\<^sup>T" by (simp add: assms(4) mult_left_isotone) also have "... \ v * top * -v\<^sup>T" by (metis mult_assoc mult_left_isotone mult_right_isotone top_greatest) also have "... = v * -v\<^sup>T" by (simp add: assms(2)) also have "... \ -1" by (simp add: pp_increasing schroeder_3_p) finally have 1: "t\<^sup>+ * e \ -1" . have 2: "e * t\<^sup>\ = e" using assms(2-4) et(1) star_absorb by blast have "e\<^sup>\ = 1 \ e \ e * e * e\<^sup>\" by (metis star.circ_loop_fixpoint star_square_2 sup_commute) also have "... = 1 \ e" using assms(2,3) ee comp_left_zero bot_least sup_absorb1 by simp finally have 3: "e\<^sup>\ = 1 \ e" . have "e \ v * -v\<^sup>T" by (simp add: assms(3)) also have "... \ -1" by (simp add: pp_increasing schroeder_3_p) finally have 4: "t\<^sup>+ * e \ e \ -1" using 1 by simp have "(t \ e)\<^sup>+ = (t \ e) * t\<^sup>\ * (e * t\<^sup>\)\<^sup>\" using star_sup_1 mult_assoc by simp also have "... = (t \ e) * t\<^sup>\ * (1 \ e)" using 2 3 by simp also have "... = t\<^sup>+ * (1 \ e) \ e * t\<^sup>\ * (1 \ e)" by (simp add: comp_right_dist_sup) also have "... = t\<^sup>+ * (1 \ e) \ e * (1 \ e)" using 2 by simp also have "... = t\<^sup>+ * (1 \ e) \ e" using 3 by (metis star_absorb assms(2,3) ee) also have "... = t\<^sup>+ \ t\<^sup>+ * e \ e" by (simp add: mult_left_dist_sup) also have "... \ -1" using 4 by (metis assms(1) sup.absorb1 sup.orderI sup_assoc) finally show ?thesis . qed text \ The following lemma shows that the extended tree is in the component reachable from the root. \ lemma mst_subgraph_inv_2: assumes "regular (v * v\<^sup>T)" and "t \ v * v\<^sup>T \ --g" and "v\<^sup>T = r\<^sup>T * t\<^sup>\" and "e \ v * -v\<^sup>T \ --g" and "vector v" and "regular ((v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T)" shows "t \ e \ (r\<^sup>T * (--((v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T \ g))\<^sup>\)\<^sup>T * (r\<^sup>T * (--((v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T \ g))\<^sup>\)" proof - let ?v = "v \ e\<^sup>T * top" let ?G = "?v * ?v\<^sup>T \ g" let ?c = "r\<^sup>T * (--?G)\<^sup>\" have "v\<^sup>T \ r\<^sup>T * (--(v * v\<^sup>T \ g))\<^sup>\" using assms(1-3) inf_pp_commute mult_right_isotone star_isotone by auto also have "... \ ?c" using comp_inf.mult_right_isotone comp_isotone conv_isotone inf.commute mult_right_isotone pp_isotone star_isotone sup.cobounded1 by presburger finally have 2: "v\<^sup>T \ ?c \ v \ ?c\<^sup>T" by (metis conv_isotone conv_involutive) have "t \ v * v\<^sup>T" using assms(2) by auto hence 3: "t \ ?c\<^sup>T * ?c" using 2 order_trans mult_isotone by blast have "e \ v * top \ --g" by (metis assms(4,5) inf.bounded_iff inf.sup_left_divisibility mult_right_isotone top.extremum) hence "e \ v * top \ top * e \ --g" by (simp add: top_left_mult_increasing inf.boundedI) hence "e \ v * top * e \ --g" by (metis comp_inf_covector inf.absorb2 mult_assoc top.extremum) hence "t \ e \ (v * v\<^sup>T \ --g) \ (v * top * e \ --g)" using assms(2) sup_mono by blast also have "... = v * ?v\<^sup>T \ --g" by (simp add: inf_sup_distrib2 mult_assoc mult_left_dist_sup conv_dist_comp conv_dist_sup) also have "... \ --?G" using assms(6) comp_left_increasing_sup inf.sup_left_isotone pp_dist_inf by auto finally have 4: "t \ e \ --?G" . have "e \ e * e\<^sup>T * e" by (simp add: ex231c) also have "... \ v * -v\<^sup>T * -v * v\<^sup>T * e" by (metis assms(4) mult_left_isotone conv_isotone conv_dist_comp mult_assoc mult_isotone conv_involutive conv_complement inf.boundedE) also have "... \ v * top * v\<^sup>T * e" by (metis mult_assoc mult_left_isotone mult_right_isotone top.extremum) also have "... = v * r\<^sup>T * t\<^sup>\ * e" using assms(3,5) by (simp add: mult_assoc) also have "... \ v * r\<^sup>T * (t \ e)\<^sup>\" by (simp add: comp_associative mult_right_isotone star.circ_mult_upper_bound star.circ_sub_dist_1 star_isotone sup_commute) also have "... \ v * ?c" using 4 by (simp add: mult_assoc mult_right_isotone star_isotone) also have "... \ ?c\<^sup>T * ?c" using 2 by (simp add: mult_left_isotone) finally show ?thesis using 3 by simp qed lemma span_inv: assumes "e \ v * -v\<^sup>T" and "vector v" and "arc e" and "t \ (v * v\<^sup>T) \ g" and "g\<^sup>T = g" and "v\<^sup>T = r\<^sup>T * t\<^sup>\" and "injective r" and "r\<^sup>T \ v\<^sup>T" and "r\<^sup>T * ((v * v\<^sup>T) \ g)\<^sup>\ \ r\<^sup>T * t\<^sup>\" shows "r\<^sup>T * (((v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T) \ g)\<^sup>\ \ r\<^sup>T * (t \ e)\<^sup>\" proof - let ?d = "(v * v\<^sup>T) \ g" have 1: "(v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T = v * v\<^sup>T \ v * v\<^sup>T * e \ e\<^sup>T * v * v\<^sup>T \ e\<^sup>T * e" using assms(1-3) ve_dist by simp have "t\<^sup>T \ ?d\<^sup>T" using assms(4) conv_isotone by simp also have "... = (v * v\<^sup>T) \ g\<^sup>T" by (simp add: conv_dist_comp conv_dist_inf) also have "... = ?d" by (simp add: assms(5)) finally have 2: "t\<^sup>T \ ?d" . have "v * v\<^sup>T = (r\<^sup>T * t\<^sup>\)\<^sup>T * (r\<^sup>T * t\<^sup>\)" by (metis assms(6) conv_involutive) also have "... = t\<^sup>T\<^sup>\ * (r * r\<^sup>T) * t\<^sup>\" by (simp add: comp_associative conv_dist_comp conv_star_commute) also have "... \ t\<^sup>T\<^sup>\ * 1 * t\<^sup>\" by (simp add: assms(7) mult_left_isotone star_right_induct_mult_iff star_sub_one) also have "... = t\<^sup>T\<^sup>\ * t\<^sup>\" by simp also have "... \ ?d\<^sup>\ * t\<^sup>\" using 2 by (simp add: comp_left_isotone star.circ_isotone) also have "... \ ?d\<^sup>\ * ?d\<^sup>\" using assms(4) mult_right_isotone star_isotone by simp also have 3: "... = ?d\<^sup>\" by (simp add: star.circ_transitive_equal) finally have 4: "v * v\<^sup>T \ ?d\<^sup>\" . have 5: "r\<^sup>T * ?d\<^sup>\ * (v * v\<^sup>T \ g) \ r\<^sup>T * ?d\<^sup>\" by (simp add: comp_associative mult_right_isotone star.circ_plus_same star.left_plus_below_circ) have "r\<^sup>T * ?d\<^sup>\ * (v * v\<^sup>T * e \ g) \ r\<^sup>T * ?d\<^sup>\ * v * v\<^sup>T * e" by (simp add: comp_associative comp_right_isotone) also have "... \ r\<^sup>T * ?d\<^sup>\ * e" using 3 4 by (metis comp_associative comp_isotone eq_refl) finally have 6: "r\<^sup>T * ?d\<^sup>\ * (v * v\<^sup>T * e \ g) \ r\<^sup>T * ?d\<^sup>\ * e" . have 7: "\x . r\<^sup>T * (1 \ v * v\<^sup>T) * e\<^sup>T * x = bot" proof fix x have "r\<^sup>T * (1 \ v * v\<^sup>T) * e\<^sup>T * x \ r\<^sup>T * (1 \ v * v\<^sup>T) * e\<^sup>T * top" by (simp add: mult_right_isotone) also have "... = r\<^sup>T * e\<^sup>T * top \ r\<^sup>T * v * v\<^sup>T * e\<^sup>T * top" by (simp add: comp_associative mult_left_dist_sup mult_right_dist_sup) also have "... = r\<^sup>T * e\<^sup>T * top" by (metis assms(1,2) mult_assoc mult_right_dist_sup mult_right_zero sup_bot_right vTeT) also have "... \ v\<^sup>T * e\<^sup>T * top" by (simp add: assms(8) comp_isotone) also have "... = bot" using vTeT assms(1,2) by simp finally show "r\<^sup>T * (1 \ v * v\<^sup>T) * e\<^sup>T * x = bot" by (simp add: le_bot) qed have "r\<^sup>T * ?d\<^sup>\ * (e\<^sup>T * v * v\<^sup>T \ g) \ r\<^sup>T * ?d\<^sup>\ * e\<^sup>T * v * v\<^sup>T" by (simp add: comp_associative comp_right_isotone) also have "... \ r\<^sup>T * (1 \ v * v\<^sup>T) * e\<^sup>T * v * v\<^sup>T" by (metis assms(2) star.circ_isotone vector_vector_star inf_le1 comp_associative comp_right_isotone comp_left_isotone) also have "... = bot" using 7 by simp finally have 8: "r\<^sup>T * ?d\<^sup>\ * (e\<^sup>T * v * v\<^sup>T \ g) = bot" by (simp add: le_bot) have "r\<^sup>T * ?d\<^sup>\ * (e\<^sup>T * e \ g) \ r\<^sup>T * ?d\<^sup>\ * e\<^sup>T * e" by (simp add: comp_associative comp_right_isotone) also have "... \ r\<^sup>T * (1 \ v * v\<^sup>T) * e\<^sup>T * e" by (metis assms(2) star.circ_isotone vector_vector_star inf_le1 comp_associative comp_right_isotone comp_left_isotone) also have "... = bot" using 7 by simp finally have 9: "r\<^sup>T * ?d\<^sup>\ * (e\<^sup>T * e \ g) = bot" by (simp add: le_bot) have "r\<^sup>T * ?d\<^sup>\ * ((v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T \ g) = r\<^sup>T * ?d\<^sup>\ * ((v * v\<^sup>T \ v * v\<^sup>T * e \ e\<^sup>T * v * v\<^sup>T \ e\<^sup>T * e) \ g)" using 1 by simp also have "... = r\<^sup>T * ?d\<^sup>\ * ((v * v\<^sup>T \ g) \ (v * v\<^sup>T * e \ g) \ (e\<^sup>T * v * v\<^sup>T \ g) \ (e\<^sup>T * e \ g))" by (simp add: inf_sup_distrib2) also have "... = r\<^sup>T * ?d\<^sup>\ * (v * v\<^sup>T \ g) \ r\<^sup>T * ?d\<^sup>\ * (v * v\<^sup>T * e \ g) \ r\<^sup>T * ?d\<^sup>\ * (e\<^sup>T * v * v\<^sup>T \ g) \ r\<^sup>T * ?d\<^sup>\ * (e\<^sup>T * e \ g)" by (simp add: comp_left_dist_sup) also have "... = r\<^sup>T * ?d\<^sup>\ * (v * v\<^sup>T \ g) \ r\<^sup>T * ?d\<^sup>\ * (v * v\<^sup>T * e \ g)" using 8 9 by simp also have "... \ r\<^sup>T * ?d\<^sup>\ \ r\<^sup>T * ?d\<^sup>\ * e" using 5 6 sup.mono by simp also have "... = r\<^sup>T * ?d\<^sup>\ * (1 \ e)" by (simp add: mult_left_dist_sup) finally have 10: "r\<^sup>T * ?d\<^sup>\ * ((v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T \ g) \ r\<^sup>T * ?d\<^sup>\ * (1 \ e)" by simp have "r\<^sup>T * ?d\<^sup>\ * e * (v * v\<^sup>T \ g) \ r\<^sup>T * ?d\<^sup>\ * e * v * v\<^sup>T" by (simp add: comp_associative comp_right_isotone) also have "... = bot" by (metis assms(1,2) comp_associative comp_right_zero ev comp_left_zero) finally have 11: "r\<^sup>T * ?d\<^sup>\ * e * (v * v\<^sup>T \ g) = bot" by (simp add: le_bot) have "r\<^sup>T * ?d\<^sup>\ * e * (v * v\<^sup>T * e \ g) \ r\<^sup>T * ?d\<^sup>\ * e * v * v\<^sup>T * e" by (simp add: comp_associative comp_right_isotone) also have "... = bot" by (metis assms(1,2) comp_associative comp_right_zero ev comp_left_zero) finally have 12: "r\<^sup>T * ?d\<^sup>\ * e * (v * v\<^sup>T * e \ g) = bot" by (simp add: le_bot) have "r\<^sup>T * ?d\<^sup>\ * e * (e\<^sup>T * v * v\<^sup>T \ g) \ r\<^sup>T * ?d\<^sup>\ * e * e\<^sup>T * v * v\<^sup>T" by (simp add: comp_associative comp_right_isotone) also have "... \ r\<^sup>T * ?d\<^sup>\ * 1 * v * v\<^sup>T" by (metis assms(3) arc_injective comp_associative comp_left_isotone comp_right_isotone) also have "... = r\<^sup>T * ?d\<^sup>\ * v * v\<^sup>T" by simp also have "... \ r\<^sup>T * ?d\<^sup>\ * ?d\<^sup>\" using 4 by (simp add: mult_right_isotone mult_assoc) also have "... = r\<^sup>T * ?d\<^sup>\" by (simp add: star.circ_transitive_equal comp_associative) finally have 13: "r\<^sup>T * ?d\<^sup>\ * e * (e\<^sup>T * v * v\<^sup>T \ g) \ r\<^sup>T * ?d\<^sup>\" . have "r\<^sup>T * ?d\<^sup>\ * e * (e\<^sup>T * e \ g) \ r\<^sup>T * ?d\<^sup>\ * e * e\<^sup>T * e" by (simp add: comp_associative comp_right_isotone) also have "... \ r\<^sup>T * ?d\<^sup>\ * 1 * e" by (metis assms(3) arc_injective comp_associative comp_left_isotone comp_right_isotone) also have "... = r\<^sup>T * ?d\<^sup>\ * e" by simp finally have 14: "r\<^sup>T * ?d\<^sup>\ * e * (e\<^sup>T * e \ g) \ r\<^sup>T * ?d\<^sup>\ * e" . have "r\<^sup>T * ?d\<^sup>\ * e * ((v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T \ g) = r\<^sup>T * ?d\<^sup>\ * e * ((v * v\<^sup>T \ v * v\<^sup>T * e \ e\<^sup>T * v * v\<^sup>T \ e\<^sup>T * e) \ g)" using 1 by simp also have "... = r\<^sup>T * ?d\<^sup>\ * e * ((v * v\<^sup>T \ g) \ (v * v\<^sup>T * e \ g) \ (e\<^sup>T * v * v\<^sup>T \ g) \ (e\<^sup>T * e \ g))" by (simp add: inf_sup_distrib2) also have "... = r\<^sup>T * ?d\<^sup>\ * e * (v * v\<^sup>T \ g) \ r\<^sup>T * ?d\<^sup>\ * e * (v * v\<^sup>T * e \ g) \ r\<^sup>T * ?d\<^sup>\ * e * (e\<^sup>T * v * v\<^sup>T \ g) \ r\<^sup>T * ?d\<^sup>\ * e * (e\<^sup>T * e \ g)" by (simp add: comp_left_dist_sup) also have "... = r\<^sup>T * ?d\<^sup>\ * e * (e\<^sup>T * v * v\<^sup>T \ g) \ r\<^sup>T * ?d\<^sup>\ * e * (e\<^sup>T * e \ g)" using 11 12 by simp also have "... \ r\<^sup>T * ?d\<^sup>\ \ r\<^sup>T * ?d\<^sup>\ * e" using 13 14 sup_mono by simp also have "... = r\<^sup>T * ?d\<^sup>\ * (1 \ e)" by (simp add: mult_left_dist_sup) finally have 15: "r\<^sup>T * ?d\<^sup>\ * e * ((v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T \ g) \ r\<^sup>T * ?d\<^sup>\ * (1 \ e)" by simp have "r\<^sup>T \ r\<^sup>T * ?d\<^sup>\" using mult_right_isotone star.circ_reflexive by fastforce also have "... \ r\<^sup>T * ?d\<^sup>\ * (1 \ e)" by (simp add: semiring.distrib_left) finally have 16: "r\<^sup>T \ r\<^sup>T * ?d\<^sup>\ * (1 \ e)" . have "r\<^sup>T * ?d\<^sup>\ * (1 \ e) * ((v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T \ g) = r\<^sup>T * ?d\<^sup>\ * ((v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T \ g) \ r\<^sup>T * ?d\<^sup>\ * e * ((v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T \ g)" by (simp add: semiring.distrib_left semiring.distrib_right) also have "... \ r\<^sup>T * ?d\<^sup>\ * (1 \ e)" using 10 15 le_supI by simp finally have "r\<^sup>T * ?d\<^sup>\ * (1 \ e) * ((v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T \ g) \ r\<^sup>T * ?d\<^sup>\ * (1 \ e)" . hence "r\<^sup>T \ r\<^sup>T * ?d\<^sup>\ * (1 \ e) * ((v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T \ g) \ r\<^sup>T * ?d\<^sup>\ * (1 \ e)" using 16 sup_least by simp hence "r\<^sup>T * ((v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T \ g)\<^sup>\ \ r\<^sup>T * ?d\<^sup>\ * (1 \ e)" by (simp add: star_right_induct) also have "... \ r\<^sup>T * t\<^sup>\ * (1 \ e)" by (simp add: assms(9) mult_left_isotone) also have "... \ r\<^sup>T * (t \ e)\<^sup>\" by (simp add: star_one_sup_below) finally show ?thesis . qed subsubsection \Exchange gives Spanning Trees\ text \ The following abbreviations are used in the spanning tree application using Prim's algorithm to construct the new tree for the exchange property. It is obtained by replacing an edge with one that has minimal weight and reversing the path connecting these edges. Here, w represents a weighted graph, v represents a set of nodes and e represents an edge. \ abbreviation prim_E :: "'a \ 'a \ 'a \ 'a" where "prim_E w v e \ w \ --v * -v\<^sup>T \ top * e * w\<^sup>T\<^sup>\" abbreviation prim_P :: "'a \ 'a \ 'a \ 'a" where "prim_P w v e \ w \ -v * -v\<^sup>T \ top * e * w\<^sup>T\<^sup>\" abbreviation prim_EP :: "'a \ 'a \ 'a \ 'a" where "prim_EP w v e \ w \ -v\<^sup>T \ top * e * w\<^sup>T\<^sup>\" abbreviation prim_W :: "'a \ 'a \ 'a \ 'a" where "prim_W w v e \ (w \ -(prim_EP w v e)) \ (prim_P w v e)\<^sup>T \ e" text \ The lemmas in this section are used to show that the relation after exchange represents a spanning tree. The results in the next section are used to show that it is a minimum spanning tree. \ lemma exchange_injective_3: assumes "e \ v * -v\<^sup>T" and "vector v" shows "(w \ -(prim_EP w v e)) * e\<^sup>T = bot" proof - have 1: "top * e \ -v\<^sup>T" by (simp add: assms schroeder_4_p vTeT) have "top * e \ top * e * w\<^sup>T\<^sup>\" using sup_right_divisibility star.circ_back_loop_fixpoint by blast hence "top * e \ -v\<^sup>T \ top * e * w\<^sup>T\<^sup>\" using 1 by simp hence "top * e \ -(w \ -prim_EP w v e)" by (metis inf.assoc inf_import_p le_infI2 p_antitone p_antitone_iff) hence "(w \ -(prim_EP w v e)) * e\<^sup>T \ bot" using p_top schroeder_4_p by blast thus ?thesis using le_bot by simp qed lemma exchange_injective_6: assumes "arc e" and "forest w" shows "(prim_P w v e)\<^sup>T * e\<^sup>T = bot" proof - have "e\<^sup>T * top * e \ --1" - by (simp add: assms(1) local.p_antitone local.p_antitone_iff local.point_injective) + by (simp add: assms(1) p_antitone p_antitone_iff point_injective) hence 1: "e * -1 * e\<^sup>T \ bot" by (metis conv_involutive p_top triple_schroeder_p) have "(prim_P w v e)\<^sup>T * e\<^sup>T \ (w \ top * e * w\<^sup>T\<^sup>\)\<^sup>T * e\<^sup>T" using comp_inf.mult_left_isotone conv_dist_inf mult_left_isotone by simp also have "... = (w\<^sup>T \ w\<^sup>T\<^sup>\\<^sup>T * e\<^sup>T * top) * e\<^sup>T" by (simp add: comp_associative conv_dist_comp conv_dist_inf) also have "... = w\<^sup>\ * e\<^sup>T * top \ w\<^sup>T * e\<^sup>T" by (simp add: conv_star_commute inf_vector_comp) also have "... \ (w\<^sup>T \ w\<^sup>\ * e\<^sup>T * top * e) * (e\<^sup>T \ w\<^sup>+ * e\<^sup>T * top)" by (metis dedekind mult_assoc conv_involutive inf_commute) also have "... \ (w\<^sup>\ * e\<^sup>T * top * e) * (w\<^sup>+ * e\<^sup>T * top)" by (simp add: mult_isotone) also have "... \ (top * e) * (w\<^sup>+ * e\<^sup>T * top)" by (simp add: mult_left_isotone) also have "... = top * e * w\<^sup>+ * e\<^sup>T * top" using mult_assoc by simp also have "... \ top * e * -1 * e\<^sup>T * top" using assms(2) mult_left_isotone mult_right_isotone by simp also have "... \ bot" using 1 by (metis le_bot semiring.mult_not_zero mult_assoc) finally show ?thesis using le_bot by simp qed text \ The graph after exchanging is injective. \ lemma exchange_injective: assumes "arc e" and "e \ v * -v\<^sup>T" and "forest w" and "vector v" shows "injective (prim_W w v e)" proof - have 1: "(w \ -(prim_EP w v e)) * (w \ -(prim_EP w v e))\<^sup>T \ 1" proof - have "(w \ -(prim_EP w v e)) * (w \ -(prim_EP w v e))\<^sup>T \ w * w\<^sup>T" by (simp add: comp_isotone conv_isotone) also have "... \ 1" by (simp add: assms(3)) finally show ?thesis . qed have 2: "(w \ -(prim_EP w v e)) * (prim_P w v e)\<^sup>T\<^sup>T \ 1" proof - have "top * (prim_P w v e)\<^sup>T = top * (w\<^sup>T \ -v * -v\<^sup>T \ w\<^sup>T\<^sup>\\<^sup>T * e\<^sup>T * top)" by (simp add: comp_associative conv_complement conv_dist_comp conv_dist_inf) also have "... = top * e * w\<^sup>T\<^sup>\ * (w\<^sup>T \ -v * -v\<^sup>T)" by (metis comp_inf_vector conv_dist_comp conv_involutive inf_top_left mult_assoc) also have "... \ top * e * w\<^sup>T\<^sup>\ * (w\<^sup>T \ top * -v\<^sup>T)" using comp_inf.mult_right_isotone mult_left_isotone mult_right_isotone by simp also have "... = top * e * w\<^sup>T\<^sup>\ * w\<^sup>T \ -v\<^sup>T" by (metis assms(4) comp_inf_covector vector_conv_compl) also have "... \ -v\<^sup>T \ top * e * w\<^sup>T\<^sup>\" by (simp add: comp_associative comp_isotone inf.coboundedI1 star.circ_plus_same star.left_plus_below_circ) finally have "top * (prim_P w v e)\<^sup>T \ -(w \ -prim_EP w v e)" by (metis inf.assoc inf_import_p le_infI2 p_antitone p_antitone_iff) hence "(w \ -(prim_EP w v e)) * (prim_P w v e)\<^sup>T\<^sup>T \ bot" using p_top schroeder_4_p by blast thus ?thesis by (simp add: bot_unique) qed have 3: "(w \ -(prim_EP w v e)) * e\<^sup>T \ 1" by (metis assms(2,4) exchange_injective_3 bot_least) have 4: "(prim_P w v e)\<^sup>T * (w \ -(prim_EP w v e))\<^sup>T \ 1" using 2 conv_dist_comp coreflexive_symmetric by fastforce have 5: "(prim_P w v e)\<^sup>T * (prim_P w v e)\<^sup>T\<^sup>T \ 1" proof - have "(prim_P w v e)\<^sup>T * (prim_P w v e)\<^sup>T\<^sup>T \ (top * e * w\<^sup>T\<^sup>\)\<^sup>T * (top * e * w\<^sup>T\<^sup>\)" by (simp add: conv_dist_inf mult_isotone) also have "... = w\<^sup>\ * e\<^sup>T * top * top * e * w\<^sup>T\<^sup>\" using conv_star_commute conv_dist_comp conv_involutive conv_top mult_assoc by presburger also have "... = w\<^sup>\ * e\<^sup>T * top * e * w\<^sup>T\<^sup>\" by (simp add: comp_associative) also have "... \ w\<^sup>\ * 1 * w\<^sup>T\<^sup>\" by (metis comp_left_isotone comp_right_isotone mult_assoc assms(1) point_injective) finally have "(prim_P w v e)\<^sup>T * (prim_P w v e)\<^sup>T\<^sup>T \ w\<^sup>\ * w\<^sup>T\<^sup>\ \ w\<^sup>T * w" by (simp add: conv_isotone inf.left_commute inf.sup_monoid.add_commute mult_isotone) also have "... \ 1" by (simp add: assms(3) forest_separate) finally show ?thesis . qed have 6: "(prim_P w v e)\<^sup>T * e\<^sup>T \ 1" using assms exchange_injective_6 bot_least by simp have 7: "e * (w \ -(prim_EP w v e))\<^sup>T \ 1" using 3 by (metis conv_dist_comp conv_involutive coreflexive_symmetric) have 8: "e * (prim_P w v e)\<^sup>T\<^sup>T \ 1" using 6 conv_dist_comp coreflexive_symmetric by fastforce have 9: "e * e\<^sup>T \ 1" by (simp add: assms(1) arc_injective) have "(prim_W w v e) * (prim_W w v e)\<^sup>T = (w \ -(prim_EP w v e)) * (w \ -(prim_EP w v e))\<^sup>T \ (w \ -(prim_EP w v e)) * (prim_P w v e)\<^sup>T\<^sup>T \ (w \ -(prim_EP w v e)) * e\<^sup>T \ (prim_P w v e)\<^sup>T * (w \ -(prim_EP w v e))\<^sup>T \ (prim_P w v e)\<^sup>T * (prim_P w v e)\<^sup>T\<^sup>T \ (prim_P w v e)\<^sup>T * e\<^sup>T \ e * (w \ -(prim_EP w v e))\<^sup>T \ e * (prim_P w v e)\<^sup>T\<^sup>T \ e * e\<^sup>T" using comp_left_dist_sup comp_right_dist_sup conv_dist_sup sup.assoc by simp also have "... \ 1" using 1 2 3 4 5 6 7 8 9 by simp finally show ?thesis . qed lemma pv: assumes "vector v" shows "(prim_P w v e)\<^sup>T * v = bot" proof - have "(prim_P w v e)\<^sup>T * v \ (-v * -v\<^sup>T)\<^sup>T * v" by (meson conv_isotone inf_le1 inf_le2 mult_left_isotone order_trans) also have "... = -v * -v\<^sup>T * v" by (simp add: conv_complement conv_dist_comp) also have "... = bot" by (simp add: assms covector_vector_comp mult_assoc) finally show ?thesis by (simp add: antisym) qed lemma vector_pred_inv: assumes "arc e" and "e \ v * -v\<^sup>T" and "forest w" and "vector v" and "w * v \ v" shows "(prim_W w v e) * (v \ e\<^sup>T * top) \ v \ e\<^sup>T * top" proof - have "(prim_W w v e) * e\<^sup>T * top = (w \ -(prim_EP w v e)) * e\<^sup>T * top \ (prim_P w v e)\<^sup>T * e\<^sup>T * top \ e * e\<^sup>T * top" by (simp add: mult_right_dist_sup) also have "... = e * e\<^sup>T * top" using assms exchange_injective_3 exchange_injective_6 comp_left_zero by simp also have "... \ v * -v\<^sup>T * e\<^sup>T * top" by (simp add: assms(2) comp_isotone) also have "... \ v * top" by (simp add: comp_associative mult_right_isotone) also have "... = v" by (simp add: assms(4)) finally have 1: "(prim_W w v e) * e\<^sup>T * top \ v" . have "(prim_W w v e) * v = (w \ -(prim_EP w v e)) * v \ (prim_P w v e)\<^sup>T * v \ e * v" by (simp add: mult_right_dist_sup) also have "... = (w \ -(prim_EP w v e)) * v" by (metis assms(2,4) pv ev sup_bot_right) also have "... \ w * v" by (simp add: mult_left_isotone) finally have 2: "(prim_W w v e) * v \ v" using assms(5) order_trans by blast have "(prim_W w v e) * (v \ e\<^sup>T * top) = (prim_W w v e) * v \ (prim_W w v e) * e\<^sup>T * top" by (simp add: semiring.distrib_left mult_assoc) also have "... \ v" using 1 2 by simp also have "... \ v \ e\<^sup>T * top" by simp finally show ?thesis . qed text \ The graph after exchanging is acyclic. \ lemma exchange_acyclic: assumes "vector v" and "e \ v * -v\<^sup>T" and "w * v \ v" and "acyclic w" shows "acyclic (prim_W w v e)" proof - have 1: "(prim_P w v e)\<^sup>T * e = bot" proof - have "(prim_P w v e)\<^sup>T * e \ (-v * -v\<^sup>T)\<^sup>T * e" by (meson conv_order dual_order.trans inf.cobounded1 inf.cobounded2 mult_left_isotone) also have "... = -v * -v\<^sup>T * e" by (simp add: conv_complement conv_dist_comp) also have "... \ -v * -v\<^sup>T * v * -v\<^sup>T" by (simp add: assms(2) comp_associative mult_right_isotone) also have "... = bot" by (simp add: assms(1) covector_vector_comp mult_assoc) finally show ?thesis by (simp add: bot_unique) qed have 2: "e * e = bot" using assms(1,2) ee by auto have 3: "(w \ -(prim_EP w v e)) * (prim_P w v e)\<^sup>T = bot" proof - have "top * (prim_P w v e) \ top * (-v * -v\<^sup>T \ top * e * w\<^sup>T\<^sup>\)" using comp_inf.mult_semi_associative mult_right_isotone by auto also have "... \ top * -v * -v\<^sup>T \ top * top * e * w\<^sup>T\<^sup>\" by (simp add: comp_inf_covector mult_assoc) also have "... \ top * -v\<^sup>T \ top * e * w\<^sup>T\<^sup>\" - using mult_left_isotone top.extremum local.inf_mono by presburger + using mult_left_isotone top.extremum inf_mono by presburger also have "... = -v\<^sup>T \ top * e * w\<^sup>T\<^sup>\" by (simp add: assms(1) vector_conv_compl) finally have "top * (prim_P w v e) \ -(w \ -prim_EP w v e)" by (metis inf.assoc inf_import_p le_infI2 p_antitone p_antitone_iff) hence "(w \ -(prim_EP w v e)) * (prim_P w v e)\<^sup>T \ bot" using p_top schroeder_4_p by blast thus ?thesis using bot_unique by blast qed hence 4: "(w \ -(prim_EP w v e)) * (prim_P w v e)\<^sup>T\<^sup>\ = w \ -(prim_EP w v e)" using star_absorb by blast hence 5: "(w \ -(prim_EP w v e))\<^sup>+ * (prim_P w v e)\<^sup>T\<^sup>\ = (w \ -(prim_EP w v e))\<^sup>+" by (metis star_plus mult_assoc) hence 6: "(w \ -(prim_EP w v e))\<^sup>\ * (prim_P w v e)\<^sup>T\<^sup>\ = (w \ -(prim_EP w v e))\<^sup>+ \ (prim_P w v e)\<^sup>T\<^sup>\" by (metis star.circ_loop_fixpoint mult_assoc) have 7: "(w \ -(prim_EP w v e))\<^sup>+ * e \ v * top" proof - have "e \ v * top" using assms(2) dual_order.trans mult_right_isotone top_greatest by blast hence 8: "e \ w * v * top \ v * top" by (simp add: assms(1,3) comp_associative) have "(w \ -(prim_EP w v e))\<^sup>+ * e \ w\<^sup>+ * e" by (simp add: comp_isotone star_isotone) also have "... \ w\<^sup>\ * e" by (simp add: mult_left_isotone star.left_plus_below_circ) also have "... \ v * top" using 8 by (simp add: comp_associative star_left_induct) finally show ?thesis . qed have 9: "(prim_P w v e)\<^sup>T * (w \ -(prim_EP w v e))\<^sup>+ * e = bot" proof - have "(prim_P w v e)\<^sup>T * (w \ -(prim_EP w v e))\<^sup>+ * e \ (prim_P w v e)\<^sup>T * v * top" using 7 by (simp add: mult_assoc mult_right_isotone) also have "... = bot" by (simp add: assms(1) pv) finally show ?thesis using bot_unique by blast qed have 10: "e * (w \ -(prim_EP w v e))\<^sup>+ * e = bot" proof - have "e * (w \ -(prim_EP w v e))\<^sup>+ * e \ e * v * top" using 7 by (simp add: mult_assoc mult_right_isotone) also have "... \ v * -v\<^sup>T * v * top" by (simp add: assms(2) mult_left_isotone) also have "... = bot" by (simp add: assms(1) covector_vector_comp mult_assoc) finally show ?thesis using bot_unique by blast qed have 11: "e * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\ \ v * -v\<^sup>T" proof - have 12: "-v\<^sup>T * w \ -v\<^sup>T" by (metis assms(3) conv_complement order_lesseq_imp pp_increasing schroeder_6_p) have "v * -v\<^sup>T * (w \ -(prim_EP w v e)) \ v * -v\<^sup>T * w" by (simp add: comp_isotone star_isotone) also have "... \ v * -v\<^sup>T" using 12 by (simp add: comp_isotone comp_associative) finally have 13: "v * -v\<^sup>T * (w \ -(prim_EP w v e)) \ v * -v\<^sup>T" . have 14: "(prim_P w v e)\<^sup>T \ -v * -v\<^sup>T" by (metis conv_complement conv_dist_comp conv_involutive conv_order inf_le1 inf_le2 order_trans) have "e * (prim_P w v e)\<^sup>T\<^sup>\ \ v * -v\<^sup>T * (prim_P w v e)\<^sup>T\<^sup>\" by (simp add: assms(2) mult_left_isotone) also have "... = v * -v\<^sup>T \ v * -v\<^sup>T * (prim_P w v e)\<^sup>T\<^sup>+" by (metis mult_assoc star.circ_back_loop_fixpoint star_plus sup_commute) also have "... = v * -v\<^sup>T \ v * -v\<^sup>T * (prim_P w v e)\<^sup>T\<^sup>\ * (prim_P w v e)\<^sup>T" by (simp add: mult_assoc star_plus) also have "... \ v * -v\<^sup>T \ v * -v\<^sup>T * (prim_P w v e)\<^sup>T\<^sup>\ * -v * -v\<^sup>T" using 14 mult_assoc mult_right_isotone sup_right_isotone by simp also have "... \ v * -v\<^sup>T \ v * top * -v\<^sup>T" by (metis top_greatest mult_right_isotone mult_left_isotone mult_assoc sup_right_isotone) also have "... = v * -v\<^sup>T" by (simp add: assms(1)) finally have "e * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\ \ v * -v\<^sup>T * (w \ -(prim_EP w v e))\<^sup>\" by (simp add: mult_left_isotone) also have "... \ v * -v\<^sup>T" using 13 by (simp add: star_right_induct_mult) finally show ?thesis . qed have 15: "(w \ -(prim_EP w v e))\<^sup>+ * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\ \ -1" proof - have "(w \ -(prim_EP w v e))\<^sup>+ * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\ = (w \ -(prim_EP w v e))\<^sup>+ * (w \ -(prim_EP w v e))\<^sup>\" using 5 by simp also have "... = (w \ -(prim_EP w v e))\<^sup>+" by (simp add: mult_assoc star.circ_transitive_equal) also have "... \ w\<^sup>+" by (simp add: comp_isotone star_isotone) finally show ?thesis using assms(4) by simp qed have 16: "(prim_P w v e)\<^sup>T * (w \ -(prim_EP w v e))\<^sup>\ * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\ \ -1" proof - have "(w \ -(prim_EP w v e))\<^sup>+ * (prim_P w v e)\<^sup>T\<^sup>+ \ (w \ -(prim_EP w v e))\<^sup>+ * (prim_P w v e)\<^sup>T\<^sup>\" by (simp add: mult_right_isotone star.left_plus_below_circ) also have "... = (w \ -(prim_EP w v e))\<^sup>+" using 5 by simp also have "... \ w\<^sup>+" by (simp add: comp_isotone star_isotone) finally have "(w \ -(prim_EP w v e))\<^sup>+ * (prim_P w v e)\<^sup>T\<^sup>+ \ -1" using assms(4) by simp hence 17: "(prim_P w v e)\<^sup>T\<^sup>+ * (w \ -(prim_EP w v e))\<^sup>+ \ -1" by (simp add: comp_commute_below_diversity) have "(prim_P w v e)\<^sup>T\<^sup>+ \ w\<^sup>T\<^sup>+" by (simp add: comp_isotone conv_dist_inf inf.left_commute inf.sup_monoid.add_commute star_isotone) also have "... = w\<^sup>+\<^sup>T" by (simp add: conv_dist_comp conv_star_commute star_plus) also have "... \ -1" using assms(4) conv_complement conv_isotone by force finally have 18: "(prim_P w v e)\<^sup>T\<^sup>+ \ -1" . have "(prim_P w v e)\<^sup>T * (w \ -(prim_EP w v e))\<^sup>\ * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\ = (prim_P w v e)\<^sup>T * ((w \ -(prim_EP w v e))\<^sup>+ \ (prim_P w v e)\<^sup>T\<^sup>\) * (w \ -(prim_EP w v e))\<^sup>\" using 6 by (simp add: comp_associative) also have "... = (prim_P w v e)\<^sup>T * (w \ -(prim_EP w v e))\<^sup>+ * (w \ -(prim_EP w v e))\<^sup>\ \ (prim_P w v e)\<^sup>T\<^sup>+ * (w \ -(prim_EP w v e))\<^sup>\" by (simp add: mult_left_dist_sup mult_right_dist_sup) also have "... = (prim_P w v e)\<^sup>T * (w \ -(prim_EP w v e))\<^sup>+ \ (prim_P w v e)\<^sup>T\<^sup>+ * (w \ -(prim_EP w v e))\<^sup>\" by (simp add: mult_assoc star.circ_transitive_equal) also have "... = (prim_P w v e)\<^sup>T * (w \ -(prim_EP w v e))\<^sup>+ \ (prim_P w v e)\<^sup>T\<^sup>+ * (1 \ (w \ -(prim_EP w v e))\<^sup>+)" using star_left_unfold_equal by simp also have "... = (prim_P w v e)\<^sup>T * (w \ -(prim_EP w v e))\<^sup>+ \ (prim_P w v e)\<^sup>T\<^sup>+ * (w \ -(prim_EP w v e))\<^sup>+ \ (prim_P w v e)\<^sup>T\<^sup>+" by (simp add: mult_left_dist_sup sup.left_commute sup_commute) also have "... = ((prim_P w v e)\<^sup>T \ (prim_P w v e)\<^sup>T\<^sup>+) * (w \ -(prim_EP w v e))\<^sup>+ \ (prim_P w v e)\<^sup>T\<^sup>+" by (simp add: mult_right_dist_sup) also have "... = (prim_P w v e)\<^sup>T\<^sup>+ * (w \ -(prim_EP w v e))\<^sup>+ \ (prim_P w v e)\<^sup>T\<^sup>+" using star.circ_mult_increasing by (simp add: le_iff_sup) also have "... \ -1" using 17 18 by simp finally show ?thesis . qed have 19: "e * (w \ -(prim_EP w v e))\<^sup>\ * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\ \ -1" proof - have "e * (w \ -(prim_EP w v e))\<^sup>\ * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\ = e * ((w \ -(prim_EP w v e))\<^sup>+ \ (prim_P w v e)\<^sup>T\<^sup>\) * (w \ -(prim_EP w v e))\<^sup>\" using 6 by (simp add: mult_assoc) also have "... = e * (w \ -(prim_EP w v e))\<^sup>+ * (w \ -(prim_EP w v e))\<^sup>\ \ e * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\" by (simp add: mult_left_dist_sup mult_right_dist_sup) also have "... = e * (w \ -(prim_EP w v e))\<^sup>+ \ e * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\" by (simp add: mult_assoc star.circ_transitive_equal) also have "... \ e * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>+ \ e * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\" by (metis mult_right_sub_dist_sup_right semiring.add_right_mono star.circ_back_loop_fixpoint) also have "... \ e * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\" using mult_right_isotone star.left_plus_below_circ by auto also have "... \ v * -v\<^sup>T" using 11 by simp also have "... \ -1" by (simp add: pp_increasing schroeder_3_p) finally show ?thesis . qed have 20: "(prim_W w v e) * (w \ -(prim_EP w v e))\<^sup>\ * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\ \ -1" using 15 16 19 by (simp add: comp_right_dist_sup) have 21: "(w \ -(prim_EP w v e))\<^sup>+ * e * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\ \ -1" proof - have "(w \ -(prim_EP w v e)) * v * -v\<^sup>T \ w * v * -v\<^sup>T" by (simp add: comp_isotone star_isotone) also have "... \ v * -v\<^sup>T" by (simp add: assms(3) mult_left_isotone) finally have 22: "(w \ -(prim_EP w v e)) * v * -v\<^sup>T \ v * -v\<^sup>T" . have "(w \ -(prim_EP w v e))\<^sup>+ * e * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\ \ (w \ -(prim_EP w v e))\<^sup>+ * v * -v\<^sup>T" using 11 by (simp add: mult_right_isotone mult_assoc) also have "... \ (w \ -(prim_EP w v e))\<^sup>\ * v * -v\<^sup>T" using mult_left_isotone star.left_plus_below_circ by blast also have "... \ v * -v\<^sup>T" using 22 by (simp add: star_left_induct_mult mult_assoc) also have "... \ -1" by (simp add: pp_increasing schroeder_3_p) finally show ?thesis . qed have 23: "(prim_P w v e)\<^sup>T * (w \ -(prim_EP w v e))\<^sup>\ * e * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\ \ -1" proof - have "(prim_P w v e)\<^sup>T * (w \ -(prim_EP w v e))\<^sup>\ * e = (prim_P w v e)\<^sup>T * e \ (prim_P w v e)\<^sup>T * (w \ -(prim_EP w v e))\<^sup>+ * e" using comp_left_dist_sup mult_assoc star.circ_loop_fixpoint sup_commute by auto also have "... = bot" using 1 9 by simp finally show ?thesis by simp qed have 24: "e * (w \ -(prim_EP w v e))\<^sup>\ * e * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\ \ -1" proof - have "e * (w \ -(prim_EP w v e))\<^sup>\ * e = e * e \ e * (w \ -(prim_EP w v e))\<^sup>+ * e" using comp_left_dist_sup mult_assoc star.circ_loop_fixpoint sup_commute by auto also have "... = bot" using 2 10 by simp finally show ?thesis by simp qed have 25: "(prim_W w v e) * (w \ -(prim_EP w v e))\<^sup>\ * e * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\ \ -1" using 21 23 24 by (simp add: comp_right_dist_sup) have "(prim_W w v e)\<^sup>\ = ((prim_P w v e)\<^sup>T \ e)\<^sup>\ * ((w \ -(prim_EP w v e)) * ((prim_P w v e)\<^sup>T \ e)\<^sup>\)\<^sup>\" by (metis star_sup_1 sup.left_commute sup_commute) also have "... = ((prim_P w v e)\<^sup>T\<^sup>\ \ e * (prim_P w v e)\<^sup>T\<^sup>\) * ((w \ -(prim_EP w v e)) * ((prim_P w v e)\<^sup>T\<^sup>\ \ e * (prim_P w v e)\<^sup>T\<^sup>\))\<^sup>\" using 1 2 star_separate by auto also have "... = ((prim_P w v e)\<^sup>T\<^sup>\ \ e * (prim_P w v e)\<^sup>T\<^sup>\) * ((w \ -(prim_EP w v e)) * (1 \ e * (prim_P w v e)\<^sup>T\<^sup>\))\<^sup>\" using 4 mult_left_dist_sup by auto also have "... = (w \ -(prim_EP w v e))\<^sup>\ * ((prim_P w v e)\<^sup>T\<^sup>\ \ e * (prim_P w v e)\<^sup>T\<^sup>\) * (w \ -(prim_EP w v e))\<^sup>\" using 3 9 10 star_separate_2 by blast also have "... = (w \ -(prim_EP w v e))\<^sup>\ * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\ \ (w \ -(prim_EP w v e))\<^sup>\ * e * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\" by (simp add: semiring.distrib_left semiring.distrib_right mult_assoc) finally have "(prim_W w v e)\<^sup>+ = (prim_W w v e) * ((w \ -(prim_EP w v e))\<^sup>\ * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\ \ (w \ -(prim_EP w v e))\<^sup>\ * e * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\)" by simp also have "... = (prim_W w v e) * (w \ -(prim_EP w v e))\<^sup>\ * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\ \ (prim_W w v e) * (w \ -(prim_EP w v e))\<^sup>\ * e * (prim_P w v e)\<^sup>T\<^sup>\ * (w \ -(prim_EP w v e))\<^sup>\" by (simp add: comp_left_dist_sup comp_associative) also have "... \ -1" using 20 25 by simp finally show ?thesis . qed text \ The following lemma shows that an edge across the cut between visited nodes and unvisited nodes does not leave the component of visited nodes. \ lemma mst_subgraph_inv: assumes "e \ v * -v\<^sup>T \ g" and "t \ g" and "v\<^sup>T = r\<^sup>T * t\<^sup>\" shows "e \ (r\<^sup>T * g\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\) \ g" proof - have "e \ v * -v\<^sup>T \ g" by (rule assms(1)) also have "... \ v * (-v\<^sup>T \ v\<^sup>T * g) \ g" by (simp add: dedekind_1) also have "... \ v * v\<^sup>T * g \ g" by (simp add: comp_associative comp_right_isotone inf_commute le_infI2) also have "... = v * (r\<^sup>T * t\<^sup>\) * g \ g" by (simp add: assms(3)) also have "... = (r\<^sup>T * t\<^sup>\)\<^sup>T * (r\<^sup>T * t\<^sup>\) * g \ g" by (metis assms(3) conv_involutive) also have "... \ (r\<^sup>T * t\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\) * g \ g" using assms(2) comp_inf.mult_left_isotone comp_isotone star_isotone by auto also have "... \ (r\<^sup>T * t\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\) \ g" using inf.sup_right_isotone inf_commute mult_assoc mult_right_isotone star.left_plus_below_circ star_plus by presburger also have "... \ (r\<^sup>T * g\<^sup>\)\<^sup>T * (r\<^sup>T * g\<^sup>\) \ g" using assms(2) comp_inf.mult_left_isotone conv_dist_comp conv_isotone mult_left_isotone star_isotone by auto finally show ?thesis . qed text \ The following lemmas show that the tree after exchanging contains the currently constructed and tree and its extension by the chosen edge. \ lemma mst_extends_old_tree: assumes "t \ w" and "t \ v * v\<^sup>T" and "vector v" shows "t \ prim_W w v e" proof - have "t \ prim_EP w v e \ t \ -v\<^sup>T" by (simp add: inf.coboundedI2 inf.sup_monoid.add_assoc) also have "... \ v * v\<^sup>T \ -v\<^sup>T" by (simp add: assms(2) inf.coboundedI1) also have "... \ bot" by (simp add: assms(3) covector_vector_comp eq_refl schroeder_2) finally have "t \ -(prim_EP w v e)" using le_bot pseudo_complement by blast hence "t \ w \ -(prim_EP w v e)" using assms(1) by simp thus ?thesis - using local.le_supI1 by blast + using le_supI1 by blast qed lemma mst_extends_new_tree: "t \ w \ t \ v * v\<^sup>T \ vector v \ t \ e \ prim_W w v e" using mst_extends_old_tree by auto +lemma forests_bot_1: + assumes "equivalence e" + and "forest f" + shows "(-e \ f) * (e \ f)\<^sup>T = bot" +proof - + have "f * f\<^sup>T \ e" + using assms dual_order.trans by blast + hence "f * (e \ f)\<^sup>T \ e" + by (metis conv_dist_inf inf.boundedE inf.cobounded2 inf.orderE mult_right_isotone) + hence "-e \ f * (e \ f)\<^sup>T = bot" + by (simp add: p_antitone pseudo_complement) + thus ?thesis + by (metis assms(1) comp_isotone conv_dist_inf equivalence_comp_right_complement inf.boundedI inf.cobounded1 inf.cobounded2 le_bot) +qed + +lemma forests_bot_2: + assumes "equivalence e" + and "forest f" + shows "(-e \ f\<^sup>T) * x \ (e \ f\<^sup>T) * y = bot" +proof - + have "(-e \ f) * (e \ f\<^sup>T) = bot" + using assms forests_bot_1 conv_dist_inf by simp + thus ?thesis + by (smt assms(1) comp_associative comp_inf.semiring.mult_not_zero conv_complement conv_dist_comp conv_dist_inf conv_involutive dedekind_1 inf.cobounded2 inf.sup_monoid.add_commute le_bot mult_right_zero p_antitone_iff pseudo_complement semiring.mult_not_zero symmetric_top_closed top.extremum) +qed + +lemma forests_bot_3: + assumes "equivalence e" + and "forest f" + shows "x * (-e \ f) \ y * (e \ f) = bot" +proof - + have "(e \ f) * (-e \ f\<^sup>T) = bot" + using assms forests_bot_1 conv_dist_inf conv_complement by (smt conv_dist_comp conv_involutive conv_order coreflexive_bot_closed coreflexive_symmetric) + hence "y * (e \ f) * (-e \ f\<^sup>T) = bot" + by (simp add: comp_associative) + hence 1: "x \ y * (e \ f) * (-e \ f\<^sup>T) = bot" + using comp_inf.semiring.mult_not_zero by blast + hence "(x \ y * (e \ f) * (-e \ f\<^sup>T)) * (-e \ f) = bot" + using semiring.mult_not_zero by blast + hence "x * (-e \ f\<^sup>T)\<^sup>T \ y * (e \ f) = bot" + using 1 dedekind_2 inf_commute schroeder_2 by auto + thus ?thesis + by (simp add: assms(1) conv_complement conv_dist_inf) +qed + end text \ We finally add the Kleene star to Stone relation algebras. Kleene star and the relational operations are reasonably independent. The only additional axiom we need in the generalisation to Stone-Kleene relation algebras is that star distributes over double complement. \ class stone_kleene_relation_algebra = stone_relation_algebra + pd_kleene_allegory + assumes pp_dist_star: "--(x\<^sup>\) = (--x)\<^sup>\" begin lemma regular_closed_star: "regular x \ regular (x\<^sup>\)" by (simp add: pp_dist_star) lemma components_idempotent: "components (components x) = components x" using pp_dist_star star_involutive by auto +lemma fc_comp_eq_fc: + "-forest_components (--f) = -forest_components f" + by (metis conv_complement p_comp_pp p_pp_comp pp_dist_star) + text \ The following lemma shows that the nodes reachable in the tree after exchange contain the nodes reachable in the tree before exchange. \ lemma mst_reachable_inv: assumes "regular (prim_EP w v e)" and "vector r" and "e \ v * -v\<^sup>T" and "vector v" and "v\<^sup>T = r\<^sup>T * t\<^sup>\" and "t \ w" and "t \ v * v\<^sup>T" and "w * v \ v" shows "r\<^sup>T * w\<^sup>\ \ r\<^sup>T * (prim_W w v e)\<^sup>\" proof - have 1: "r\<^sup>T \ r\<^sup>T * (prim_W w v e)\<^sup>\" using sup.bounded_iff star.circ_back_loop_prefixpoint by blast have "top * e * (w\<^sup>T \ -v\<^sup>T)\<^sup>\ * w\<^sup>T \ -v\<^sup>T = top * e * (w\<^sup>T \ -v\<^sup>T)\<^sup>\ * (w\<^sup>T \ -v\<^sup>T)" by (simp add: assms(4) covector_comp_inf vector_conv_compl) also have "... \ top * e * (w\<^sup>T \ -v\<^sup>T)\<^sup>\" by (simp add: comp_isotone mult_assoc star.circ_plus_same star.left_plus_below_circ) finally have 2: "top * e * (w\<^sup>T \ -v\<^sup>T)\<^sup>\ * w\<^sup>T \ top * e * (w\<^sup>T \ -v\<^sup>T)\<^sup>\ \ --v\<^sup>T" by (simp add: shunting_var_p) have 3: "--v\<^sup>T * w\<^sup>T \ top * e * (w\<^sup>T \ -v\<^sup>T)\<^sup>\ \ --v\<^sup>T" by (metis assms(8) conv_dist_comp conv_order mult_assoc order.trans pp_comp_semi_commute pp_isotone sup.coboundedI1 sup_commute) have 4: "top * e \ top * e * (w\<^sup>T \ -v\<^sup>T)\<^sup>\ \ --v\<^sup>T" using sup_right_divisibility star.circ_back_loop_fixpoint le_supI1 by blast have "(top * e * (w\<^sup>T \ -v\<^sup>T)\<^sup>\ \ --v\<^sup>T) * w\<^sup>T = top * e * (w\<^sup>T \ -v\<^sup>T)\<^sup>\ * w\<^sup>T \ --v\<^sup>T * w\<^sup>T" by (simp add: comp_right_dist_sup) also have "... \ top * e * (w\<^sup>T \ -v\<^sup>T)\<^sup>\ \ --v\<^sup>T" using 2 3 by simp finally have "top * e \ (top * e * (w\<^sup>T \ -v\<^sup>T)\<^sup>\ \ --v\<^sup>T) * w\<^sup>T \ top * e * (w\<^sup>T \ -v\<^sup>T)\<^sup>\ \ --v\<^sup>T" using 4 by simp hence 5: "top * e * w\<^sup>T\<^sup>\ \ top * e * (w\<^sup>T \ -v\<^sup>T)\<^sup>\ \ --v\<^sup>T" by (simp add: star_right_induct) have 6: "top * e \ top * e * (w\<^sup>T \ -v * -v\<^sup>T \ w\<^sup>\ * e\<^sup>T * top)\<^sup>\" using sup_right_divisibility star.circ_back_loop_fixpoint by blast have "(top * e * (w\<^sup>T \ -v * -v\<^sup>T \ w\<^sup>\ * e\<^sup>T * top)\<^sup>\)\<^sup>T \ (top * e * w\<^sup>T\<^sup>\)\<^sup>T" by (simp add: star_isotone mult_right_isotone conv_isotone inf_assoc) also have "... = w\<^sup>\ * e\<^sup>T * top" by (simp add: conv_dist_comp conv_star_commute mult_assoc) finally have 7: "(top * e * (w\<^sup>T \ -v * -v\<^sup>T \ w\<^sup>\ * e\<^sup>T * top)\<^sup>\)\<^sup>T \ w\<^sup>\ * e\<^sup>T * top" . have "(top * e * (w\<^sup>T \ -v * -v\<^sup>T \ w\<^sup>\ * e\<^sup>T * top)\<^sup>\)\<^sup>T \ (top * e * (-v * -v\<^sup>T)\<^sup>\)\<^sup>T" by (simp add: conv_isotone inf_commute mult_right_isotone star_isotone le_infI2) also have "... \ (top * v * -v\<^sup>T * (-v * -v\<^sup>T)\<^sup>\)\<^sup>T" by (metis assms(3) conv_isotone mult_left_isotone mult_right_isotone mult_assoc) also have "... = (top * v * (-v\<^sup>T * -v)\<^sup>\ * -v\<^sup>T)\<^sup>T" by (simp add: mult_assoc star_slide) also have "... \ (top * -v\<^sup>T)\<^sup>T" using conv_order mult_left_isotone by auto also have "... = -v" by (simp add: assms(4) conv_complement vector_conv_compl) finally have 8: "(top * e * (w\<^sup>T \ -v * -v\<^sup>T \ w\<^sup>\ * e\<^sup>T * top)\<^sup>\)\<^sup>T \ w\<^sup>\ * e\<^sup>T * top \ -v" using 7 by simp have "covector (top * e * (w\<^sup>T \ -v * -v\<^sup>T \ w\<^sup>\ * e\<^sup>T * top)\<^sup>\)" by (simp add: covector_mult_closed) hence "top * e * (w\<^sup>T \ -v * -v\<^sup>T \ w\<^sup>\ * e\<^sup>T * top)\<^sup>\ * (w\<^sup>T \ -v\<^sup>T) = top * e * (w\<^sup>T \ -v * -v\<^sup>T \ w\<^sup>\ * e\<^sup>T * top)\<^sup>\ * (w\<^sup>T \ -v\<^sup>T \ (top * e * (w\<^sup>T \ -v * -v\<^sup>T \ w\<^sup>\ * e\<^sup>T * top)\<^sup>\)\<^sup>T)" by (metis comp_inf_vector_1 inf.idem) also have "... \ top * e * (w\<^sup>T \ -v * -v\<^sup>T \ w\<^sup>\ * e\<^sup>T * top)\<^sup>\ * (w\<^sup>T \ -v\<^sup>T \ w\<^sup>\ * e\<^sup>T * top \ -v)" using 8 mult_right_isotone inf.sup_right_isotone inf_assoc by simp also have "... = top * e * (w\<^sup>T \ -v * -v\<^sup>T \ w\<^sup>\ * e\<^sup>T * top)\<^sup>\ * (w\<^sup>T \ (-v \ -v\<^sup>T) \ w\<^sup>\ * e\<^sup>T * top)" using inf_assoc inf_commute by (simp add: inf_assoc) also have "... = top * e * (w\<^sup>T \ -v * -v\<^sup>T \ w\<^sup>\ * e\<^sup>T * top)\<^sup>\ * (w\<^sup>T \ -v * -v\<^sup>T \ w\<^sup>\ * e\<^sup>T * top)" using assms(4) conv_complement vector_complement_closed vector_covector by fastforce also have "... \ top * e * (w\<^sup>T \ -v * -v\<^sup>T \ w\<^sup>\ * e\<^sup>T * top)\<^sup>\" by (simp add: comp_associative comp_isotone star.circ_plus_same star.left_plus_below_circ) finally have 9: "top * e \ top * e * (w\<^sup>T \ -v * -v\<^sup>T \ w\<^sup>\ * e\<^sup>T * top)\<^sup>\ * (w\<^sup>T \ -v\<^sup>T) \ top * e * (w\<^sup>T \ -v * -v\<^sup>T \ w\<^sup>\ * e\<^sup>T * top)\<^sup>\" using 6 by simp have "prim_EP w v e \ -v\<^sup>T \ top * e * w\<^sup>T\<^sup>\" using inf.sup_left_isotone by auto also have "... \ top * e * (w\<^sup>T \ -v\<^sup>T)\<^sup>\" using 5 by (metis inf_commute shunting_var_p) also have "... \ top * e * (w\<^sup>T \ -v * -v\<^sup>T \ w\<^sup>\ * e\<^sup>T * top)\<^sup>\" using 9 by (simp add: star_right_induct) finally have 10: "prim_EP w v e \ top * e * (prim_P w v e)\<^sup>T\<^sup>\" by (simp add: conv_complement conv_dist_comp conv_dist_inf conv_star_commute mult_assoc) have "top * e = top * (v * -v\<^sup>T \ e)" by (simp add: assms(3) inf.absorb2) also have "... \ top * (v * top \ e)" using inf.sup_right_isotone inf_commute mult_right_isotone top_greatest by presburger also have "... = (top \ (v * top)\<^sup>T) * e" using assms(4) covector_inf_comp_3 by presburger also have "... = top * v\<^sup>T * e" by (simp add: conv_dist_comp) also have "... = top * r\<^sup>T * t\<^sup>\ * e" by (simp add: assms(5) comp_associative) also have "... \ top * r\<^sup>T * (prim_W w v e)\<^sup>\ * e" by (metis assms(4,6,7) mst_extends_old_tree star_isotone mult_left_isotone mult_right_isotone) finally have 11: "top * e \ top * r\<^sup>T * (prim_W w v e)\<^sup>\ * e" . have "r\<^sup>T * (prim_W w v e)\<^sup>\ * (prim_EP w v e) \ r\<^sup>T * (prim_W w v e)\<^sup>\ * (top * e * (prim_P w v e)\<^sup>T\<^sup>\)" using 10 mult_right_isotone by blast also have "... = r\<^sup>T * (prim_W w v e)\<^sup>\ * top * e * (prim_P w v e)\<^sup>T\<^sup>\" by (simp add: mult_assoc) also have "... \ top * e * (prim_P w v e)\<^sup>T\<^sup>\" by (metis comp_associative comp_inf_covector inf.idem inf.sup_right_divisibility) also have "... \ top * r\<^sup>T * (prim_W w v e)\<^sup>\ * e * (prim_P w v e)\<^sup>T\<^sup>\" using 11 by (simp add: mult_left_isotone) also have "... = r\<^sup>T * (prim_W w v e)\<^sup>\ * e * (prim_P w v e)\<^sup>T\<^sup>\" using assms(2) vector_conv_covector by auto also have "... \ r\<^sup>T * (prim_W w v e)\<^sup>\ * (prim_W w v e) * (prim_P w v e)\<^sup>T\<^sup>\" by (simp add: mult_left_isotone mult_right_isotone) also have "... \ r\<^sup>T * (prim_W w v e)\<^sup>\ * (prim_W w v e) * (prim_W w v e)\<^sup>\" by (meson dual_order.trans mult_right_isotone star_isotone sup_ge1 sup_ge2) also have "... \ r\<^sup>T * (prim_W w v e)\<^sup>\" by (metis mult_assoc mult_right_isotone star.circ_transitive_equal star.left_plus_below_circ) finally have 12: "r\<^sup>T * (prim_W w v e)\<^sup>\ * (prim_EP w v e) \ r\<^sup>T * (prim_W w v e)\<^sup>\" . have "r\<^sup>T * (prim_W w v e)\<^sup>\ * w \ r\<^sup>T * (prim_W w v e)\<^sup>\ * (w \ prim_EP w v e)" by (simp add: inf_assoc) also have "... = r\<^sup>T * (prim_W w v e)\<^sup>\ * ((w \ prim_EP w v e) \ (-(prim_EP w v e) \ prim_EP w v e))" by (metis assms(1) inf_top_right stone) also have "... = r\<^sup>T * (prim_W w v e)\<^sup>\ * ((w \ -(prim_EP w v e)) \ prim_EP w v e)" by (simp add: sup_inf_distrib2) also have "... = r\<^sup>T * (prim_W w v e)\<^sup>\ * (w \ -(prim_EP w v e)) \ r\<^sup>T * (prim_W w v e)\<^sup>\ * (prim_EP w v e)" by (simp add: comp_left_dist_sup) also have "... \ r\<^sup>T * (prim_W w v e)\<^sup>\ * (prim_W w v e) \ r\<^sup>T * (prim_W w v e)\<^sup>\ * (prim_EP w v e)" using mult_right_isotone sup_left_isotone by auto also have "... \ r\<^sup>T * (prim_W w v e)\<^sup>\ \ r\<^sup>T * (prim_W w v e)\<^sup>\ * (prim_EP w v e)" using mult_assoc mult_right_isotone star.circ_plus_same star.left_plus_below_circ sup_left_isotone by auto also have "... = r\<^sup>T * (prim_W w v e)\<^sup>\" using 12 sup.absorb1 by blast finally have "r\<^sup>T \ r\<^sup>T * (prim_W w v e)\<^sup>\ * w \ r\<^sup>T * (prim_W w v e)\<^sup>\" using 1 by simp thus ?thesis by (simp add: star_right_induct) qed text \ Some of the following lemmas already hold in pseudocomplemented distributive Kleene allegories. \ subsubsection \Exchange gives Minimum Spanning Trees\ text \ The lemmas in this section are used to show that the after exchange we obtain a minimum spanning tree. The following lemmas show various interactions between the three constituents of the tree after exchange. \ lemma epm_1: "vector v \ prim_E w v e \ prim_P w v e = prim_EP w v e" by (metis inf_commute inf_sup_distrib1 mult_assoc mult_right_dist_sup regular_closed_p regular_complement_top vector_conv_compl) lemma epm_2: assumes "regular (prim_EP w v e)" and "vector v" shows "(w \ -(prim_EP w v e)) \ prim_P w v e \ prim_E w v e = w" proof - have "(w \ -(prim_EP w v e)) \ prim_P w v e \ prim_E w v e = (w \ -(prim_EP w v e)) \ prim_EP w v e" using epm_1 sup_assoc sup_commute assms(2) by (simp add: inf_sup_distrib1) also have "... = w \ prim_EP w v e" by (metis assms(1) inf_top.right_neutral regular_complement_top sup_inf_distrib2) also have "... = w" by (simp add: sup_inf_distrib1) finally show ?thesis . qed lemma epm_4: assumes "e \ w" and "injective w" and "w * v \ v" and "e \ v * -v\<^sup>T" shows "top * e * w\<^sup>T\<^sup>+ \ top * v\<^sup>T" proof - have "w\<^sup>\ * v \ v" by (simp add: assms(3) star_left_induct_mult) hence 1: "v\<^sup>T * w\<^sup>T\<^sup>\ \ v\<^sup>T" using conv_star_commute conv_dist_comp conv_isotone by fastforce have "e * w\<^sup>T \ w * w\<^sup>T \ e * w\<^sup>T" by (simp add: assms(1) mult_left_isotone) also have "... \ 1 \ e * w\<^sup>T" using assms(2) inf.sup_left_isotone by auto also have "... = 1 \ w * e\<^sup>T" using calculation conv_dist_comp conv_involutive coreflexive_symmetric by fastforce also have "... \ w * e\<^sup>T" by simp also have "... \ w * -v * v\<^sup>T" by (metis assms(4) conv_complement conv_dist_comp conv_involutive conv_order mult_assoc mult_right_isotone) also have "... \ top * v\<^sup>T" by (simp add: mult_left_isotone) finally have "top * e * w\<^sup>T\<^sup>+ \ top * v\<^sup>T * w\<^sup>T\<^sup>\" by (metis antisym comp_associative comp_isotone dense_top_closed mult_left_isotone transitive_top_closed) also have "... \ top * v\<^sup>T" using 1 by (simp add: mult_assoc mult_right_isotone) finally show ?thesis . qed lemma epm_5: assumes "e \ w" and "injective w" and "w * v \ v" and "e \ v * -v\<^sup>T" and "vector v" shows "prim_P w v e = bot" proof - have 1: "e = w \ top * e" by (simp add: assms(1,2) epm_3) have 2: "top * e * w\<^sup>T\<^sup>+ \ top * v\<^sup>T" by (simp add: assms(1-4) epm_4) have 3: "-v * -v\<^sup>T \ top * v\<^sup>T = bot" by (simp add: assms(5) comp_associative covector_vector_comp inf.sup_monoid.add_commute schroeder_2) have "prim_P w v e = (w \ -v * -v\<^sup>T \ top * e) \ (w \ -v * -v\<^sup>T \ top * e * w\<^sup>T\<^sup>+)" by (metis inf_sup_distrib1 mult_assoc star.circ_back_loop_fixpoint star_plus sup_commute) also have "... \ (e \ -v * -v\<^sup>T) \ (w \ -v * -v\<^sup>T \ top * e * w\<^sup>T\<^sup>+)" using 1 by (metis comp_inf.mult_semi_associative inf.sup_monoid.add_commute semiring.add_right_mono) also have "... \ (e \ -v * -v\<^sup>T) \ (w \ -v * -v\<^sup>T \ top * v\<^sup>T)" using 2 by (metis sup_right_isotone inf.sup_right_isotone) also have "... \ (e \ -v * -v\<^sup>T) \ (-v * -v\<^sup>T \ top * v\<^sup>T)" using inf.assoc le_infI2 by auto also have "... \ v * -v\<^sup>T \ -v * -v\<^sup>T" using 3 assms(4) inf.sup_left_isotone by auto also have "... \ v * top \ -v * top" using inf.sup_mono mult_right_isotone top_greatest by blast also have "... = bot" using assms(5) inf_compl_bot vector_complement_closed by auto finally show ?thesis by (simp add: le_iff_inf) qed lemma epm_6: assumes "e \ w" and "injective w" and "w * v \ v" and "e \ v * -v\<^sup>T" and "vector v" shows "prim_E w v e = e" proof - have 1: "e \ --v * -v\<^sup>T" using assms(4) mult_isotone order_lesseq_imp pp_increasing by blast have 2: "top * e * w\<^sup>T\<^sup>+ \ top * v\<^sup>T" by (simp add: assms(1-4) epm_4) have 3: "e = w \ top * e" by (simp add: assms(1,2) epm_3) hence "e \ top * e * w\<^sup>T\<^sup>\" by (metis le_infI2 star.circ_back_loop_fixpoint sup.commute sup_ge1) hence 4: "e \ prim_E w v e" using 1 by (simp add: assms(1)) have 5: "--v * -v\<^sup>T \ top * v\<^sup>T = bot" by (simp add: assms(5) comp_associative covector_vector_comp inf.sup_monoid.add_commute schroeder_2) have "prim_E w v e = (w \ --v * -v\<^sup>T \ top * e) \ (w \ --v * -v\<^sup>T \ top * e * w\<^sup>T\<^sup>+)" by (metis inf_sup_distrib1 mult_assoc star.circ_back_loop_fixpoint star_plus sup_commute) also have "... \ (e \ --v * -v\<^sup>T) \ (w \ --v * -v\<^sup>T \ top * e * w\<^sup>T\<^sup>+)" using 3 by (metis comp_inf.mult_semi_associative inf.sup_monoid.add_commute semiring.add_right_mono) also have "... \ (e \ --v * -v\<^sup>T) \ (w \ --v * -v\<^sup>T \ top * v\<^sup>T)" using 2 by (metis sup_right_isotone inf.sup_right_isotone) also have "... \ (e \ --v * -v\<^sup>T) \ (--v * -v\<^sup>T \ top * v\<^sup>T)" using inf.assoc le_infI2 by auto also have "... \ e" by (simp add: "5") finally show ?thesis using 4 by (simp add: antisym) qed lemma epm_7: "regular (prim_EP w v e) \ e \ w \ injective w \ w * v \ v \ e \ v * -v\<^sup>T \ vector v \ prim_W w v e = w" by (metis conv_bot epm_2 epm_5 epm_6) lemma epm_8: assumes "acyclic w" shows "(w \ -(prim_EP w v e)) \ (prim_P w v e)\<^sup>T = bot" proof - have "(w \ -(prim_EP w v e)) \ (prim_P w v e)\<^sup>T \ w \ w\<^sup>T" by (meson conv_isotone inf_le1 inf_mono order_trans) thus ?thesis by (metis assms acyclic_asymmetric inf.commute le_bot) qed lemma epm_9: assumes "e \ v * -v\<^sup>T" and "vector v" shows "(w \ -(prim_EP w v e)) \ e = bot" proof - have 1: "e \ -v\<^sup>T" by (metis assms complement_conv_sub vector_conv_covector ev p_antitone_iff p_bot) have "(w \ -(prim_EP w v e)) \ e = (w \ --v\<^sup>T \ e) \ (w \ -(top * e * w\<^sup>T\<^sup>\) \ e)" by (simp add: inf_commute inf_sup_distrib1) also have "... \ (--v\<^sup>T \ e) \ (-(top * e * w\<^sup>T\<^sup>\) \ e)" using comp_inf.mult_left_isotone inf.cobounded2 semiring.add_mono by blast also have "... = -(top * e * w\<^sup>T\<^sup>\) \ e" using 1 by (metis inf.sup_relative_same_increasing inf_commute inf_sup_distrib1 maddux_3_13 regular_closed_p) also have "... = bot" by (metis inf.sup_relative_same_increasing inf_bot_right inf_commute inf_p mult_left_isotone star_outer_increasing top_greatest) finally show ?thesis by (simp add: le_iff_inf) qed lemma epm_10: assumes "e \ v * -v\<^sup>T" and "vector v" shows "(prim_P w v e)\<^sup>T \ e = bot" proof - have "(prim_P w v e)\<^sup>T \ -v * -v\<^sup>T" by (simp add: conv_complement conv_dist_comp conv_dist_inf inf.absorb_iff1 inf.left_commute inf_commute) hence "(prim_P w v e)\<^sup>T \ e \ -v * -v\<^sup>T \ v * -v\<^sup>T" using assms(1) inf_mono by blast also have "... \ -v * top \ v * top" using inf.sup_mono mult_right_isotone top_greatest by blast also have "... = bot" using assms(2) inf_compl_bot vector_complement_closed by auto finally show ?thesis by (simp add: le_iff_inf) qed lemma epm_11: assumes "vector v" shows "(w \ -(prim_EP w v e)) \ prim_P w v e = bot" proof - have "prim_P w v e \ prim_EP w v e" by (metis assms comp_isotone inf.sup_left_isotone inf.sup_right_isotone order.refl top_greatest vector_conv_compl) thus ?thesis using inf_le2 order_trans p_antitone pseudo_complement by blast qed lemma epm_12: assumes "vector v" shows "(w \ -(prim_EP w v e)) \ prim_E w v e = bot" proof - have "prim_E w v e \ prim_EP w v e" by (metis assms comp_isotone inf.sup_left_isotone inf.sup_right_isotone order.refl top_greatest vector_conv_compl) thus ?thesis using inf_le2 order_trans p_antitone pseudo_complement by blast qed lemma epm_13: assumes "vector v" shows "prim_P w v e \ prim_E w v e = bot" proof - have "prim_P w v e \ prim_E w v e \ -v * -v\<^sup>T \ --v * -v\<^sup>T" by (meson dual_order.trans inf.cobounded1 inf.sup_mono inf_le2) also have "... \ -v * top \ --v * top" using inf.sup_mono mult_right_isotone top_greatest by blast also have "... = bot" using assms inf_compl_bot vector_complement_closed by auto finally show ?thesis by (simp add: le_iff_inf) qed text \ The following lemmas show that the relation characterising the edge across the cut is an arc. \ lemma arc_edge_1: assumes "e \ v * -v\<^sup>T \ g" and "vector v" and "v\<^sup>T = r\<^sup>T * t\<^sup>\" and "t \ g" and "r\<^sup>T * g\<^sup>\ \ r\<^sup>T * w\<^sup>\" shows "top * e \ v\<^sup>T * w\<^sup>\" proof - have "top * e \ top * (v * -v\<^sup>T \ g)" using assms(1) mult_right_isotone by auto also have "... \ top * (v * top \ g)" using inf.sup_right_isotone inf_commute mult_right_isotone top_greatest by presburger also have "... = v\<^sup>T * g" by (metis assms(2) covector_inf_comp_3 inf_top.left_neutral) also have "... = r\<^sup>T * t\<^sup>\ * g" by (simp add: assms(3)) also have "... \ r\<^sup>T * g\<^sup>\ * g" by (simp add: assms(4) mult_left_isotone mult_right_isotone star_isotone) also have "... \ r\<^sup>T * g\<^sup>\" by (simp add: mult_assoc mult_right_isotone star.right_plus_below_circ) also have "... \ r\<^sup>T * w\<^sup>\" by (simp add: assms(5)) also have "... \ v\<^sup>T * w\<^sup>\" by (metis assms(3) mult_left_isotone mult_right_isotone mult_1_right star.circ_reflexive) finally show ?thesis . qed lemma arc_edge_2: assumes "e \ v * -v\<^sup>T \ g" and "vector v" and "v\<^sup>T = r\<^sup>T * t\<^sup>\" and "t \ g" and "r\<^sup>T * g\<^sup>\ \ r\<^sup>T * w\<^sup>\" and "w * v \ v" and "injective w" shows "top * e * w\<^sup>T\<^sup>\ \ v\<^sup>T * w\<^sup>\" proof - have 1: "top * e \ v\<^sup>T * w\<^sup>\" using assms(1-5) arc_edge_1 by blast have "v\<^sup>T * w\<^sup>\ * w\<^sup>T = v\<^sup>T * w\<^sup>T \ v\<^sup>T * w\<^sup>+ * w\<^sup>T" by (metis mult_assoc mult_left_dist_sup star.circ_loop_fixpoint sup_commute) also have "... \ v\<^sup>T \ v\<^sup>T * w\<^sup>+ * w\<^sup>T" by (metis assms(6) conv_dist_comp conv_isotone sup_left_isotone) also have "... = v\<^sup>T \ v\<^sup>T * w\<^sup>\ * (w * w\<^sup>T)" by (metis mult_assoc star_plus) also have "... \ v\<^sup>T \ v\<^sup>T * w\<^sup>\" by (metis assms(7) mult_right_isotone mult_1_right sup_right_isotone) also have "... = v\<^sup>T * w\<^sup>\" by (metis star.circ_back_loop_fixpoint sup_absorb2 sup_ge2) finally show ?thesis using 1 star_right_induct by auto qed lemma arc_edge_3: assumes "e \ v * -v\<^sup>T \ g" and "vector v" and "v\<^sup>T = r\<^sup>T * t\<^sup>\" and "t \ g" and "r\<^sup>T * g\<^sup>\ \ r\<^sup>T * w\<^sup>\" and "w * v \ v" and "injective w" and "prim_E w v e = bot" shows "e = bot" proof - have "bot = prim_E w v e" by (simp add: assms(8)) also have "... = w \ --v * top \ top * -v\<^sup>T \ top * e * w\<^sup>T\<^sup>\" by (metis assms(2) comp_inf_covector inf.assoc inf_top.left_neutral vector_conv_compl) also have "... = w \ top * e * w\<^sup>T\<^sup>\ \ -v\<^sup>T \ --v" using assms(2) inf.assoc inf.commute vector_conv_compl vector_complement_closed by (simp add: inf_assoc) finally have 1: "w \ top * e * w\<^sup>T\<^sup>\ \ -v\<^sup>T \ -v" using shunting_1_pp by force have "w\<^sup>\ * e\<^sup>T * top = (top * e * w\<^sup>T\<^sup>\)\<^sup>T" by (simp add: conv_star_commute comp_associative conv_dist_comp) also have "... \ (v\<^sup>T * w\<^sup>\)\<^sup>T" using assms(1-7) arc_edge_2 by (simp add: conv_isotone) also have "... = w\<^sup>T\<^sup>\ * v" by (simp add: conv_star_commute conv_dist_comp) finally have 2: "w\<^sup>\ * e\<^sup>T * top \ w\<^sup>T\<^sup>\ * v" . have "(w\<^sup>T \ w\<^sup>\ * e\<^sup>T * top)\<^sup>T * -v = (w \ top * e * w\<^sup>T\<^sup>\) * -v" by (simp add: conv_dist_comp conv_dist_inf conv_star_commute mult_assoc) also have "... = (w \ top * e * w\<^sup>T\<^sup>\ \ -v\<^sup>T) * top" by (metis assms(2) conv_complement covector_inf_comp_3 inf_top.right_neutral vector_complement_closed) also have "... \ -v * top" using 1 by (simp add: comp_isotone) also have "... = -v" using assms(2) vector_complement_closed by auto finally have "(w\<^sup>T \ w\<^sup>\ * e\<^sup>T * top) * --v \ --v" using p_antitone_iff schroeder_3_p by auto hence "w\<^sup>\ * e\<^sup>T * top \ w\<^sup>T * --v \ --v" by (simp add: inf_vector_comp) hence 3: "w\<^sup>T * --v \ --v \ -(w\<^sup>\ * e\<^sup>T * top)" by (simp add: inf.commute shunting_p) have "w\<^sup>T * -(w\<^sup>\ * e\<^sup>T * top) \ -(w\<^sup>\ * e\<^sup>T * top)" by (metis mult_assoc p_antitone p_antitone_iff schroeder_3_p star.circ_loop_fixpoint sup_commute sup_right_divisibility) also have "... \ --v \ -(w\<^sup>\ * e\<^sup>T * top)" by simp finally have "w\<^sup>T * (--v \ -(w\<^sup>\ * e\<^sup>T * top)) \ --v \ -(w\<^sup>\ * e\<^sup>T * top)" using 3 by (simp add: mult_left_dist_sup) hence "w\<^sup>T\<^sup>\ * (--v \ -(w\<^sup>\ * e\<^sup>T * top)) \ --v \ -(w\<^sup>\ * e\<^sup>T * top)" using star_left_induct_mult_iff by blast hence "w\<^sup>T\<^sup>\ * --v \ --v \ -(w\<^sup>\ * e\<^sup>T * top)" by (simp add: semiring.distrib_left) hence "w\<^sup>\ * e\<^sup>T * top \ w\<^sup>T\<^sup>\ * --v \ --v" by (simp add: inf_commute shunting_p) hence "w\<^sup>\ * e\<^sup>T * top \ --v" using 2 by (metis inf.absorb1 p_antitone_iff p_comp_pp vector_export_comp) hence 4: "e\<^sup>T * top \ --v" by (metis mult_assoc star.circ_loop_fixpoint sup.bounded_iff) have "e\<^sup>T * top \ (v * -v\<^sup>T)\<^sup>T * top" using assms(1) comp_isotone conv_isotone by auto also have "... \ -v * top" by (simp add: conv_complement conv_dist_comp mult_assoc mult_right_isotone) also have "... = -v" using assms(2) vector_complement_closed by auto finally have "e\<^sup>T * top \ bot" using 4 shunting_1_pp by auto hence "e\<^sup>T = bot" using antisym bot_least top_right_mult_increasing by blast thus ?thesis using conv_bot by fastforce qed lemma arc_edge_4: assumes "e \ v * -v\<^sup>T \ g" and "vector v" and "v\<^sup>T = r\<^sup>T * t\<^sup>\" and "t \ g" and "r\<^sup>T * g\<^sup>\ \ r\<^sup>T * w\<^sup>\" and "arc e" shows "top * prim_E w v e * top = top" proof - have "--v\<^sup>T * w = (--v\<^sup>T * w \ -v\<^sup>T) \ (--v\<^sup>T * w \ --v\<^sup>T)" by (simp add: maddux_3_11_pp) also have "... \ (--v\<^sup>T * w \ -v\<^sup>T) \ --v\<^sup>T" using sup_right_isotone by auto also have "... = --v\<^sup>T * (w \ -v\<^sup>T) \ --v\<^sup>T" using assms(2) covector_comp_inf covector_complement_closed vector_conv_covector by auto also have "... \ --v\<^sup>T * (w \ -v\<^sup>T) * w\<^sup>\ \ --v\<^sup>T" by (metis star.circ_back_loop_fixpoint sup.cobounded2 sup_left_isotone) finally have 1: "--v\<^sup>T * w \ --v\<^sup>T * (w \ -v\<^sup>T) * w\<^sup>\ \ --v\<^sup>T" . have "--v\<^sup>T * (w \ -v\<^sup>T) * w\<^sup>\ * w \ --v\<^sup>T * (w \ -v\<^sup>T) * w\<^sup>\ \ --v\<^sup>T" by (simp add: le_supI1 mult_assoc mult_right_isotone star.circ_plus_same star.left_plus_below_circ) hence 2: "(--v\<^sup>T * (w \ -v\<^sup>T) * w\<^sup>\ \ --v\<^sup>T) * w \ --v\<^sup>T * (w \ -v\<^sup>T) * w\<^sup>\ \ --v\<^sup>T" using 1 by (simp add: inf.orderE mult_right_dist_sup) have "v\<^sup>T \ --v\<^sup>T * (w \ -v\<^sup>T) * w\<^sup>\ \ --v\<^sup>T" by (simp add: pp_increasing sup.coboundedI2) hence "v\<^sup>T * w\<^sup>\ \ --v\<^sup>T * (w \ -v\<^sup>T) * w\<^sup>\ \ --v\<^sup>T" using 2 by (simp add: star_right_induct) hence 3: "-v\<^sup>T \ v\<^sup>T * w\<^sup>\ \ --v\<^sup>T * (w \ -v\<^sup>T) * w\<^sup>\" by (metis inf_commute shunting_var_p) have "top * e = top * e \ v\<^sup>T * w\<^sup>\" by (meson assms(1-5) arc_edge_1 inf.orderE) also have "... \ top * v * -v\<^sup>T \ v\<^sup>T * w\<^sup>\" using assms(1) inf.sup_left_isotone mult_assoc mult_right_isotone by auto also have "... \ top * -v\<^sup>T \ v\<^sup>T * w\<^sup>\" using inf.sup_left_isotone mult_left_isotone top_greatest by blast also have "... = -v\<^sup>T \ v\<^sup>T * w\<^sup>\" by (simp add: assms(2) vector_conv_compl) also have "... \ --v\<^sup>T * (w \ -v\<^sup>T) * w\<^sup>\" using 3 by simp also have "... = (top \ (--v)\<^sup>T) * (w \ -v\<^sup>T) * w\<^sup>\" by (simp add: conv_complement) also have "... = top * (w \ --v \ -v\<^sup>T) * w\<^sup>\" using assms(2) covector_inf_comp_3 inf_assoc inf_left_commute vector_complement_closed by presburger also have "... = top * (w \ --v * -v\<^sup>T) * w\<^sup>\" by (metis assms(2) vector_complement_closed conv_complement inf_assoc vector_covector) finally have "top * (e\<^sup>T * top)\<^sup>T \ top * (w \ --v * -v\<^sup>T) * w\<^sup>\" by (metis conv_dist_comp conv_involutive conv_top mult_assoc top_mult_top) hence "top \ top * (w \ --v * -v\<^sup>T) * w\<^sup>\ * (e\<^sup>T * top)" using assms(6) shunt_bijective by blast also have "... = top * (w \ --v * -v\<^sup>T) * (top * e * w\<^sup>\\<^sup>T)\<^sup>T" by (simp add: conv_dist_comp mult_assoc) also have "... = top * (w \ --v * -v\<^sup>T \ top * e * w\<^sup>\\<^sup>T) * top" by (simp add: comp_inf_vector_1 mult_assoc) finally show ?thesis by (simp add: conv_star_commute top_le) qed lemma arc_edge_5: assumes "vector v" and "w * v \ v" and "injective w" and "arc e" shows "(prim_E w v e)\<^sup>T * top * prim_E w v e \ 1" proof - have 1: "e\<^sup>T * top * e \ 1" by (simp add: assms(4) point_injective) have "prim_E w v e \ --v * top" by (simp add: inf_commute le_infI2 mult_right_isotone) hence 2: "prim_E w v e \ --v" by (simp add: assms(1) vector_complement_closed) have 3: "w * --v \ --v" by (simp add: assms(2) p_antitone p_antitone_iff) have "w \ top * prim_E w v e \ w * (prim_E w v e)\<^sup>T * prim_E w v e" by (metis dedekind_2 inf.commute inf_top.left_neutral) also have "... \ w * w\<^sup>T * prim_E w v e" by (simp add: conv_isotone le_infI1 mult_left_isotone mult_right_isotone) also have "... \ prim_E w v e" by (metis assms(3) mult_left_isotone mult_left_one) finally have 4: "w \ top * prim_E w v e \ prim_E w v e" . have "w\<^sup>+ \ top * prim_E w v e = w\<^sup>\ * (w \ top * prim_E w v e)" by (simp add: comp_inf_covector star_plus) also have "... \ w\<^sup>\ * prim_E w v e" using 4 by (simp add: mult_right_isotone) also have "... \ --v" using 2 3 star_left_induct sup.bounded_iff by blast finally have 5: "w\<^sup>+ \ top * prim_E w v e \ -v = bot" using shunting_1_pp by blast hence 6: "w\<^sup>+\<^sup>T \ (prim_E w v e)\<^sup>T * top \ -v\<^sup>T = bot" using conv_complement conv_dist_comp conv_dist_inf conv_top conv_bot by force have "(prim_E w v e)\<^sup>T * top * prim_E w v e \ (top * e * w\<^sup>T\<^sup>\)\<^sup>T * top * (top * e * w\<^sup>T\<^sup>\)" by (simp add: conv_isotone mult_isotone) also have "... = w\<^sup>\ * e\<^sup>T * top * e * w\<^sup>T\<^sup>\" by (metis conv_star_commute conv_dist_comp conv_involutive conv_top mult_assoc top_mult_top) also have "... \ w\<^sup>\ * w\<^sup>T\<^sup>\" using 1 by (metis mult_assoc mult_1_right mult_right_isotone mult_left_isotone) also have "... = w\<^sup>\ \ w\<^sup>T\<^sup>\" by (metis assms(3) cancel_separate inf.eq_iff star.circ_sup_sub_sup_one_1 star.circ_plus_one star_involutive) also have "... = w\<^sup>+ \ w\<^sup>T\<^sup>+ \ 1" by (metis star.circ_plus_one star_left_unfold_equal sup.assoc sup.commute) finally have 7: "(prim_E w v e)\<^sup>T * top * prim_E w v e \ w\<^sup>+ \ w\<^sup>T\<^sup>+ \ 1" . have "prim_E w v e \ --v * -v\<^sup>T" by (simp add: le_infI1) also have "... \ top * -v\<^sup>T" by (simp add: mult_left_isotone) also have "... = -v\<^sup>T" by (simp add: assms(1) vector_conv_compl) finally have 8: "prim_E w v e \ -v\<^sup>T" . hence 9: "(prim_E w v e)\<^sup>T \ -v" by (metis conv_complement conv_involutive conv_isotone) have "(prim_E w v e)\<^sup>T * top * prim_E w v e = (w\<^sup>+ \ w\<^sup>T\<^sup>+ \ 1) \ (prim_E w v e)\<^sup>T * top * prim_E w v e" using 7 by (simp add: inf.absorb_iff2) also have "... = (1 \ (prim_E w v e)\<^sup>T * top * prim_E w v e) \ (w\<^sup>+ \ (prim_E w v e)\<^sup>T * top * prim_E w v e) \ (w\<^sup>T\<^sup>+ \ (prim_E w v e)\<^sup>T * top * prim_E w v e)" using comp_inf.mult_right_dist_sup sup_assoc sup_commute by auto also have "... \ 1 \ (w\<^sup>+ \ (prim_E w v e)\<^sup>T * top * prim_E w v e) \ (w\<^sup>T\<^sup>+ \ (prim_E w v e)\<^sup>T * top * prim_E w v e)" using inf_le1 sup_left_isotone by blast also have "... \ 1 \ (w\<^sup>+ \ (prim_E w v e)\<^sup>T * top * prim_E w v e) \ (w\<^sup>T\<^sup>+ \ (prim_E w v e)\<^sup>T * top * -v\<^sup>T)" using 8 inf.sup_right_isotone mult_right_isotone sup_right_isotone by blast also have "... \ 1 \ (w\<^sup>+ \ -v * top * prim_E w v e) \ (w\<^sup>T\<^sup>+ \ (prim_E w v e)\<^sup>T * top * -v\<^sup>T)" using 9 by (metis inf.sup_right_isotone mult_left_isotone sup.commute sup_right_isotone) also have "... = 1 \ (w\<^sup>+ \ -v * top \ top * prim_E w v e) \ (w\<^sup>T\<^sup>+ \ (prim_E w v e)\<^sup>T * top \ top * -v\<^sup>T)" by (metis (no_types) vector_export_comp inf_top_right inf_assoc) also have "... = 1 \ (w\<^sup>+ \ -v \ top * prim_E w v e) \ (w\<^sup>T\<^sup>+ \ (prim_E w v e)\<^sup>T * top \ -v\<^sup>T)" using assms(1) vector_complement_closed vector_conv_compl by auto also have "... = 1" using 5 6 by (simp add: conv_star_commute conv_dist_comp inf.commute inf_assoc star.circ_plus_same) finally show ?thesis . qed lemma arc_edge_6: assumes "vector v" and "w * v \ v" and "injective w" and "arc e" shows "prim_E w v e * top * (prim_E w v e)\<^sup>T \ 1" proof - have "prim_E w v e * 1 * (prim_E w v e)\<^sup>T \ w * w\<^sup>T" using comp_isotone conv_order inf.coboundedI1 mult_one_associative by auto also have "... \ 1" by (simp add: assms(3)) finally have 1: "prim_E w v e * 1 * (prim_E w v e)\<^sup>T \ 1" . have "(prim_E w v e)\<^sup>T * top * prim_E w v e \ 1" by (simp add: assms arc_edge_5) also have "... \ --1" by (simp add: pp_increasing) finally have 2: "prim_E w v e * -1 * (prim_E w v e)\<^sup>T \ bot" by (metis conv_involutive regular_closed_bot regular_dense_top triple_schroeder_p) have "prim_E w v e * top * (prim_E w v e)\<^sup>T = prim_E w v e * 1 * (prim_E w v e)\<^sup>T \ prim_E w v e * -1 * (prim_E w v e)\<^sup>T" by (metis mult_left_dist_sup mult_right_dist_sup regular_complement_top regular_one_closed) also have "... \ 1" using 1 2 by (simp add: bot_unique) finally show ?thesis . qed lemma arc_edge: assumes "e \ v * -v\<^sup>T \ g" and "vector v" and "v\<^sup>T = r\<^sup>T * t\<^sup>\" and "t \ g" and "r\<^sup>T * g\<^sup>\ \ r\<^sup>T * w\<^sup>\" and "w * v \ v" and "injective w" and "arc e" shows "arc (prim_E w v e)" proof (intro conjI) have "prim_E w v e * top * (prim_E w v e)\<^sup>T \ 1" using assms(2,6-8) arc_edge_6 by simp thus "injective (prim_E w v e * top)" by (metis conv_dist_comp conv_top mult_assoc top_mult_top) next show "surjective (prim_E w v e * top)" using assms(1-5,8) arc_edge_4 mult_assoc by simp next have "(prim_E w v e)\<^sup>T * top * prim_E w v e \ 1" using assms(2,6-8) arc_edge_5 by simp thus "injective ((prim_E w v e)\<^sup>T * top)" by (metis conv_dist_comp conv_involutive conv_top mult_assoc top_mult_top) next have "top * prim_E w v e * top = top" using assms(1-5,8) arc_edge_4 by simp thus "surjective ((prim_E w v e)\<^sup>T * top)" by (metis mult_assoc conv_dist_comp conv_top) qed subsubsection \Invariant implies Postcondition\ text \ The lemmas in this section are used to show that the invariant implies the postcondition at the end of the algorithm. The following lemma shows that the nodes reachable in the graph are the same as those reachable in the constructed tree. \ lemma span_post: assumes "regular v" and "vector v" and "v\<^sup>T = r\<^sup>T * t\<^sup>\" and "v * -v\<^sup>T \ g = bot" and "t \ v * v\<^sup>T \ g" and "r\<^sup>T * (v * v\<^sup>T \ g)\<^sup>\ \ r\<^sup>T * t\<^sup>\" shows "v\<^sup>T = r\<^sup>T * g\<^sup>\" proof - let ?vv = "v * v\<^sup>T \ g" have 1: "r\<^sup>T \ v\<^sup>T" using assms(3) mult_right_isotone mult_1_right star.circ_reflexive by fastforce have "v * top \ g = (v * v\<^sup>T \ v * -v\<^sup>T) \ g" by (metis assms(1) conv_complement mult_left_dist_sup regular_complement_top) also have "... = ?vv \ (v * -v\<^sup>T \ g)" by (simp add: inf_sup_distrib2) also have "... = ?vv" by (simp add: assms(4)) finally have 2: "v * top \ g = ?vv" by simp have "r\<^sup>T * ?vv\<^sup>\ \ v\<^sup>T * ?vv\<^sup>\" using 1 by (simp add: comp_left_isotone) also have "... \ v\<^sup>T * (v * v\<^sup>T)\<^sup>\" by (simp add: comp_right_isotone star.circ_isotone) also have "... \ v\<^sup>T" by (simp add: assms(2) vector_star_1) finally have "r\<^sup>T * ?vv\<^sup>\ \ v\<^sup>T" by simp hence "r\<^sup>T * ?vv\<^sup>\ * g = (r\<^sup>T * ?vv\<^sup>\ \ v\<^sup>T) * g" by (simp add: inf.absorb1) also have "... = r\<^sup>T * ?vv\<^sup>\ * (v * top \ g)" by (simp add: assms(2) covector_inf_comp_3) also have "... = r\<^sup>T * ?vv\<^sup>\ * ?vv" using 2 by simp also have "... \ r\<^sup>T * ?vv\<^sup>\" by (simp add: comp_associative comp_right_isotone star.left_plus_below_circ star_plus) finally have "r\<^sup>T \ r\<^sup>T * ?vv\<^sup>\ * g \ r\<^sup>T * ?vv\<^sup>\" using star.circ_back_loop_prefixpoint by auto hence "r\<^sup>T * g\<^sup>\ \ r\<^sup>T * ?vv\<^sup>\" using star_right_induct by blast hence "r\<^sup>T * g\<^sup>\ = r\<^sup>T * ?vv\<^sup>\" by (simp add: antisym mult_right_isotone star_isotone) also have "... = r\<^sup>T * t\<^sup>\" using assms(5,6) antisym mult_right_isotone star_isotone by auto also have "... = v\<^sup>T" by (simp add: assms(3)) finally show ?thesis by simp qed text \ The following lemma shows that the minimum spanning tree extending a tree is the same as the tree at the end of the algorithm. \ lemma mst_post: assumes "vector r" and "injective r" and "v\<^sup>T = r\<^sup>T * t\<^sup>\" and "forest w" and "t \ w" and "w \ v * v\<^sup>T" shows "w = t" proof - have 1: "vector v" using assms(1,3) covector_mult_closed vector_conv_covector by auto have "w * v \ v * v\<^sup>T * v" by (simp add: assms(6) mult_left_isotone) also have "... \ v" using 1 by (metis mult_assoc mult_right_isotone top_greatest) finally have 2: "w * v \ v" . have 3: "r \ v" by (metis assms(3) conv_order mult_right_isotone mult_1_right star.circ_reflexive) have 4: "v \ -r = t\<^sup>T\<^sup>\ * r \ -r" by (metis assms(3) conv_dist_comp conv_involutive conv_star_commute) also have "... = (r \ t\<^sup>T\<^sup>+ * r) \ -r" using mult_assoc star.circ_loop_fixpoint sup_commute by auto also have "... \ t\<^sup>T\<^sup>+ * r" by (simp add: shunting) also have "... \ t\<^sup>T * top" by (simp add: comp_isotone mult_assoc) finally have "1 \ (v \ -r) * (v \ -r)\<^sup>T \ 1 \ t\<^sup>T * top * (t\<^sup>T * top)\<^sup>T" using conv_order inf.sup_right_isotone mult_isotone by auto also have "... = 1 \ t\<^sup>T * top * t" by (metis conv_dist_comp conv_involutive conv_top mult_assoc top_mult_top) also have "... \ t\<^sup>T * (top * t \ t * 1)" by (metis conv_involutive dedekind_1 inf.commute mult_assoc) also have "... \ t\<^sup>T * t" by (simp add: mult_right_isotone) finally have 5: "1 \ (v \ -r) * (v \ -r)\<^sup>T \ t\<^sup>T * t" . have "w * w\<^sup>+ \ -1" by (metis assms(4) mult_right_isotone order_trans star.circ_increasing star.left_plus_circ) hence 6: "w\<^sup>T\<^sup>+ \ -w" by (metis conv_star_commute mult_assoc mult_1_left triple_schroeder_p) have "w * r \ w\<^sup>T\<^sup>+ * r = (w \ w\<^sup>T\<^sup>+) * r" using assms(2) by (simp add: injective_comp_right_dist_inf) also have "... = bot" using 6 p_antitone pseudo_complement_pp semiring.mult_not_zero by blast finally have 7: "w * r \ w\<^sup>T\<^sup>+ * r = bot" . have "-1 * r \ -r" - using assms(2) local.dual_order.trans local.pp_increasing local.schroeder_4_p by blast + using assms(2) dual_order.trans pp_increasing schroeder_4_p by blast hence "-1 * r * top \ -r" by (simp add: assms(1) comp_associative) hence 8: "r\<^sup>T * -1 * r \ bot" by (simp add: mult_assoc schroeder_6_p) have "r\<^sup>T * w * r \ r\<^sup>T * w\<^sup>+ * r" by (simp add: mult_left_isotone mult_right_isotone star.circ_mult_increasing) also have "... \ r\<^sup>T * -1 * r" by (simp add: assms(4) comp_isotone) finally have "r\<^sup>T * w * r \ bot" using 8 by simp hence "w * r * top \ -r" by (simp add: mult_assoc schroeder_6_p) hence "w * r \ -r" by (simp add: assms(1) comp_associative) hence "w * r \ -r \ w * v" using 3 by (simp add: mult_right_isotone) also have "... \ -r \ v" using 2 by (simp add: le_infI2) also have "... = -r \ t\<^sup>T\<^sup>\ * r" using 4 by (simp add: inf_commute) also have "... \ -r \ w\<^sup>T\<^sup>\ * r" using assms(5) comp_inf.mult_right_isotone conv_isotone mult_left_isotone star_isotone by auto also have "... = -r \ (r \ w\<^sup>T\<^sup>+ * r)" using mult_assoc star.circ_loop_fixpoint sup_commute by auto also have "... \ w\<^sup>T\<^sup>+ * r" using inf.commute maddux_3_13 by auto finally have "w * r = bot" using 7 by (simp add: le_iff_inf) hence "w = w \ top * -r\<^sup>T" by (metis complement_conv_sub conv_dist_comp conv_involutive conv_bot inf.assoc inf.orderE regular_closed_bot regular_dense_top top_left_mult_increasing) also have "... = w \ v * v\<^sup>T \ top * -r\<^sup>T" by (simp add: assms(6) inf_absorb1) also have "... \ w \ top * v\<^sup>T \ top * -r\<^sup>T" using comp_inf.mult_left_isotone comp_inf.mult_right_isotone mult_left_isotone by auto also have "... = w \ top * (v\<^sup>T \ -r\<^sup>T)" using 1 assms(1) covector_inf_closed inf_assoc vector_conv_compl vector_conv_covector by auto also have "... = w * (1 \ (v \ -r) * top)" by (simp add: comp_inf_vector conv_complement conv_dist_inf) also have "... = w * (1 \ (v \ -r) * (v \ -r)\<^sup>T)" by (metis conv_top dedekind_eq inf_commute inf_top_left mult_1_left mult_1_right) also have "... \ w * t\<^sup>T * t" using 5 by (simp add: comp_isotone mult_assoc) also have "... \ w * w\<^sup>T * t" by (simp add: assms(5) comp_isotone conv_isotone) also have "... \ t" using assms(4) mult_left_isotone mult_1_left by fastforce finally show ?thesis by (simp add: assms(5) antisym) qed subsection \Kruskal's Algorithm\ text \ The following results are used for proving the correctness of Kruskal's minimum spanning tree algorithm. \ subsubsection \Preservation of Invariant\ text \ We first treat the preservation of the invariant. The following lemmas show conditions necessary for preserving that \f\ is a forest. \ lemma kruskal_injective_inv_2: assumes "arc e" and "acyclic f" shows "top * e * f\<^sup>T\<^sup>\ * f\<^sup>T \ -e" proof - have "f \ -f\<^sup>T\<^sup>\" using assms(2) acyclic_star_below_complement p_antitone_iff by simp hence "e * f \ top * e * -f\<^sup>T\<^sup>\" by (simp add: comp_isotone top_left_mult_increasing) also have "... = -(top * e * f\<^sup>T\<^sup>\)" by (metis assms(1) comp_mapping_complement conv_dist_comp conv_involutive conv_top) finally show ?thesis using schroeder_4_p by simp qed lemma kruskal_injective_inv_3: assumes "arc e" and "forest f" shows "(top * e * f\<^sup>T\<^sup>\)\<^sup>T * (top * e * f\<^sup>T\<^sup>\) \ f\<^sup>T * f \ 1" proof - have "(top * e * f\<^sup>T\<^sup>\)\<^sup>T * (top * e * f\<^sup>T\<^sup>\) = f\<^sup>\ * e\<^sup>T * top * e * f\<^sup>T\<^sup>\" by (metis conv_dist_comp conv_involutive conv_star_commute conv_top vector_top_closed mult_assoc) also have "... \ f\<^sup>\ * f\<^sup>T\<^sup>\" by (metis assms(1) arc_expanded mult_left_isotone mult_right_isotone mult_1_left mult_assoc) finally have "(top * e * f\<^sup>T\<^sup>\)\<^sup>T * (top * e * f\<^sup>T\<^sup>\) \ f\<^sup>T * f \ f\<^sup>\ * f\<^sup>T\<^sup>\ \ f\<^sup>T * f" using inf.sup_left_isotone by simp also have "... \ 1" using assms(2) forest_separate by simp finally show ?thesis by simp qed lemma kruskal_acyclic_inv: assumes "acyclic f" and "covector q" and "(f \ q)\<^sup>T * f\<^sup>\ * e = bot" and "e * f\<^sup>\ * e = bot" and "f\<^sup>T\<^sup>\ * f\<^sup>\ \ -e" shows "acyclic ((f \ -q) \ (f \ q)\<^sup>T \ e)" proof - have "(f \ -q) * (f \ q)\<^sup>T = (f \ -q) * (f\<^sup>T \ q\<^sup>T)" by (simp add: conv_dist_inf) hence 1: "(f \ -q) * (f \ q)\<^sup>T = bot" by (metis assms(2) comp_inf.semiring.mult_zero_right comp_inf_vector_1 conv_bot covector_bot_closed inf.sup_monoid.add_assoc p_inf) hence 2: "(f \ -q)\<^sup>\ * (f \ q)\<^sup>T = (f \ q)\<^sup>T" using mult_right_zero star_absorb star_simulation_right_equal by fastforce hence 3: "((f \ -q) \ (f \ q)\<^sup>T)\<^sup>+ = (f \ q)\<^sup>T\<^sup>\ * (f \ -q)\<^sup>+ \ (f \ q)\<^sup>T\<^sup>+" by (simp add: plus_sup) have 4: "((f \ -q) \ (f \ q)\<^sup>T)\<^sup>\ = (f \ q)\<^sup>T\<^sup>\ * (f \ -q)\<^sup>\" using 2 by (simp add: star.circ_sup_9) have "(f \ q)\<^sup>T * (f \ -q)\<^sup>\ * e \ (f \ q)\<^sup>T * f\<^sup>\ * e" by (simp add: mult_left_isotone mult_right_isotone star_isotone) hence "(f \ q)\<^sup>T * (f \ -q)\<^sup>\ * e = bot" using assms(3) le_bot by simp hence 5: "(f \ q)\<^sup>T\<^sup>\ * (f \ -q)\<^sup>\ * e = (f \ -q)\<^sup>\ * e" by (metis comp_associative conv_bot conv_dist_comp conv_involutive conv_star_commute star_absorb) have "e * (f \ -q)\<^sup>\ * e \ e * f\<^sup>\ * e" by (simp add: mult_left_isotone mult_right_isotone star_isotone) hence "e * (f \ -q)\<^sup>\ * e = bot" using assms(4) le_bot by simp hence 6: "((f \ -q)\<^sup>\ * e)\<^sup>+ = (f \ -q)\<^sup>\ * e" by (simp add: comp_associative star_absorb) have "f\<^sup>T\<^sup>\ * 1 * f\<^sup>T\<^sup>\ * f\<^sup>\ \ -e" by (simp add: assms(5) star.circ_transitive_equal) hence 7: "f\<^sup>\ * e * f\<^sup>T\<^sup>\ * f\<^sup>\ \ -1" by (metis comp_right_one conv_involutive conv_one conv_star_commute triple_schroeder_p) have "(f \ -q)\<^sup>+ * (f \ q)\<^sup>T\<^sup>+ \ -1" using 1 2 by (metis forest_bot mult_left_zero mult_assoc) hence 8: "(f \ q)\<^sup>T\<^sup>+ * (f \ -q)\<^sup>+ \ -1" using comp_commute_below_diversity by simp have 9: "f\<^sup>T\<^sup>+ \ -1" using assms(1) acyclic_star_below_complement schroeder_5_p by force have "((f \ -q) \ (f \ q)\<^sup>T \ e)\<^sup>+ = (((f \ -q) \ (f \ q)\<^sup>T)\<^sup>\ * e)\<^sup>\ * ((f \ -q) \ (f \ q)\<^sup>T)\<^sup>+ \ (((f \ -q) \ (f \ q)\<^sup>T)\<^sup>\ * e)\<^sup>+" by (simp add: plus_sup) also have "... = ((f \ q)\<^sup>T\<^sup>\ * (f \ -q)\<^sup>\ * e)\<^sup>\ * ((f \ q)\<^sup>T\<^sup>\ * (f \ -q)\<^sup>+ \ (f \ q)\<^sup>T\<^sup>+) \ ((f \ q)\<^sup>T\<^sup>\ * (f \ -q)\<^sup>\ * e)\<^sup>+" using 3 4 by simp also have "... = ((f \ -q)\<^sup>\ * e)\<^sup>\ * ((f \ q)\<^sup>T\<^sup>\ * (f \ -q)\<^sup>+ \ (f \ q)\<^sup>T\<^sup>+) \ ((f \ -q)\<^sup>\ * e)\<^sup>+" using 5 by simp also have "... = ((f \ -q)\<^sup>\ * e \ 1) * ((f \ q)\<^sup>T\<^sup>\ * (f \ -q)\<^sup>+ \ (f \ q)\<^sup>T\<^sup>+) \ (f \ -q)\<^sup>\ * e" using 6 by (metis star_left_unfold_equal sup_monoid.add_commute) also have "... = (f \ -q)\<^sup>\ * e \ (f \ -q)\<^sup>\ * e * (f \ q)\<^sup>T\<^sup>+ \ (f \ -q)\<^sup>\ * e * (f \ q)\<^sup>T\<^sup>\ * (f \ -q)\<^sup>+ \ (f \ q)\<^sup>T\<^sup>\ * (f \ -q)\<^sup>+ \ (f \ q)\<^sup>T\<^sup>+" using comp_associative mult_left_dist_sup mult_right_dist_sup sup_assoc sup_commute by simp also have "... = (f \ -q)\<^sup>\ * e * (f \ q)\<^sup>T\<^sup>\ * (f \ -q)\<^sup>\ \ (f \ q)\<^sup>T\<^sup>\ * (f \ -q)\<^sup>+ \ (f \ q)\<^sup>T\<^sup>+" by (metis star.circ_back_loop_fixpoint star_plus sup_monoid.add_commute mult_assoc) also have "... \ f\<^sup>\ * e * f\<^sup>T\<^sup>\ * (f \ -q)\<^sup>\ \ (f \ q)\<^sup>T\<^sup>\ * (f \ -q)\<^sup>+ \ (f \ q)\<^sup>T\<^sup>+" using mult_left_isotone mult_right_isotone star_isotone sup_left_isotone conv_isotone order_trans inf_le1 by meson also have "... \ f\<^sup>\ * e * f\<^sup>T\<^sup>\ * f\<^sup>\ \ (f \ q)\<^sup>T\<^sup>\ * (f \ -q)\<^sup>+ \ f\<^sup>T\<^sup>+" using mult_left_isotone mult_right_isotone star_isotone sup_left_isotone sup_right_isotone conv_isotone order_trans inf_le1 by meson also have "... = f\<^sup>\ * e * f\<^sup>T\<^sup>\ * f\<^sup>\ \ (f \ q)\<^sup>T\<^sup>+ * (f \ -q)\<^sup>+ \ (f \ -q)\<^sup>+ \ f\<^sup>T\<^sup>+" by (simp add: star.circ_loop_fixpoint sup_monoid.add_assoc mult_assoc) also have "... \ f\<^sup>\ * e * f\<^sup>T\<^sup>\ * f\<^sup>\ \ (f \ q)\<^sup>T\<^sup>+ * (f \ -q)\<^sup>+ \ f\<^sup>+ \ f\<^sup>T\<^sup>+" using mult_left_isotone mult_right_isotone star_isotone sup_left_isotone sup_right_isotone order_trans inf_le1 by meson also have "... \ -1" using 7 8 9 assms(1) by simp finally show ?thesis by simp qed lemma kruskal_exchange_acyclic_inv_1: assumes "acyclic f" and "covector q" shows "acyclic ((f \ -q) \ (f \ q)\<^sup>T)" using kruskal_acyclic_inv[where e=bot] by (simp add: assms) lemma kruskal_exchange_acyclic_inv_2: assumes "acyclic w" and "injective w" and "d \ w" and "bijective (d\<^sup>T * top)" and "bijective (e * top)" and "d \ top * e\<^sup>T * w\<^sup>T\<^sup>\" and "w * e\<^sup>T * top = bot" shows "acyclic ((w \ -d) \ e)" proof - let ?v = "w \ -d" let ?w = "?v \ e" have "d\<^sup>T * top \ w\<^sup>\ * e * top" by (metis assms(6) comp_associative comp_inf.star.circ_decompose_9 comp_inf.star_star_absorb comp_isotone conv_dist_comp conv_involutive conv_order conv_star_commute conv_top inf.cobounded1 vector_top_closed) hence 1: "e * top \ w\<^sup>T\<^sup>\ * d\<^sup>T * top" by (metis assms(4,5) bijective_reverse comp_associative conv_star_commute) have 2: "?v * d\<^sup>T * top = bot" by (simp add: assms(2,3) kruskal_exchange_acyclic_inv_3) have "?v * w\<^sup>T\<^sup>+ * d\<^sup>T * top \ w * w\<^sup>T\<^sup>+ * d\<^sup>T * top" by (simp add: mult_left_isotone) also have "... \ w\<^sup>T\<^sup>\ * d\<^sup>T * top" by (metis assms(2) mult_left_isotone mult_1_left mult_assoc) finally have "?v * w\<^sup>T\<^sup>\ * d\<^sup>T * top \ w\<^sup>T\<^sup>\ * d\<^sup>T * top" using 2 by (metis bot_least comp_associative mult_right_dist_sup star.circ_back_loop_fixpoint star.circ_plus_same sup_least) hence 3: "?v\<^sup>\ * e * top \ w\<^sup>T\<^sup>\ * d\<^sup>T * top" using 1 by (simp add: comp_associative star_left_induct sup_least) have "d * e\<^sup>T \ bot" by (metis assms(3,7) conv_bot conv_dist_comp conv_involutive conv_top order.trans inf.absorb2 inf.cobounded2 inf_commute le_bot p_antitone_iff p_top schroeder_4_p top_left_mult_increasing) hence 4: "e\<^sup>T * top \ -(d\<^sup>T * top)" by (metis (no_types) comp_associative inf.cobounded2 le_bot p_antitone_iff schroeder_3_p semiring.mult_zero_left) have "?v\<^sup>T * -(d\<^sup>T * top) \ -(d\<^sup>T * top)" using schroeder_3_p mult_assoc 2 by simp hence "?v\<^sup>T\<^sup>\ * e\<^sup>T * top \ -(d\<^sup>T * top)" using 4 by (simp add: comp_associative star_left_induct sup_least) hence 5: "d\<^sup>T * top \ -(?v\<^sup>T\<^sup>\ * e\<^sup>T * top)" by (simp add: p_antitone_iff) have "w * ?v\<^sup>T\<^sup>\ * e\<^sup>T * top = w * e\<^sup>T * top \ w * ?v\<^sup>T\<^sup>+ * e\<^sup>T * top" by (metis star_left_unfold_equal mult_right_dist_sup mult_left_dist_sup mult_1_right mult_assoc) also have "... = w * ?v\<^sup>T\<^sup>+ * e\<^sup>T * top" using assms(7) by simp also have "... \ w * w\<^sup>T * ?v\<^sup>T\<^sup>\ * e\<^sup>T * top" by (simp add: comp_associative conv_isotone mult_left_isotone mult_right_isotone) also have "... \ ?v\<^sup>T\<^sup>\ * e\<^sup>T * top" by (metis assms(2) mult_1_left mult_left_isotone) finally have "w * ?v\<^sup>T\<^sup>\ * e\<^sup>T * top \ --(?v\<^sup>T\<^sup>\ * e\<^sup>T * top)" by (simp add: p_antitone p_antitone_iff) hence "w\<^sup>T * -(?v\<^sup>T\<^sup>\ * e\<^sup>T * top) \ -(?v\<^sup>T\<^sup>\ * e\<^sup>T * top)" using comp_associative schroeder_3_p by simp hence 6: "w\<^sup>T\<^sup>\ * d\<^sup>T * top \ -(?v\<^sup>T\<^sup>\ * e\<^sup>T * top)" using 5 by (simp add: comp_associative star_left_induct sup_least) have "e * ?v\<^sup>\ * e \ e * ?v\<^sup>\ * e * top" by (simp add: top_right_mult_increasing) also have "... \ e * w\<^sup>T\<^sup>\ * d\<^sup>T * top" using 3 by (simp add: comp_associative mult_right_isotone) also have "... \ e * -(?v\<^sup>T\<^sup>\ * e\<^sup>T * top)" using 6 by (simp add: comp_associative mult_right_isotone) also have "... \ bot" by (metis conv_complement_sub_leq conv_dist_comp conv_involutive conv_star_commute le_bot mult_right_sub_dist_sup_right p_bot regular_closed_bot star.circ_back_loop_fixpoint) finally have 7: "e * ?v\<^sup>\ * e = bot" by (simp add: antisym) hence "?v\<^sup>\ * e \ -1" by (metis bot_least comp_associative comp_commute_below_diversity ex231d order_lesseq_imp semiring.mult_zero_left star.circ_left_top) hence 8: "?v\<^sup>\ * e * ?v\<^sup>\ \ -1" by (metis comp_associative comp_commute_below_diversity star.circ_transitive_equal) have "1 \ ?w\<^sup>+ = 1 \ ?w * ?v\<^sup>\ * (e * ?v\<^sup>\)\<^sup>\" by (simp add: star_sup_1 mult_assoc) also have "... = 1 \ ?w * ?v\<^sup>\ * (e * ?v\<^sup>\ \ 1)" using 7 by (metis star.circ_mult_1 star_absorb sup_monoid.add_commute mult_assoc) also have "... = 1 \ (?v\<^sup>+ * e * ?v\<^sup>\ \ ?v\<^sup>+ \ e * ?v\<^sup>\ * e * ?v\<^sup>\ \ e * ?v\<^sup>\)" by (simp add: comp_associative mult_left_dist_sup mult_right_dist_sup sup_assoc sup_commute sup_left_commute) also have "... = 1 \ (?v\<^sup>+ * e * ?v\<^sup>\ \ ?v\<^sup>+ \ e * ?v\<^sup>\)" using 7 by simp also have "... = 1 \ (?v\<^sup>\ * e * ?v\<^sup>\ \ ?v\<^sup>+)" by (metis (mono_tags, hide_lams) comp_associative star.circ_loop_fixpoint sup_assoc sup_commute) also have "... \ 1 \ (?v\<^sup>\ * e * ?v\<^sup>\ \ w\<^sup>+)" using comp_inf.mult_right_isotone comp_isotone semiring.add_right_mono star_isotone sup_commute by simp also have "... = (1 \ ?v\<^sup>\ * e * ?v\<^sup>\) \ (1 \ w\<^sup>+)" by (simp add: inf_sup_distrib1) also have "... = 1 \ ?v\<^sup>\ * e * ?v\<^sup>\" by (metis assms(1) inf_commute pseudo_complement sup_bot_right) also have "... = bot" using 8 p_antitone_iff pseudo_complement by simp finally show ?thesis using le_bot p_antitone_iff pseudo_complement by auto qed subsubsection \Exchange gives Spanning Trees\ text \ The lemmas in this section are used to show that the relation after exchange represents a spanning tree. \ lemma inf_star_import: assumes "x \ z" and "univalent z" and "reflexive y" and "regular z" shows "x\<^sup>\ * y \ z\<^sup>\ \ x\<^sup>\ * (y \ z\<^sup>\)" proof - have 1: "y \ x\<^sup>\ * (y \ z\<^sup>\) \ -z\<^sup>\" by (metis assms(4) pp_dist_star shunting_var_p star.circ_loop_fixpoint sup.cobounded2) have "x * -z\<^sup>\ \ z\<^sup>+ \ x * (-z\<^sup>\ \ x\<^sup>T * z\<^sup>+)" by (simp add: dedekind_1) also have "... \ x * (-z\<^sup>\ \ z\<^sup>T * z\<^sup>+)" using assms(1) comp_inf.mult_right_isotone conv_isotone mult_left_isotone mult_right_isotone by simp also have "... \ x * (-z\<^sup>\ \ 1 * z\<^sup>\)" by (metis assms(2) comp_associative comp_inf.mult_right_isotone mult_left_isotone mult_right_isotone) finally have 2: "x * -z\<^sup>\ \ z\<^sup>+ = bot" by (simp add: antisym) have "x * -z\<^sup>\ \ z\<^sup>\ = (x * -z\<^sup>\ \ z\<^sup>+) \ (x * -z\<^sup>\ \ 1)" by (metis comp_inf.semiring.distrib_left star_left_unfold_equal sup_commute) also have "... \ x\<^sup>\ * (y \ z\<^sup>\)" using 2 by (simp add: assms(3) inf.coboundedI2 reflexive_mult_closed star.circ_reflexive) finally have "x * -z\<^sup>\ \ x\<^sup>\ * (y \ z\<^sup>\) \ -z\<^sup>\" by (metis assms(4) pp_dist_star shunting_var_p) hence "x * (x\<^sup>\ * (y \ z\<^sup>\) \ -z\<^sup>\) \ x\<^sup>\ * (y \ z\<^sup>\) \ -z\<^sup>\" by (metis le_supE le_supI mult_left_dist_sup star.circ_loop_fixpoint sup.cobounded1) hence "x\<^sup>\ * y \ x\<^sup>\ * (y \ z\<^sup>\) \ -z\<^sup>\" using 1 by (simp add: star_left_induct) hence "x\<^sup>\ * y \ --z\<^sup>\ \ x\<^sup>\ * (y \ z\<^sup>\)" using shunting_var_p by simp thus ?thesis using order.trans inf.sup_right_isotone pp_increasing by blast qed lemma kruskal_exchange_forest_components_inv: assumes "injective ((w \ -d) \ e)" and "regular d" and "e * top * e = e" and "d \ top * e\<^sup>T * w\<^sup>T\<^sup>\" and "w * e\<^sup>T * top = bot" and "injective w" and "d \ w" and "d \ (w \ -d)\<^sup>T\<^sup>\ * e\<^sup>T * top" shows "forest_components w \ forest_components ((w \ -d) \ e)" proof - let ?v = "w \ -d" let ?w = "?v \ e" let ?f = "forest_components ?w" have 1: "?v * d\<^sup>T * top = bot" by (simp add: assms(6,7) kruskal_exchange_acyclic_inv_3) have 2: "d * e\<^sup>T \ bot" by (metis assms(5,7) conv_bot conv_dist_comp conv_involutive conv_top order.trans inf.absorb2 inf.cobounded2 inf_commute le_bot p_antitone_iff p_top schroeder_4_p top_left_mult_increasing) have "w\<^sup>\ * e\<^sup>T * top = e\<^sup>T * top" by (metis assms(5) conv_bot conv_dist_comp conv_involutive conv_star_commute star.circ_top star_absorb) hence "w\<^sup>\ * e\<^sup>T * top \ -(d\<^sup>T * top)" using 2 by (metis (no_types) comp_associative inf.cobounded2 le_bot p_antitone_iff schroeder_3_p semiring.mult_zero_left) hence 3: "e\<^sup>T * top \ -(w\<^sup>T\<^sup>\ * d\<^sup>T * top)" by (metis conv_star_commute p_antitone_iff schroeder_3_p mult_assoc) have "?v * w\<^sup>T\<^sup>\ * d\<^sup>T * top = ?v * d\<^sup>T * top \ ?v * w\<^sup>T\<^sup>+ * d\<^sup>T * top" by (metis comp_associative mult_left_dist_sup star.circ_loop_fixpoint sup_commute) also have "... \ w * w\<^sup>T\<^sup>+ * d\<^sup>T * top" using 1 by (simp add: mult_left_isotone) also have "... \ w\<^sup>T\<^sup>\ * d\<^sup>T * top" by (metis assms(6) mult_assoc mult_1_left mult_left_isotone) finally have "?v * w\<^sup>T\<^sup>\ * d\<^sup>T * top \ --(w\<^sup>T\<^sup>\ * d\<^sup>T * top)" using p_antitone p_antitone_iff by auto hence 4: "?v\<^sup>T * -(w\<^sup>T\<^sup>\ * d\<^sup>T * top) \ -(w\<^sup>T\<^sup>\ * d\<^sup>T * top)" using comp_associative schroeder_3_p by simp have 5: "injective ?v" using assms(1) conv_dist_sup mult_left_dist_sup mult_right_dist_sup by simp have "?v * ?v\<^sup>T\<^sup>\ * e\<^sup>T * top = ?v * e\<^sup>T * top \ ?v * ?v\<^sup>T\<^sup>+ * e\<^sup>T * top" by (metis comp_associative mult_left_dist_sup star.circ_loop_fixpoint sup_commute) also have "... \ w * e\<^sup>T * top \ ?v * ?v\<^sup>T\<^sup>+ * e\<^sup>T * top" using mult_left_isotone sup_left_isotone by simp also have "... \ w * e\<^sup>T * top \ ?v\<^sup>T\<^sup>\ * e\<^sup>T * top" using 5 by (metis mult_assoc mult_1_left mult_left_isotone sup_right_isotone) finally have "?v * ?v\<^sup>T\<^sup>\ * e\<^sup>T * top \ ?v\<^sup>T\<^sup>\ * e\<^sup>T * top" by (simp add: assms(5)) hence "?v\<^sup>\ * d * top \ ?v\<^sup>T\<^sup>\ * e\<^sup>T * top" by (metis assms(8) star_left_induct sup_least comp_associative mult_right_sub_dist_sup_right sup.orderE vector_top_closed) also have "... \ -(w\<^sup>T\<^sup>\ * d\<^sup>T * top)" using 3 4 by (simp add: comp_associative star_left_induct) also have "... \ -(d\<^sup>T * top)" by (metis p_antitone star.circ_left_top star_outer_increasing mult_assoc) finally have 6: "?v\<^sup>\ * d * top \ -(d\<^sup>T * top)" by simp have "d\<^sup>T * top \ w\<^sup>\ * e * top" by (metis assms(4) comp_associative comp_inf.star.circ_sup_2 comp_isotone conv_dist_comp conv_involutive conv_order conv_star_commute conv_top vector_top_closed) also have "... \ (?v \ d)\<^sup>\ * e * top" by (metis assms(2) comp_inf.semiring.distrib_left maddux_3_11_pp mult_left_isotone star_isotone sup.cobounded2 sup_commute sup_inf_distrib1) also have "... = ?v\<^sup>\ * (d * ?v\<^sup>\)\<^sup>\ * e * top" by (simp add: star_sup_1) also have "... = ?v\<^sup>\ * e * top \ ?v\<^sup>\ * d * ?v\<^sup>\ * (d * ?v\<^sup>\)\<^sup>\ * e * top" by (metis semiring.distrib_right star.circ_unfold_sum star_decompose_1 star_decompose_3 mult_assoc) also have "... \ ?v\<^sup>\ * e * top \ ?v\<^sup>\ * d * top" by (metis comp_associative comp_isotone le_supI mult_left_dist_sup mult_right_dist_sup mult_right_isotone star.circ_decompose_5 star_decompose_3 sup.cobounded1 sup_commute top.extremum) finally have "d\<^sup>T * top \ ?v\<^sup>\ * e * top \ (d\<^sup>T * top \ ?v\<^sup>\ * d * top)" using sup_inf_distrib2 sup_monoid.add_commute by simp hence "d\<^sup>T * top \ ?v\<^sup>\ * e * top" using 6 by (metis inf_commute pseudo_complement sup_monoid.add_0_right) hence 7: "d \ top * e\<^sup>T * ?v\<^sup>T\<^sup>\" by (metis comp_associative conv_dist_comp conv_involutive conv_isotone conv_star_commute conv_top order.trans top_right_mult_increasing) have 8: "?v \ ?f" using forest_components_increasing le_supE by blast have "d \ ?v\<^sup>T\<^sup>\ * e\<^sup>T * top \ top * e\<^sup>T * ?v\<^sup>T\<^sup>\" using 7 assms(8) by simp also have "... = ?v\<^sup>T\<^sup>\ * e\<^sup>T * top * e\<^sup>T * ?v\<^sup>T\<^sup>\" by (metis inf_top_right vector_inf_comp vector_top_closed mult_assoc) also have "... = ?v\<^sup>T\<^sup>\ * e\<^sup>T * ?v\<^sup>T\<^sup>\" by (metis assms(3) comp_associative conv_dist_comp conv_top) also have "... \ ?v\<^sup>T\<^sup>\ * e\<^sup>T * ?f" using 8 by (metis assms(1) forest_components_equivalence cancel_separate_1 conv_dist_comp conv_order mult_left_isotone star_involutive star_isotone) also have "... \ ?v\<^sup>T\<^sup>\ * ?f * ?f" by (metis assms(1) forest_components_equivalence forest_components_increasing conv_isotone le_supE mult_left_isotone mult_right_isotone) also have "... \ ?f * ?f * ?f" by (metis comp_associative comp_isotone conv_dist_sup star.circ_loop_fixpoint star_isotone sup.cobounded1 sup.cobounded2) also have "... = ?f" by (simp add: assms(1) forest_components_equivalence preorder_idempotent) finally have "w \ ?f" using 8 by (metis assms(2) shunting_var_p sup.orderE) thus ?thesis using assms(1) forest_components_idempotent forest_components_isotone by fastforce qed lemma kruskal_spanning_inv: assumes "injective ((f \ -q) \ (f \ q)\<^sup>T \ e)" and "regular q" and "regular e" and "(-h \ --g)\<^sup>\ \ forest_components f" shows "components (-(h \ -e \ -e\<^sup>T) \ g) \ forest_components ((f \ -q) \ (f \ q)\<^sup>T \ e)" proof - let ?f = "(f \ -q) \ (f \ q)\<^sup>T \ e" let ?h = "h \ -e \ -e\<^sup>T" let ?F = "forest_components f" let ?FF = "forest_components ?f" have 1: "equivalence ?FF" using assms(1) forest_components_equivalence by simp hence 2: "?f * ?FF \ ?FF" using order.trans forest_components_increasing mult_left_isotone by blast have 3: "?f\<^sup>T * ?FF \ ?FF" using 1 by (metis forest_components_increasing mult_left_isotone conv_isotone preorder_idempotent) have "(f \ q) * ?FF \ ?f\<^sup>T * ?FF" using conv_dist_sup conv_involutive sup_assoc sup_left_commute mult_left_isotone by simp hence 4: "(f \ q) * ?FF \ ?FF" using 3 order.trans by blast have "(f \ -q) * ?FF \ ?f * ?FF" using le_supI1 mult_left_isotone by simp hence "(f \ -q) * ?FF \ ?FF" using 2 order.trans by blast hence "((f \ q) \ (f \ -q)) * ?FF \ ?FF" using 4 mult_right_dist_sup by simp hence "f * ?FF \ ?FF" by (metis assms(2) maddux_3_11_pp) hence 5: "f\<^sup>\ * ?FF \ ?FF" using star_left_induct_mult_iff by simp have "(f \ -q)\<^sup>T * ?FF \ ?f\<^sup>T * ?FF" by (meson conv_isotone order.trans mult_left_isotone sup.cobounded1) hence 6: "(f \ -q)\<^sup>T * ?FF \ ?FF" using 3 order.trans by blast have "(f \ q)\<^sup>T * ?FF \ ?f * ?FF" by (simp add: mult_left_isotone sup.left_commute sup_assoc) hence "(f \ q)\<^sup>T * ?FF \ ?FF" using 2 order.trans by blast hence "((f \ -q)\<^sup>T \ (f \ q)\<^sup>T) * ?FF \ ?FF" using 6 mult_right_dist_sup by simp hence "f\<^sup>T * ?FF \ ?FF" by (metis assms(2) conv_dist_sup maddux_3_11_pp) hence 7: "?F * ?FF \ ?FF" using 5 star_left_induct mult_assoc by simp have 8: "e * ?FF \ ?FF" using 2 by (simp add: mult_right_dist_sup mult_left_isotone) have "e\<^sup>T * ?FF \ ?f\<^sup>T * ?FF" by (simp add: mult_left_isotone conv_isotone) also have "... \ ?FF * ?FF" using 1 by (metis forest_components_increasing mult_left_isotone conv_isotone) finally have "e\<^sup>T * ?FF \ ?FF" using 1 preorder_idempotent by auto hence 9: "(?F \ e \ e\<^sup>T) * ?FF \ ?FF" using 7 8 mult_right_dist_sup by simp have "components (-?h \ g) \ ((-h \ --g) \ e \ e\<^sup>T)\<^sup>\" by (metis assms(3) comp_inf.mult_left_sub_dist_sup_left conv_complement p_dist_inf pp_dist_inf regular_closed_p star_isotone sup_inf_distrib2 sup_monoid.add_assoc) also have "... \ ((-h \ --g)\<^sup>\ \ e \ e\<^sup>T)\<^sup>\" using star.circ_increasing star_isotone sup_left_isotone by simp also have "... \ (?F \ e \ e\<^sup>T)\<^sup>\" using assms(4) sup_left_isotone star_isotone by simp also have "... \ ?FF" using 1 9 star_left_induct by force finally show ?thesis by simp qed lemma kruskal_exchange_spanning_inv_1: assumes "injective ((w \ -q) \ (w \ q)\<^sup>T)" and "regular (w \ q)" and "components g \ forest_components w" shows "components g \ forest_components ((w \ -q) \ (w \ q)\<^sup>T)" proof - let ?p = "w \ q" let ?w = "(w \ -q) \ ?p\<^sup>T" have 1: "w \ -?p \ forest_components ?w" by (metis forest_components_increasing inf_import_p le_supE) have "w \ ?p \ ?w\<^sup>T" by (simp add: conv_dist_sup) also have "... \ forest_components ?w" by (metis assms(1) conv_isotone forest_components_equivalence forest_components_increasing) finally have "w \ (?p \ -?p) \ forest_components ?w" using 1 inf_sup_distrib1 by simp hence "w \ forest_components ?w" by (metis assms(2) inf_top_right stone) hence 2: "w\<^sup>\ \ forest_components ?w" using assms(1) star_isotone forest_components_star by force hence 3: "w\<^sup>T\<^sup>\ \ forest_components ?w" using assms(1) conv_isotone conv_star_commute forest_components_equivalence by force have "components g \ forest_components w" using assms(3) by simp also have "... \ forest_components ?w * forest_components ?w" using 2 3 mult_isotone by simp also have "... = forest_components ?w" using assms(1) forest_components_equivalence preorder_idempotent by simp finally show ?thesis by simp qed lemma kruskal_exchange_spanning_inv_2: assumes "injective w" and "w\<^sup>\ * e\<^sup>T = e\<^sup>T" and "f \ f\<^sup>T \ (w \ -d \ -d\<^sup>T) \ (w\<^sup>T \ -d \ -d\<^sup>T)" and "d \ forest_components f * e\<^sup>T * top" shows "d \ (w \ -d)\<^sup>T\<^sup>\ * e\<^sup>T * top" proof - have 1: "(w \ -d \ -d\<^sup>T) * (w\<^sup>T \ -d \ -d\<^sup>T) \ 1" using assms(1) comp_isotone order.trans inf.cobounded1 by blast have "d \ forest_components f * e\<^sup>T * top" using assms(4) by simp also have "... \ (f \ f\<^sup>T)\<^sup>\ * (f \ f\<^sup>T)\<^sup>\ * e\<^sup>T * top" by (simp add: comp_isotone star_isotone) also have "... = (f \ f\<^sup>T)\<^sup>\ * e\<^sup>T * top" by (simp add: star.circ_transitive_equal) also have "... \ ((w \ -d \ -d\<^sup>T) \ (w\<^sup>T \ -d \ -d\<^sup>T))\<^sup>\ * e\<^sup>T * top" using assms(3) by (simp add: comp_isotone star_isotone) also have "... = (w\<^sup>T \ -d \ -d\<^sup>T)\<^sup>\ * (w \ -d \ -d\<^sup>T)\<^sup>\ * e\<^sup>T * top" using 1 cancel_separate_1 by simp also have "... \ (w\<^sup>T \ -d \ -d\<^sup>T)\<^sup>\ * w\<^sup>\ * e\<^sup>T * top" by (simp add: inf_assoc mult_left_isotone mult_right_isotone star_isotone) also have "... = (w\<^sup>T \ -d \ -d\<^sup>T)\<^sup>\ * e\<^sup>T * top" using assms(2) mult_assoc by simp also have "... \ (w\<^sup>T \ -d\<^sup>T)\<^sup>\ * e\<^sup>T * top" using mult_left_isotone conv_isotone star_isotone comp_inf.mult_right_isotone inf.cobounded2 inf.left_commute inf.sup_monoid.add_commute by presburger also have "... = (w \ -d)\<^sup>T\<^sup>\ * e\<^sup>T * top" using conv_complement conv_dist_inf by presburger finally show ?thesis by simp qed lemma kruskal_spanning_inv_1: assumes "e \ F" and "regular e" and "components (-h \ g) \ F" and "equivalence F" shows "components (-(h \ -e \ -e\<^sup>T) \ g) \ F" proof - have 1: "F * F \ F" using assms(4) by simp hence 2: "e * F \ F" using assms(1) mult_left_isotone order_lesseq_imp by blast have "e\<^sup>T * F \ F" by (metis assms(1,4) conv_isotone mult_left_isotone preorder_idempotent) hence 3: "(F \ e \ e\<^sup>T) * F \ F" using 1 2 mult_right_dist_sup by simp have "components (-(h \ -e \ -e\<^sup>T) \ g) \ ((-h \ --g) \ e \ e\<^sup>T)\<^sup>\" by (metis assms(2) comp_inf.mult_left_sub_dist_sup_left conv_complement p_dist_inf pp_dist_inf regular_closed_p star_isotone sup_inf_distrib2 sup_monoid.add_assoc) also have "... \ ((-h \ --g)\<^sup>\ \ e \ e\<^sup>T)\<^sup>\" using sup_left_isotone star.circ_increasing star_isotone by simp also have "... \ (F \ e \ e\<^sup>T)\<^sup>\" using assms(3) sup_left_isotone star_isotone by simp also have "... \ F" using 3 assms(4) star_left_induct by force finally show ?thesis by simp qed lemma kruskal_reroot_edge: assumes "injective (e\<^sup>T * top)" and "acyclic w" shows "((w \ -(top * e * w\<^sup>T\<^sup>\)) \ (w \ top * e * w\<^sup>T\<^sup>\)\<^sup>T) * e\<^sup>T = bot" proof - let ?q = "top * e * w\<^sup>T\<^sup>\" let ?p = "w \ ?q" let ?w = "(w \ -?q) \ ?p\<^sup>T" have "(w \ -?q) * e\<^sup>T * top = w * (e\<^sup>T * top \ -?q\<^sup>T)" by (metis comp_associative comp_inf_vector_1 conv_complement covector_complement_closed vector_top_closed) also have "... = w * (e\<^sup>T * top \ -(w\<^sup>\ * e\<^sup>T * top))" by (simp add: conv_dist_comp conv_star_commute mult_assoc) also have "... = bot" by (metis comp_associative comp_inf.semiring.mult_not_zero inf.sup_relative_same_increasing inf_p mult_right_zero star.circ_loop_fixpoint sup_commute sup_left_divisibility) finally have 1: "(w \ -?q) * e\<^sup>T * top = bot" by simp have "?p\<^sup>T * e\<^sup>T * top = (w\<^sup>T \ w\<^sup>\ * e\<^sup>T * top) * e\<^sup>T * top" by (simp add: conv_dist_comp conv_star_commute mult_assoc conv_dist_inf) also have "... = w\<^sup>\ * e\<^sup>T * top \ w\<^sup>T * e\<^sup>T * top" by (simp add: inf_vector_comp vector_export_comp) also have "... = (w\<^sup>\ \ w\<^sup>T) * e\<^sup>T * top" using assms(1) injective_comp_right_dist_inf mult_assoc by simp also have "... = bot" using assms(2) acyclic_star_below_complement_1 semiring.mult_not_zero by blast finally have "?w * e\<^sup>T * top = bot" using 1 mult_right_dist_sup by simp thus ?thesis by (metis star.circ_top star_absorb) qed subsubsection \Exchange gives Minimum Spanning Trees\ text \ The lemmas in this section are used to show that the after exchange we obtain a minimum spanning tree. The following lemmas show that the relation characterising the edge across the cut is an arc. \ lemma kruskal_edge_arc: assumes "equivalence F" and "forest w" and "arc e" and "regular F" and "F \ forest_components (F \ w)" and "regular w" and "w * e\<^sup>T = bot" and "e * F * e = bot" and "e\<^sup>T \ w\<^sup>\" shows "arc (w \ top * e\<^sup>T * w\<^sup>T\<^sup>\ \ F * e\<^sup>T * top \ top * e * -F)" proof (unfold arc_expanded, intro conjI) let ?E = "top * e\<^sup>T * w\<^sup>T\<^sup>\" let ?F = "F * e\<^sup>T * top" let ?G = "top * e * -F" let ?FF = "F * e\<^sup>T * e * F" let ?GG = "-F * e\<^sup>T * e * -F" let ?w = "forest_components (F \ w)" have "F \ w\<^sup>T\<^sup>\ \ forest_components (F \ w) \ w\<^sup>T\<^sup>\" by (simp add: assms(5) inf.coboundedI1) also have "... \ (F \ w)\<^sup>T\<^sup>\ * ((F \ w)\<^sup>\ \ w\<^sup>T\<^sup>\)" apply (rule inf_star_import) apply (simp add: conv_isotone) apply (simp add: assms(2)) apply (simp add: star.circ_reflexive) by (metis assms(6) conv_complement) also have "... \ (F \ w)\<^sup>T\<^sup>\ * (w\<^sup>\ \ w\<^sup>T\<^sup>\)" using comp_inf.mult_left_isotone mult_right_isotone star_isotone by simp also have "... = (F \ w)\<^sup>T\<^sup>\" by (simp add: assms(2) acyclic_star_inf_conv) finally have "w * (F \ w\<^sup>T\<^sup>\) * e\<^sup>T * e \ w * (F \ w)\<^sup>T\<^sup>\ * e\<^sup>T * e" by (simp add: mult_left_isotone mult_right_isotone) also have "... = w * e\<^sup>T * e \ w * (F \ w)\<^sup>T\<^sup>+ * e\<^sup>T * e" by (metis comp_associative mult_left_dist_sup star.circ_loop_fixpoint sup_commute) also have "... = w * (F \ w)\<^sup>T\<^sup>+ * e\<^sup>T * e" by (simp add: assms(7)) also have "... \ w * (F \ w)\<^sup>T\<^sup>+" by (metis assms(3) arc_univalent mult_assoc mult_1_right mult_right_isotone) also have "... \ w * w\<^sup>T * (F \ w)\<^sup>T\<^sup>\" by (simp add: comp_associative conv_isotone mult_left_isotone mult_right_isotone) also have "... \ (F \ w)\<^sup>T\<^sup>\" using assms(2) coreflexive_comp_top_inf inf.sup_right_divisibility by auto also have "... \ F\<^sup>T\<^sup>\" by (simp add: conv_dist_inf star_isotone) finally have 1: "w * (F \ w\<^sup>T\<^sup>\) * e\<^sup>T * e \ F" by (metis assms(1) antisym mult_1_left mult_left_isotone star.circ_plus_same star.circ_reflexive star.left_plus_below_circ star_left_induct_mult_iff) have "F * e\<^sup>T * e \ forest_components (F \ w) * e\<^sup>T * e" by (simp add: assms(5) mult_left_isotone) also have "... \ forest_components w * e\<^sup>T * e" by (simp add: comp_isotone conv_dist_inf star_isotone) also have "... = w\<^sup>T\<^sup>\ * e\<^sup>T * e" by (metis (no_types) assms(7) comp_associative conv_bot conv_dist_comp conv_involutive conv_star_commute star_absorb) also have "... \ w\<^sup>T\<^sup>\" by (metis assms(3) arc_univalent mult_assoc mult_1_right mult_right_isotone) finally have 2: "F * e\<^sup>T * e \ w\<^sup>T\<^sup>\" by simp have "w * F * e\<^sup>T * e \ w * F * e\<^sup>T * e * e\<^sup>T * e" using comp_associative ex231c mult_right_isotone by simp also have "... = w * (F * e\<^sup>T * e \ w\<^sup>T\<^sup>\) * e\<^sup>T * e" using 2 by (simp add: comp_associative inf.absorb1) also have "... \ w * (F \ w\<^sup>T\<^sup>\) * e\<^sup>T * e" by (metis assms(3) arc_univalent mult_assoc mult_1_right mult_right_isotone mult_left_isotone inf.sup_left_isotone) also have "... \ F" using 1 by simp finally have 3: "w * F * e\<^sup>T * e \ F" by simp hence "e\<^sup>T * e * F * w\<^sup>T \ F" by (metis assms(1) conv_dist_comp conv_dist_inf conv_involutive inf.absorb_iff1 mult_assoc) hence "e\<^sup>T * e * F * w\<^sup>T \ e\<^sup>T * top \ F" by (simp add: comp_associative mult_right_isotone) also have "... \ e\<^sup>T * e * F" by (metis conv_involutive dedekind_1 inf_top_left mult_assoc) finally have 4: "e\<^sup>T * e * F * w\<^sup>T \ e\<^sup>T * e * F" by simp have "(top * e)\<^sup>T * (?F \ w\<^sup>T\<^sup>\) = e\<^sup>T * top * e * F * w\<^sup>T\<^sup>\" by (metis assms(1) comp_inf.star.circ_decompose_9 comp_inf.star_star_absorb conv_dist_comp conv_involutive conv_top covector_inf_comp_3 vector_top_closed mult_assoc) also have "... = e\<^sup>T * e * F * w\<^sup>T\<^sup>\" by (simp add: assms(3) arc_top_edge) also have "... \ e\<^sup>T * e * F" using 4 star_right_induct_mult by simp also have "... \ F" by (metis assms(3) arc_injective conv_involutive mult_1_left mult_left_isotone) finally have 5: "(top * e)\<^sup>T * (?F \ w\<^sup>T\<^sup>\) \ F" by simp have "(?F \ w) * w\<^sup>T\<^sup>+ = ?F \ w * w\<^sup>T\<^sup>+" by (simp add: vector_export_comp) also have "... \ ?F \ w\<^sup>T\<^sup>\" by (metis assms(2) comp_associative inf.sup_right_isotone mult_left_isotone star.circ_transitive_equal star_left_unfold_equal sup.absorb_iff2 sup_monoid.add_assoc) also have 6: "... \ top * e * F" using 5 by (metis assms(3) shunt_mapping conv_dist_comp conv_involutive conv_top) finally have 7: "(?F \ w) * w\<^sup>T\<^sup>+ \ top * e * F" by simp have "e\<^sup>T * top * e \ 1" by (simp add: assms(3) point_injective) also have "... \ F" by (simp add: assms(1)) finally have 8: "e * -F * e\<^sup>T \ bot" by (metis p_antitone p_antitone_iff p_bot regular_closed_bot schroeder_3_p schroeder_4_p mult_assoc) have "?FF \ w * (w\<^sup>T\<^sup>+ \ ?GG) * w\<^sup>T \ ?F \ w * (w\<^sup>T\<^sup>+ \ ?GG) * w\<^sup>T" using comp_inf.mult_left_isotone mult_isotone mult_assoc by simp also have "... \ ?F \ w * (w\<^sup>T\<^sup>+ \ ?G) * w\<^sup>T" by (metis assms(3) arc_top_edge comp_inf.star.circ_decompose_9 comp_inf_covector inf.sup_right_isotone inf_le2 mult_left_isotone mult_right_isotone vector_top_closed mult_assoc) also have "... = (?F \ w) * (w\<^sup>T\<^sup>+ \ ?G) * w\<^sup>T" by (simp add: vector_export_comp) also have "... = (?F \ w) * w\<^sup>T\<^sup>+ * (?G\<^sup>T \ w\<^sup>T)" by (simp add: covector_comp_inf covector_comp_inf_1 covector_mult_closed) also have "... \ top * e * F * (?G\<^sup>T \ w\<^sup>T)" using 7 mult_left_isotone by simp also have "... \ top * e * F * ?G\<^sup>T" by (simp add: mult_right_isotone) also have "... = top * e * -F * e\<^sup>T * top" by (metis assms(1) conv_complement conv_dist_comp conv_top equivalence_comp_left_complement mult_assoc) finally have 9: "?FF \ w * (w\<^sup>T\<^sup>+ \ ?GG) * w\<^sup>T = bot" using 8 by (metis comp_associative covector_bot_closed le_bot vector_bot_closed) hence 10: "?FF \ w * (w\<^sup>+ \ ?GG) * w\<^sup>T = bot" using assms(1) comp_associative conv_bot conv_complement conv_dist_comp conv_dist_inf conv_star_commute star.circ_plus_same by fastforce have "(w \ ?E \ ?F \ ?G) * top * (w \ ?E \ ?F \ ?G)\<^sup>T = (?F \ (w \ ?E \ ?G)) * top * ((w \ ?E \ ?G)\<^sup>T \ ?F\<^sup>T)" by (simp add: conv_dist_inf inf_commute inf_left_commute) also have "... = (?F \ (w \ ?E \ ?G)) * top * (w \ ?E \ ?G)\<^sup>T \ ?F\<^sup>T" using covector_comp_inf vector_conv_covector vector_mult_closed vector_top_closed by simp also have "... = ?F \ (w \ ?E \ ?G) * top * (w \ ?E \ ?G)\<^sup>T \ ?F\<^sup>T" by (simp add: vector_export_comp) also have "... = ?F \ top * e * F \ (w \ ?E \ ?G) * top * (w \ ?E \ ?G)\<^sup>T" by (simp add: assms(1) conv_dist_comp inf_assoc inf_commute mult_assoc) also have "... = ?F * e * F \ (w \ ?E \ ?G) * top * (w \ ?E \ ?G)\<^sup>T" by (metis comp_associative comp_inf_covector inf_top.left_neutral) also have "... = ?FF \ (w \ ?E \ ?G) * (top * (w \ ?E \ ?G)\<^sup>T)" using assms(3) arc_top_edge comp_associative by simp also have "... = ?FF \ (w \ ?E \ ?G) * (top * (?G\<^sup>T \ (?E\<^sup>T \ w\<^sup>T)))" by (simp add: conv_dist_inf inf_assoc inf_commute inf_left_commute) also have "... = ?FF \ (w \ ?E \ ?G) * (?G * (?E\<^sup>T \ w\<^sup>T))" by (metis covector_comp_inf_1 covector_top_closed covector_mult_closed inf_top_left) also have "... = ?FF \ (w \ ?E \ ?G) * (?G \ ?E) * w\<^sup>T" by (metis covector_comp_inf_1 covector_top_closed mult_assoc) also have "... = ?FF \ (w \ ?E) * (?G\<^sup>T \ ?G \ ?E) * w\<^sup>T" by (metis covector_comp_inf_1 covector_mult_closed inf.sup_monoid.add_assoc vector_top_closed) also have "... = ?FF \ w * (?E\<^sup>T \ ?G\<^sup>T \ ?G \ ?E) * w\<^sup>T" by (metis covector_comp_inf_1 covector_mult_closed inf.sup_monoid.add_assoc vector_top_closed) also have "... = ?FF \ w * (?E\<^sup>T \ ?E \ (?G\<^sup>T \ ?G)) * w\<^sup>T" by (simp add: inf_commute inf_left_commute) also have "... = ?FF \ w * (?E\<^sup>T \ ?E \ (-F * e\<^sup>T * top \ ?G)) * w\<^sup>T" by (simp add: assms(1) conv_complement conv_dist_comp mult_assoc) also have "... = ?FF \ w * (?E\<^sup>T \ ?E \ (-F * e\<^sup>T * ?G)) * w\<^sup>T" by (metis comp_associative comp_inf_covector inf_top.left_neutral) also have "... = ?FF \ w * (?E\<^sup>T \ ?E \ ?GG) * w\<^sup>T" by (metis assms(3) arc_top_edge comp_associative) also have "... = ?FF \ w * (w\<^sup>\ * e * top \ ?E \ ?GG) * w\<^sup>T" by (simp add: comp_associative conv_dist_comp conv_star_commute) also have "... = ?FF \ w * (w\<^sup>\ * e * ?E \ ?GG) * w\<^sup>T" by (metis comp_associative comp_inf_covector inf_top.left_neutral) also have "... \ ?FF \ w * (w\<^sup>\ * w\<^sup>T\<^sup>\ \ ?GG) * w\<^sup>T" by (metis assms(3) mult_assoc mult_1_right mult_left_isotone mult_right_isotone inf.sup_left_isotone inf.sup_right_isotone arc_expanded) also have "... = ?FF \ w * ((w\<^sup>+ \ 1 \ w\<^sup>T\<^sup>\) \ ?GG) * w\<^sup>T" by (simp add: assms(2) cancel_separate_eq star_left_unfold_equal sup_monoid.add_commute) also have "... = ?FF \ w * ((w\<^sup>+ \ 1 \ w\<^sup>T\<^sup>+) \ ?GG) * w\<^sup>T" using star.circ_plus_one star_left_unfold_equal sup_assoc by presburger also have "... = (?FF \ w * (w\<^sup>+ \ ?GG) * w\<^sup>T) \ (?FF \ w * (1 \ ?GG) * w\<^sup>T) \ (?FF \ w * (w\<^sup>T\<^sup>+ \ ?GG) * w\<^sup>T)" by (simp add: inf_sup_distrib1 inf_sup_distrib2 semiring.distrib_left semiring.distrib_right) also have "... \ w * (1 \ ?GG) * w\<^sup>T" using 9 10 by simp also have "... \ w * w\<^sup>T" by (metis inf.cobounded1 mult_1_right mult_left_isotone mult_right_isotone) also have "... \ 1" by (simp add: assms(2)) finally show "(w \ ?E \ ?F \ ?G) * top * (w \ ?E \ ?F \ ?G)\<^sup>T \ 1" by simp have "w\<^sup>T\<^sup>+ \ -F * e\<^sup>T * e * -F \ w\<^sup>T * F * e\<^sup>T * e * F * w \ w\<^sup>T\<^sup>+ \ ?G \ w\<^sup>T * F * e\<^sup>T * e * F * w" using top_greatest inf.sup_left_isotone inf.sup_right_isotone mult_left_isotone by simp also have "... \ w\<^sup>T\<^sup>+ \ ?G \ w\<^sup>T * ?F" using comp_associative inf.sup_right_isotone mult_right_isotone top.extremum by presburger also have "... = w\<^sup>T * (w\<^sup>T\<^sup>\ \ ?F) \ ?G" using assms(2) inf_assoc inf_commute inf_left_commute univalent_comp_left_dist_inf by simp also have "... \ w\<^sup>T * (top * e * F) \ ?G" using 6 by (metis inf.sup_monoid.add_commute inf.sup_right_isotone mult_right_isotone) also have "... \ top * e * F \ ?G" by (metis comp_associative comp_inf_covector mult_left_isotone top.extremum) also have "... = bot" by (metis assms(3) conv_dist_comp conv_involutive conv_top inf_p mult_right_zero univalent_comp_left_dist_inf) finally have 11: "w\<^sup>T\<^sup>+ \ -F * e\<^sup>T * e * -F \ w\<^sup>T * F * e\<^sup>T * e * F * w = bot" by (simp add: antisym) hence 12: "w\<^sup>+ \ -F * e\<^sup>T * e * -F \ w\<^sup>T * F * e\<^sup>T * e * F * w = bot" using assms(1) comp_associative conv_bot conv_complement conv_dist_comp conv_dist_inf conv_star_commute star.circ_plus_same by fastforce have "(w \ ?E \ ?F \ ?G)\<^sup>T * top * (w \ ?E \ ?F \ ?G) = ((w \ ?E \ ?G)\<^sup>T \ ?F\<^sup>T) * top * (?F \ (w \ ?E \ ?G))" by (simp add: conv_dist_inf inf_commute inf_left_commute) also have "... = (w \ ?E \ ?G)\<^sup>T * ?F * (?F \ (w \ ?E \ ?G))" by (simp add: covector_inf_comp_3 vector_mult_closed) also have "... = (w \ ?E \ ?G)\<^sup>T * (?F \ ?F\<^sup>T) * (w \ ?E \ ?G)" using covector_comp_inf covector_inf_comp_3 vector_conv_covector vector_mult_closed by simp also have "... = (w \ ?E \ ?G)\<^sup>T * (?F \ ?F\<^sup>T) * (w \ ?E) \ ?G" by (simp add: comp_associative comp_inf_covector) also have "... = (w \ ?E \ ?G)\<^sup>T * (?F \ ?F\<^sup>T) * w \ ?E \ ?G" by (simp add: comp_associative comp_inf_covector) also have "... = (?G\<^sup>T \ (?E\<^sup>T \ w\<^sup>T)) * (?F \ ?F\<^sup>T) * w \ ?E \ ?G" by (simp add: conv_dist_inf inf.left_commute inf.sup_monoid.add_commute) also have "... = ?G\<^sup>T \ (?E\<^sup>T \ w\<^sup>T) * (?F \ ?F\<^sup>T) * w \ ?E \ ?G" by (metis (no_types) comp_associative conv_dist_comp conv_top vector_export_comp) also have "... = ?G\<^sup>T \ ?E\<^sup>T \ w\<^sup>T * (?F \ ?F\<^sup>T) * w \ ?E \ ?G" by (metis (no_types) comp_associative inf_assoc conv_dist_comp conv_top vector_export_comp) also have "... = ?E\<^sup>T \ ?E \ (?G\<^sup>T \ ?G) \ w\<^sup>T * (?F \ ?F\<^sup>T) * w" by (simp add: inf_assoc inf.left_commute inf.sup_monoid.add_commute) also have "... = w\<^sup>\ * e * top \ ?E \ (?G\<^sup>T \ ?G) \ w\<^sup>T * (?F \ ?F\<^sup>T) * w" by (simp add: comp_associative conv_dist_comp conv_star_commute) also have "... = w\<^sup>\ * e * ?E \ (?G\<^sup>T \ ?G) \ w\<^sup>T * (?F \ ?F\<^sup>T) * w" by (metis comp_associative comp_inf_covector inf_top.left_neutral) also have "... \ w\<^sup>\ * w\<^sup>T\<^sup>\ \ (?G\<^sup>T \ ?G) \ w\<^sup>T * (?F \ ?F\<^sup>T) * w" by (metis assms(3) mult_assoc mult_1_right mult_left_isotone mult_right_isotone inf.sup_left_isotone arc_expanded) also have "... = w\<^sup>\ * w\<^sup>T\<^sup>\ \ (-F * e\<^sup>T * top \ ?G) \ w\<^sup>T * (?F \ ?F\<^sup>T) * w" by (simp add: assms(1) conv_complement conv_dist_comp mult_assoc) also have "... = w\<^sup>\ * w\<^sup>T\<^sup>\ \ -F * e\<^sup>T * ?G \ w\<^sup>T * (?F \ ?F\<^sup>T) * w" by (metis comp_associative comp_inf_covector inf_top.left_neutral) also have "... = w\<^sup>\ * w\<^sup>T\<^sup>\ \ -F * e\<^sup>T * e * -F \ w\<^sup>T * (?F \ ?F\<^sup>T) * w" by (metis assms(3) arc_top_edge mult_assoc) also have "... = w\<^sup>\ * w\<^sup>T\<^sup>\ \ -F * e\<^sup>T * e * -F \ w\<^sup>T * (?F \ top * e * F) * w" by (simp add: assms(1) conv_dist_comp mult_assoc) also have "... = w\<^sup>\ * w\<^sup>T\<^sup>\ \ -F * e\<^sup>T * e * -F \ w\<^sup>T * (?F * e * F) * w" by (metis comp_associative comp_inf_covector inf_top.left_neutral) also have "... = w\<^sup>\ * w\<^sup>T\<^sup>\ \ -F * e\<^sup>T * e * -F \ w\<^sup>T * F * e\<^sup>T * e * F * w" by (metis assms(3) arc_top_edge mult_assoc) also have "... = (w\<^sup>+ \ 1 \ w\<^sup>T\<^sup>\) \ -F * e\<^sup>T * e * -F \ w\<^sup>T * F * e\<^sup>T * e * F * w" by (simp add: assms(2) cancel_separate_eq star_left_unfold_equal sup_monoid.add_commute) also have "... = (w\<^sup>+ \ 1 \ w\<^sup>T\<^sup>+) \ -F * e\<^sup>T * e * -F \ w\<^sup>T * F * e\<^sup>T * e * F * w" using star.circ_plus_one star_left_unfold_equal sup_assoc by presburger also have "... = (w\<^sup>+ \ -F * e\<^sup>T * e * -F \ w\<^sup>T * F * e\<^sup>T * e * F * w) \ (1 \ -F * e\<^sup>T * e * -F \ w\<^sup>T * F * e\<^sup>T * e * F * w) \ (w\<^sup>T\<^sup>+ \ -F * e\<^sup>T * e * -F \ w\<^sup>T * F * e\<^sup>T * e * F * w)" by (simp add: inf_sup_distrib2) also have "... \ 1" using 11 12 by (simp add: inf.coboundedI1) finally show "(w \ ?E \ ?F \ ?G)\<^sup>T * top * (w \ ?E \ ?F \ ?G) \ 1" by simp have "(w \ -F) * (F \ w\<^sup>T) \ w * w\<^sup>T \ -F * F" by (simp add: mult_isotone) also have "... \ 1 \ -F" using assms(1,2) comp_inf.comp_isotone equivalence_comp_right_complement by auto also have "... = bot" using assms(1) bot_unique pp_isotone pseudo_complement_pp by blast finally have 13: "(w \ -F) * (F \ w\<^sup>T) = bot" by (simp add: antisym) have "w \ ?G \ F * (w \ ?G)" by (metis assms(1) mult_1_left mult_right_dist_sup sup.absorb_iff2) also have "... \ F * (w \ ?G) * w\<^sup>\" by (metis eq_refl le_supE star.circ_back_loop_fixpoint) finally have 14: "w \ ?G \ F * (w \ ?G) * w\<^sup>\" by simp have "w \ top * e * F \ w * (e * F)\<^sup>T * e * F" by (metis (no_types) comp_inf.star_slide dedekind_2 inf_left_commute inf_top_right mult_assoc) also have "... \ F" using 3 assms(1) by (metis comp_associative conv_dist_comp mult_left_isotone preorder_idempotent) finally have "w \ -F \ -(top * e * F)" using order.trans p_shunting_swap pp_increasing by blast also have "... = ?G" by (metis assms(3) comp_mapping_complement conv_dist_comp conv_involutive conv_top) finally have "(w \ -F) * F * (w \ ?G) = (w \ -F \ ?G) * F * (w \ ?G)" by (simp add: inf.absorb1) also have "... \ (w \ -F \ ?G) * F * w" by (simp add: comp_isotone) also have "... \ (w \ -F \ ?G) * forest_components (F \ w) * w" by (simp add: assms(5) mult_left_isotone mult_right_isotone) also have "... \ (w \ -F \ ?G) * (F \ w)\<^sup>T\<^sup>\ * w\<^sup>\ * w" by (simp add: mult_left_isotone mult_right_isotone star_isotone mult_assoc) also have "... \ (w \ -F \ ?G) * (F \ w)\<^sup>T\<^sup>\ * w\<^sup>\" by (simp add: comp_associative mult_right_isotone star.circ_plus_same star.left_plus_below_circ) also have "... = (w \ -F \ ?G) * w\<^sup>\ \ (w \ -F \ ?G) * (F \ w)\<^sup>T\<^sup>+ * w\<^sup>\" by (metis comp_associative inf.sup_monoid.add_assoc mult_left_dist_sup star.circ_loop_fixpoint sup_commute) also have "... \ (w \ -F \ ?G) * w\<^sup>\ \ (w \ -F \ ?G) * (F \ w)\<^sup>T * top" by (metis mult_assoc top_greatest mult_right_isotone sup_right_isotone) also have "... \ (w \ -F \ ?G) * w\<^sup>\ \ (w \ -F) * (F \ w)\<^sup>T * top" using inf.cobounded1 mult_left_isotone sup_right_isotone by blast also have "... \ (w \ ?G) * w\<^sup>\ \ (w \ -F) * (F \ w)\<^sup>T * top" using inf.sup_monoid.add_assoc inf.sup_right_isotone mult_left_isotone sup_commute sup_right_isotone by simp also have "... = (w \ ?G) * w\<^sup>\ \ (w \ -F) * (F \ w\<^sup>T) * top" by (simp add: assms(1) conv_dist_inf) also have "... \ 1 * (w \ ?G) * w\<^sup>\" using 13 by simp also have "... \ F * (w \ ?G) * w\<^sup>\" using assms(1) mult_left_isotone by blast finally have 15: "(w \ -F) * F * (w \ ?G) \ F * (w \ ?G) * w\<^sup>\" by simp have "(w \ F) * F * (w \ ?G) \ F * F * (w \ ?G)" by (simp add: mult_left_isotone) also have "... = F * (w \ ?G)" by (simp add: assms(1) preorder_idempotent) also have "... \ F * (w \ ?G) * w\<^sup>\" by (metis eq_refl le_supE star.circ_back_loop_fixpoint) finally have "(w \ F) * F * (w \ ?G) \ F * (w \ ?G) * w\<^sup>\" by simp hence "((w \ F) \ (w \ -F)) * F * (w \ ?G) \ F * (w \ ?G) * w\<^sup>\" using 15 by (simp add: semiring.distrib_right) hence "w * F * (w \ ?G) \ F * (w \ ?G) * w\<^sup>\" by (metis assms(4) maddux_3_11_pp) hence "w * F * (w \ ?G) * w\<^sup>\ \ F * (w \ ?G) * w\<^sup>\" by (metis (full_types) comp_associative mult_left_isotone star.circ_transitive_equal) hence "w\<^sup>\ * (w \ ?G) \ F * (w \ ?G) * w\<^sup>\" using 14 by (simp add: mult_assoc star_left_induct) hence 16: "w\<^sup>+ \ ?G \ F * (w \ ?G) * w\<^sup>\" by (simp add: covector_comp_inf covector_mult_closed star.circ_plus_same) have 17: "e\<^sup>T * top * e\<^sup>T \ -F" using assms(8) le_bot triple_schroeder_p by simp hence "(top * e)\<^sup>T * e\<^sup>T \ -F" by (simp add: conv_dist_comp) hence 18: "e\<^sup>T \ ?G" by (metis assms(3) shunt_mapping conv_dist_comp conv_involutive conv_top) have "e\<^sup>T \ -F" using 17 by (simp add: assms(3) arc_top_arc) also have "... \ -1" by (simp add: assms(1) p_antitone) finally have "e\<^sup>T \ w\<^sup>\ \ -1" using assms(9) by simp also have "... \ w\<^sup>+" using shunting_var_p star_left_unfold_equal sup_commute by simp finally have "e\<^sup>T \ w\<^sup>+ \ ?G" using 18 by simp hence "e\<^sup>T \ F * (w \ ?G) * w\<^sup>\" using 16 order_trans by blast also have "... = (F * w \ ?G) * w\<^sup>\" by (simp add: comp_associative comp_inf_covector) finally have "e\<^sup>T * top * e\<^sup>T \ (F * w \ ?G) * w\<^sup>\" by (simp add: assms(3) arc_top_arc) hence "e\<^sup>T * top * (e * top)\<^sup>T \ (F * w \ ?G) * w\<^sup>\" by (metis conv_dist_comp conv_top vector_top_closed mult_assoc) hence "e\<^sup>T * top \ (F * w \ ?G) * w\<^sup>\ * e * top" by (metis assms(3) shunt_bijective mult_assoc) hence "(top * e)\<^sup>T * top \ (F * w \ ?G) * w\<^sup>\ * e * top" by (simp add: conv_dist_comp mult_assoc) hence "top \ top * e * (F * w \ ?G) * w\<^sup>\ * e * top" by (metis assms(3) shunt_mapping conv_dist_comp conv_involutive conv_top mult_assoc) also have "... = top * e * F * w * (w\<^sup>\ * e * top \ ?G\<^sup>T)" by (metis comp_associative comp_inf_vector_1) also have "... = top * (w \ (top * e * F)\<^sup>T) * (w\<^sup>\ * e * top \ ?G\<^sup>T)" by (metis comp_inf_vector_1 inf_top.left_neutral) also have "... = top * (w \ ?F) * (w\<^sup>\ * e * top \ ?G\<^sup>T)" by (simp add: assms(1) conv_dist_comp mult_assoc) also have "... = top * (w \ ?F) * (?E\<^sup>T \ ?G\<^sup>T)" by (simp add: comp_associative conv_dist_comp conv_star_commute) also have "... = top * (w \ ?F \ ?G) * ?E\<^sup>T" by (simp add: comp_associative comp_inf_vector_1) also have "... = top * (w \ ?F \ ?G \ ?E) * top" using comp_inf_vector_1 mult_assoc by simp finally show "top * (w \ ?E \ ?F \ ?G) * top = top" by (simp add: inf_commute inf_left_commute top_le) qed lemma kruskal_edge_arc_1: assumes "e \ --h" and "h \ g" and "symmetric g" and "components g \ forest_components w" and "w * e\<^sup>T = bot" shows "e\<^sup>T \ w\<^sup>\" proof - have "w\<^sup>T * top \ -(e\<^sup>T * top)" using assms(5) schroeder_3_p vector_bot_closed mult_assoc by fastforce hence 1: "w\<^sup>T * top \ e\<^sup>T * top = bot" using pseudo_complement by simp have "e\<^sup>T \ e\<^sup>T * top \ --h\<^sup>T" using assms(1) conv_complement conv_isotone top_right_mult_increasing by fastforce also have "... \ e\<^sup>T * top \ --g" by (metis assms(2,3) inf.sup_right_isotone pp_isotone conv_isotone) also have "... \ e\<^sup>T * top \ components g" using inf.sup_right_isotone star.circ_increasing by simp also have "... \ e\<^sup>T * top \ forest_components w" using assms(4) comp_inf.mult_right_isotone by simp also have "... = (e\<^sup>T * top \ w\<^sup>T\<^sup>\) * w\<^sup>\" by (simp add: inf_assoc vector_export_comp) also have "... = (e\<^sup>T * top \ 1) * w\<^sup>\ \ (e\<^sup>T * top \ w\<^sup>T\<^sup>+) * w\<^sup>\" by (metis inf_sup_distrib1 semiring.distrib_right star_left_unfold_equal) also have "... \ w\<^sup>\ \ (e\<^sup>T * top \ w\<^sup>T\<^sup>+) * w\<^sup>\" by (metis inf_le2 mult_1_left mult_left_isotone sup_left_isotone) also have "... \ w\<^sup>\ \ (e\<^sup>T * top \ w\<^sup>T) * top" using comp_associative comp_inf.mult_right_isotone sup_right_isotone mult_right_isotone top.extremum vector_export_comp by presburger also have "... = w\<^sup>\" using 1 inf.sup_monoid.add_commute inf_vector_comp by simp finally show ?thesis by simp qed lemma kruskal_edge_between_components_1: assumes "equivalence F" and "mapping (top * e)" shows "F \ -(w \ top * e\<^sup>T * w\<^sup>T\<^sup>\ \ F * e\<^sup>T * top \ top * e * -F)" proof - let ?d = "w \ top * e\<^sup>T * w\<^sup>T\<^sup>\ \ F * e\<^sup>T * top \ top * e * -F" have "?d \ F \ F * e\<^sup>T * top \ F" by (meson inf_le1 inf_le2 le_infI order_trans) also have "... \ (F * e\<^sup>T * top)\<^sup>T * F" by (simp add: mult_assoc vector_restrict_comp_conv) also have "... = top * e * F * F" by (simp add: assms(1) comp_associative conv_dist_comp conv_star_commute) also have "... = top * e * F" using assms(1) preorder_idempotent mult_assoc by fastforce finally have "?d \ F \ top * e * F \ top * e * -F" by (simp add: le_infI1) also have "... = top * e * F \ -(top * e * F)" using assms(2) conv_dist_comp total_conv_surjective comp_mapping_complement by simp finally show ?thesis by (metis inf_p le_bot p_antitone_iff pseudo_complement) qed lemma kruskal_edge_between_components_2: assumes "forest_components f \ -d" and "injective f" and "f \ f\<^sup>T \ w \ w\<^sup>T" shows "f \ f\<^sup>T \ (w \ -d \ -d\<^sup>T) \ (w\<^sup>T \ -d \ -d\<^sup>T)" proof - let ?F = "forest_components f" have "?F\<^sup>T \ -d\<^sup>T" using assms(1) conv_complement conv_order by fastforce hence 1: "?F \ -d\<^sup>T" by (simp add: conv_dist_comp conv_star_commute) have "equivalence ?F" using assms(2) forest_components_equivalence by simp hence "f \ f\<^sup>T \ ?F" by (metis conv_dist_inf forest_components_increasing inf.absorb_iff2 sup.boundedI) also have "... \ -d \ -d\<^sup>T" using 1 assms(1) by simp finally have "f \ f\<^sup>T \ -d \ -d\<^sup>T" by simp thus ?thesis by (metis assms(3) inf_sup_distrib2 le_inf_iff) qed end subsection \Related Structures\ text \ Stone algebras can be expanded to Stone-Kleene relation algebras by reusing some operations. \ sublocale stone_algebra < comp_inf: stone_kleene_relation_algebra where star = "\x . top" and one = top and times = inf and conv = id apply unfold_locales by simp text \ Every bounded linear order can be expanded to a Stone algebra, which can be expanded to a Stone relation algebra, which can be expanded to a Stone-Kleene relation algebra. \ class linorder_stone_kleene_relation_algebra_expansion = linorder_stone_relation_algebra_expansion + star + assumes star_def [simp]: "x\<^sup>\ = top" begin subclass kleene_algebra apply unfold_locales apply simp apply (simp add: min.coboundedI1 min.commute) by (simp add: min.coboundedI1) subclass stone_kleene_relation_algebra apply unfold_locales by simp end text \ A Kleene relation algebra is based on a relation algebra. \ class kleene_relation_algebra = relation_algebra + stone_kleene_relation_algebra end diff --git a/thys/Stone_Relation_Algebras/Relation_Algebras.thy b/thys/Stone_Relation_Algebras/Relation_Algebras.thy --- a/thys/Stone_Relation_Algebras/Relation_Algebras.thy +++ b/thys/Stone_Relation_Algebras/Relation_Algebras.thy @@ -1,1817 +1,1906 @@ (* Title: Relation Algebras Author: Walter Guttmann Maintainer: Walter Guttmann *) section \Relation Algebras\ text \ The main structures introduced in this theory are Stone relation algebras. They generalise Tarski's relation algebras \cite{Tarski1941} by weakening the Boolean algebra lattice structure to a Stone algebra. Our motivation is to generalise relation-algebraic methods from unweighted graphs to weighted graphs. Unlike unweighted graphs, weighted graphs do not form a Boolean algebra because there is no complement operation on the edge weights. However, edge weights form a Stone algebra, and matrices over edge weights (that is, weighted graphs) form a Stone relation algebra. The development in this theory is described in our papers \cite{Guttmann2016c,Guttmann2017b}. Our main application there is the verification of Prim's minimum spanning tree algorithm. Related work about fuzzy relations \cite{Goguen1967,Winter2001b}, Dedekind categories \cite{KawaharaFurusawa2001} and rough relations \cite{Comer1993,Pawlak1996} is also discussed in these papers. In particular, Stone relation algebras do not assume that the underlying lattice is complete or a Heyting algebra, and they do not assume that composition has residuals. We proceed in two steps. First, we study the positive fragment in the form of single-object bounded distributive allegories \cite{FreydScedrov1990}. Second, we extend these structures by a pseudocomplement operation with additional axioms to obtain Stone relation algebras. Tarski's relation algebras are then obtained by a simple extension that imposes a Boolean algebra. See, for example, \cite{BirdMoor1997,HirschHodkinson2002,Maddux1996,Maddux2006,Schmidt2011,SchmidtStroehlein1993} for further details about relations and relation algebras, and \cite{AndrekaMikulas2011,BredihinSchein1978} for algebras of relations with a smaller signature. \ theory Relation_Algebras imports Stone_Algebras.P_Algebras Semirings begin subsection \Single-Object Bounded Distributive Allegories\ text \ We start with developing bounded distributive allegories. The following definitions concern properties of relations that require converse in addition to lattice and semiring operations. \ class conv = fixes conv :: "'a \ 'a" ("_\<^sup>T" [100] 100) class bounded_distrib_allegory_signature = inf + sup + times + conv + bot + top + one + ord begin subclass times_one_ord . subclass times_top . abbreviation total_var :: "'a \ bool" where "total_var x \ 1 \ x * x\<^sup>T" abbreviation surjective_var :: "'a \ bool" where "surjective_var x \ 1 \ x\<^sup>T * x" abbreviation univalent :: "'a \ bool" where "univalent x \ x\<^sup>T * x \ 1" abbreviation injective :: "'a \ bool" where "injective x \ x * x\<^sup>T \ 1" abbreviation mapping :: "'a \ bool" where "mapping x \ univalent x \ total x" abbreviation bijective :: "'a \ bool" where "bijective x \ injective x \ surjective x" abbreviation point :: "'a \ bool" where "point x \ vector x \ bijective x" abbreviation arc :: "'a \ bool" where "arc x \ bijective (x * top) \ bijective (x\<^sup>T * top)" (* earlier name: atom *) abbreviation symmetric :: "'a \ bool" where "symmetric x \ x\<^sup>T = x" abbreviation antisymmetric :: "'a \ bool" where "antisymmetric x \ x \ x\<^sup>T \ 1" abbreviation asymmetric :: "'a \ bool" where "asymmetric x \ x \ x\<^sup>T = bot" abbreviation linear :: "'a \ bool" where "linear x \ x \ x\<^sup>T = top" abbreviation equivalence :: "'a \ bool" where "equivalence x \ preorder x \ symmetric x" abbreviation order :: "'a \ bool" where "order x \ preorder x \ antisymmetric x" abbreviation linear_order :: "'a \ bool" where "linear_order x \ order x \ linear x" end text \ We reuse the relation algebra axioms given in \cite{Maddux1996} except for one -- see lemma \conv_complement_sub\ below -- which we replace with the Dedekind rule (or modular law) \dedekind_1\. The Dedekind rule or variants of it are known from \cite{BirdMoor1997,FreydScedrov1990,KawaharaFurusawaMori1999,SchmidtStroehlein1993}. We add \comp_left_zero\, which follows in relation algebras but not in the present setting. The main change is that only a bounded distributive lattice is required, not a Boolean algebra. \ class bounded_distrib_allegory = bounded_distrib_lattice + times + one + conv + assumes comp_associative : "(x * y) * z = x * (y * z)" assumes comp_right_dist_sup : "(x \ y) * z = (x * z) \ (y * z)" assumes comp_left_zero [simp]: "bot * x = bot" assumes comp_left_one [simp]: "1 * x = x" assumes conv_involutive [simp]: "x\<^sup>T\<^sup>T = x" assumes conv_dist_sup : "(x \ y)\<^sup>T = x\<^sup>T \ y\<^sup>T" assumes conv_dist_comp : "(x * y)\<^sup>T = y\<^sup>T * x\<^sup>T" assumes dedekind_1 : "x * y \ z \ x * (y \ (x\<^sup>T * z))" begin subclass bounded_distrib_allegory_signature . text \ Many properties of relation algebras already follow in bounded distributive allegories. \ lemma conv_isotone: "x \ y \ x\<^sup>T \ y\<^sup>T" by (metis conv_dist_sup le_iff_sup) lemma conv_order: "x \ y \ x\<^sup>T \ y\<^sup>T" using conv_isotone by fastforce lemma conv_bot [simp]: "bot\<^sup>T = bot" using conv_order bot_unique by force lemma conv_top [simp]: "top\<^sup>T = top" by (metis conv_involutive conv_order eq_iff top_greatest) lemma conv_dist_inf: "(x \ y)\<^sup>T = x\<^sup>T \ y\<^sup>T" apply (rule antisym) using conv_order apply simp by (metis conv_order conv_involutive inf.boundedI inf.cobounded1 inf.cobounded2) lemma conv_inf_bot_iff: "bot = x\<^sup>T \ y \ bot = x \ y\<^sup>T" using conv_dist_inf conv_bot by fastforce lemma conv_one [simp]: "1\<^sup>T = 1" by (metis comp_left_one conv_dist_comp conv_involutive) lemma comp_left_dist_sup: "(x * y) \ (x * z) = x * (y \ z)" by (metis comp_right_dist_sup conv_involutive conv_dist_sup conv_dist_comp) lemma comp_right_isotone: "x \ y \ z * x \ z * y" by (simp add: comp_left_dist_sup sup.absorb_iff1) lemma comp_left_isotone: "x \ y \ x * z \ y * z" by (metis comp_right_dist_sup le_iff_sup) lemma comp_isotone: "x \ y \ w \ z \ x * w \ y * z" using comp_left_isotone comp_right_isotone order.trans by blast lemma comp_left_subdist_inf: "(x \ y) * z \ x * z \ y * z" by (simp add: comp_left_isotone) lemma comp_left_increasing_sup: "x * y \ (x \ z) * y" by (simp add: comp_left_isotone) lemma comp_right_subdist_inf: "x * (y \ z) \ x * y \ x * z" by (simp add: comp_right_isotone) lemma comp_right_increasing_sup: "x * y \ x * (y \ z)" by (simp add: comp_right_isotone) lemma comp_right_zero [simp]: "x * bot = bot" by (metis comp_left_zero conv_dist_comp conv_involutive) lemma comp_right_one [simp]: "x * 1 = x" by (metis comp_left_one conv_dist_comp conv_involutive) lemma comp_left_conjugate: "conjugate (\y . x * y) (\y . x\<^sup>T * y)" apply (unfold conjugate_def, intro allI) by (metis comp_right_zero bot.extremum_unique conv_involutive dedekind_1 inf.commute) lemma comp_right_conjugate: "conjugate (\y . y * x) (\y . y * x\<^sup>T)" apply (unfold conjugate_def, intro allI) by (metis comp_left_conjugate[unfolded conjugate_def] conv_inf_bot_iff conv_dist_comp conv_involutive) text \ We still obtain a semiring structure. \ subclass bounded_idempotent_semiring by (unfold_locales) (auto simp: comp_right_isotone comp_right_dist_sup comp_associative comp_left_dist_sup) sublocale inf: semiring_0 sup bot inf by (unfold_locales, auto simp: inf_sup_distrib2 inf_sup_distrib1 inf_assoc) lemma schroeder_1: "x * y \ z = bot \ x\<^sup>T * z \ y = bot" using abel_semigroup.commute comp_left_conjugate conjugate_def inf.abel_semigroup_axioms by fastforce lemma schroeder_2: "x * y \ z = bot \ z * y\<^sup>T \ x = bot" by (metis comp_right_conjugate conjugate_def inf_commute) lemma comp_additive: "additive (\y . x * y) \ additive (\y . x\<^sup>T * y) \ additive (\y . y * x) \ additive (\y . y * x\<^sup>T)" by (simp add: comp_left_dist_sup additive_def comp_right_dist_sup) lemma dedekind_2: "y * x \ z \ (y \ (z * x\<^sup>T)) * x" by (metis conv_dist_inf conv_order conv_dist_comp dedekind_1) text \ The intersection with a vector can still be exported from the first argument of a composition, and many other properties of vectors and covectors continue to hold. \ lemma vector_inf_comp: "vector x \ (x \ y) * z = x \ (y * z)" apply (rule antisym) apply (metis comp_left_subdist_inf comp_right_isotone inf.sup_left_isotone order_lesseq_imp top_greatest) by (metis comp_left_isotone comp_right_isotone dedekind_2 inf_commute inf_mono order_refl order_trans top_greatest) lemma vector_inf_closed: "vector x \ vector y \ vector (x \ y)" by (simp add: vector_inf_comp) lemma vector_inf_one_comp: "vector x \ (x \ 1) * y = x \ y" by (simp add: vector_inf_comp) lemma covector_inf_comp_1: assumes "vector x" shows "(y \ x\<^sup>T) * z = (y \ x\<^sup>T) * (x \ z)" proof - have "(y \ x\<^sup>T) * z \ (y \ x\<^sup>T) * (z \ ((y\<^sup>T \ x) * top))" by (metis inf_top_right dedekind_1 conv_dist_inf conv_involutive) also have "... \ (y \ x\<^sup>T) * (x \ z)" by (metis assms comp_left_isotone comp_right_isotone inf_le2 inf_mono order_refl inf_commute) finally show ?thesis by (simp add: comp_right_isotone antisym) qed lemma covector_inf_comp_2: assumes "vector x" shows "y * (x \ z) = (y \ x\<^sup>T) * (x \ z)" proof - have "y * (x \ z) \ (y \ (top * (x \ z)\<^sup>T)) * (x \ z)" by (metis dedekind_2 inf_top_right) also have "... \ (y \ x\<^sup>T) * (x \ z)" by (metis assms comp_left_isotone conv_dist_comp conv_order conv_top eq_refl inf_le1 inf_mono) finally show ?thesis using comp_left_subdist_inf antisym by auto qed lemma covector_inf_comp_3: "vector x \ (y \ x\<^sup>T) * z = y * (x \ z)" by (metis covector_inf_comp_1 covector_inf_comp_2) lemma covector_inf_closed: "covector x \ covector y \ covector (x \ y)" by (metis comp_right_subdist_inf inf.antisym top_left_mult_increasing) lemma vector_conv_covector: "vector v \ covector (v\<^sup>T)" by (metis conv_dist_comp conv_involutive conv_top) lemma covector_conv_vector: "covector v \ vector (v\<^sup>T)" by (simp add: vector_conv_covector) lemma covector_comp_inf: "covector z \ x * (y \ z) = x * y \ z" apply (rule antisym) apply (metis comp_isotone comp_right_subdist_inf inf.boundedE inf.boundedI inf.cobounded2 top.extremum) by (metis comp_left_isotone comp_right_isotone dedekind_1 inf_commute inf_mono order_refl order_trans top_greatest) lemma vector_restrict_comp_conv: "vector x \ x \ y \ x\<^sup>T * y" by (metis covector_inf_comp_3 eq_refl inf.sup_monoid.add_commute inf_top_right le_supE sup.orderE top_left_mult_increasing) lemma covector_restrict_comp_conv: "covector x \ y \ x \ y * x\<^sup>T" by (metis conv_dist_comp conv_dist_inf conv_order conv_top inf.sup_monoid.add_commute vector_restrict_comp_conv) lemma covector_comp_inf_1: "covector x \ (y \ x) * z = y * (x\<^sup>T \ z)" using covector_conv_vector covector_inf_comp_3 by fastforce text \ We still have two ways to represent surjectivity and totality. \ lemma surjective_var: "surjective x \ surjective_var x" proof assume "surjective x" thus "surjective_var x" by (metis dedekind_2 comp_left_one inf_absorb2 top_greatest) next assume "surjective_var x" hence "x\<^sup>T * (x * top) = top" by (metis comp_left_isotone comp_associative comp_left_one top_le) thus "surjective x" by (metis comp_right_isotone conv_top conv_dist_comp conv_involutive top_greatest top_le) qed lemma total_var: "total x \ total_var x" by (metis conv_top conv_dist_comp conv_involutive surjective_var) lemma surjective_conv_total: "surjective x \ total (x\<^sup>T)" by (metis conv_top conv_dist_comp conv_involutive) lemma total_conv_surjective: "total x \ surjective (x\<^sup>T)" by (simp add: surjective_conv_total) lemma injective_conv_univalent: "injective x \ univalent (x\<^sup>T)" by simp lemma univalent_conv_injective: "univalent x \ injective (x\<^sup>T)" by simp text \ We continue with studying further closure properties. \ lemma univalent_bot_closed: "univalent bot" by simp lemma univalent_one_closed: "univalent 1" by simp lemma univalent_inf_closed: "univalent x \ univalent (x \ y)" by (metis comp_left_subdist_inf comp_right_subdist_inf conv_dist_inf inf.cobounded1 order_lesseq_imp) lemma univalent_mult_closed: assumes "univalent x" and "univalent y" shows "univalent (x * y)" proof - have "(x * y)\<^sup>T * x \ y\<^sup>T" by (metis assms(1) comp_left_isotone comp_right_one conv_one conv_order comp_associative conv_dist_comp conv_involutive) thus ?thesis by (metis assms(2) comp_left_isotone comp_associative dual_order.trans) qed lemma injective_bot_closed: "injective bot" by simp lemma injective_one_closed: "injective 1" by simp lemma injective_inf_closed: "injective x \ injective (x \ y)" by (metis conv_dist_inf injective_conv_univalent univalent_inf_closed) lemma injective_mult_closed: "injective x \ injective y \ injective (x * y)" by (metis injective_conv_univalent conv_dist_comp univalent_mult_closed) lemma mapping_one_closed: "mapping 1" by simp lemma mapping_mult_closed: "mapping x \ mapping y \ mapping (x * y)" by (simp add: comp_associative univalent_mult_closed) lemma bijective_one_closed: "bijective 1" by simp lemma bijective_mult_closed: "bijective x \ bijective y \ bijective (x * y)" by (metis injective_mult_closed comp_associative) lemma bijective_conv_mapping: "bijective x \ mapping (x\<^sup>T)" by (simp add: surjective_conv_total) lemma mapping_conv_bijective: "mapping x \ bijective (x\<^sup>T)" by (simp add: total_conv_surjective) lemma reflexive_inf_closed: "reflexive x \ reflexive y \ reflexive (x \ y)" by simp lemma reflexive_conv_closed: "reflexive x \ reflexive (x\<^sup>T)" using conv_isotone by force lemma coreflexive_inf_closed: "coreflexive x \ coreflexive (x \ y)" by (simp add: le_infI1) lemma coreflexive_conv_closed: "coreflexive x \ coreflexive (x\<^sup>T)" using conv_order by force lemma coreflexive_symmetric: "coreflexive x \ symmetric x" by (metis comp_right_one comp_right_subdist_inf conv_dist_inf conv_dist_comp conv_involutive dedekind_1 inf.absorb1 inf_absorb2) lemma transitive_inf_closed: "transitive x \ transitive y \ transitive (x \ y)" by (meson comp_left_subdist_inf inf.cobounded1 inf.sup_mono inf_le2 mult_right_isotone order.trans) lemma transitive_conv_closed: "transitive x \ transitive (x\<^sup>T)" using conv_order conv_dist_comp by fastforce lemma dense_conv_closed: "dense_rel x \ dense_rel (x\<^sup>T)" using conv_order conv_dist_comp by fastforce lemma idempotent_conv_closed: "idempotent x \ idempotent (x\<^sup>T)" by (metis conv_dist_comp) lemma preorder_inf_closed: "preorder x \ preorder y \ preorder (x \ y)" using transitive_inf_closed by auto lemma preorder_conv_closed: "preorder x \ preorder (x\<^sup>T)" by (simp add: reflexive_conv_closed transitive_conv_closed) lemma symmetric_bot_closed: "symmetric bot" by simp lemma symmetric_one_closed: "symmetric 1" by simp lemma symmetric_top_closed: "symmetric top" by simp lemma symmetric_inf_closed: "symmetric x \ symmetric y \ symmetric (x \ y)" by (simp add: conv_dist_inf) lemma symmetric_sup_closed: "symmetric x \ symmetric y \ symmetric (x \ y)" by (simp add: conv_dist_sup) lemma symmetric_conv_closed: "symmetric x \ symmetric (x\<^sup>T)" by simp lemma one_inf_conv: "1 \ x = 1 \ x\<^sup>T" by (metis conv_dist_inf coreflexive_symmetric inf.cobounded1 symmetric_one_closed) lemma antisymmetric_bot_closed: "antisymmetric bot" by simp lemma antisymmetric_one_closed: "antisymmetric 1" by simp lemma antisymmetric_inf_closed: "antisymmetric x \ antisymmetric (x \ y)" by (rule order_trans[where y="x \ x\<^sup>T"]) (simp_all add: conv_isotone inf.coboundedI2 inf.sup_assoc) lemma antisymmetric_conv_closed: "antisymmetric x \ antisymmetric (x\<^sup>T)" by (simp add: inf_commute) lemma asymmetric_bot_closed: "asymmetric bot" by simp lemma asymmetric_inf_closed: "asymmetric x \ asymmetric (x \ y)" by (metis conv_dist_inf inf.mult_zero_left inf.left_commute inf_assoc) lemma asymmetric_conv_closed: "asymmetric x \ asymmetric (x\<^sup>T)" by (simp add: inf_commute) lemma linear_top_closed: "linear top" by simp lemma linear_sup_closed: "linear x \ linear (x \ y)" by (metis conv_dist_sup sup_assoc sup_commute sup_top_right) lemma linear_reflexive: "linear x \ reflexive x" by (metis one_inf_conv inf.distrib_left inf.cobounded2 inf.orderE reflexive_top_closed sup.idem) lemma linear_conv_closed: "linear x \ linear (x\<^sup>T)" by (simp add: sup_commute) lemma linear_comp_closed: assumes "linear x" and "linear y" shows "linear (x * y)" proof - have "reflexive y" by (simp add: assms(2) linear_reflexive) hence "x \ x\<^sup>T \ x * y \ y\<^sup>T * x\<^sup>T" by (metis case_split_left case_split_right le_supI sup.cobounded1 sup.cobounded2 sup.idem reflexive_conv_closed) thus ?thesis by (simp add: assms(1) conv_dist_comp top_le) qed lemma equivalence_one_closed: "equivalence 1" by simp lemma equivalence_top_closed: "equivalence top" by simp lemma equivalence_inf_closed: "equivalence x \ equivalence y \ equivalence (x \ y)" using conv_dist_inf preorder_inf_closed by auto lemma equivalence_conv_closed: "equivalence x \ equivalence (x\<^sup>T)" by simp lemma order_one_closed: "order 1" by simp lemma order_inf_closed: "order x \ order y \ order (x \ y)" using antisymmetric_inf_closed transitive_inf_closed by auto lemma order_conv_closed: "order x \ order (x\<^sup>T)" by (simp add: inf_commute reflexive_conv_closed transitive_conv_closed) lemma linear_order_conv_closed: "linear_order x \ linear_order (x\<^sup>T)" using equivalence_top_closed conv_dist_sup inf_commute reflexive_conv_closed transitive_conv_closed by force text \ We show a fact about equivalences. \ lemma equivalence_comp_dist_inf: "equivalence x \ x * y \ x * z = x * (y \ x * z)" by (metis antisym comp_right_subdist_inf dedekind_1 eq_iff inf.absorb1 inf.absorb2 mult_1_right mult_assoc) text \ The following result generalises the fact that composition with a test amounts to intersection with the corresponding vector. Both tests and vectors can be used to represent sets as relations. \ lemma coreflexive_comp_top_inf: "coreflexive x \ x * top \ y = x * y" apply (rule antisym) apply (metis comp_left_isotone comp_left_one coreflexive_symmetric dedekind_1 inf_top_left order_trans) using comp_left_isotone comp_right_isotone by fastforce lemma coreflexive_comp_top_inf_one: "coreflexive x \ x * top \ 1 = x" by (simp add: coreflexive_comp_top_inf) lemma coreflexive_comp_inf: "coreflexive x \ coreflexive y \ x * y = x \ y" by (metis (full_types) coreflexive_comp_top_inf coreflexive_comp_top_inf_one inf.mult_assoc inf.absorb2) lemma coreflexive_comp_inf_comp: assumes "coreflexive x" and "coreflexive y" shows "(x * z) \ (y * z) = (x \ y) * z" proof - have "(x * z) \ (y * z) = x * top \ z \ y * top \ z" using assms coreflexive_comp_top_inf inf_assoc by auto also have "... = x * top \ y * top \ z" by (simp add: inf.commute inf.left_commute) also have "... = (x \ y) * top \ z" by (metis assms coreflexive_comp_inf coreflexive_comp_top_inf mult_assoc) also have "... = (x \ y) * z" by (simp add: assms(1) coreflexive_comp_top_inf coreflexive_inf_closed) finally show ?thesis . qed lemma coreflexive_idempotent: "coreflexive x \ idempotent x" by (simp add: coreflexive_comp_inf) lemma coreflexive_univalent: "coreflexive x \ univalent x" by (simp add: coreflexive_idempotent coreflexive_symmetric) lemma coreflexive_injective: "coreflexive x \ injective x" by (simp add: coreflexive_idempotent coreflexive_symmetric) lemma coreflexive_commutative: "coreflexive x \ coreflexive y \ x * y = y * x" by (simp add: coreflexive_comp_inf inf.commute) lemma coreflexive_dedekind: "coreflexive x \ coreflexive y \ coreflexive z \ x * y \ z \ x * (y \ x * z)" by (simp add: coreflexive_comp_inf inf.coboundedI1 inf.left_commute) text \ Also the equational version of the Dedekind rule continues to hold. \ lemma dedekind_eq: "x * y \ z = (x \ (z * y\<^sup>T)) * (y \ (x\<^sup>T * z)) \ z" proof (rule antisym) have "x * y \ z \ x * (y \ (x\<^sup>T * z)) \ z" by (simp add: dedekind_1) also have "... \ (x \ (z * (y \ (x\<^sup>T * z))\<^sup>T)) * (y \ (x\<^sup>T * z)) \ z" by (simp add: dedekind_2) also have "... \ (x \ (z * y\<^sup>T)) * (y \ (x\<^sup>T * z)) \ z" by (metis comp_left_isotone comp_right_isotone inf_mono conv_order inf.cobounded1 order_refl) finally show "x * y \ z \ (x \ (z * y\<^sup>T)) * (y \ (x\<^sup>T * z)) \ z" . next show "(x \ (z * y\<^sup>T)) * (y \ (x\<^sup>T * z)) \ z \ x * y \ z" using comp_isotone inf.sup_left_isotone by auto qed lemma dedekind: "x * y \ z \ (x \ (z * y\<^sup>T)) * (y \ (x\<^sup>T * z))" by (metis dedekind_eq inf.cobounded1) lemma vector_export_comp: "(x * top \ y) * z = x * top \ y * z" proof - have "vector (x * top)" by (simp add: comp_associative) thus ?thesis by (simp add: vector_inf_comp) qed lemma vector_export_comp_unit: "(x * top \ 1) * y = x * top \ y" by (simp add: vector_export_comp) text \ We solve a few exercises from \cite{SchmidtStroehlein1993}. \ lemma ex231a [simp]: "(1 \ x * x\<^sup>T) * x = x" by (metis inf.cobounded1 inf.idem inf_right_idem comp_left_one conv_one coreflexive_comp_top_inf dedekind_eq) lemma ex231b [simp]: "x * (1 \ x\<^sup>T * x) = x" by (metis conv_dist_comp conv_dist_inf conv_involutive conv_one ex231a) lemma ex231c: "x \ x * x\<^sup>T * x" by (metis comp_left_isotone ex231a inf_le2) lemma ex231d: "x \ x * top * x" by (metis comp_left_isotone comp_right_isotone top_greatest order_trans ex231c) lemma ex231e [simp]: "x * top * x * top = x * top" by (metis ex231d antisym comp_associative mult_right_isotone top.extremum) lemma arc_injective: "arc x \ injective x" by (metis conv_dist_inf conv_involutive inf.absorb2 top_right_mult_increasing univalent_inf_closed) lemma arc_conv_closed: "arc x \ arc (x\<^sup>T)" by simp lemma arc_univalent: "arc x \ univalent x" using arc_conv_closed arc_injective univalent_conv_injective by blast text \ The following result generalises \cite[Exercise 2]{Oliveira2009}. It is used to show that the while-loop preserves injectivity of the constructed tree. \ lemma injective_sup: assumes "injective t" and "e * t\<^sup>T \ 1" and "injective e" shows "injective (t \ e)" proof - have "(t \ e) * (t \ e)\<^sup>T = t * t\<^sup>T \ t * e\<^sup>T \ e * t\<^sup>T \ e * e\<^sup>T" by (simp add: comp_left_dist_sup conv_dist_sup semiring.distrib_right sup.assoc) thus ?thesis using assms coreflexive_symmetric conv_dist_comp by fastforce qed lemma injective_inv: "injective t \ e * t\<^sup>T = bot \ arc e \ injective (t \ e)" using arc_injective injective_sup bot_least by blast lemma univalent_sup: "univalent t \ e\<^sup>T * t \ 1 \ univalent e \ univalent (t \ e)" by (metis injective_sup conv_dist_sup conv_involutive) lemma point_injective: "arc x \ x\<^sup>T * top * x \ 1" by (metis conv_top comp_associative conv_dist_comp conv_involutive vector_top_closed) lemma vv_transitive: "vector v \ (v * v\<^sup>T) * (v * v\<^sup>T) \ v * v\<^sup>T" by (metis comp_associative comp_left_isotone comp_right_isotone top_greatest) lemma epm_3: assumes "e \ w" and "injective w" shows "e = w \ top * e" proof - have "w \ top * e \ w * e\<^sup>T * e" by (metis (no_types, lifting) inf.absorb2 top.extremum dedekind_2 inf.commute) also have "... \ w * w\<^sup>T * e" by (simp add: assms(1) conv_isotone mult_left_isotone mult_right_isotone) also have "... \ e" using assms(2) coreflexive_comp_top_inf inf.sup_right_divisibility by blast finally show ?thesis by (simp add: assms(1) top_left_mult_increasing antisym) qed lemma comp_inf_vector: "x * (y \ z * top) = (x \ top * z\<^sup>T) * y" by (metis conv_top covector_inf_comp_3 comp_associative conv_dist_comp inf.commute vector_top_closed) lemma inf_vector_comp: "(x \ y * top) * z = y * top \ x * z" using inf.commute vector_export_comp by auto lemma comp_inf_covector: "x * (y \ top * z) = x * y \ top * z" by (simp add: covector_comp_inf covector_mult_closed) text \ Well-known distributivity properties of univalent and injective relations over meet continue to hold. \ lemma univalent_comp_left_dist_inf: assumes "univalent x" shows "x * (y \ z) = x * y \ x * z" proof (rule antisym) show "x * (y \ z) \ x * y \ x * z" by (simp add: comp_right_isotone) next have "x * y \ x * z \ (x \ x * z * y\<^sup>T) * (y \ x\<^sup>T * x * z)" by (metis comp_associative dedekind) also have "... \ x * (y \ x\<^sup>T * x * z)" by (simp add: comp_left_isotone) also have "... \ x * (y \ 1 * z)" using assms comp_left_isotone comp_right_isotone inf.sup_right_isotone by blast finally show "x * y \ x * z \ x * (y \ z)" by simp qed lemma injective_comp_right_dist_inf: "injective z \ (x \ y) * z = x * z \ y * z" by (metis univalent_comp_left_dist_inf conv_dist_comp conv_involutive conv_dist_inf) lemma vector_covector: "vector v \ vector w \ v \ w\<^sup>T = v * w\<^sup>T" by (metis covector_comp_inf inf_top_left vector_conv_covector) lemma comp_inf_vector_1: "(x \ top * y) * z = x * (z \ (top * y)\<^sup>T)" by (simp add: comp_inf_vector conv_dist_comp) text \ The shunting properties for bijective relations and mappings continue to hold. \ lemma shunt_bijective: assumes "bijective z" shows "x \ y * z \ x * z\<^sup>T \ y" proof assume "x \ y * z" hence "x * z\<^sup>T \ y * z * z\<^sup>T" by (simp add: mult_left_isotone) also have "... \ y" using assms comp_associative mult_right_isotone by fastforce finally show "x * z\<^sup>T \ y" . next assume 1: "x * z\<^sup>T \ y" have "x = x \ top * z" by (simp add: assms) also have "... \ x * z\<^sup>T * z" by (metis dedekind_2 inf_commute inf_top.right_neutral) also have "... \ y * z" using 1 by (simp add: mult_left_isotone) finally show "x \ y * z" . qed lemma shunt_mapping: "mapping z \ x \ z * y \ z\<^sup>T * x \ y" by (metis shunt_bijective mapping_conv_bijective conv_order conv_dist_comp conv_involutive) lemma bijective_reverse: assumes "bijective p" and "bijective q" shows "p \ r * q \ q \ r\<^sup>T * p" proof - have "p \ r * q \ p * q\<^sup>T \ r" by (simp add: assms(2) shunt_bijective) also have "... \ q\<^sup>T \ p\<^sup>T * r" by (metis assms(1) conv_dist_comp conv_involutive conv_order shunt_bijective) also have "... \ q \ r\<^sup>T * p" using conv_dist_comp conv_isotone by fastforce finally show ?thesis by simp qed lemma arc_expanded: "arc x \ x * top * x\<^sup>T \ 1 \ x\<^sup>T * top * x \ 1 \ top * x * top = top" by (metis conv_top comp_associative conv_dist_comp conv_involutive vector_top_closed) lemma arc_top_arc: assumes "arc x" shows "x * top * x = x" by (metis assms epm_3 top_right_mult_increasing vector_inf_comp vector_mult_closed vector_top_closed) lemma arc_top_edge: assumes "arc x" shows "x\<^sup>T * top * x = x\<^sup>T * x" proof - have "x\<^sup>T = x\<^sup>T * top \ top * x\<^sup>T" using assms epm_3 top_right_mult_increasing by simp thus ?thesis by (metis comp_inf_vector_1 conv_dist_comp conv_involutive conv_top inf.absorb1 top_right_mult_increasing) qed +lemma arc_eq_1: + assumes "arc x" + shows "x = x * x\<^sup>T * x" +proof - + have "x * x\<^sup>T * x \ x * top * x" + by (simp add: mult_left_isotone mult_right_isotone) + also have "... \ x" + by (simp add: assms arc_top_arc) + finally have "x * x\<^sup>T * x \ x" + by simp + thus ?thesis + by (simp add: antisym ex231c) +qed + +lemma arc_eq_2: + assumes "arc x" + shows "x\<^sup>T = x\<^sup>T * x * x\<^sup>T" + using arc_eq_1 assms conv_involutive by fastforce + end subsection \Single-Object Pseudocomplemented Distributive Allegories\ text \ We extend single-object bounded distributive allegories by a pseudocomplement operation. The following definitions concern properties of relations that require a pseudocomplement. \ class relation_algebra_signature = bounded_distrib_allegory_signature + uminus begin abbreviation irreflexive :: "'a \ bool" where "irreflexive x \ x \ -1" abbreviation strict_linear :: "'a \ bool" where "strict_linear x \ x \ x\<^sup>T = -1" abbreviation strict_order :: "'a \ bool" where "strict_order x \ irreflexive x \ transitive x" abbreviation linear_strict_order :: "'a \ bool" where "linear_strict_order x \ strict_order x \ strict_linear x" text \ The following variants are useful for the graph model. \ abbreviation pp_mapping :: "'a \ bool" where "pp_mapping x \ univalent x \ total (--x)" abbreviation pp_bijective :: "'a \ bool" where "pp_bijective x \ injective x \ surjective (--x)" abbreviation pp_point :: "'a \ bool" where "pp_point x \ vector x \ pp_bijective x" abbreviation pp_arc :: "'a \ bool" where "pp_arc x \ pp_bijective (x * top) \ pp_bijective (x\<^sup>T * top)" end class pd_allegory = bounded_distrib_allegory + p_algebra begin subclass relation_algebra_signature . subclass pd_algebra .. lemma conv_complement_1: "-(x\<^sup>T) \ (-x)\<^sup>T = (-x)\<^sup>T" by (metis conv_dist_inf conv_order bot_least conv_involutive pseudo_complement sup.absorb2 sup.cobounded2) lemma conv_complement: "(-x)\<^sup>T = -(x\<^sup>T)" by (metis conv_complement_1 conv_dist_sup conv_involutive sup_commute) lemma conv_complement_sub_inf [simp]: "x\<^sup>T * -(x * y) \ y = bot" by (metis comp_left_zero conv_dist_comp conv_involutive dedekind_1 inf_import_p inf_p inf_right_idem ppp pseudo_complement regular_closed_bot) lemma conv_complement_sub_leq: "x\<^sup>T * -(x * y) \ -y" using pseudo_complement conv_complement_sub_inf by blast lemma conv_complement_sub [simp]: "x\<^sup>T * -(x * y) \ -y = -y" by (simp add: conv_complement_sub_leq sup.absorb2) lemma complement_conv_sub: "-(y * x) * x\<^sup>T \ -y" by (metis conv_complement conv_complement_sub_leq conv_order conv_dist_comp) text \ The following so-called Schr\"oder equivalences, or De Morgan's Theorem K, hold only with a pseudocomplemented element on both right-hand sides. \ lemma schroeder_3_p: "x * y \ -z \ x\<^sup>T * z \ -y" using pseudo_complement schroeder_1 by auto lemma schroeder_4_p: "x * y \ -z \ z * y\<^sup>T \ -x" using pseudo_complement schroeder_2 by auto lemma comp_pp_semi_commute: "x * --y \ --(x * y)" using conv_complement_sub_leq schroeder_3_p by fastforce text \ The following result looks similar to a property of (anti)domain. \ lemma p_comp_pp [simp]: "-(x * --y) = -(x * y)" using comp_pp_semi_commute comp_right_isotone inf.eq_iff p_antitone pp_increasing by fastforce lemma pp_comp_semi_commute: "--x * y \ --(x * y)" using complement_conv_sub schroeder_4_p by fastforce lemma p_pp_comp [simp]: "-(--x * y) = -(x * y)" using pp_comp_semi_commute comp_left_isotone inf.eq_iff p_antitone pp_increasing by fastforce lemma pp_comp_subdist: "--x * --y \ --(x * y)" by (simp add: p_antitone_iff) lemma theorem24xxiii: "x * y \ -(x * z) = x * (y \ -z) \ -(x * z)" proof - have "x * y \ -(x * z) \ x * (y \ (x\<^sup>T * -(x * z)))" by (simp add: dedekind_1) also have "... \ x * (y \ -z)" using comp_right_isotone conv_complement_sub_leq inf.sup_right_isotone by auto finally show ?thesis using comp_right_subdist_inf antisym inf.coboundedI2 inf.commute by auto qed text \ Even in Stone relation algebras, we do not obtain the backward implication in the following result. \ lemma vector_complement_closed: "vector x \ vector (-x)" by (metis complement_conv_sub conv_top eq_iff top_right_mult_increasing) lemma covector_complement_closed: "covector x \ covector (-x)" by (metis conv_complement_sub_leq conv_top eq_iff top_left_mult_increasing) lemma covector_vector_comp: "vector v \ -v\<^sup>T * v = bot" by (metis conv_bot conv_complement conv_complement_sub_inf conv_dist_comp conv_involutive inf_top.right_neutral) lemma irreflexive_bot_closed: "irreflexive bot" by simp lemma irreflexive_inf_closed: "irreflexive x \ irreflexive (x \ y)" by (simp add: le_infI1) lemma irreflexive_sup_closed: "irreflexive x \ irreflexive y \ irreflexive (x \ y)" by simp lemma irreflexive_conv_closed: "irreflexive x \ irreflexive (x\<^sup>T)" using conv_complement conv_isotone by fastforce lemma reflexive_complement_irreflexive: "reflexive x \ irreflexive (-x)" by (simp add: p_antitone) lemma irreflexive_complement_reflexive: "irreflexive x \ reflexive (-x)" by (simp add: p_antitone_iff) lemma symmetric_complement_closed: "symmetric x \ symmetric (-x)" by (simp add: conv_complement) lemma asymmetric_irreflexive: "asymmetric x \ irreflexive x" by (metis inf.mult_not_zero inf.left_commute inf.right_idem inf.sup_monoid.add_commute pseudo_complement one_inf_conv) lemma linear_asymmetric: "linear x \ asymmetric (-x)" using conv_complement p_top by force lemma strict_linear_sup_closed: "strict_linear x \ strict_linear y \ strict_linear (x \ y)" by (metis (mono_tags, hide_lams) conv_dist_sup sup.right_idem sup_assoc sup_commute) lemma strict_linear_irreflexive: "strict_linear x \ irreflexive x" using sup_left_divisibility by blast lemma strict_linear_conv_closed: "strict_linear x \ strict_linear (x\<^sup>T)" by (simp add: sup_commute) lemma strict_order_var: "strict_order x \ asymmetric x \ transitive x" by (metis asymmetric_irreflexive comp_right_one irreflexive_conv_closed conv_dist_comp dual_order.trans pseudo_complement schroeder_3_p) lemma strict_order_bot_closed: "strict_order bot" by simp lemma strict_order_inf_closed: "strict_order x \ strict_order y \ strict_order (x \ y)" using inf.coboundedI1 transitive_inf_closed by auto lemma strict_order_conv_closed: "strict_order x \ strict_order (x\<^sup>T)" using irreflexive_conv_closed transitive_conv_closed by blast lemma order_strict_order: assumes "order x" shows "strict_order (x \ -1)" proof (rule conjI) show 1: "irreflexive (x \ -1)" by simp have "antisymmetric (x \ -1)" using antisymmetric_inf_closed assms by blast hence "(x \ -1) * (x \ -1) \ 1 \ (x \ -1 \ (x \ -1)\<^sup>T) * (x \ -1 \ (x \ -1)\<^sup>T)" using 1 by (metis (no_types) coreflexive_symmetric irreflexive_inf_closed coreflexive_transitive dedekind_1 inf_idem mult_1_right semiring.mult_not_zero strict_order_var) also have "... = (x \ x\<^sup>T \ -1) * (x \ x\<^sup>T \ -1)" by (simp add: conv_complement conv_dist_inf inf.absorb2 inf.sup_monoid.add_assoc) also have "... = bot" using assms inf.antisym reflexive_conv_closed by fastforce finally have "(x \ -1) * (x \ -1) \ -1" using le_bot pseudo_complement by blast thus "transitive (x \ -1)" by (meson assms comp_isotone inf.boundedI inf.cobounded1 inf.order_lesseq_imp) qed lemma strict_order_order: "strict_order x \ order (x \ 1)" apply (unfold strict_order_var, intro conjI) apply simp apply (simp add: mult_left_dist_sup mult_right_dist_sup sup.absorb2) using conv_dist_sup coreflexive_bot_closed sup.absorb2 sup_inf_distrib2 by fastforce lemma linear_strict_order_conv_closed: "linear_strict_order x \ linear_strict_order (x\<^sup>T)" by (simp add: irreflexive_conv_closed sup_monoid.add_commute transitive_conv_closed) lemma linear_order_strict_order: "linear_order x \ linear_strict_order (x \ -1)" apply (rule conjI) using order_strict_order apply simp by (metis conv_complement conv_dist_inf coreflexive_symmetric eq_iff inf.absorb2 inf.distrib_left inf.sup_monoid.add_commute top.extremum) lemma regular_conv_closed: "regular x \ regular (x\<^sup>T)" by (metis conv_complement) text \ We show a number of facts about equivalences. \ lemma equivalence_comp_left_complement: "equivalence x \ x * -x = -x" apply (rule antisym) apply (metis conv_complement_sub_leq preorder_idempotent) using mult_left_isotone by fastforce lemma equivalence_comp_right_complement: "equivalence x \ -x * x = -x" by (metis equivalence_comp_left_complement conv_complement conv_dist_comp) text \ The pseudocomplement of tests is given by the following operation. \ abbreviation coreflexive_complement :: "'a \ 'a" ("_ ''" [80] 80) where "x ' \ -x \ 1" lemma coreflexive_comp_top_coreflexive_complement: "coreflexive x \ (x * top)' = x '" by (metis coreflexive_comp_top_inf_one inf.commute inf_import_p) lemma coreflexive_comp_inf_complement: "coreflexive x \ (x * y) \ -z = (x * y) \ -(x * z)" by (metis coreflexive_comp_top_inf inf.sup_relative_same_increasing inf_import_p inf_le1) lemma double_coreflexive_complement: "x '' = (-x)'" using inf.sup_monoid.add_commute inf_import_p by auto lemma coreflexive_pp_dist_comp: assumes "coreflexive x" and "coreflexive y" shows "(x * y)'' = x '' * y ''" proof - have "(x * y)'' = --(x * y) \ 1" by (simp add: double_coreflexive_complement) also have "... = --x \ --y \ 1" by (simp add: assms coreflexive_comp_inf) also have "... = (--x \ 1) * (--y \ 1)" by (simp add: coreflexive_comp_inf inf.left_commute inf.sup_monoid.add_assoc) also have "... = x '' * y ''" by (simp add: double_coreflexive_complement) finally show ?thesis . qed lemma coreflexive_pseudo_complement: "coreflexive x \ x \ y = bot \ x \ y '" by (simp add: pseudo_complement) lemma pp_bijective_conv_mapping: "pp_bijective x \ pp_mapping (x\<^sup>T)" by (simp add: conv_complement surjective_conv_total) lemma pp_arc_expanded: "pp_arc x \ x * top * x\<^sup>T \ 1 \ x\<^sup>T * top * x \ 1 \ top * --x * top = top" proof assume 1: "pp_arc x" have 2: "x * top * x\<^sup>T \ 1" using 1 by (metis comp_associative conv_dist_comp equivalence_top_closed vector_top_closed) have 3: "x\<^sup>T * top * x \ 1" using 1 by (metis conv_dist_comp conv_involutive equivalence_top_closed vector_top_closed mult_assoc) have 4: "x\<^sup>T \ x\<^sup>T * x * x\<^sup>T" by (metis conv_involutive ex231c) have "top = --(top * x) * top" using 1 by (metis conv_complement conv_dist_comp conv_involutive equivalence_top_closed) also have "... \ --(top * x\<^sup>T * top * x) * top" using 1 by (metis eq_refl mult_assoc p_comp_pp p_pp_comp) also have "... = (top * --(x * top) \ --(top * x\<^sup>T * top * x)) * top" using 1 by simp also have "... = top * (--(x * top) \ --(top * x\<^sup>T * top * x)) * top" by (simp add: covector_complement_closed covector_comp_inf covector_mult_closed) also have "... = top * --(x * top \ top * x\<^sup>T * top * x) * top" by simp also have "... = top * --(x * top * x\<^sup>T * top * x) * top" by (metis comp_associative comp_inf_covector inf_top.left_neutral) also have "... \ top * --(x * top * x\<^sup>T * x * x\<^sup>T * top * x) * top" using 4 by (metis comp_associative comp_left_isotone comp_right_isotone pp_isotone) also have "... \ top * --(x * x\<^sup>T * top * x) * top" using 2 by (metis comp_associative comp_left_isotone comp_right_isotone pp_isotone comp_left_one) also have "... \ top * --x * top" using 3 by (metis comp_associative comp_left_isotone comp_right_isotone pp_isotone comp_right_one) finally show "x * top * x\<^sup>T \ 1 \ x\<^sup>T * top * x \ 1 \ top * --x * top = top" using 2 3 top_le by blast next assume "x * top * x\<^sup>T \ 1 \ x\<^sup>T * top * x \ 1 \ top * --x * top = top" thus "pp_arc x" apply (intro conjI) apply (metis comp_associative conv_dist_comp equivalence_top_closed vector_top_closed) apply (metis comp_associative mult_right_isotone top_le pp_comp_semi_commute) apply (metis conv_dist_comp coreflexive_symmetric vector_conv_covector vector_top_closed mult_assoc) by (metis conv_complement conv_dist_comp equivalence_top_closed inf.orderE inf_top.left_neutral mult_right_isotone pp_comp_semi_commute) qed text \ The following operation represents states with infinite executions of non-strict computations. \ abbreviation N :: "'a \ 'a" where "N x \ -(-x * top) \ 1" lemma N_comp: "N x * y = -(-x * top) \ y" by (simp add: vector_mult_closed vector_complement_closed vector_inf_one_comp) lemma N_comp_top [simp]: "N x * top = -(-x * top)" by (simp add: N_comp) lemma vector_N_pp: "vector x \ N x = --x \ 1" by (simp add: vector_complement_closed) lemma N_vector_pp [simp]: "N (x * top) = --(x * top) \ 1" by (simp add: comp_associative vector_complement_closed) lemma N_vector_top_pp [simp]: "N (x * top) * top = --(x * top)" by (metis N_comp_top comp_associative vector_top_closed vector_complement_closed) lemma N_below_inf_one_pp: "N x \ --x \ 1" using inf.sup_left_isotone p_antitone top_right_mult_increasing by auto lemma N_below_pp: "N x \ --x" using N_below_inf_one_pp by auto lemma N_comp_N: "N x * N y = -(-x * top) \ -(-y * top) \ 1" by (simp add: N_comp inf.mult_assoc) lemma N_bot [simp]: "N bot = bot" by simp lemma N_top [simp]: "N top = 1" by simp lemma n_split_omega_mult_pp: "xs * --xo = xo \ vector xo \ N top * xo = xs * N xo * top" by (metis N_top N_vector_top_pp comp_associative comp_left_one) text \ Many of the following results have been derived for verifying Prim's minimum spanning tree algorithm. \ lemma ee: assumes "vector v" and "e \ v * -v\<^sup>T" shows "e * e = bot" proof - have "e * v \ bot" by (metis assms covector_vector_comp comp_associative mult_left_isotone mult_right_zero) thus ?thesis by (metis assms(2) bot_unique comp_associative mult_right_isotone semiring.mult_not_zero) qed lemma et: assumes "vector v" and "e \ v * -v\<^sup>T" and "t \ v * v\<^sup>T" shows "e * t = bot" and "e * t\<^sup>T = bot" proof - have "e * t \ v * -v\<^sup>T * v * v\<^sup>T" using assms(2-3) comp_isotone mult_assoc by fastforce thus "e * t = bot" by (simp add: assms(1) covector_vector_comp le_bot mult_assoc) next have "t\<^sup>T \ v * v\<^sup>T" using assms(3) conv_order conv_dist_comp by fastforce hence "e * t\<^sup>T \ v * -v\<^sup>T * v * v\<^sup>T" by (metis assms(2) comp_associative comp_isotone) thus "e * t\<^sup>T = bot" by (simp add: assms(1) covector_vector_comp le_bot mult_assoc) qed lemma ve_dist: assumes "e \ v * -v\<^sup>T" and "vector v" and "arc e" shows "(v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T = v * v\<^sup>T \ v * v\<^sup>T * e \ e\<^sup>T * v * v\<^sup>T \ e\<^sup>T * e" proof - have "e \ v * top" using assms(1) comp_right_isotone dual_order.trans top_greatest by blast hence "v * top * e = v * top * (v * top \ e)" by (simp add: inf.absorb2) also have "... = (v * top \ top * v\<^sup>T) * e" using assms(2) covector_inf_comp_3 vector_conv_covector by force also have "... = v * top * v\<^sup>T * e" by (metis assms(2) inf_top_right vector_inf_comp) also have "... = v * v\<^sup>T * e" by (simp add: assms(2)) finally have 1: "v * top * e = v * v\<^sup>T * e" . have "e\<^sup>T * top * e \ e\<^sup>T * top * e * e\<^sup>T * e" using ex231c comp_associative mult_right_isotone by auto also have "... \ e\<^sup>T * e" by (metis assms(3) coreflexive_comp_top_inf le_infE mult_semi_associative point_injective) finally have 2: "e\<^sup>T * top * e = e\<^sup>T * e" by (simp add: inf.antisym mult_left_isotone top_right_mult_increasing) have "(v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T = (v \ e\<^sup>T * top) * (v\<^sup>T \ top * e)" by (simp add: conv_dist_comp conv_dist_sup) also have "... = v * v\<^sup>T \ v * top * e \ e\<^sup>T * top * v\<^sup>T \ e\<^sup>T * top * top * e" by (metis semiring.distrib_left semiring.distrib_right sup_assoc mult_assoc) also have "... = v * v\<^sup>T \ v * top * e \ (v * top * e)\<^sup>T \ e\<^sup>T * top * e" by (simp add: comp_associative conv_dist_comp) also have "... = v * v\<^sup>T \ v * v\<^sup>T * e \ (v * v\<^sup>T * e)\<^sup>T \ e\<^sup>T * e" using 1 2 by simp finally show ?thesis by (simp add: comp_associative conv_dist_comp) qed lemma ev: "vector v \ e \ v * -v\<^sup>T \ e * v = bot" by (metis covector_vector_comp antisym bot_least comp_associative mult_left_isotone mult_right_zero) lemma vTeT: "vector v \ e \ v * -v\<^sup>T \ v\<^sup>T * e\<^sup>T = bot" using conv_bot ev conv_dist_comp by fastforce text \ The following result is used to show that the while-loop of Prim's algorithm preserves that the constructed tree is a subgraph of g. \ lemma prim_subgraph_inv: assumes "e \ v * -v\<^sup>T \ g" and "t \ v * v\<^sup>T \ g" shows "t \ e \ ((v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T) \ g" proof (rule sup_least) have "t \ ((v \ e\<^sup>T * top) * v\<^sup>T) \ g" using assms(2) le_supI1 mult_right_dist_sup by auto also have "... \ ((v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T) \ g" using comp_right_isotone conv_dist_sup inf.sup_left_isotone by auto finally show "t \ ((v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T) \ g" . next have "e \ v * top" by (meson assms(1) inf.boundedE mult_right_isotone order.trans top.extremum) hence "e \ v * top \ top * e" by (simp add: top_left_mult_increasing) also have "... = v * top * e" by (metis inf_top_right vector_export_comp) finally have "e \ v * top * e \ g" using assms(1) by auto also have "... = v * (e\<^sup>T * top)\<^sup>T \ g" by (simp add: comp_associative conv_dist_comp) also have "... \ v * (v \ e\<^sup>T * top)\<^sup>T \ g" by (simp add: conv_dist_sup mult_left_dist_sup sup.assoc sup.orderI) also have "... \ (v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T \ g" using inf.sup_left_isotone mult_right_sub_dist_sup_left by auto finally show "e \ ((v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T) \ g" . qed text \ The following result shows how to apply the Schr\"oder equivalence to the middle factor in a composition of three relations. Again the elements on the right-hand side need to be pseudocomplemented. \ lemma triple_schroeder_p: "x * y * z \ -w \ x\<^sup>T * w * z\<^sup>T \ -y" using mult_assoc p_antitone_iff schroeder_3_p schroeder_4_p by auto text \ The rotation versions of the Schr\"oder equivalences continue to hold, again with pseudocomplemented elements on the right-hand side. \ lemma schroeder_5_p: "x * y \ -z \ y * z\<^sup>T \ -x\<^sup>T" using schroeder_3_p schroeder_4_p by auto lemma schroeder_6_p: "x * y \ -z \ z\<^sup>T * x \ -y\<^sup>T" using schroeder_3_p schroeder_4_p by auto lemma vector_conv_compl: "vector v \ top * -v\<^sup>T = -v\<^sup>T" by (simp add: covector_complement_closed vector_conv_covector) text \ Composition commutes, relative to the diversity relation. \ lemma comp_commute_below_diversity: "x * y \ -1 \ y * x \ -1" by (metis comp_right_one conv_dist_comp conv_one schroeder_3_p schroeder_4_p) lemma comp_injective_below_complement: "injective y \ -x * y \ -(x * y)" by (metis p_antitone_iff comp_associative comp_right_isotone comp_right_one schroeder_4_p) lemma comp_univalent_below_complement: "univalent x \ x * -y \ -(x * y)" by (metis p_inf pseudo_complement semiring.mult_zero_right univalent_comp_left_dist_inf) text \ Bijective relations and mappings can be exported from a pseudocomplement. \ lemma comp_bijective_complement: "bijective y \ -x * y = -(x * y)" using comp_injective_below_complement complement_conv_sub antisym shunt_bijective by blast lemma comp_mapping_complement: "mapping x \ x * -y = -(x * y)" by (metis (full_types) comp_bijective_complement conv_complement conv_dist_comp conv_involutive total_conv_surjective) text \ The following facts are used in the correctness proof of Kruskal's minimum spanning tree algorithm. \ lemma kruskal_injective_inv: assumes "injective f" and "covector q" and "q * f\<^sup>T \ q" and "e \ q" and "q * f\<^sup>T \ -e" and "injective e" and "q\<^sup>T * q \ f\<^sup>T * f \ 1" shows "injective ((f \ -q) \ (f \ q)\<^sup>T \ e)" proof - have 1: "(f \ -q) * (f \ -q)\<^sup>T \ 1" by (simp add: assms(1) injective_inf_closed) have 2: "(f \ -q) * (f \ q) \ 1" proof - have 21: "bot = q * f\<^sup>T \ - q" by (metis assms(3) inf.sup_monoid.add_assoc inf.sup_right_divisibility inf_import_p inf_p) have "(f \ -q) * (f \ q) \ -q * f \ q" by (metis assms(2) comp_inf_covector comp_isotone inf.cobounded2 inf.left_idem) also have "... = bot" using 21 schroeder_2 by auto finally show ?thesis by (simp add: bot_unique) qed have 3: "(f \ -q) * e\<^sup>T \ 1" proof - have "(f \ -q) * e\<^sup>T \ -q * e\<^sup>T" by (simp add: mult_left_isotone) also have "... = bot" by (metis assms(2,4) bot_unique conv_bot conv_complement covector_complement_closed p_antitone p_bot regular_closed_bot schroeder_5_p) finally show ?thesis by (simp add: bot_unique) qed have 4: "(f \ q)\<^sup>T * (f \ -q)\<^sup>T \ 1" using 2 conv_dist_comp conv_isotone by force have 5: "(f \ q)\<^sup>T * (f \ q) \ 1" proof - have "(f \ q)\<^sup>T * (f \ q) \ q\<^sup>T * q \ f\<^sup>T * f" by (simp add: conv_isotone mult_isotone) also have "... \ 1" by (simp add: assms(7)) finally show ?thesis by simp qed have 6: "(f \ q)\<^sup>T * e\<^sup>T \ 1" proof - have "f\<^sup>T * e\<^sup>T \ -q\<^sup>T" using assms(5) schroeder_5_p by simp hence "(f \ q)\<^sup>T * e\<^sup>T = bot" by (metis assms(2,5) conv_bot conv_dist_comp covector_comp_inf inf.absorb1 inf.cobounded2 inf.sup_monoid.add_commute inf_left_commute inf_p schroeder_4_p) thus ?thesis by (simp add: bot_unique) qed have 7: "e * (f \ -q)\<^sup>T \ 1" using 3 conv_dist_comp coreflexive_symmetric by fastforce have 8: "e * (f \ q) \ 1" using 6 conv_dist_comp coreflexive_symmetric by fastforce have 9: "e * e\<^sup>T \ 1" by (simp add: assms(6)) have "((f \ -q) \ (f \ q)\<^sup>T \ e) * ((f \ -q) \ (f \ q)\<^sup>T \ e)\<^sup>T = (f \ -q) * (f \ -q)\<^sup>T \ (f \ -q) * (f \ q) \ (f \ -q) * e\<^sup>T \ (f \ q)\<^sup>T * (f \ -q)\<^sup>T \ (f \ q)\<^sup>T * (f \ q) \ (f \ q)\<^sup>T * e\<^sup>T \ e * (f \ -q)\<^sup>T \ e * (f \ q) \ e * e\<^sup>T" using comp_left_dist_sup comp_right_dist_sup conv_dist_sup sup.assoc by simp also have "... \ 1" using 1 2 3 4 5 6 7 8 9 by simp finally show ?thesis by simp qed lemma kruskal_exchange_injective_inv_1: assumes "injective f" and "covector q" and "q * f\<^sup>T \ q" and "q\<^sup>T * q \ f\<^sup>T * f \ 1" shows "injective ((f \ -q) \ (f \ q)\<^sup>T)" using kruskal_injective_inv[where e=bot] by (simp add: assms) lemma kruskal_exchange_acyclic_inv_3: assumes "injective w" and "d \ w" shows "(w \ -d) * d\<^sup>T * top = bot" proof - have "(w \ -d) * d\<^sup>T * top = (w \ -d \ (d\<^sup>T * top)\<^sup>T) * top" by (simp add: comp_associative comp_inf_vector_1 conv_dist_comp) also have "... = (w \ top * d \ -d) * top" by (simp add: conv_dist_comp inf_commute inf_left_commute) finally show ?thesis using assms epm_3 by simp qed lemma kruskal_subgraph_inv: assumes "f \ --(-h \ g)" and "e \ --g" and "symmetric h" and "symmetric g" shows "(f \ -q) \ (f \ q)\<^sup>T \ e \ --(-(h \ -e \ -e\<^sup>T) \ g)" proof - let ?f = "(f \ -q) \ (f \ q)\<^sup>T \ e" let ?h = "h \ -e \ -e\<^sup>T" have 1: "f \ -q \ -h \ --g" using assms(1) inf.coboundedI1 by simp have "(f \ q)\<^sup>T \ (-h \ --g)\<^sup>T" using assms(1) inf.coboundedI1 conv_isotone by simp also have "... = -h \ --g" using assms(3,4) conv_complement conv_dist_inf by simp finally have "?f \ (-h \ --g) \ (e \ --g)" using 1 assms(2) inf.absorb1 semiring.add_right_mono by simp also have "... \ (-h \ --e) \ --g" by (simp add: inf.coboundedI1 le_supI2 pp_increasing) also have "... \ -?h \ --g" using inf.sup_left_isotone order_trans p_antitone_inf p_supdist_inf by blast finally show "?f \ --(-?h \ g)" using inf_pp_semi_commute order_lesseq_imp by blast qed end subsection \Stone Relation Algebras\ text \ We add \pp_dist_comp\ and \pp_one\, which follow in relation algebras but not in the present setting. The main change is that only a Stone algebra is required, not a Boolean algebra. \ class stone_relation_algebra = pd_allegory + stone_algebra + assumes pp_dist_comp : "--(x * y) = --x * --y" assumes pp_one [simp]: "--1 = 1" begin text \ The following property is a simple consequence of the Stone axiom. We cannot hope to remove the double complement in it. \ lemma conv_complement_0_p [simp]: "(-x)\<^sup>T \ (--x)\<^sup>T = top" by (metis conv_top conv_dist_sup stone) lemma theorem24xxiv_pp: "-(x * y) \ --(x * z) = -(x * (y \ -z)) \ --(x * z)" by (metis p_dist_inf theorem24xxiii) lemma asymmetric_linear: "asymmetric x \ linear (-x)" by (metis conv_complement inf.distrib_left inf_p maddux_3_11_pp p_bot p_dist_inf) lemma strict_linear_asymmetric: "strict_linear x \ antisymmetric (-x)" by (metis conv_complement eq_refl p_dist_sup pp_one) lemma regular_complement_top: "regular x \ x \ -x = top" by (metis stone) lemma regular_mult_closed: "regular x \ regular y \ regular (x * y)" by (simp add: pp_dist_comp) lemma regular_one_closed: "regular 1" by simp text \ The following variants of total and surjective are useful for graphs. \ lemma pp_total: "total (--x) \ -(x*top) = bot" by (simp add: dense_pp pp_dist_comp) lemma pp_surjective: "surjective (--x) \ -(top*x) = bot" by (metis p_bot p_comp_pp p_top pp_dist_comp) text \ Bijective elements and mappings are necessarily regular, that is, invariant under double-complement. This implies that points are regular. Moreover, also arcs are regular. \ lemma bijective_regular: "bijective x \ regular x" by (metis comp_bijective_complement mult_left_one regular_one_closed) lemma mapping_regular: "mapping x \ regular x" by (metis bijective_regular conv_complement conv_involutive total_conv_surjective) lemma arc_regular: assumes "arc x" shows "regular x" proof - have "--x \ --(x * top \ top * x)" by (simp add: pp_isotone top_left_mult_increasing top_right_mult_increasing) also have "... = --(x * top) \ --(top * x)" by simp also have "... = x * top \ top * x" by (metis assms bijective_regular conv_top conv_dist_comp conv_involutive mapping_regular) also have "... \ x * x\<^sup>T * top * x" by (metis comp_associative dedekind_1 inf.commute inf_top.right_neutral) also have "... \ x" by (metis assms comp_right_one conv_top comp_associative conv_dist_comp conv_involutive mult_right_isotone vector_top_closed) finally show ?thesis by (simp add: antisym pp_increasing) qed (* lemma conv_complement_0 [simp]: "x\<^sup>T \ (-x)\<^sup>T = top" nitpick [expect=genuine] oops lemma schroeder_3: "x * y \ z \ x\<^sup>T * -z \ -y" nitpick [expect=genuine] oops lemma schroeder_4: "x * y \ z \ -z * y\<^sup>T \ -x" nitpick [expect=genuine] oops lemma theorem24xxiv: "-(x * y) \ (x * z) = -(x * (y \ -z)) \ (x * z)" nitpick [expect=genuine] oops lemma vector_N: "x = x * top \ N(x) = x \ 1" nitpick [expect=genuine] oops lemma N_vector [simp]: "N(x * top) = x * top \ 1" nitpick [expect=genuine] oops lemma N_vector_top [simp]: "N(x * top) * top = x * top" nitpick [expect=genuine] oops lemma N_below_inf_one: "N(x) \ x \ 1" nitpick [expect=genuine] oops lemma N_below: "N(x) \ x" nitpick [expect=genuine] oops lemma n_split_omega_mult: "xs * xo = xo \ xo * top = xo \ N(top) * xo = xs * N(xo) * top" nitpick [expect=genuine] oops lemma complement_vector: "vector v \ vector (-v)" nitpick [expect=genuine] oops lemma complement_covector: "covector v \ covector (-v)" nitpick [expect=genuine] oops lemma triple_schroeder: "x * y * z \ w \ x\<^sup>T * -w * z\<^sup>T \ -y" nitpick [expect=genuine] oops lemma schroeder_5: "x * y \ z \ y * -z\<^sup>T \ -x\<^sup>T" nitpick [expect=genuine] oops lemma schroeder_6: "x * y \ z \ -z\<^sup>T * x \ -y\<^sup>T" nitpick [expect=genuine] oops *) end text \ Every Stone algebra can be expanded to a Stone relation algebra by identifying the semiring and lattice structures and taking identity as converse. \ sublocale stone_algebra < comp_inf: stone_relation_algebra where one = top and times = inf and conv = id proof (unfold_locales, goal_cases) case 7 show ?case by (simp add: inf_commute) qed (auto simp: inf.assoc inf_sup_distrib2 inf_left_commute) text \ Every bounded linear order can be expanded to a Stone algebra, which can be expanded to a Stone relation algebra by reusing some of the operations. In particular, composition is meet, its identity is \top\ and converse is the identity function. \ class linorder_stone_relation_algebra_expansion = linorder_stone_algebra_expansion + times + conv + one + assumes times_def [simp]: "x * y = min x y" assumes conv_def [simp]: "x\<^sup>T = x" assumes one_def [simp]: "1 = top" begin lemma times_inf [simp]: "x * y = x \ y" by simp subclass stone_relation_algebra apply unfold_locales using comp_inf.mult_right_dist_sup inf_commute inf_assoc inf_left_commute pp_dist_inf min_def by simp_all lemma times_dense: "x \ bot \ y \ bot \ x * y \ bot" using inf_dense min_inf times_def by presburger end subsection \Relation Algebras\ text \ For a relation algebra, we only require that the underlying lattice is a Boolean algebra. In fact, the only missing axiom is that double-complement is the identity. \ class relation_algebra = boolean_algebra + stone_relation_algebra begin lemma conv_complement_0 [simp]: "x\<^sup>T \ (-x)\<^sup>T = top" by (simp add: conv_complement) text \ We now obtain the original formulations of the Schr\"oder equivalences. \ lemma schroeder_3: "x * y \ z \ x\<^sup>T * -z \ -y" by (simp add: schroeder_3_p) lemma schroeder_4: "x * y \ z \ -z * y\<^sup>T \ -x" by (simp add: schroeder_4_p) lemma theorem24xxiv: "-(x * y) \ (x * z) = -(x * (y \ -z)) \ (x * z)" using theorem24xxiv_pp by auto lemma vector_N: "vector x \ N(x) = x \ 1" by (simp add: vector_N_pp) lemma N_vector [simp]: "N(x * top) = x * top \ 1" by simp lemma N_vector_top [simp]: "N(x * top) * top = x * top" using N_vector_top_pp by simp lemma N_below_inf_one: "N(x) \ x \ 1" using N_below_inf_one_pp by simp lemma N_below: "N(x) \ x" using N_below_pp by simp lemma n_split_omega_mult: "xs * xo = xo \ xo * top = xo \ N(top) * xo = xs * N(xo) * top" using n_split_omega_mult_pp by simp lemma complement_vector: "vector v \ vector (-v)" using vector_complement_closed by fastforce lemma complement_covector: "covector v \ covector (-v)" using covector_complement_closed by force lemma triple_schroeder: "x * y * z \ w \ x\<^sup>T * -w * z\<^sup>T \ -y" by (simp add: triple_schroeder_p) lemma schroeder_5: "x * y \ z \ y * -z\<^sup>T \ -x\<^sup>T" by (simp add: conv_complement schroeder_5_p) lemma schroeder_6: "x * y \ z \ -z\<^sup>T * x \ -y\<^sup>T" by (simp add: conv_complement schroeder_5_p) end text \ We briefly look at the so-called Tarski rule. In some models of Stone relation algebras it only holds for regular elements, so we add this as an assumption. \ class stone_relation_algebra_tarski = stone_relation_algebra + assumes tarski: "regular x \ x \ bot \ top * x * top = top" begin text \ We can then show, for example, that every arc is contained in a pseudocomplemented relation or its pseudocomplement. \ lemma arc_in_partition: assumes "arc x" shows "x \ -y \ x \ --y" proof - have 1: "x * top * x\<^sup>T \ 1 \ x\<^sup>T * top * x \ 1" using assms arc_expanded by auto have "\ x \ --y \ x \ -y" proof assume "\ x \ --y" hence "x \ -y \ bot" using pseudo_complement by simp hence "top * (x \ -y) * top = top" using assms arc_regular tarski by auto hence "x = x \ top * (x \ -y) * top" by simp also have "... \ x \ x * ((x \ -y) * top)\<^sup>T * (x \ -y) * top" by (metis dedekind_2 inf.cobounded1 inf.boundedI inf_commute mult_assoc inf.absorb2 top.extremum) also have "... = x \ x * top * (x\<^sup>T \ -y\<^sup>T) * (x \ -y) * top" by (simp add: comp_associative conv_complement conv_dist_comp conv_dist_inf) also have "... \ x \ x * top * x\<^sup>T * (x \ -y) * top" using inf.sup_right_isotone mult_left_isotone mult_right_isotone by auto also have "... \ x \ 1 * (x \ -y) * top" using 1 by (metis comp_associative comp_isotone inf.sup_right_isotone mult_1_left mult_semi_associative) also have "... = x \ (x \ -y) * top" by simp also have "... \ (x \ -y) * ((x \ -y)\<^sup>T * x)" by (metis dedekind_1 inf_commute inf_top_right) also have "... \ (x \ -y) * (x\<^sup>T * x)" by (simp add: conv_dist_inf mult_left_isotone mult_right_isotone) also have "... \ (x \ -y) * (x\<^sup>T * top * x)" by (simp add: mult_assoc mult_right_isotone top_left_mult_increasing) also have "... \ x \ -y" using 1 by (metis mult_right_isotone mult_1_right) finally show "x \ -y" by simp qed thus ?thesis by auto qed lemma non_bot_arc_in_partition_xor: assumes "arc x" and "x \ bot" shows "(x \ -y \ \ x \ --y) \ (\ x \ -y \ x \ --y)" proof - have "x \ -y \ x \ --y \ False" by (simp add: assms(2) inf_absorb1 shunting_1_pp) thus ?thesis using assms(1) arc_in_partition by auto qed +lemma point_in_vector_or_pseudo_complement: + assumes "point p" + and "vector v" + shows "p \ --v \ p \ -v" +proof (rule disjCI) + assume "\(p \ -v)" + hence "top * (p \ --v) = top" + by (smt assms bijective_regular regular_closed_inf regular_closed_p shunting_1_pp tarski vector_complement_closed vector_inf_closed vector_mult_closed) + thus "p \ --v" + by (metis assms(1) epm_3 inf.absorb_iff1 inf.cobounded1 inf_top.right_neutral) +qed + +lemma point_in_vector_or_complement: + assumes "point p" + and "vector v" + and "regular v" + shows "p \ v \ p \ -v" + using assms point_in_vector_or_pseudo_complement by fastforce + +lemma point_in_vector_sup: + assumes "point p" + and "vector v" + and "regular v" + and "p \ v \ w" + shows "p \ v \ p \ w" + by (metis assms inf.absorb1 shunting_var_p sup_commute point_in_vector_or_complement) + +lemma arc_in_arc_or_complement: + assumes "arc x" + and "arc y" + and "\ x \ y" + shows "x \ -y" + using assms arc_in_partition arc_regular by force + +lemma arc_in_sup_arc: + assumes "arc x" + and "arc y" + and "x \ z \ y" + shows "x \ z \ x \ y" +proof (cases "x \ y") + case True + thus ?thesis + by simp +next + case False + hence "x \ -y" + using assms(1,2) arc_in_arc_or_complement by blast + hence "x \ -y \ (z \ y)" + using assms(3) by simp + hence "x \ z" + by (metis inf.boundedE inf.sup_monoid.add_commute maddux_3_13 sup_commute) + thus ?thesis + by simp +qed + +lemma different_arc_in_sup_arc: + assumes "arc x" + and "arc y" + and "x \ z \ y" + and "x \ y" + shows "x \ z" +proof - + have "x \ -y" + using arc_in_arc_or_complement assms(1,2,4) eq_iff p_antitone_iff by blast + hence "x \ -y \ (z \ y)" + using assms arc_in_sup_arc by simp + thus ?thesis + by (metis order_lesseq_imp p_inf_sup_below sup_commute) +qed + end class relation_algebra_tarski = relation_algebra + stone_relation_algebra_tarski text \ Finally, the above axioms of relation algebras do not imply that they contain at least two elements. This is necessary, for example, to show that arcs are not empty. \ class stone_relation_algebra_consistent = stone_relation_algebra + assumes consistent: "bot \ top" begin lemma arc_not_bot: "arc x \ x \ bot" using consistent mult_right_zero by auto end class relation_algebra_consistent = relation_algebra + stone_relation_algebra_consistent class stone_relation_algebra_tarski_consistent = stone_relation_algebra_tarski + stone_relation_algebra_consistent begin lemma arc_in_partition_xor: "arc x \ (x \ -y \ \ x \ --y) \ (\ x \ -y \ x \ --y)" by (simp add: non_bot_arc_in_partition_xor arc_not_bot) end end