diff --git a/thys/Relational_Disjoint_Set_Forests/Disjoint_Set_Forests.thy b/thys/Relational_Disjoint_Set_Forests/Disjoint_Set_Forests.thy --- a/thys/Relational_Disjoint_Set_Forests/Disjoint_Set_Forests.thy +++ b/thys/Relational_Disjoint_Set_Forests/Disjoint_Set_Forests.thy @@ -1,2280 +1,2312 @@ (* Title: Disjoint-Set Forests Author: Walter Guttmann Maintainer: Walter Guttmann *) theory Disjoint_Set_Forests imports "HOL-Hoare.Hoare_Logic" Stone_Kleene_Relation_Algebras.Kleene_Relation_Algebras begin no_notation trancl ("(_\<^sup>+)" [1000] 999) text \ An arc in a Stone relation algebra corresponds to an atom in a relation algebra and represents a single edge in a graph. A point represents a set of nodes. A rectangle represents the Cartesian product of two sets of nodes \cite{BerghammerStruth2010}. \ context times_top begin abbreviation rectangle :: "'a \ bool" where "rectangle x \ x * top * x = x" end context stone_relation_algebra begin lemma arc_rectangle: "arc x \ rectangle x" using arc_top_arc by blast section \Relation-Algebraic Semantics of Associative Array Access\ text \ The following two operations model updating array $x$ at index $y$ to value $z$, and reading the content of array $x$ at index $y$, respectively. The read operation uses double brackets to avoid ambiguity with list syntax. The remainder of this section shows basic properties of these operations. \ abbreviation rel_update :: "'a \ 'a \ 'a \ 'a" ("(_[_\_])" [70, 65, 65] 61) where "x[y\z] \ (y \ z\<^sup>T) \ (-y \ x)" abbreviation rel_access :: "'a \ 'a \ 'a" ("(2_[[_]])" [70, 65] 65) where "x[[y]] \ x\<^sup>T * y" text \Theorem 1.1\ lemma update_univalent: assumes "univalent x" and "vector y" and "injective z" shows "univalent (x[y\z])" proof - have 1: "univalent (y \ z\<^sup>T)" using assms(3) inf_commute univalent_inf_closed by force have "(y \ z\<^sup>T)\<^sup>T * (-y \ x) = (y\<^sup>T \ z) * (-y \ x)" by (simp add: conv_dist_inf) also have "... = z * (y \ -y \ x)" by (metis assms(2) covector_inf_comp_3 inf.sup_monoid.add_assoc inf.sup_monoid.add_commute) finally have 2: "(y \ z\<^sup>T)\<^sup>T * (-y \ x) = bot" by simp have 3: "vector (-y)" using assms(2) vector_complement_closed by simp have "(-y \ x)\<^sup>T * (y \ z\<^sup>T) = (-y\<^sup>T \ x\<^sup>T) * (y \ z\<^sup>T)" by (simp add: conv_complement conv_dist_inf) also have "... = x\<^sup>T * (-y \ y \ z\<^sup>T)" using 3 by (metis (mono_tags, hide_lams) conv_complement covector_inf_comp_3 inf.sup_monoid.add_assoc inf.sup_monoid.add_commute) finally have 4: "(-y \ x)\<^sup>T * (y \ z\<^sup>T) = bot" by simp have 5: "univalent (-y \ x)" using assms(1) inf_commute univalent_inf_closed by fastforce have "(x[y\z])\<^sup>T * (x[y\z]) = (y \ z\<^sup>T)\<^sup>T * (x[y\z]) \ (-y \ x)\<^sup>T * (x[y\z])" by (simp add: conv_dist_sup mult_right_dist_sup) also have "... = (y \ z\<^sup>T)\<^sup>T * (y \ z\<^sup>T) \ (y \ z\<^sup>T)\<^sup>T * (-y \ x) \ (-y \ x)\<^sup>T * (y \ z\<^sup>T) \ (-y \ x)\<^sup>T * (-y \ x)" by (simp add: mult_left_dist_sup sup_assoc) finally show ?thesis using 1 2 4 5 by simp qed text \Theorem 1.2\ lemma update_total: assumes "total x" and "vector y" and "regular y" and "surjective z" shows "total (x[y\z])" proof - have "(x[y\z]) * top = x*top[y\top*z]" by (simp add: assms(2) semiring.distrib_right vector_complement_closed vector_inf_comp conv_dist_comp) also have "... = top[y\top]" using assms(1) assms(4) by simp also have "... = top" using assms(3) regular_complement_top by auto finally show ?thesis by simp qed text \Theorem 1.3\ lemma update_mapping: assumes "mapping x" and "vector y" and "regular y" and "bijective z" shows "mapping (x[y\z])" using assms update_univalent update_total by simp text \Theorem 1.4\ lemma read_injective: assumes "injective y" and "univalent x" shows "injective (x[[y]])" using assms injective_mult_closed univalent_conv_injective by blast text \Theorem 1.5\ lemma read_surjective: assumes "surjective y" and "total x" shows "surjective (x[[y]])" using assms surjective_mult_closed total_conv_surjective by blast text \Theorem 1.6\ lemma read_bijective: assumes "bijective y" and "mapping x" shows "bijective (x[[y]])" by (simp add: assms read_injective read_surjective) text \Theorem 1.7\ lemma read_point: assumes "point p" and "mapping x" shows "point (x[[p]])" using assms comp_associative read_injective read_surjective by auto text \Theorem 1.8\ lemma update_postcondition: assumes "point x" "point y" shows "x \ p = x * y\<^sup>T \ p[[x]] = y" apply (rule iffI) subgoal by (metis assms comp_associative conv_dist_comp conv_involutive covector_inf_comp_3 equivalence_top_closed vector_covector) subgoal apply (rule order.antisym) subgoal by (metis assms conv_dist_comp conv_involutive inf.boundedI inf.cobounded1 vector_covector vector_restrict_comp_conv) subgoal by (smt assms comp_associative conv_dist_comp conv_involutive covector_restrict_comp_conv dense_conv_closed equivalence_top_closed inf.boundedI shunt_mapping vector_covector preorder_idempotent) done done text \Back and von Wright's array independence requirements \cite{BackWright1998}, later also lens laws \cite{FosterGreenwaldMoorePierceSchmitt2007}\ +text \Theorem 2.1\ + lemma put_get: assumes "vector y" "surjective y" "vector z" shows "(x[y\z])[[y]] = z" proof - have "(x[y\z])[[y]] = (y\<^sup>T \ z) * y \ (-y\<^sup>T \ x\<^sup>T) * y" by (simp add: conv_complement conv_dist_inf conv_dist_sup mult_right_dist_sup) also have "... = z * y" proof - have "(-y\<^sup>T \ x\<^sup>T) * y = bot" by (metis assms(1) covector_inf_comp_3 inf_commute conv_complement mult_right_zero p_inf vector_complement_closed) thus ?thesis by (simp add: assms covector_inf_comp_3 inf_commute) qed also have "... = z" by (metis assms(2,3) mult_assoc) finally show ?thesis . qed +text \Theorem 2.3\ + lemma put_put: "(x[y\z])[y\w] = x[y\w]" by (metis inf_absorb2 inf_commute inf_le1 inf_sup_distrib1 maddux_3_13 sup_inf_absorb) +text \Theorem 2.5\ + lemma get_put: assumes "point y" shows "x[y\x[[y]]] = x" proof - have "x[y\x[[y]]] = (y \ y\<^sup>T * x) \ (-y \ x)" by (simp add: conv_dist_comp) also have "... = (y \ x) \ (-y \ x)" proof - have "y \ y\<^sup>T * x = y \ x" proof (rule order.antisym) have "y \ y\<^sup>T * x = (y \ y\<^sup>T) * x" by (simp add: assms vector_inf_comp) also have "(y \ y\<^sup>T) * x = y * y\<^sup>T * x" by (simp add: assms vector_covector) also have "... \ x" using assms comp_isotone by fastforce finally show "y \ y\<^sup>T * x \ y \ x" by simp have "y \ x \ y\<^sup>T * x" by (simp add: assms vector_restrict_comp_conv) thus "y \ x \ y \ y\<^sup>T * x" by simp qed thus ?thesis by simp qed also have "... = x" proof - have "regular y" using assms bijective_regular by blast thus ?thesis by (metis inf.sup_monoid.add_commute maddux_3_11_pp) qed finally show ?thesis . qed lemma update_inf: "u \ y \ (x[y\z]) \ u = z\<^sup>T \ u" by (smt comp_inf.mult_right_dist_sup comp_inf.semiring.mult_zero_right inf.left_commute inf.sup_monoid.add_assoc inf_absorb2 p_inf sup_bot_right inf.sup_monoid.add_commute) lemma update_inf_same: "(x[y\z]) \ y = z\<^sup>T \ y" by (simp add: update_inf) lemma update_inf_different: "u \ -y \ (x[y\z]) \ u = x \ u" by (smt inf.right_idem inf.sup_monoid.add_commute inf.sup_relative_same_increasing inf_import_p maddux_3_13 sup.cobounded2 update_inf_same) end section \Relation-Algebraic Semantics of Disjoint-Set Forests\ text \ A disjoint-set forest represents a partition of a set into equivalence classes. We take the represented equivalence relation as the semantics of a forest. It is obtained by operation \fc\ below. Additionally, operation \wcc\ giving the weakly connected components of a graph will be used for the semantics of the union of two disjoint sets. Finally, operation \root\ yields the root of a component tree, that is, the representative of a set containing a given element. This section defines these operations and derives their properties. \ context stone_kleene_relation_algebra begin -text \Theorem 4.2\ +text \Theorem 5.2\ lemma omit_redundant_points: assumes "point p" shows "p \ x\<^sup>\ = (p \ 1) \ (p \ x) * (-p \ x)\<^sup>\" proof (rule order.antisym) let ?p = "p \ 1" have "?p * x * (-p \ x)\<^sup>\ * ?p \ ?p * top * ?p" by (metis comp_associative mult_left_isotone mult_right_isotone top.extremum) also have "... \ ?p" by (simp add: assms injective_codomain vector_inf_one_comp) finally have "?p * x * (-p \ x)\<^sup>\ * ?p * x \ ?p * x" using mult_left_isotone by blast hence "?p * x * (-p \ x)\<^sup>\ * (p \ x) \ ?p * x" by (simp add: assms comp_associative vector_inf_one_comp) also have 1: "... \ ?p * x * (-p \ x)\<^sup>\" using mult_right_isotone star.circ_reflexive by fastforce finally have "?p * x * (-p \ x)\<^sup>\ * (p \ x) \ ?p * x * (-p \ x)\<^sup>\ * (-p \ x) \ ?p * x * (-p \ x)\<^sup>\" by (simp add: mult_right_isotone star.circ_plus_same star.left_plus_below_circ mult_assoc) hence "?p * x * (-p \ x)\<^sup>\ * ((p \ -p) \ x) \ ?p * x * (-p \ x)\<^sup>\" by (simp add: comp_inf.mult_right_dist_sup mult_left_dist_sup) hence "?p * x * (-p \ x)\<^sup>\ * x \ ?p * x * (-p \ x)\<^sup>\" by (metis assms bijective_regular inf.absorb2 inf.cobounded1 inf.sup_monoid.add_commute shunting_p) hence "?p * x * (-p \ x)\<^sup>\ * x \ ?p * x \ ?p * x * (-p \ x)\<^sup>\" using 1 by simp hence "?p * (1 \ x * (-p \ x)\<^sup>\) * x \ ?p * x * (-p \ x)\<^sup>\" by (simp add: comp_associative mult_left_dist_sup mult_right_dist_sup) also have "... \ ?p * (1 \ x * (-p \ x)\<^sup>\)" by (simp add: comp_associative mult_right_isotone) finally have "?p * x\<^sup>\ \ ?p * (1 \ x * (-p \ x)\<^sup>\)" using star_right_induct by (meson dual_order.trans le_supI mult_left_sub_dist_sup_left mult_sub_right_one) also have "... = ?p \ ?p * x * (-p \ x)\<^sup>\" by (simp add: comp_associative semiring.distrib_left) finally show "p \ x\<^sup>\ \ ?p \ (p \ x) * (-p \ x)\<^sup>\" by (simp add: assms vector_inf_one_comp) show "?p \ (p \ x) * (-p \ x)\<^sup>\ \ p \ x\<^sup>\" by (metis assms comp_isotone inf.boundedI inf.cobounded1 inf.coboundedI2 inf.sup_monoid.add_commute le_supI star.circ_increasing star.circ_transitive_equal star_isotone star_left_unfold_equal sup.cobounded1 vector_export_comp) qed text \Weakly connected components\ abbreviation "wcc x \ (x \ x\<^sup>T)\<^sup>\" -text \Theorem 5.1\ +text \Theorem 7.1\ lemma wcc_equivalence: "equivalence (wcc x)" apply (intro conjI) subgoal by (simp add: star.circ_reflexive) subgoal by (simp add: star.circ_transitive_equal) subgoal by (simp add: conv_dist_sup conv_star_commute sup_commute) done -text \Theorem 5.2\ +text \Theorem 7.2\ lemma wcc_increasing: "x \ wcc x" by (simp add: star.circ_sub_dist_1) lemma wcc_isotone: "x \ y \ wcc x \ wcc y" using conv_isotone star_isotone sup_mono by blast lemma wcc_idempotent: "wcc (wcc x) = wcc x" using star_involutive wcc_equivalence by auto -text \Theorem 5.3\ +text \Theorem 7.3\ lemma wcc_below_wcc: "x \ wcc y \ wcc x \ wcc y" using wcc_idempotent wcc_isotone by fastforce -text \Theorem 5.4\ +text \Theorem 7.4\ lemma wcc_bot: "wcc bot = 1" by (simp add: star.circ_zero) lemma wcc_one: "wcc 1 = 1" by (simp add: star_one) -text \Theorem 5.5\ +text \Theorem 7.5\ lemma wcc_top: "wcc top = top" by (simp add: star.circ_top) -text \Theorem 5.6\ +text \Theorem 7.6\ lemma wcc_with_loops: "wcc x = wcc (x \ 1)" by (metis conv_dist_sup star_decompose_1 star_sup_one sup_commute symmetric_one_closed) lemma wcc_without_loops: "wcc x = wcc (x \ -1)" by (metis conv_star_commute star_sum reachable_without_loops) lemma forest_components_wcc: "injective x \ wcc x = forest_components x" by (simp add: cancel_separate_1) +text \Theorem 7.8\ + +lemma wcc_sup_wcc: + "wcc (x \ y) = wcc (x \ wcc y)" + by (smt (verit, del_insts) le_sup_iff order.antisym sup_right_divisibility wcc_below_wcc wcc_increasing wcc_isotone) + text \Components of a forest, which is represented using edges directed towards the roots\ abbreviation "fc x \ x\<^sup>\ * x\<^sup>T\<^sup>\" -text \Theorem 2.1\ +text \Theorem 3.1\ lemma fc_equivalence: "univalent x \ equivalence (fc x)" apply (intro conjI) subgoal by (simp add: reflexive_mult_closed star.circ_reflexive) subgoal by (metis cancel_separate_1 order.eq_iff star.circ_transitive_equal) subgoal by (simp add: conv_dist_comp conv_star_commute) done -text \Theorem 2.2\ +text \Theorem 3.2\ lemma fc_increasing: "x \ fc x" by (metis le_supE mult_left_isotone star.circ_back_loop_fixpoint star.circ_increasing) -text \Theorem 2.3\ +text \Theorem 3.3\ lemma fc_isotone: "x \ y \ fc x \ fc y" by (simp add: comp_isotone conv_isotone star_isotone) -text \Theorem 2.4\ +text \Theorem 3.4\ lemma fc_idempotent: "univalent x \ fc (fc x) = fc x" by (metis fc_equivalence cancel_separate_1 star.circ_transitive_equal star_involutive) -text \Theorem 2.5\ +text \Theorem 3.5\ lemma fc_star: "univalent x \ (fc x)\<^sup>\ = fc x" using fc_equivalence fc_idempotent star.circ_transitive_equal by simp lemma fc_plus: "univalent x \ (fc x)\<^sup>+ = fc x" by (metis fc_star star.circ_decompose_9) -text \Theorem 2.6\ +text \Theorem 3.6\ lemma fc_bot: "fc bot = 1" by (simp add: star.circ_zero) lemma fc_one: "fc 1 = 1" by (simp add: star_one) -text \Theorem 2.7\ +text \Theorem 3.7\ lemma fc_top: "fc top = top" by (simp add: star.circ_top) -text \Theorem 5.7\ +text \Theorem 7.7\ lemma fc_wcc: "univalent x \ wcc x = fc x" by (simp add: fc_star star_decompose_1) lemma fc_via_root: assumes "total (p\<^sup>\ * (p \ 1))" shows "fc p = p\<^sup>\ * (p \ 1) * p\<^sup>T\<^sup>\" proof (rule order.antisym) have "1 \ p\<^sup>\ * (p \ 1) * p\<^sup>T\<^sup>\" by (smt assms comp_associative conv_dist_comp conv_star_commute coreflexive_idempotent coreflexive_symmetric inf.cobounded2 total_var) hence "fc p \ p\<^sup>\ * p\<^sup>\ * (p \ 1) * p\<^sup>T\<^sup>\ * p\<^sup>T\<^sup>\" by (metis comp_right_one mult_left_isotone mult_right_isotone mult_assoc) thus "fc p \ p\<^sup>\ * (p \ 1) * p\<^sup>T\<^sup>\" by (simp add: star.circ_transitive_equal mult_assoc) show "p\<^sup>\ * (p \ 1) * p\<^sup>T\<^sup>\ \ fc p" by (metis comp_isotone inf.cobounded2 mult_1_right order.refl) qed -text \Theorem 4.1\ +text \Theorem 5.1\ lemma update_acyclic_1: assumes "acyclic (p \ -1)" and "point y" and "vector w" and "w \ p\<^sup>\ * y" shows "acyclic ((p[w\y]) \ -1)" proof - let ?p = "p[w\y]" have "w * y\<^sup>T \ p\<^sup>\" using assms(2,4) shunt_bijective by blast hence "w * y\<^sup>T \ (p \ -1)\<^sup>\" using reachable_without_loops by auto hence "w * y\<^sup>T \ -1 \ (p \ -1)\<^sup>\ \ -1" by (simp add: inf.coboundedI2 inf.sup_monoid.add_commute) also have "... \ (p \ -1)\<^sup>+" by (simp add: star_plus_without_loops) finally have 1: "w \ y\<^sup>T \ -1 \ (p \ -1)\<^sup>+" using assms(2,3) vector_covector by auto have "?p \ -1 = (w \ y\<^sup>T \ -1) \ (-w \ p \ -1)" by (simp add: inf_sup_distrib2) also have "... \ (p \ -1)\<^sup>+ \ (-w \ p \ -1)" using 1 sup_left_isotone by blast also have "... \ (p \ -1)\<^sup>+ \ (p \ -1)" using comp_inf.mult_semi_associative sup_right_isotone by auto also have "... = (p \ -1)\<^sup>+" by (metis star.circ_back_loop_fixpoint sup.right_idem) finally have "(?p \ -1)\<^sup>+ \ (p \ -1)\<^sup>+" by (metis comp_associative comp_isotone star.circ_transitive_equal star.left_plus_circ star_isotone) also have "... \ -1" using assms(1) by blast finally show ?thesis by simp qed lemma update_acyclic_2: assumes "acyclic (p \ -1)" and "point y" and "point x" and "y \ p\<^sup>T\<^sup>\ * x" and "univalent p" and "p\<^sup>T * y \ y" shows "acyclic ((p[p\<^sup>T\<^sup>\*x\y]) \ -1)" proof - have "p\<^sup>T * p\<^sup>\ * y = p\<^sup>T * p * p\<^sup>\ * y \ p\<^sup>T * y" by (metis comp_associative mult_left_dist_sup star.circ_loop_fixpoint) also have "... \ p\<^sup>\ * y" by (metis assms(5,6) comp_right_one le_supI le_supI2 mult_left_isotone star.circ_loop_fixpoint star.circ_transitive_equal) finally have "p\<^sup>T\<^sup>\ * x \ p\<^sup>\ * y" by (simp add: assms(2-4) bijective_reverse conv_star_commute comp_associative star_left_induct) thus ?thesis by (simp add: assms(1-3) vector_mult_closed update_acyclic_1) qed lemma update_acyclic_3: assumes "acyclic (p \ -1)" and "point y" and "point w" and "y \ p\<^sup>T\<^sup>\ * w" shows "acyclic ((p[w\y]) \ -1)" by (simp add: assms bijective_reverse conv_star_commute update_acyclic_1) lemma rectangle_star_rectangle: "rectangle a \ a * x\<^sup>\ * a \ a" by (metis mult_left_isotone mult_right_isotone top.extremum) lemma arc_star_arc: "arc a \ a * x\<^sup>\ * a \ a" using arc_top_arc rectangle_star_rectangle by blast lemma star_rectangle_decompose: assumes "rectangle a" shows "(a \ x)\<^sup>\ = x\<^sup>\ \ x\<^sup>\ * a * x\<^sup>\" proof (rule order.antisym) have 1: "1 \ x\<^sup>\ \ x\<^sup>\ * a * x\<^sup>\" by (simp add: star.circ_reflexive sup.coboundedI1) have "(a \ x) * (x\<^sup>\ \ x\<^sup>\ * a * x\<^sup>\) = a * x\<^sup>\ \ a * x\<^sup>\ * a * x\<^sup>\ \ x\<^sup>+ \ x\<^sup>+ * a * x\<^sup>\" by (metis comp_associative semiring.combine_common_factor semiring.distrib_left sup_commute) also have "... = a * x\<^sup>\ \ x\<^sup>+ \ x\<^sup>+ * a * x\<^sup>\" using assms rectangle_star_rectangle by (simp add: mult_left_isotone sup_absorb1) also have "... = x\<^sup>+ \ x\<^sup>\ * a * x\<^sup>\" by (metis comp_associative star.circ_loop_fixpoint sup_assoc sup_commute) also have "... \ x\<^sup>\ \ x\<^sup>\ * a * x\<^sup>\" using star.left_plus_below_circ sup_left_isotone by auto finally show "(a \ x)\<^sup>\ \ x\<^sup>\ \ x\<^sup>\ * a * x\<^sup>\" using 1 by (metis comp_right_one le_supI star_left_induct) next show "x\<^sup>\ \ x\<^sup>\ * a * x\<^sup>\ \ (a \ x)\<^sup>\" by (metis comp_isotone le_supE le_supI star.circ_increasing star.circ_transitive_equal star_isotone sup_ge2) qed lemma star_arc_decompose: "arc a \ (a \ x)\<^sup>\ = x\<^sup>\ \ x\<^sup>\ * a * x\<^sup>\" using arc_top_arc star_rectangle_decompose by blast lemma plus_rectangle_decompose: assumes "rectangle a" shows "(a \ x)\<^sup>+ = x\<^sup>+ \ x\<^sup>\ * a * x\<^sup>\" proof - have "(a \ x)\<^sup>+ = (a \ x) * (x\<^sup>\ \ x\<^sup>\ * a * x\<^sup>\)" by (simp add: assms star_rectangle_decompose) also have "... = a * x\<^sup>\ \ a * x\<^sup>\ * a * x\<^sup>\ \ x\<^sup>+ \ x\<^sup>+ * a * x\<^sup>\" by (metis comp_associative semiring.combine_common_factor semiring.distrib_left sup_commute) also have "... = a * x\<^sup>\ \ x\<^sup>+ \ x\<^sup>+ * a * x\<^sup>\" using assms rectangle_star_rectangle by (simp add: mult_left_isotone sup_absorb1) also have "... = x\<^sup>+ \ x\<^sup>\ * a * x\<^sup>\" by (metis comp_associative star.circ_loop_fixpoint sup_assoc sup_commute) finally show ?thesis by simp qed -text \Theorem 6.1\ +text \Theorem 8.1\ lemma plus_arc_decompose: "arc a \ (a \ x)\<^sup>+ = x\<^sup>+ \ x\<^sup>\ * a * x\<^sup>\" using arc_top_arc plus_rectangle_decompose by blast -text \Theorem 6.2\ +text \Theorem 8.2\ lemma update_acyclic_4: assumes "acyclic (p \ -1)" and "point y" and "point w" and "y \ p\<^sup>\ * w = bot" shows "acyclic ((p[w\y]) \ -1)" proof - let ?p = "p[w\y]" have "y\<^sup>T * p\<^sup>\ * w \ -1" using assms(4) comp_associative pseudo_complement schroeder_3_p by auto hence 1: "p\<^sup>\ * w * y\<^sup>T * p\<^sup>\ \ -1" by (metis comp_associative comp_commute_below_diversity star.circ_transitive_equal) have "?p \ -1 \ (w \ y\<^sup>T) \ (p \ -1)" by (metis comp_inf.mult_right_dist_sup dual_order.trans inf.cobounded1 inf.coboundedI2 inf.sup_monoid.add_assoc le_supI sup.cobounded1 sup_ge2) also have "... = w * y\<^sup>T \ (p \ -1)" using assms(2,3) by (simp add: vector_covector) finally have "(?p \ -1)\<^sup>+ \ (w * y\<^sup>T \ (p \ -1))\<^sup>+" by (simp add: comp_isotone star_isotone) also have "... = (p \ -1)\<^sup>+ \ (p \ -1)\<^sup>\ * w * y\<^sup>T * (p \ -1)\<^sup>\" using assms(2,3) plus_arc_decompose points_arc by (simp add: comp_associative) also have "... \ (p \ -1)\<^sup>+ \ p\<^sup>\ * w * y\<^sup>T * p\<^sup>\" using reachable_without_loops by auto also have "... \ -1" using 1 assms(1) by simp finally show ?thesis by simp qed -text \Theorem 6.3\ +text \Theorem 8.3\ lemma update_acyclic_5: assumes "acyclic (p \ -1)" and "point w" shows "acyclic ((p[w\w]) \ -1)" proof - let ?p = "p[w\w]" have "?p \ -1 \ (w \ w\<^sup>T \ -1) \ (p \ -1)" by (metis comp_inf.mult_right_dist_sup inf.cobounded2 inf.sup_monoid.add_assoc sup_right_isotone) also have "... = p \ -1" using assms(2) by (metis comp_inf.covector_complement_closed equivalence_top_closed inf_top.right_neutral maddux_3_13 pseudo_complement regular_closed_top regular_one_closed vector_covector vector_top_closed) finally show ?thesis using assms(1) acyclic_down_closed by blast qed text \Root of the tree containing point $x$ in the disjoint-set forest $p$\ abbreviation "root p x \ p\<^sup>T\<^sup>\ * x \ (p \ 1) * top" -text \Theorem 3.1\ +text \Theorem 4.1\ lemma root_var: "root p x = (p \ 1) * p\<^sup>T\<^sup>\ * x" by (simp add: coreflexive_comp_top_inf inf_commute mult_assoc) -text \Theorem 3.2\ +text \Theorem 4.2\ lemma root_successor_loop: "univalent p \ root p x = p[[root p x]]" by (metis root_var injective_codomain comp_associative conv_dist_inf coreflexive_symmetric equivalence_one_closed inf.cobounded2 univalent_conv_injective) lemma root_transitive_successor_loop: "univalent p \ root p x = p\<^sup>T\<^sup>\ * (root p x)" by (metis mult_1_right star_one star_simulation_right_equal root_successor_loop) text \The root of a tree of a node belongs to the same component as the node.\ lemma root_same_component: "injective x \ root p x * x\<^sup>T \ fc p" by (metis comp_associative coreflexive_comp_top_inf eq_refl inf.sup_left_divisibility inf.sup_monoid.add_commute mult_isotone star.circ_circ_mult star.circ_right_top star.circ_transitive_equal star_one star_outer_increasing test_preserves_equation top_greatest) lemma root_vector: "vector x \ vector (root p x)" by (simp add: vector_mult_closed root_var) lemma root_vector_inf: "vector x \ root p x * x\<^sup>T = root p x \ x\<^sup>T" by (simp add: vector_covector root_vector) lemma root_same_component_vector: "injective x \ vector x \ root p x \ x\<^sup>T \ fc p" using root_same_component root_vector_inf by fastforce lemma univalent_root_successors: assumes "univalent p" shows "(p \ 1) * p\<^sup>\ = p \ 1" proof (rule order.antisym) have "(p \ 1) * p \ p \ 1" by (smt assms(1) comp_inf.mult_semi_associative conv_dist_comp conv_dist_inf conv_order equivalence_one_closed inf.absorb1 inf.sup_monoid.add_assoc injective_codomain) thus "(p \ 1) * p\<^sup>\ \ p \ 1" using star_right_induct_mult by blast show "p \ 1 \ (p \ 1) * p\<^sup>\" by (metis coreflexive_idempotent inf_le1 inf_le2 mult_right_isotone order_trans star.circ_increasing) qed lemma same_component_same_root_sub: assumes "univalent p" and "bijective y" and "x * y\<^sup>T \ fc p" shows "root p x \ root p y" proof - have "root p x * y\<^sup>T \ (p \ 1) * p\<^sup>T\<^sup>\" by (smt assms(1,3) mult_isotone mult_assoc root_var fc_plus fc_star order.eq_iff univalent_root_successors) thus ?thesis by (simp add: assms(2) shunt_bijective root_var) qed lemma same_component_same_root: assumes "univalent p" and "bijective x" and "bijective y" and "x * y\<^sup>T \ fc p" shows "root p x = root p y" proof (rule order.antisym) show "root p x \ root p y" using assms(1,3,4) same_component_same_root_sub by blast have "y * x\<^sup>T \ fc p" using assms(1,4) fc_equivalence conv_dist_comp conv_isotone by fastforce thus "root p y \ root p x" using assms(1,2) same_component_same_root_sub by blast qed lemma same_roots_sub: assumes "univalent q" and "p \ 1 \ q \ 1" and "fc p \ fc q" shows "p\<^sup>\ * (p \ 1) \ q\<^sup>\ * (q \ 1)" proof - have "p\<^sup>\ * (p \ 1) \ p\<^sup>\ * (q \ 1)" using assms(2) mult_right_isotone by auto also have "... \ fc p * (q \ 1)" using mult_left_isotone mult_right_isotone star.circ_reflexive by fastforce also have "... \ fc q * (q \ 1)" by (simp add: assms(3) mult_left_isotone) also have "... = q\<^sup>\ * (q \ 1)" by (metis assms(1) conv_dist_comp conv_dist_inf conv_star_commute inf_commute one_inf_conv symmetric_one_closed mult_assoc univalent_root_successors) finally show ?thesis . qed lemma same_roots: assumes "univalent p" and "univalent q" and "p \ 1 = q \ 1" and "fc p = fc q" shows "p\<^sup>\ * (p \ 1) = q\<^sup>\ * (q \ 1)" by (smt assms conv_dist_comp conv_dist_inf conv_involutive conv_star_commute inf_commute one_inf_conv symmetric_one_closed root_var univalent_root_successors) lemma same_root: assumes "univalent p" and "univalent q" and "p \ 1 = q \ 1" and "fc p = fc q" shows "root p x = root q x" by (metis assms mult_assoc root_var univalent_root_successors) lemma loop_root: assumes "injective x" and "x = p[[x]]" shows "x = root p x" proof (rule order.antisym) have "x \ p * x" by (metis assms comp_associative comp_right_one conv_order equivalence_one_closed ex231c inf.orderE inf.sup_monoid.add_commute mult_left_isotone mult_right_isotone one_inf_conv) hence "x = (p \ 1) * x" by (simp add: assms(1) inf_absorb2 injective_comp_right_dist_inf) thus "x \ root p x" by (metis assms(2) coreflexive_comp_top_inf inf.boundedI inf.cobounded1 inf.cobounded2 mult_isotone star.circ_increasing) next show "root p x \ x" using assms(2) le_infI1 star_left_induct_mult by auto qed lemma one_loop: assumes "acyclic (p \ -1)" and "univalent p" shows "(p \ 1) * (p\<^sup>T \ -1)\<^sup>+ * (p \ 1) = bot" proof - have "p\<^sup>T\<^sup>+ \ (p \ 1) * top * (p \ 1) = (p \ 1) * p\<^sup>T\<^sup>+ * (p \ 1)" by (simp add: test_comp_test_top) also have "... \ p\<^sup>T\<^sup>\ * (p \ 1)" by (simp add: inf.coboundedI2 mult_left_isotone star.circ_mult_upper_bound star.circ_reflexive star.left_plus_below_circ) also have "... = p \ 1" by (metis assms(2) conv_dist_comp conv_dist_inf conv_star_commute inf_commute one_inf_conv symmetric_one_closed univalent_root_successors) also have "... \ 1" by simp finally have "(p \ 1) * top * (p \ 1) \ -(p\<^sup>T\<^sup>+ \ -1)" using p_antitone p_antitone_iff p_shunting_swap by blast hence "(p \ 1)\<^sup>T * (p\<^sup>T\<^sup>+ \ -1) * (p \ 1)\<^sup>T \ bot" using triple_schroeder_p p_top by blast hence "(p \ 1) * (p\<^sup>T\<^sup>+ \ -1) * (p \ 1) = bot" by (simp add: coreflexive_symmetric le_bot) thus ?thesis by (smt assms(1) conv_complement conv_dist_comp conv_dist_inf conv_star_commute inf_absorb1 star.circ_plus_same symmetric_one_closed reachable_without_loops star_plus_without_loops) qed lemma root_root: "root p x = root p (root p x)" by (smt comp_associative comp_inf.mult_right_sub_dist_sup_right dual_order.eq_iff inf.cobounded1 inf.cobounded2 inf.orderE mult_right_isotone star.circ_loop_fixpoint star.circ_transitive_equal root_var) lemma loop_root_2: assumes "acyclic (p \ -1)" and "univalent p" and "injective x" and "x \ p\<^sup>T\<^sup>+ * x" shows "x = root p x" proof (rule order.antisym) have 1: "x = x \ -(-1 * x)" by (metis assms(3) comp_injective_below_complement inf.orderE mult_1_left regular_one_closed) have "x \ (p\<^sup>T \ -1)\<^sup>+ * x \ (p \ 1) * x" by (metis assms(4) inf_commute mult_right_dist_sup one_inf_conv plus_reachable_without_loops) also have "... \ -1 * x \ (p \ 1) * x" by (metis assms(1) conv_complement conv_dist_inf conv_isotone conv_plus_commute mult_left_isotone semiring.add_right_mono symmetric_one_closed) also have "... \ -1 * x \ root p x" using comp_isotone inf.coboundedI2 star.circ_reflexive sup_right_isotone by auto finally have "x \ (-1 * x \ root p x) \ -(-1 * x)" using 1 inf.boundedI inf.order_iff by blast also have "... \ root p x" using inf.sup_left_divisibility by auto finally show 2: "x \ root p x" . have "root p x = (p \ 1) * x \ (p \ 1) * (p\<^sup>T \ -1)\<^sup>+ * x" by (metis comp_associative mult_left_dist_sup star.circ_loop_fixpoint sup_commute reachable_without_loops root_var) also have "... \ x \ (p \ 1) * (p\<^sup>T \ -1)\<^sup>+ * root p x" using 2 by (metis coreflexive_comp_top_inf inf.cobounded2 mult_right_isotone semiring.add_mono) also have "... = x" by (metis assms(1,2) one_loop root_var mult_assoc semiring.mult_not_zero sup_bot_right) finally show "root p x \ x" . qed +lemma path_compression_invariant_simplify: + assumes "point w" + and "p\<^sup>T\<^sup>+ * w \ -w" + and "w \ y" + shows "p[[w]] \ w" +proof + assume "p[[w]] = w" + hence "w \ p\<^sup>T\<^sup>+ * w" + by (metis comp_isotone eq_refl star.circ_mult_increasing) + also have "... \ -w" + by (simp add: assms(2)) + finally have "w = bot" + using inf.orderE by fastforce + thus False + using assms(1,3) le_bot by force +qed + end context stone_relation_algebra_tarski begin -text \Theorem 4.3 \distinct_points\ has been moved to theory \Relation_Algebras\ in entry \Stone_Relation_Algebras\\ +text \Theorem 5.4 \distinct_points\ has been moved to theory \Relation_Algebras\ in entry \Stone_Relation_Algebras\\ text \Back and von Wright's array independence requirements \cite{BackWright1998}\ +text \Theorem 2.2\ + lemma put_get_different_vector: - assumes "vector y" "vector w" "w \ -y" + assumes "vector y" "w \ -y" shows "(x[y\z])[[w]] = x[[w]]" proof - have "(x[y\z])[[w]] = (y\<^sup>T \ z) * w \ (-y\<^sup>T \ x\<^sup>T) * w" by (simp add: conv_complement conv_dist_inf conv_dist_sup mult_right_dist_sup) also have "... = z * (w \ y) \ x\<^sup>T * (w \ -y)" by (metis assms(1) conv_complement covector_inf_comp_3 inf_commute vector_complement_closed) also have "... = z * (w \ y) \ x\<^sup>T * w" - by (simp add: assms(3) inf.absorb1) + by (simp add: assms(2) inf.absorb1) also have "... = z * bot \ x\<^sup>T * w" - by (metis assms(3) comp_inf.semiring.mult_zero_right inf.absorb1 inf.sup_monoid.add_assoc p_inf) + by (metis assms(2) comp_inf.semiring.mult_zero_right inf.absorb1 inf.sup_monoid.add_assoc p_inf) also have "... = x\<^sup>T * w" by simp finally show ?thesis . qed lemma put_get_different: assumes "point y" "point w" "w \ y" shows "(x[y\z])[[w]] = x[[w]]" proof - have "w \ y = bot" using assms distinct_points by simp hence "w \ -y" using pseudo_complement by simp thus ?thesis by (simp add: assms(1) assms(2) put_get_different_vector) qed +text \Theorem 2.4\ + lemma put_put_different_vector: assumes "vector y" "vector v" "v \ y = bot" shows "(x[y\z])[v\w] = (x[v\w])[y\z]" proof - have "(x[y\z])[v\w] = (v \ w\<^sup>T) \ (-v \ y \ z\<^sup>T) \ (-v \ -y \ x)" by (simp add: comp_inf.semiring.distrib_left inf_assoc sup_assoc) also have "... = (v \ w\<^sup>T) \ (y \ z\<^sup>T) \ (-v \ -y \ x)" by (metis assms(3) inf_commute inf_import_p p_inf selection_closed_id) also have "... = (y \ z\<^sup>T) \ (v \ w\<^sup>T) \ (-y \ -v \ x)" by (simp add: inf_commute sup_commute) also have "... = (y \ z\<^sup>T) \ (-y \ v \ w\<^sup>T) \ (-y \ -v \ x)" using assms distinct_points pseudo_complement inf.absorb2 by simp also have "... = (x[v\w])[y\z]" by (simp add: comp_inf.semiring.distrib_left inf_assoc sup_assoc) finally show ?thesis . qed lemma put_put_different: assumes "point y" "point v" "v \ y" shows "(x[y\z])[v\w] = (x[v\w])[y\z]" using assms distinct_points put_put_different_vector by blast end section \Verifying Operations on Disjoint-Set Forests\ text \ In this section we verify the make-set, find-set and union-sets operations of disjoint-set forests. We start by introducing syntax for updating arrays in programs. Updating the value at a given array index means updating the whole array. \ syntax "_rel_update" :: "idt \ 'a \ 'a \ 'b com" ("(2_[_] :=/ _)" [70, 65, 65] 61) translations "x[y] := z" => "(x := (y \ z\<^sup>T) \ (CONST uminus y \ x))" text \ The finiteness requirement in the following class is used for proving that the operations terminate. \ class finite_regular_p_algebra = p_algebra + assumes finite_regular: "finite { x . regular x }" -class stone_kleene_relation_algebra_tarski = stone_kleene_relation_algebra + stone_relation_algebra_tarski - class stone_kleene_relation_algebra_tarski_finite_regular = stone_kleene_relation_algebra_tarski + finite_regular_p_algebra begin subsection \Make-Set\ text \ We prove two correctness results about make-set. The first shows that the forest changes only to the extent of making one node the root of a tree. The second result adds that only singleton sets are created. \ definition "make_set_postcondition p x p0 \ x \ p = x * x\<^sup>T \ -x \ p = -x \ p0" theorem make_set: "VARS p [ point x \ p0 = p ] p[x] := x [ make_set_postcondition p x p0 ]" apply vcg_tc_simp by (simp add: make_set_postcondition_def inf_sup_distrib1 inf_assoc[THEN sym] vector_covector[THEN sym]) theorem make_set_2: "VARS p [ point x \ p0 = p \ p \ 1 ] p[x] := x [ make_set_postcondition p x p0 \ p \ 1 ]" proof vcg_tc fix p assume 1: "point x \ p0 = p \ p \ 1" show "make_set_postcondition (p[x\x]) x p0 \ p[x\x] \ 1" proof (rule conjI) show "make_set_postcondition (p[x\x]) x p0" using 1 by (simp add: make_set_postcondition_def inf_sup_distrib1 inf_assoc[THEN sym] vector_covector[THEN sym]) show "p[x\x] \ 1" using 1 by (metis coreflexive_sup_closed dual_order.trans inf.cobounded2 vector_covector) qed qed text \ The above total-correctness proof allows us to extract a function, which can be used in other implementations below. This is a technique of \cite{Guttmann2018c}. \ lemma make_set_exists: "point x \ \p' . make_set_postcondition p' x p" using tc_extract_function make_set by blast definition "make_set p x \ (SOME p' . make_set_postcondition p' x p)" lemma make_set_function: assumes "point x" and "p' = make_set p x" shows "make_set_postcondition p' x p" proof - let ?P = "\p' . make_set_postcondition p' x p" have "?P (SOME z . ?P z)" using assms(1) make_set_exists by (meson someI) thus ?thesis using assms(2) make_set_def by auto qed +end + subsection \Find-Set\ text \ Disjoint-set forests are represented by their parent mapping. It is a forest except each root of a component tree points to itself. We prove that find-set returns the root of the component tree of the given node. \ +context pd_kleene_allegory +begin + abbreviation "disjoint_set_forest p \ mapping p \ acyclic (p \ -1)" +end + +context stone_kleene_relation_algebra_tarski_finite_regular +begin + definition "find_set_precondition p x \ disjoint_set_forest p \ point x" definition "find_set_invariant p x y \ find_set_precondition p x \ point y \ y \ p\<^sup>T\<^sup>\ * x" definition "find_set_postcondition p x y \ point y \ y = root p x" lemma find_set_1: "find_set_precondition p x \ find_set_invariant p x x" apply (unfold find_set_invariant_def) using mult_left_isotone star.circ_reflexive find_set_precondition_def by fastforce lemma find_set_2: "find_set_invariant p x y \ y \ p[[y]] \ find_set_invariant p x (p[[y]]) \ card { z . regular z \ z \ p\<^sup>T\<^sup>\ * (p[[y]]) } < card { z . regular z \ z \ p\<^sup>T\<^sup>\ * y }" proof - let ?s = "{ z . regular z \ z \ p\<^sup>T\<^sup>\ * y }" let ?t = "{ z . regular z \ z \ p\<^sup>T\<^sup>\ * (p[[y]]) }" assume 1: "find_set_invariant p x y \ y \ p[[y]]" have 2: "point (p[[y]])" using 1 read_point find_set_invariant_def find_set_precondition_def by simp show "find_set_invariant p x (p[[y]]) \ card ?t < card ?s" proof (unfold find_set_invariant_def, intro conjI) show "find_set_precondition p x" using 1 find_set_invariant_def by simp show "vector (p[[y]])" using 2 by simp show "injective (p[[y]])" using 2 by simp show "surjective (p[[y]])" using 2 by simp show "p[[y]] \ p\<^sup>T\<^sup>\ * x" using 1 by (metis (hide_lams) find_set_invariant_def comp_associative comp_isotone star.circ_increasing star.circ_transitive_equal) show "card ?t < card ?s" proof - have 3: "(p\<^sup>T \ -1) * (p\<^sup>T \ -1)\<^sup>+ * y \ (p\<^sup>T \ -1)\<^sup>+ * y" by (simp add: mult_left_isotone mult_right_isotone star.left_plus_below_circ) have "p[[y]] = (p\<^sup>T \ 1) * y \ (p\<^sup>T \ -1) * y" by (metis maddux_3_11_pp mult_right_dist_sup regular_one_closed) also have "... \ ((p[[y]]) \ y) \ (p\<^sup>T \ -1) * y" by (metis comp_left_subdist_inf mult_1_left semiring.add_right_mono) also have "... = (p\<^sup>T \ -1) * y" using 1 2 find_set_invariant_def distinct_points by auto finally have 4: "(p\<^sup>T \ -1)\<^sup>\ * (p[[y]]) \ (p\<^sup>T \ -1)\<^sup>+ * y" using 3 by (metis inf.antisym_conv inf.eq_refl inf_le1 mult_left_isotone star_plus mult_assoc) hence "p\<^sup>T\<^sup>\ * (p[[y]]) \ p\<^sup>T\<^sup>\ * y" by (metis mult_isotone order_refl star.left_plus_below_circ star_plus mult_assoc) hence 5: "?t \ ?s" using order_trans by auto have 6: "y \ ?s" using 1 find_set_invariant_def bijective_regular mult_left_isotone star.circ_reflexive by fastforce have 7: "\ y \ ?t" proof assume "y \ ?t" hence "y \ (p\<^sup>T \ -1)\<^sup>+ * y" using 4 by (metis reachable_without_loops mem_Collect_eq order_trans) hence "y * y\<^sup>T \ (p\<^sup>T \ -1)\<^sup>+" using 1 find_set_invariant_def shunt_bijective by simp also have "... \ -1" using 1 by (metis (mono_tags, lifting) find_set_invariant_def find_set_precondition_def conv_dist_comp conv_dist_inf conv_isotone conv_star_commute equivalence_one_closed star.circ_plus_same symmetric_complement_closed) finally have "y \ -y" using schroeder_4_p by auto thus False using 1 by (metis find_set_invariant_def comp_inf.coreflexive_idempotent conv_complement covector_vector_comp inf.absorb1 inf.sup_monoid.add_commute pseudo_complement surjective_conv_total top.extremum vector_top_closed regular_closed_top) qed show "card ?t < card ?s" apply (rule psubset_card_mono) subgoal using finite_regular by simp subgoal using 5 6 7 by auto done qed qed qed lemma find_set_3: "find_set_invariant p x y \ y = p[[y]] \ find_set_postcondition p x y" proof - assume 1: "find_set_invariant p x y \ y = p[[y]]" show "find_set_postcondition p x y" proof (unfold find_set_postcondition_def, rule conjI) show "point y" using 1 find_set_invariant_def by simp show "y = root p x" proof (rule order.antisym) have "y * y\<^sup>T \ p" using 1 by (metis find_set_invariant_def find_set_precondition_def shunt_bijective shunt_mapping top_right_mult_increasing) hence "y * y\<^sup>T \ p \ 1" using 1 find_set_invariant_def le_infI by blast hence "y \ (p \ 1) * top" using 1 by (metis find_set_invariant_def order_lesseq_imp shunt_bijective top_right_mult_increasing mult_assoc) thus "y \ root p x" using 1 find_set_invariant_def by simp next have 2: "x \ p\<^sup>\ * y" using 1 find_set_invariant_def find_set_precondition_def bijective_reverse conv_star_commute by auto have "p\<^sup>T * p\<^sup>\ * y = p\<^sup>T * p * p\<^sup>\ * y \ (p[[y]])" by (metis comp_associative mult_left_dist_sup star.circ_loop_fixpoint) also have "... \ p\<^sup>\ * y \ y" using 1 by (metis find_set_invariant_def find_set_precondition_def comp_isotone mult_left_sub_dist_sup semiring.add_right_mono star.circ_back_loop_fixpoint star.circ_circ_mult star.circ_top star.circ_transitive_equal star_involutive star_one) also have "... = p\<^sup>\ * y" by (metis star.circ_loop_fixpoint sup.left_idem sup_commute) finally have 3: "p\<^sup>T\<^sup>\ * x \ p\<^sup>\ * y" using 2 by (simp add: comp_associative star_left_induct) have "p * y \ (p \ 1) * top = (p \ 1) * p * y" using comp_associative coreflexive_comp_top_inf inf_commute by auto also have "... \ p\<^sup>T * p * y" by (metis inf.cobounded2 inf.sup_monoid.add_commute mult_left_isotone one_inf_conv) also have "... \ y" using 1 find_set_invariant_def find_set_precondition_def mult_left_isotone by fastforce finally have 4: "p * y \ y \ -((p \ 1) * top)" using 1 by (metis find_set_invariant_def shunting_p bijective_regular) have "p\<^sup>T * (p \ 1) \ p\<^sup>T \ 1" using 1 by (metis find_set_invariant_def find_set_precondition_def N_top comp_isotone coreflexive_idempotent inf.cobounded2 inf.sup_monoid.add_commute inf_assoc one_inf_conv shunt_mapping) hence "p\<^sup>T * (p \ 1) * top \ (p \ 1) * top" using inf_commute mult_isotone one_inf_conv by auto hence "p * -((p \ 1) * top) \ -((p \ 1) * top)" by (metis comp_associative inf.sup_monoid.add_commute p_antitone p_antitone_iff schroeder_3_p) hence "p * y \ p * -((p \ 1) * top) \ y \ -((p \ 1) * top)" using 4 dual_order.trans le_supI sup_ge2 by blast hence "p * (y \ -((p \ 1) * top)) \ y \ -((p \ 1) * top)" by (simp add: mult_left_dist_sup) hence "p\<^sup>\ * y \ y \ -((p \ 1) * top)" by (simp add: star_left_induct) hence "p\<^sup>T\<^sup>\ * x \ y \ -((p \ 1) * top)" using 3 dual_order.trans by blast thus "root p x \ y" using 1 by (metis find_set_invariant_def shunting_p bijective_regular) qed qed qed theorem find_set: "VARS y [ find_set_precondition p x ] y := x; WHILE y \ p[[y]] INV { find_set_invariant p x y } VAR { card { z . regular z \ z \ p\<^sup>T\<^sup>\ * y } } DO y := p[[y]] OD [ find_set_postcondition p x y ]" apply vcg_tc_simp apply (fact find_set_1) apply (fact find_set_2) by (fact find_set_3) lemma find_set_exists: "find_set_precondition p x \ \y . find_set_postcondition p x y" using tc_extract_function find_set by blast text \ The root of a component tree is a point, that is, represents a singleton set of nodes. This could be proved from the definitions using Kleene-relation algebraic calculations. But they can be avoided because the property directly follows from the postcondition of the previous correctness proof. The corresponding algorithm shows how to obtain the root. We therefore have an essentially constructive proof of the following result. \ -text \Theorem 3.3\ +text \Theorem 4.3\ lemma root_point: "disjoint_set_forest p \ point x \ point (root p x)" using find_set_exists find_set_precondition_def find_set_postcondition_def by simp definition "find_set p x \ (SOME y . find_set_postcondition p x y)" lemma find_set_function: assumes "find_set_precondition p x" and "y = find_set p x" shows "find_set_postcondition p x y" by (metis assms find_set_def find_set_exists someI) subsection \Path Compression\ text \ The path-compression technique is frequently implemented in recursive implementations of find-set modifying the tree on the way out from recursive calls. Here we implement it using a second while-loop, which iterates over the same path to the root and changes edges to point to the root of the component, which is known after the while-loop in find-set completes. We prove that path compression preserves the equivalence-relational semantics of the disjoint-set forest and also preserves the roots of the component trees. Additionally we prove the exact effect of path compression. \ definition "path_compression_precondition p x y \ disjoint_set_forest p \ point x \ point y \ y = root p x" definition "path_compression_invariant p x y p0 w \ path_compression_precondition p x y \ point w \ y \ p\<^sup>T\<^sup>\ * w \ (w \ x \ p[[x]] = y \ y \ x \ p\<^sup>T\<^sup>+ * w \ -x) \ p \ 1 = p0 \ 1 \ fc p = fc p0 \ - root p w = y \ (w \ y \ p[[w]] \ w \ p\<^sup>T\<^sup>+ * w \ -w) \ p[[w]] = p0[[w]] \ p0[p0\<^sup>T\<^sup>\ * x \ -(p0\<^sup>T\<^sup>\ * w)\y] = p \ + root p w = y \ (w \ y \ p\<^sup>T\<^sup>+ * w \ -w) \ p[[w]] = p0[[w]] \ p0[p0\<^sup>T\<^sup>\ * x \ -(p0\<^sup>T\<^sup>\ * w)\y] = p \ disjoint_set_forest p0 \ y = root p0 x \ w \ p0\<^sup>T\<^sup>\ * x" definition "path_compression_postcondition p x y p0 \ path_compression_precondition p x y \ p \ 1 = p0 \ 1 \ fc p = fc p0 \ p0[p0\<^sup>T\<^sup>\ * x\y] = p" text \ We first consider a variant that achieves the effect as a single update. The parents of all nodes reachable from x are simultaneously updated to the root of the component of x. \ lemma path_compression_exact: assumes "path_compression_precondition p0 x y" and "p0[p0\<^sup>T\<^sup>\ * x\y] = p" shows "p \ 1 = p0 \ 1" "fc p = fc p0" proof - have a1: "disjoint_set_forest p0" and a2: "point x" and a3: "point y" and a4: "y = root p0 x" using path_compression_precondition_def assms(1) by auto have 1: "regular (p0\<^sup>T\<^sup>\ * x)" using a1 a2 bijective_regular mapping_regular regular_closed_star regular_conv_closed regular_mult_closed by auto have "p \ 1 = (p0\<^sup>T\<^sup>\ * x \ y\<^sup>T \ 1) \ (-(p0\<^sup>T\<^sup>\ * x) \ p0 \ 1)" using assms(2) inf_sup_distrib2 by auto also have "... = (p0\<^sup>T\<^sup>\ * x \ p0 \ 1) \ (-(p0\<^sup>T\<^sup>\ * x) \ p0 \ 1)" proof - have "p0\<^sup>T\<^sup>\ * x \ y\<^sup>T \ 1 = p0\<^sup>T\<^sup>\ * x \ p0 \ 1" proof (rule order.antisym) have "(p0 \ 1) * p0\<^sup>T\<^sup>\ * x \ 1 \ p0" by (smt coreflexive_comp_top_inf_one inf.absorb_iff2 inf.cobounded2 inf.sup_monoid.add_assoc root_var) hence "p0\<^sup>T\<^sup>\ * x \ y\<^sup>T \ 1 \ p0" by (metis inf_le1 a4 conv_dist_inf coreflexive_symmetric inf.absorb2 inf.cobounded2 inf.sup_monoid.add_assoc root_var symmetric_one_closed) thus "p0\<^sup>T\<^sup>\ * x \ y\<^sup>T \ 1 \ p0\<^sup>T\<^sup>\ * x \ p0 \ 1" by (meson inf.le_sup_iff order.refl) have "p0\<^sup>T\<^sup>\ * x \ p0 \ 1 \ y" by (metis a4 coreflexive_comp_top_inf_one inf.cobounded1 inf_assoc inf_le2) thus "p0\<^sup>T\<^sup>\ * x \ p0 \ 1 \ p0\<^sup>T\<^sup>\ * x \ y\<^sup>T \ 1" by (smt conv_dist_inf coreflexive_symmetric inf.absorb_iff2 inf.cobounded2 inf.sup_monoid.add_assoc) qed thus ?thesis by simp qed also have "... = p0 \ 1" using 1 by (metis inf.sup_monoid.add_commute inf_sup_distrib1 maddux_3_11_pp) finally show "p \ 1 = p0 \ 1" . show "fc p = fc p0" proof (rule order.antisym) have 2: "univalent (p0[p0\<^sup>T\<^sup>\ * x\y])" by (simp add: a1 a2 a3 update_univalent mult_assoc) have 3: "-(p0\<^sup>T\<^sup>\ * x) \ p0 \ (p0[p0\<^sup>T\<^sup>\ * x\y])\<^sup>\ * (p0[p0\<^sup>T\<^sup>\ * x\y])\<^sup>T\<^sup>\" using fc_increasing inf.order_trans sup.cobounded2 by blast have "p0\<^sup>T\<^sup>\ * x \ p0 \ (p0\<^sup>T\<^sup>\ \ p0 * x\<^sup>T) * (x \ p0\<^sup>\ * p0)" by (metis conv_involutive conv_star_commute dedekind) also have "... \ p0\<^sup>T\<^sup>\ * x \ p0 * x\<^sup>T * p0\<^sup>\ * p0" by (metis comp_associative inf.boundedI inf.cobounded2 inf_le1 mult_isotone) also have "... \ p0\<^sup>T\<^sup>\ * x \ top * x\<^sup>T * p0\<^sup>\" using comp_associative comp_inf.mult_right_isotone mult_isotone star.right_plus_below_circ by auto also have "... = p0\<^sup>T\<^sup>\ * x * x\<^sup>T * p0\<^sup>\" by (metis a2 symmetric_top_closed vector_covector vector_inf_comp vector_mult_closed) also have "... \ (p0\<^sup>T\<^sup>\ * x * y\<^sup>T) * (y * x\<^sup>T * p0\<^sup>\)" by (metis a3 order.antisym comp_inf.top_right_mult_increasing conv_involutive dedekind_1 inf.sup_left_divisibility inf.sup_monoid.add_commute mult_right_isotone surjective_conv_total mult_assoc) also have "... = (p0\<^sup>T\<^sup>\ * x \ y\<^sup>T) * (y \ x\<^sup>T * p0\<^sup>\)" by (metis a2 a3 vector_covector vector_inf_comp vector_mult_closed) also have "... = (p0\<^sup>T\<^sup>\ * x \ y\<^sup>T) * (p0\<^sup>T\<^sup>\ * x \ y\<^sup>T)\<^sup>T" by (simp add: conv_dist_comp conv_dist_inf conv_star_commute inf_commute) also have "... \ (p0[p0\<^sup>T\<^sup>\ * x\y])\<^sup>\ * (p0[p0\<^sup>T\<^sup>\ * x\y])\<^sup>T\<^sup>\" by (meson conv_isotone dual_order.trans mult_isotone star.circ_increasing sup.cobounded1) finally have "p0\<^sup>T\<^sup>\ * x \ p0 \ (p0[p0\<^sup>T\<^sup>\ * x\y])\<^sup>\ * (p0[p0\<^sup>T\<^sup>\ * x\y])\<^sup>T\<^sup>\" . hence "(p0\<^sup>T\<^sup>\ * x \ p0) \ (-(p0\<^sup>T\<^sup>\ * x) \ p0) \ (p0[p0\<^sup>T\<^sup>\ * x\y])\<^sup>\ * (p0[p0\<^sup>T\<^sup>\ * x\y])\<^sup>T\<^sup>\" using 3 le_supI by blast hence "p0 \ (p0[p0\<^sup>T\<^sup>\ * x\y])\<^sup>\ * (p0[p0\<^sup>T\<^sup>\ * x\y])\<^sup>T\<^sup>\" using 1 by (metis inf_commute maddux_3_11_pp) hence "fc p0 \ (p0[p0\<^sup>T\<^sup>\ * x\y])\<^sup>\ * (p0[p0\<^sup>T\<^sup>\ * x\y])\<^sup>T\<^sup>\" using 2 fc_idempotent fc_isotone by fastforce thus "fc p0 \ fc p" by (simp add: assms(2)) have "((p0\<^sup>T\<^sup>\ * x \ y\<^sup>T) \ (-(p0\<^sup>T\<^sup>\ * x) \ p0))\<^sup>\ = (-(p0\<^sup>T\<^sup>\ * x) \ p0)\<^sup>\ * ((p0\<^sup>T\<^sup>\ * x \ y\<^sup>T) \ 1)" proof (rule star_sup_2) have 4: "transitive (p0\<^sup>T\<^sup>\ * x)" using a2 comp_associative mult_right_isotone rectangle_star_rectangle by auto have "transitive (y\<^sup>T)" by (metis a3 conv_dist_comp inf.eq_refl mult_assoc) thus "transitive (p0\<^sup>T\<^sup>\ * x \ y\<^sup>T)" using 4 transitive_inf_closed by auto have 5: "p0\<^sup>T\<^sup>\ * x * (-(p0\<^sup>T\<^sup>\ * x) \ p0) \ p0\<^sup>T\<^sup>\ * x" by (metis a2 mult_right_isotone top_greatest mult_assoc) have "(-(p0\<^sup>T\<^sup>\ * x) \ p0)\<^sup>T * y \ p0\<^sup>T * y" by (simp add: conv_dist_inf mult_left_isotone) also have "... \ y" using a1 a4 root_successor_loop by auto finally have "y\<^sup>T * (-(p0\<^sup>T\<^sup>\ * x) \ p0) \ y\<^sup>T" using conv_dist_comp conv_isotone by fastforce thus "(p0\<^sup>T\<^sup>\ * x \ y\<^sup>T) * (-(p0\<^sup>T\<^sup>\ * x) \ p0) \ p0\<^sup>T\<^sup>\ * x \ y\<^sup>T" using 5 comp_left_subdist_inf inf_mono order_trans by blast qed hence "p\<^sup>\ = (-(p0\<^sup>T\<^sup>\ * x) \ p0)\<^sup>\ * ((p0\<^sup>T\<^sup>\ * x \ y\<^sup>T) \ 1)" by (simp add: assms(2)) also have "... \ p0\<^sup>\ * ((p0\<^sup>T\<^sup>\ * x \ y\<^sup>T) \ 1)" by (simp add: mult_left_isotone star_isotone) also have "... = p0\<^sup>\ * (p0\<^sup>T\<^sup>\ * x * y\<^sup>T \ 1)" by (simp add: a2 a3 vector_covector vector_mult_closed) also have "... = p0\<^sup>\ * (p0\<^sup>T\<^sup>\ * (x * x\<^sup>T) * p0\<^sup>\ * (p0 \ 1) \ 1)" by (metis a4 coreflexive_symmetric inf.cobounded2 root_var comp_associative conv_dist_comp conv_involutive conv_star_commute) also have "... \ p0\<^sup>\ * (p0\<^sup>T\<^sup>\ * 1 * p0\<^sup>\ * (p0 \ 1) \ 1)" by (metis a2 mult_left_isotone mult_right_isotone semiring.add_left_mono sup_commute) also have "... = p0\<^sup>\ * (p0\<^sup>T\<^sup>\ * (p0 \ 1) \ p0\<^sup>\ * (p0 \ 1) \ 1)" by (simp add: a1 cancel_separate_eq mult_right_dist_sup) also have "... = p0\<^sup>\ * ((p0 \ 1) \ p0\<^sup>\ * (p0 \ 1) \ 1)" by (smt univalent_root_successors a1 conv_dist_comp conv_dist_inf coreflexive_idempotent coreflexive_symmetric inf.cobounded2 injective_codomain loop_root root_transitive_successor_loop symmetric_one_closed) also have "... = p0\<^sup>\ * (p0\<^sup>\ * (p0 \ 1) \ 1)" by (metis inf.sup_left_divisibility inf_commute sup.left_idem sup_commute sup_relative_same_increasing) also have "... \ p0\<^sup>\ * p0\<^sup>\" by (metis inf.cobounded2 inf_commute order.refl order_lesseq_imp star.circ_mult_upper_bound star.circ_reflexive star.circ_transitive_equal sup.boundedI sup_monoid.add_commute) also have "... = p0\<^sup>\" by (simp add: star.circ_transitive_equal) finally show "fc p \ fc p0" by (metis conv_order conv_star_commute mult_isotone) qed qed lemma update_acyclic_6: assumes "disjoint_set_forest p" and "point x" shows "acyclic ((p[p\<^sup>T\<^sup>\*x\root p x]) \ -1)" using assms root_point root_successor_loop update_acyclic_2 by auto theorem path_compression_assign: - "VARS p t w + "VARS p [ path_compression_precondition p x y \ p0 = p ] p[p\<^sup>T\<^sup>\ * x] := y [ path_compression_postcondition p x y p0 ]" apply vcg_tc_simp apply (unfold path_compression_precondition_def path_compression_postcondition_def) apply (intro conjI) subgoal using update_univalent mult_assoc by auto subgoal using bijective_regular mapping_regular regular_closed_star regular_conv_closed regular_mult_closed update_mapping mult_assoc by auto subgoal using update_acyclic_6 by blast subgoal by blast subgoal by blast subgoal by blast subgoal by blast subgoal by blast subgoal by blast subgoal by (smt same_root path_compression_exact path_compression_precondition_def update_univalent vector_mult_closed) subgoal using path_compression_exact(1) path_compression_precondition_def by blast subgoal using path_compression_exact(2) path_compression_precondition_def by blast by blast text \ We next look at implementing these updates using a loop. \ lemma path_compression_1a: assumes "point x" and "disjoint_set_forest p" and "x \ root p x" shows "p\<^sup>T\<^sup>+ * x \ - x" by (meson assms bijective_regular mapping_regular regular_closed_star regular_conv_closed regular_mult_closed vector_mult_closed point_in_vector_or_complement_2 loop_root_2) lemma path_compression_1b: "x \ p\<^sup>T\<^sup>\ * x" using mult_left_isotone star.circ_reflexive by fastforce lemma path_compression_1: "path_compression_precondition p x y \ path_compression_invariant p x y p x" using path_compression_invariant_def path_compression_precondition_def loop_root path_compression_1a path_compression_1b by auto lemma path_compression_2: "path_compression_invariant p x y p0 w \ y \ p[[w]] \ path_compression_invariant (p[w\y]) x y p0 (p[[w]]) \ card { z . regular z \ z \ (p[w\y])\<^sup>T\<^sup>\ * (p[[w]]) } < card { z . regular z \ z \ p\<^sup>T\<^sup>\ * w }" proof - let ?p = "p[w\y]" let ?s = "{ z . regular z \ z \ p\<^sup>T\<^sup>\ * w }" let ?t = "{ z . regular z \ z \ ?p\<^sup>T\<^sup>\ * (p[[w]]) }" assume 1: "path_compression_invariant p x y p0 w \ y \ p[[w]]" have i1: "disjoint_set_forest p" and i2: "point x" and i3: "point y" and i4: "y = root p x" using 1 path_compression_invariant_def path_compression_precondition_def by meson+ have i5: "point w" and i6: "y \ p\<^sup>T\<^sup>\ * w" and i7: "w \ x \ p[[x]] = y \ y \ x \ p\<^sup>T\<^sup>+ * w \ -x" and i8: "p \ 1 = p0 \ 1" and i9: "fc p = fc p0" - and i10: "root p w = y" and i11: "w \ y \ p[[w]] \ w" and i12: "p[[w]] = p0[[w]]" and i13: "p0[p0\<^sup>T\<^sup>\ * x \ -(p0\<^sup>T\<^sup>\ * w)\y] = p" + and i10: "root p w = y" and i11: "p[[w]] = p0[[w]]" and i12: "p0[p0\<^sup>T\<^sup>\ * x \ -(p0\<^sup>T\<^sup>\ * w)\y] = p" using 1 path_compression_invariant_def by blast+ - have i14: "disjoint_set_forest p0" and i15: "y = root p0 x" and i16: "w \ p0\<^sup>T\<^sup>\ * x" + have i13: "disjoint_set_forest p0" and i14: "y = root p0 x" and i15: "w \ p0\<^sup>T\<^sup>\ * x" using 1 path_compression_invariant_def by auto have 2: "point (p[[w]])" using i1 i5 read_point by blast show "path_compression_invariant ?p x y p0 (p[[w]]) \ card ?t < card ?s" proof (unfold path_compression_invariant_def, intro conjI) have 3: "mapping ?p" by (simp add: i1 i3 i5 bijective_regular update_total update_univalent) have 4: "w \ y" using 1 i1 i4 root_successor_loop by blast hence 5: "w \ y = bot" by (simp add: i3 i5 distinct_points) hence "y * w\<^sup>T \ -1" using pseudo_complement schroeder_4_p by auto hence "y * w\<^sup>T \ p\<^sup>T\<^sup>\ \ -1" using i5 i6 shunt_bijective by auto also have "... \ p\<^sup>T\<^sup>+" by (simp add: star_plus_without_loops) finally have 6: "y \ p\<^sup>T\<^sup>+ * w" using i5 shunt_bijective by auto have 7: "w * w\<^sup>T \ -p\<^sup>T\<^sup>+" proof (rule ccontr) assume "\ w * w\<^sup>T \ -p\<^sup>T\<^sup>+" hence "w * w\<^sup>T \ --p\<^sup>T\<^sup>+" using i5 point_arc arc_in_partition by blast hence "w * w\<^sup>T \ p\<^sup>T\<^sup>+ \ 1" using i1 i5 mapping_regular regular_conv_closed regular_closed_star regular_mult_closed by simp also have "... = ((p\<^sup>T \ 1) * p\<^sup>T\<^sup>\ \ 1) \ ((p\<^sup>T \ -1) * p\<^sup>T\<^sup>\ \ 1)" by (metis comp_inf.mult_right_dist_sup maddux_3_11_pp mult_right_dist_sup regular_one_closed) also have "... = ((p\<^sup>T \ 1) * p\<^sup>T\<^sup>\ \ 1) \ ((p \ -1)\<^sup>+ \ 1)\<^sup>T" by (metis conv_complement conv_dist_inf conv_plus_commute equivalence_one_closed reachable_without_loops) also have "... \ ((p\<^sup>T \ 1) * p\<^sup>T\<^sup>\ \ 1) \ (-1 \ 1)\<^sup>T" by (metis (no_types, hide_lams) i1 sup_right_isotone inf.sup_left_isotone conv_isotone) also have "... = (p\<^sup>T \ 1) * p\<^sup>T\<^sup>\ \ 1" by simp also have "... \ (p\<^sup>T \ 1) * top \ 1" by (metis comp_inf.comp_isotone coreflexive_comp_top_inf equivalence_one_closed inf.cobounded1 inf.cobounded2) also have "... \ p\<^sup>T" by (simp add: coreflexive_comp_top_inf_one) finally have "w * w\<^sup>T \ p\<^sup>T" by simp hence "w \ p[[w]]" using i5 shunt_bijective by blast hence "w = p[[w]]" using 2 by (metis i5 epm_3 mult_semi_associative) thus False - using 4 i11 by auto + using 2 4 i10 loop_root by auto qed hence 8: "w \ p\<^sup>T\<^sup>+ * w = bot" using p_antitone_iff pseudo_complement schroeder_4_p by blast show "y \ ?p\<^sup>T\<^sup>\ * (p[[w]])" proof - have "(w \ y\<^sup>T)\<^sup>T * (-w \ p)\<^sup>T\<^sup>\ * p\<^sup>T * w \ w\<^sup>T * (-w \ p)\<^sup>T\<^sup>\ * p\<^sup>T * w" by (simp add: conv_isotone mult_left_isotone) also have "... \ w\<^sup>T * p\<^sup>T\<^sup>\ * p\<^sup>T * w" by (simp add: conv_isotone mult_left_isotone star_isotone mult_right_isotone) also have "... = w\<^sup>T * p\<^sup>T\<^sup>+ * w" by (simp add: star_plus mult_assoc) also have "... = bot" using 8 by (smt i5 covector_inf_comp_3 mult_assoc conv_dist_comp conv_star_commute covector_bot_closed equivalence_top_closed inf.le_iff_sup mult_left_isotone) finally have "((w \ y\<^sup>T)\<^sup>T \ (-w \ p)\<^sup>T) * (-w \ p)\<^sup>T\<^sup>\ * p\<^sup>T * w \ (-w \ p)\<^sup>T * (-w \ p)\<^sup>T\<^sup>\ * p\<^sup>T * w" by (simp add: bot_unique mult_right_dist_sup) also have "... \ (-w \ p)\<^sup>T\<^sup>\ * p\<^sup>T * w" by (simp add: mult_left_isotone star.left_plus_below_circ) finally have "?p\<^sup>T * (-w \ p)\<^sup>T\<^sup>\ * p\<^sup>T * w \ (-w \ p)\<^sup>T\<^sup>\ * p\<^sup>T * w" by (simp add: conv_dist_sup) hence "?p\<^sup>T\<^sup>\ * p\<^sup>T * w \ (-w \ p)\<^sup>T\<^sup>\ * p\<^sup>T * w" by (metis comp_associative star.circ_loop_fixpoint star_left_induct sup_commute sup_least sup_left_divisibility) hence "w \ ?p\<^sup>T\<^sup>\ * p\<^sup>T * w \ w \ (-w \ p)\<^sup>T\<^sup>\ * p\<^sup>T * w" using inf.sup_right_isotone by blast also have "... \ w \ p\<^sup>T\<^sup>\ * p\<^sup>T * w" using conv_isotone mult_left_isotone star_isotone inf.sup_right_isotone by simp also have "... = bot" using 8 by (simp add: star_plus) finally have 9: "w\<^sup>T * ?p\<^sup>T\<^sup>\ * p\<^sup>T * w = bot" by (smt i5 covector_inf_comp_3 mult_assoc conv_dist_comp covector_bot_closed equivalence_top_closed inf.le_iff_sup mult_left_isotone bot_least inf.absorb1) have "p\<^sup>T * ?p\<^sup>T\<^sup>\ * p\<^sup>T * w = ((w \ p)\<^sup>T \ (-w \ p)\<^sup>T) * ?p\<^sup>T\<^sup>\ * p\<^sup>T * w" by (metis i5 bijective_regular conv_dist_sup inf.sup_monoid.add_commute maddux_3_11_pp) also have "... = (w \ p)\<^sup>T * ?p\<^sup>T\<^sup>\ * p\<^sup>T * w \ (-w \ p)\<^sup>T * ?p\<^sup>T\<^sup>\ * p\<^sup>T * w" by (simp add: mult_right_dist_sup) also have "... \ w\<^sup>T * ?p\<^sup>T\<^sup>\ * p\<^sup>T * w \ (-w \ p)\<^sup>T * ?p\<^sup>T\<^sup>\ * p\<^sup>T * w" using semiring.add_right_mono comp_isotone conv_isotone by auto also have "... = (-w \ p)\<^sup>T * ?p\<^sup>T\<^sup>\ * p\<^sup>T * w" using 9 by simp also have "... \ ?p\<^sup>T\<^sup>+ * p\<^sup>T * w" by (simp add: conv_isotone mult_left_isotone) also have "... \ ?p\<^sup>T\<^sup>\ * p\<^sup>T * w" by (simp add: comp_isotone star.left_plus_below_circ) finally have "p\<^sup>T\<^sup>\ * p\<^sup>T * w \ ?p\<^sup>T\<^sup>\ * p\<^sup>T * w" by (metis comp_associative star.circ_loop_fixpoint star_left_induct sup_commute sup_least sup_left_divisibility) thus "y \ ?p\<^sup>T\<^sup>\ * (p[[w]])" using 6 by (simp add: star_simulation_right_equal mult_assoc) qed have 10: "acyclic (?p \ -1)" using i1 i10 i3 i5 inf_le1 update_acyclic_3 by blast have "?p[[p\<^sup>T\<^sup>+ * w]] \ p\<^sup>T\<^sup>+ * w" proof - have "(w\<^sup>T \ y) * p\<^sup>T\<^sup>+ * w = y \ w\<^sup>T * p\<^sup>T\<^sup>+ * w" by (metis i3 inf_vector_comp vector_inf_comp) hence "?p[[p\<^sup>T\<^sup>+ * w]] = (y \ w\<^sup>T * p\<^sup>T\<^sup>+ * w) \ (-w\<^sup>T \ p\<^sup>T) * p\<^sup>T\<^sup>+ * w" by (simp add: comp_associative conv_complement conv_dist_inf conv_dist_sup mult_right_dist_sup) also have "... \ y \ (-w\<^sup>T \ p\<^sup>T) * p\<^sup>T\<^sup>+ * w" using sup_left_isotone by auto also have "... \ y \ p\<^sup>T * p\<^sup>T\<^sup>+ * w" using mult_left_isotone sup_right_isotone by auto also have "... \ y \ p\<^sup>T\<^sup>+ * w" using semiring.add_left_mono mult_left_isotone mult_right_isotone star.left_plus_below_circ by auto also have "... = p\<^sup>T\<^sup>+ * w" using 6 by (simp add: sup_absorb2) finally show ?thesis by simp qed hence 11: "?p\<^sup>T\<^sup>\ * (p[[w]]) \ p\<^sup>T\<^sup>+ * w" using star_left_induct by (simp add: mult_left_isotone star.circ_mult_increasing) hence 12: "?p\<^sup>T\<^sup>+ * (p[[w]]) \ p\<^sup>T\<^sup>+ * w" using dual_order.trans mult_left_isotone star.left_plus_below_circ by blast have 13: "?p[[x]] = y \ y \ x \ ?p\<^sup>T\<^sup>+ * (p[[w]]) \ -x" proof (cases "w = x") case True hence "?p[[x]] = (w\<^sup>T \ y) * w \ (-w\<^sup>T \ p\<^sup>T) * w" by (simp add: conv_complement conv_dist_inf conv_dist_sup mult_right_dist_sup) also have "... = (w\<^sup>T \ y) * w \ p\<^sup>T * (-w \ w)" by (metis i5 conv_complement covector_inf_comp_3 inf.sup_monoid.add_commute vector_complement_closed) also have "... = (w\<^sup>T \ y) * w" by simp also have "... = y * w" by (simp add: i5 covector_inf_comp_3 inf.sup_monoid.add_commute) also have "... = y" by (metis i3 i5 comp_associative) finally show ?thesis using 4 8 12 True pseudo_complement inf.sup_monoid.add_commute order.trans by blast next case False have "?p[[x]] = (w\<^sup>T \ y) * x \ (-w\<^sup>T \ p\<^sup>T) * x" by (simp add: conv_complement conv_dist_inf conv_dist_sup mult_right_dist_sup) also have "... = y * (w \ x) \ p\<^sup>T * (-w \ x)" by (metis i5 conv_complement covector_inf_comp_3 inf_commute vector_complement_closed) also have "... = p\<^sup>T * (-w \ x)" using i2 i5 False distinct_points by auto also have "... = y" using i2 i5 i7 False distinct_points inf.absorb2 pseudo_complement by auto finally show ?thesis using 12 False i7 dual_order.trans by blast qed thus "p[[w]] \ x \ ?p[[x]] = y \ y \ x \ ?p\<^sup>T\<^sup>+ * (p[[w]]) \ -x" by simp have 14: "?p\<^sup>T\<^sup>\ * x = x \ y" proof (rule order.antisym) have "?p\<^sup>T * (x \ y) = y \ ?p\<^sup>T * y" using 13 by (simp add: mult_left_dist_sup) also have "... = y \ (w\<^sup>T \ y) * y \ (-w\<^sup>T \ p\<^sup>T) * y" by (simp add: conv_complement conv_dist_inf conv_dist_sup mult_right_dist_sup sup_assoc) also have "... \ y \ (w\<^sup>T \ y) * y \ p\<^sup>T * y" using mult_left_isotone sup_right_isotone by auto also have "... = y \ (w\<^sup>T \ y) * y" using i1 i10 root_successor_loop sup_commute by auto also have "... \ y \ y * y" using mult_left_isotone sup_right_isotone by auto also have "... = y" by (metis i3 comp_associative sup.idem) also have "... \ x \ y" by simp finally show "?p\<^sup>T\<^sup>\ * x \ x \ y" by (simp add: star_left_induct) next show "x \ y \ ?p\<^sup>T\<^sup>\ * x" using 13 by (metis mult_left_isotone star.circ_increasing star.circ_loop_fixpoint sup.boundedI sup_ge2) qed have 15: "y = root ?p x" proof - have "(p \ 1) * y = (p \ 1) * (p \ 1) * p\<^sup>T\<^sup>\ * x" by (simp add: i4 comp_associative root_var) also have "... = (p \ 1) * p\<^sup>T\<^sup>\ * x" using coreflexive_idempotent by auto finally have 16: "(p \ 1) * y = y" by (simp add: i4 root_var) have 17: "(p \ 1) * x \ y" by (metis (no_types, lifting) i4 comp_right_one mult_left_isotone mult_right_isotone star.circ_reflexive root_var) have "root ?p x = (?p \ 1) * (x \ y)" using 14 by (metis mult_assoc root_var) also have "... = (w \ y\<^sup>T \ 1) * (x \ y) \ (-w \ p \ 1) * (x \ y)" by (simp add: inf_sup_distrib2 semiring.distrib_right) also have "... = (w \ 1 \ y\<^sup>T) * (x \ y) \ (-w \ p \ 1) * (x \ y)" by (simp add: inf.left_commute inf.sup_monoid.add_commute) also have "... = (w \ 1) * (y \ (x \ y)) \ (-w \ p \ 1) * (x \ y)" by (simp add: i3 covector_inf_comp_3) also have "... = (w \ 1) * y \ (-w \ p \ 1) * (x \ y)" by (simp add: inf.absorb1) also have "... = (w \ 1 * y) \ (-w \ (p \ 1) * (x \ y))" by (simp add: i5 inf_assoc vector_complement_closed vector_inf_comp) also have "... = (w \ y) \ (-w \ ((p \ 1) * x \ y))" using 16 by (simp add: mult_left_dist_sup) also have "... = (w \ y) \ (-w \ y)" using 17 by (simp add: sup.absorb2) also have "... = y" using 5 inf.sup_monoid.add_commute le_iff_inf pseudo_complement sup_monoid.add_0_left by fastforce finally show ?thesis by simp qed show "path_compression_precondition ?p x y" using 3 10 15 i2 i3 path_compression_precondition_def by blast show "vector (p[[w]])" using 2 by simp show "injective (p[[w]])" using 2 by simp show "surjective (p[[w]])" using 2 by simp have "w \ p \ 1 \ w \ w\<^sup>T \ p" by (metis inf.boundedE inf.boundedI inf.cobounded1 inf.cobounded2 one_inf_conv) also have "... = w * w\<^sup>T \ p" by (simp add: i5 vector_covector) also have "... \ -p\<^sup>T\<^sup>+ \ p" using 7 by (simp add: inf.coboundedI2 inf.sup_monoid.add_commute) finally have "w \ p \ 1 = bot" by (metis (no_types, hide_lams) conv_dist_inf coreflexive_symmetric inf.absorb1 inf.boundedE inf.cobounded2 pseudo_complement star.circ_mult_increasing) also have "w \ y\<^sup>T \ 1 = bot" using 5 antisymmetric_bot_closed asymmetric_bot_closed comp_inf.schroeder_2 inf.absorb1 one_inf_conv by fastforce finally have "w \ p \ 1 = w \ y\<^sup>T \ 1" by simp thus 18: "?p \ 1 = p0 \ 1" by (metis i5 i8 bijective_regular inf.sup_monoid.add_commute inf_sup_distrib2 maddux_3_11_pp) show 19: "fc ?p = fc p0" proof - have "p[[w]] = p\<^sup>T * (w \ p\<^sup>\ * y)" by (metis i3 i5 i6 bijective_reverse conv_star_commute inf.absorb1) also have "... = p\<^sup>T * (w \ p\<^sup>\) * y" by (simp add: i5 vector_inf_comp mult_assoc) also have "... = p\<^sup>T * ((w \ 1) \ (w \ p) * (-w \ p)\<^sup>\) * y" by (simp add: i5 omit_redundant_points) also have "... = p\<^sup>T * (w \ 1) * y \ p\<^sup>T * (w \ p) * (-w \ p)\<^sup>\ * y" by (simp add: comp_associative mult_left_dist_sup mult_right_dist_sup) also have "... \ p\<^sup>T * y \ p\<^sup>T * (w \ p) * (-w \ p)\<^sup>\ * y" by (metis semiring.add_right_mono comp_isotone order.eq_iff inf.cobounded1 inf.sup_monoid.add_commute mult_1_right) also have "... = y \ p\<^sup>T * (w \ p) * (-w \ p)\<^sup>\ * y" using i1 i4 root_successor_loop by auto also have "... \ y \ p\<^sup>T * p * (-w \ p)\<^sup>\ * y" using comp_isotone sup_right_isotone by auto also have "... \ y \ (-w \ p)\<^sup>\ * y" by (metis i1 comp_associative eq_refl shunt_mapping sup_right_isotone) also have "... = (-w \ p)\<^sup>\ * y" by (metis star.circ_loop_fixpoint sup.left_idem sup_commute) finally have 20: "p[[w]] \ (-w \ p)\<^sup>\ * y" by simp have "p\<^sup>T * (-w \ p)\<^sup>\ * y = p\<^sup>T * y \ p\<^sup>T * (-w \ p) * (-w \ p)\<^sup>\ * y" by (metis comp_associative mult_left_dist_sup star.circ_loop_fixpoint sup_commute) also have "... = y \ p\<^sup>T * (-w \ p) * (-w \ p)\<^sup>\ * y" using i1 i4 root_successor_loop by auto also have "... \ y \ p\<^sup>T * p * (-w \ p)\<^sup>\ * y" using comp_isotone sup_right_isotone by auto also have "... \ y \ (-w \ p)\<^sup>\ * y" by (metis i1 comp_associative eq_refl shunt_mapping sup_right_isotone) also have "... = (-w \ p)\<^sup>\ * y" by (metis star.circ_loop_fixpoint sup.left_idem sup_commute) finally have 21: "p\<^sup>T\<^sup>\ * p\<^sup>T * w \ (-w \ p)\<^sup>\ * y" using 20 by (simp add: comp_associative star_left_induct) have "w\<^sup>T \ p\<^sup>T = p\<^sup>T * (w\<^sup>T \ 1)" by (metis i5 comp_right_one covector_inf_comp_3 inf.sup_monoid.add_commute one_inf_conv) also have "... \ p[[w]]" by (metis comp_right_subdist_inf inf.boundedE inf.sup_monoid.add_commute one_inf_conv) also have "... \ p\<^sup>T\<^sup>\ * p\<^sup>T * w" by (simp add: mult_left_isotone star.circ_mult_increasing_2) also have "... \ (-w \ p)\<^sup>\ * y" using 21 by simp finally have "w \ p \ y\<^sup>T * (-w \ p)\<^sup>T\<^sup>\" by (metis conv_dist_comp conv_dist_inf conv_involutive conv_isotone conv_star_commute) hence "w \ p \ (w \ y\<^sup>T) * (-w \ p)\<^sup>T\<^sup>\" by (simp add: i5 vector_inf_comp) also have "... \ (w \ y\<^sup>T) * ?p\<^sup>T\<^sup>\" by (simp add: conv_isotone mult_right_isotone star_isotone) also have "... \ ?p * ?p\<^sup>T\<^sup>\" by (simp add: mult_left_isotone) also have "... \ fc ?p" by (simp add: mult_left_isotone star.circ_increasing) finally have 22: "w \ p \ fc ?p" by simp have "-w \ p \ ?p" by simp also have "... \ fc ?p" by (simp add: fc_increasing) finally have "(w \ -w) \ p \ fc ?p" using 22 by (simp add: comp_inf.semiring.distrib_left inf.sup_monoid.add_commute) hence "p \ fc ?p" by (metis i5 bijective_regular inf.sup_monoid.add_commute inf_sup_distrib1 maddux_3_11_pp) hence 23: "fc p \ fc ?p" using 3 fc_idempotent fc_isotone by fastforce have "?p \ (w \ y\<^sup>T) \ p" using sup_right_isotone by auto also have "... = w * y\<^sup>T \ p" by (simp add: i3 i5 vector_covector) also have "... \ p\<^sup>\ \ p" by (smt i5 i6 conv_dist_comp conv_involutive conv_isotone conv_star_commute le_supI shunt_bijective star.circ_increasing sup_absorb1) also have "... \ fc p" using fc_increasing star.circ_back_loop_prefixpoint by auto finally have "fc ?p \ fc p" using i1 fc_idempotent fc_isotone by fastforce thus ?thesis using 23 i9 by auto qed show "?p[[p[[w]]]] = p0[[p[[w]]]]" proof - have "?p[[p[[w]]]] = p[[p[[w]]]]" - using 2 4 i5 i11 by (simp add: put_get_different) + using 2 4 by (metis i5 i10 loop_root put_get_different) also have "... = p[[p0[[w]]]]" - by (simp add: i12) + by (simp add: i11) also have "... = (p0[p0\<^sup>T\<^sup>\ * x \ -(p0\<^sup>T\<^sup>\ * w)\y])[[p0[[w]]]]" - using i13 by auto + using i12 by auto also have "... = p0[[p0[[w]]]]" proof - have "p0[[w]] \ -(p0\<^sup>T\<^sup>\ * x \ -(p0\<^sup>T\<^sup>\ * w))" by (meson inf.coboundedI2 mult_left_isotone p_antitone p_antitone_iff star.circ_increasing) thus ?thesis by (meson i2 i5 put_get_different_vector vector_complement_closed vector_inf_closed vector_mult_closed) qed also have "... = p0[[p[[w]]]]" - by (simp add: i12) + by (simp add: i11) finally show ?thesis . qed have 24: "root ?p (p[[w]]) = root p0 (p[[w]])" - using 3 18 19 i14 same_root by blast + using 3 18 19 i13 same_root by blast also have "... = root p0 (p0[[w]])" - by (simp add: i12) + by (simp add: i11) also have 25: "... = root p0 w" - by (metis i5 i14 conv_involutive forest_components_increasing mult_left_isotone shunt_bijective injective_mult_closed read_surjective same_component_same_root) + by (metis i5 i13 conv_involutive forest_components_increasing mult_left_isotone shunt_bijective injective_mult_closed read_surjective same_component_same_root) finally show 26: "root ?p (p[[w]]) = y" - by (metis i1 i10 i14 i8 i9 same_root) - show "p[[w]] \ y \ ?p[[p[[w]]]] \ p[[w]] \ ?p\<^sup>T\<^sup>+ * (p[[w]]) \ -(p[[w]])" - proof - assume 27: "p[[w]] \ y" - show "?p[[p[[w]]]] \ p[[w]] \ ?p\<^sup>T\<^sup>+ * (p[[w]]) \ -(p[[w]])" - proof - show "?p[[p[[w]]]] \ p[[w]]" - using 2 26 27 loop_root by auto - show "?p\<^sup>T\<^sup>+ * (p[[w]]) \ -(p[[w]])" - using 2 3 10 26 27 by (simp add: path_compression_1a) - qed - qed + by (metis i1 i10 i13 i8 i9 same_root) + thus "p[[w]] \ y \ ?p\<^sup>T\<^sup>+ * (p[[w]]) \ -(p[[w]])" + using 2 3 10 by (simp add: path_compression_1a) show "univalent p0" "total p0" "acyclic (p0 \ -1)" - by (simp_all add: i14) + by (simp_all add: i13) show "y = root p0 x" - by (simp add: i15) + by (simp add: i14) show "p[[w]] \ p0\<^sup>T\<^sup>\ * x" - by (metis i12 i16 mult_isotone star.circ_increasing star.circ_transitive_equal mult_assoc) + by (metis i11 i15 mult_isotone star.circ_increasing star.circ_transitive_equal mult_assoc) let ?q = "p0[p0\<^sup>T\<^sup>\ * x \ -(p0\<^sup>T\<^sup>\ * (p[[w]]))\y]" show "?q = ?p" proof - - have 28: "w \ p0\<^sup>T\<^sup>+ * w = p0\<^sup>T\<^sup>\ * w" + have 27: "w \ p0\<^sup>T\<^sup>+ * w = p0\<^sup>T\<^sup>\ * w" using comp_associative star.circ_loop_fixpoint sup_commute by auto - hence 29: "p0\<^sup>T\<^sup>+ * w = p0\<^sup>T\<^sup>\ * w \ -w" - using 4 24 25 26 by (metis i12 i14 i5 inf.orderE maddux_3_13 path_compression_1a) + hence 28: "p0\<^sup>T\<^sup>+ * w = p0\<^sup>T\<^sup>\ * w \ -w" + using 4 24 25 26 by (metis i11 i13 i5 inf.orderE maddux_3_13 path_compression_1a) hence "p0\<^sup>T\<^sup>\ * (p[[w]]) \ -w" - by (metis i12 inf_le2 star_plus mult.assoc) + by (metis i11 inf_le2 star_plus mult.assoc) hence "w \ -(p0\<^sup>T\<^sup>\ * (p[[w]]))" by (simp add: p_antitone_iff) hence "w \ p0\<^sup>T\<^sup>\ * x \ -(p0\<^sup>T\<^sup>\ * (p[[w]]))" - by (simp add: i16) - hence 30: "?q \ w = ?p \ w" + by (simp add: i15) + hence 29: "?q \ w = ?p \ w" by (metis update_inf update_inf_same) - have 31: "?q \ p0\<^sup>T\<^sup>+ * w = ?p \ p0\<^sup>T\<^sup>+ * w" + have 30: "?q \ p0\<^sup>T\<^sup>+ * w = ?p \ p0\<^sup>T\<^sup>+ * w" proof - have "?q \ p0\<^sup>T\<^sup>+ * w = p0 \ p0\<^sup>T\<^sup>+ * w" - by (metis i12 comp_associative inf.cobounded2 p_antitone_iff star.circ_plus_same update_inf_different) + by (metis i11 comp_associative inf.cobounded2 p_antitone_iff star.circ_plus_same update_inf_different) also have "... = p \ p0\<^sup>T\<^sup>+ * w" - using 29 by (metis i13 inf.cobounded2 inf.sup_monoid.add_assoc p_antitone_iff update_inf_different) + using 28 by (metis i12 inf.cobounded2 inf.sup_monoid.add_assoc p_antitone_iff update_inf_different) also have "... = ?p \ p0\<^sup>T\<^sup>+ * w" - using 29 by (simp add: update_inf_different) + using 28 by (simp add: update_inf_different) finally show ?thesis . qed - have 32: "?q \ p0\<^sup>T\<^sup>\ * w = ?p \ p0\<^sup>T\<^sup>\ * w" - using 28 30 31 by (metis inf_sup_distrib1) - have 33: "?q \ (p0\<^sup>T\<^sup>\ * x \ -(p0\<^sup>T\<^sup>\ * w)) = ?p \ (p0\<^sup>T\<^sup>\ * x \ -(p0\<^sup>T\<^sup>\ * w))" + have 31: "?q \ p0\<^sup>T\<^sup>\ * w = ?p \ p0\<^sup>T\<^sup>\ * w" + using 27 29 30 by (metis inf_sup_distrib1) + have 32: "?q \ (p0\<^sup>T\<^sup>\ * x \ -(p0\<^sup>T\<^sup>\ * w)) = ?p \ (p0\<^sup>T\<^sup>\ * x \ -(p0\<^sup>T\<^sup>\ * w))" proof - have "p0\<^sup>T\<^sup>\ * x \ -(p0\<^sup>T\<^sup>\ * w) \ p0\<^sup>T\<^sup>\ * x \ -(p0\<^sup>T\<^sup>\ * (p[[w]]))" - using 29 by (metis i12 inf.sup_right_isotone mult.semigroup_axioms p_antitone_inf star_plus semigroup.assoc) + using 28 by (metis i11 inf.sup_right_isotone mult.semigroup_axioms p_antitone_inf star_plus semigroup.assoc) hence "?q \ (p0\<^sup>T\<^sup>\ * x \ -(p0\<^sup>T\<^sup>\ * w)) = y\<^sup>T \ p0\<^sup>T\<^sup>\ * x \ -(p0\<^sup>T\<^sup>\ * w)" by (metis inf_assoc update_inf) also have "... = p \ (p0\<^sup>T\<^sup>\ * x \ -(p0\<^sup>T\<^sup>\ * w))" - by (metis i13 inf_assoc update_inf_same) + by (metis i12 inf_assoc update_inf_same) also have "... = ?p \ (p0\<^sup>T\<^sup>\ * x \ -(p0\<^sup>T\<^sup>\ * w))" by (simp add: inf.coboundedI2 p_antitone path_compression_1b inf_assoc update_inf_different) finally show ?thesis . qed have "p0\<^sup>T\<^sup>\ * w \ (p0\<^sup>T\<^sup>\ * x \ -(p0\<^sup>T\<^sup>\ * w)) = p0\<^sup>T\<^sup>\ * x" proof - - have 34: "regular (p0\<^sup>T\<^sup>\ * w)" - using i14 i5 bijective_regular mapping_regular regular_closed_star regular_conv_closed regular_mult_closed by auto + have 33: "regular (p0\<^sup>T\<^sup>\ * w)" + using i13 i5 bijective_regular mapping_regular regular_closed_star regular_conv_closed regular_mult_closed by auto have "p0\<^sup>T\<^sup>\ * w \ p0\<^sup>T\<^sup>\ * x" - by (metis i16 comp_associative mult_right_isotone star.circ_transitive_equal) + by (metis i15 comp_associative mult_right_isotone star.circ_transitive_equal) hence "p0\<^sup>T\<^sup>\ * w \ (p0\<^sup>T\<^sup>\ * x \ -(p0\<^sup>T\<^sup>\ * w)) = p0\<^sup>T\<^sup>\ * x \ (p0\<^sup>T\<^sup>\ * w \ -(p0\<^sup>T\<^sup>\ * w))" by (simp add: comp_inf.semiring.distrib_left inf.absorb2) also have "... = p0\<^sup>T\<^sup>\ * x" - using 34 by (metis inf_sup_distrib1 maddux_3_11_pp) + using 33 by (metis inf_sup_distrib1 maddux_3_11_pp) finally show ?thesis . qed - hence 35: "?q \ p0\<^sup>T\<^sup>\ * x = ?p \ p0\<^sup>T\<^sup>\ * x" - using 32 33 by (metis inf_sup_distrib1) - have 36: "regular (p0\<^sup>T\<^sup>\ * x)" - using i14 i2 bijective_regular mapping_regular regular_closed_star regular_conv_closed regular_mult_closed by auto + hence 34: "?q \ p0\<^sup>T\<^sup>\ * x = ?p \ p0\<^sup>T\<^sup>\ * x" + using 31 32 by (metis inf_sup_distrib1) + have 35: "regular (p0\<^sup>T\<^sup>\ * x)" + using i13 i2 bijective_regular mapping_regular regular_closed_star regular_conv_closed regular_mult_closed by auto have "-(p0\<^sup>T\<^sup>\ * x) \ -w" - by (simp add: i16 p_antitone) + by (simp add: i15 p_antitone) hence "?q \ -(p0\<^sup>T\<^sup>\ * x) = ?p \ -(p0\<^sup>T\<^sup>\ * x)" - by (metis i13 p_antitone_inf update_inf_different) + by (metis i12 p_antitone_inf update_inf_different) thus ?thesis - using 35 36 by (metis maddux_3_11_pp) + using 34 35 by (metis maddux_3_11_pp) qed show "card ?t < card ?s" proof - have "?p\<^sup>T * p\<^sup>T\<^sup>\ * w = (w\<^sup>T \ y) * p\<^sup>T\<^sup>\ * w \ (-w\<^sup>T \ p\<^sup>T) * p\<^sup>T\<^sup>\ * w" by (simp add: conv_complement conv_dist_inf conv_dist_sup mult_right_dist_sup) also have "... \ (w\<^sup>T \ y) * p\<^sup>T\<^sup>\ * w \ p\<^sup>T * p\<^sup>T\<^sup>\ * w" using mult_left_isotone sup_right_isotone by auto also have "... \ (w\<^sup>T \ y) * p\<^sup>T\<^sup>\ * w \ p\<^sup>T\<^sup>\ * w" using mult_left_isotone star.left_plus_below_circ sup_right_isotone by blast also have "... \ y * p\<^sup>T\<^sup>\ * w \ p\<^sup>T\<^sup>\ * w" using semiring.add_right_mono mult_left_isotone by auto also have "... \ y * top \ p\<^sup>T\<^sup>\ * w" by (simp add: comp_associative le_supI1 mult_right_isotone) also have "... = p\<^sup>T\<^sup>\ * w" by (simp add: i3 i6 sup_absorb2) finally have "?p\<^sup>T\<^sup>\ * p\<^sup>T * w \ p\<^sup>T\<^sup>\ * w" using 11 by (metis dual_order.trans star.circ_loop_fixpoint sup_commute sup_ge2 mult_assoc) - hence 37: "?t \ ?s" + hence 36: "?t \ ?s" using order_lesseq_imp mult_assoc by auto - have 38: "w \ ?s" + have 37: "w \ ?s" by (simp add: i5 bijective_regular path_compression_1b) - have 39: "\ w \ ?t" + have 38: "\ w \ ?t" proof assume "w \ ?t" - hence 40: "w \ (?p\<^sup>T \ -1)\<^sup>\ * (p[[w]])" + hence 39: "w \ (?p\<^sup>T \ -1)\<^sup>\ * (p[[w]])" using reachable_without_loops by auto hence "p[[w]] \ (?p \ -1)\<^sup>\ * w" using 2 by (smt i5 bijective_reverse conv_star_commute reachable_without_loops) also have "... \ p\<^sup>\ * w" proof - have "p\<^sup>T\<^sup>\ * y = y" using i1 i4 root_transitive_successor_loop by auto hence "y\<^sup>T * p\<^sup>\ * w = y\<^sup>T * w" by (metis conv_dist_comp conv_involutive conv_star_commute) also have "... = bot" using 5 by (metis i5 inf.idem inf.sup_monoid.add_commute mult_left_zero schroeder_1 vector_inf_comp) - finally have 41: "y\<^sup>T * p\<^sup>\ * w = bot" + finally have 40: "y\<^sup>T * p\<^sup>\ * w = bot" by simp have "(?p \ -1) * p\<^sup>\ * w = (w \ y\<^sup>T \ -1) * p\<^sup>\ * w \ (-w \ p \ -1) * p\<^sup>\ * w" by (simp add: comp_inf.mult_right_dist_sup mult_right_dist_sup) also have "... \ (w \ y\<^sup>T \ -1) * p\<^sup>\ * w \ p * p\<^sup>\ * w" by (meson inf_le1 inf_le2 mult_left_isotone order_trans sup_right_isotone) also have "... \ (w \ y\<^sup>T \ -1) * p\<^sup>\ * w \ p\<^sup>\ * w" using mult_left_isotone star.left_plus_below_circ sup_right_isotone by blast also have "... \ y\<^sup>T * p\<^sup>\ * w \ p\<^sup>\ * w" by (meson inf_le1 inf_le2 mult_left_isotone order_trans sup_left_isotone) also have "... = p\<^sup>\ * w" - using 41 by simp + using 40 by simp finally show ?thesis by (metis comp_associative le_supI star.circ_loop_fixpoint sup_ge2 star_left_induct) qed finally have "w \ p\<^sup>T\<^sup>\ * p\<^sup>T * w" - using 11 40 reachable_without_loops star_plus by auto + using 11 39 reachable_without_loops star_plus by auto thus False using 4 i1 i10 i5 loop_root_2 star.circ_plus_same by auto qed show "card ?t < card ?s" apply (rule psubset_card_mono) subgoal using finite_regular by simp - subgoal using 37 38 39 by auto + subgoal using 36 37 38 by auto done qed qed qed lemma path_compression_3a: assumes "path_compression_invariant p x (p[[w]]) p0 w" shows "p0[p0\<^sup>T\<^sup>\ * x\p[[w]]] = p" proof - let ?y = "p[[w]]" let ?p = "p0[p0\<^sup>T\<^sup>\ * x\?y]" have i1: "disjoint_set_forest p" and i2: "point x" and i3: "point ?y" and i4: "?y = root p x" using assms path_compression_invariant_def path_compression_precondition_def by meson+ have i5: "point w" and i6: "?y \ p\<^sup>T\<^sup>\ * w" and i7: "w \ x \ p[[x]] = ?y \ ?y \ x \ p\<^sup>T\<^sup>+ * w \ -x" and i8: "p \ 1 = p0 \ 1" and i9: "fc p = fc p0" - and i10: "root p w = ?y" and i11: "w \ ?y \ p[[w]] \ w" and i12: "p[[w]] = p0[[w]]" and i13: "p0[p0\<^sup>T\<^sup>\ * x \ -(p0\<^sup>T\<^sup>\ * w)\?y] = p" - and i14: "disjoint_set_forest p0" and i15: "?y = root p0 x" and i16: "w \ p0\<^sup>T\<^sup>\ * x" + and i10: "root p w = ?y" and i11: "p[[w]] = p0[[w]]" and i12: "p0[p0\<^sup>T\<^sup>\ * x \ -(p0\<^sup>T\<^sup>\ * w)\?y] = p" + and i13: "disjoint_set_forest p0" and i14: "?y = root p0 x" and i15: "w \ p0\<^sup>T\<^sup>\ * x" using assms path_compression_invariant_def by blast+ have 1: "?p \ ?y = p \ ?y" - by (metis i1 i15 i3 i4 get_put inf_le1 root_successor_loop update_inf update_inf_same) + by (metis i1 i14 i3 i4 get_put inf_le1 root_successor_loop update_inf update_inf_same) have 2: "?p \ w = p \ w" - by (metis i5 i12 i16 get_put update_inf update_inf_same) + by (metis i5 i11 i15 get_put update_inf update_inf_same) have "?y = root p0 w" - by (metis i1 i10 i14 i8 i9 same_root) + by (metis i1 i10 i13 i8 i9 same_root) hence "p0\<^sup>T\<^sup>\ * w = w \ ?y" - by (metis i12 i14 root_transitive_successor_loop star.circ_loop_fixpoint star_plus sup_monoid.add_commute mult_assoc) + by (metis i11 i13 root_transitive_successor_loop star.circ_loop_fixpoint star_plus sup_monoid.add_commute mult_assoc) hence 3: "?p \ p0\<^sup>T\<^sup>\ * w = p \ p0\<^sup>T\<^sup>\ * w" using 1 2 by (simp add: inf_sup_distrib1) have "p0\<^sup>T\<^sup>\ * w \ p0\<^sup>T\<^sup>\ * x" - by (metis i16 comp_associative mult_right_isotone star.circ_transitive_equal) + by (metis i15 comp_associative mult_right_isotone star.circ_transitive_equal) hence 4: "?p \ (p0\<^sup>T\<^sup>\ * x \ p0\<^sup>T\<^sup>\ * w) = p \ (p0\<^sup>T\<^sup>\ * x \ p0\<^sup>T\<^sup>\ * w)" using 3 by (simp add: inf.absorb2) have 5: "?p \ (p0\<^sup>T\<^sup>\ * x \ -(p0\<^sup>T\<^sup>\ * w)) = p \ (p0\<^sup>T\<^sup>\ * x \ -(p0\<^sup>T\<^sup>\ * w))" - by (metis i13 inf_le1 update_inf update_inf_same) + by (metis i12 inf_le1 update_inf update_inf_same) have "regular (p0\<^sup>T\<^sup>\ * w)" - using i14 i5 bijective_regular mapping_regular regular_closed_star regular_conv_closed regular_mult_closed by auto + using i13 i5 bijective_regular mapping_regular regular_closed_star regular_conv_closed regular_mult_closed by auto hence 6: "?p \ p0\<^sup>T\<^sup>\ * x = p \ p0\<^sup>T\<^sup>\ * x" using 4 5 by (smt inf_sup_distrib1 maddux_3_11_pp) have 7: "?p \ -(p0\<^sup>T\<^sup>\ * x) = p \ -(p0\<^sup>T\<^sup>\ * x)" - by (smt i13 inf.sup_monoid.add_commute inf_import_p inf_sup_absorb le_iff_inf p_dist_inf update_inf_different inf.idem p_antitone_inf) + by (smt i12 inf.sup_monoid.add_commute inf_import_p inf_sup_absorb le_iff_inf p_dist_inf update_inf_different inf.idem p_antitone_inf) have "regular (p0\<^sup>T\<^sup>\ * x)" - using i14 i2 bijective_regular mapping_regular regular_closed_star regular_conv_closed regular_mult_closed by auto + using i13 i2 bijective_regular mapping_regular regular_closed_star regular_conv_closed regular_mult_closed by auto thus "?p = p" using 6 7 by (smt inf_sup_distrib1 maddux_3_11_pp) qed lemma path_compression_3: "path_compression_invariant p x (p[[w]]) p0 w \ path_compression_postcondition p x (p[[w]]) p0" using path_compression_invariant_def path_compression_postcondition_def path_compression_precondition_def path_compression_3a by blast theorem path_compression: "VARS p t w [ path_compression_precondition p x y \ p0 = p ] w := x; WHILE y \ p[[w]] INV { path_compression_invariant p x y p0 w } VAR { card { z . regular z \ z \ p\<^sup>T\<^sup>\ * w } } DO t := w; w := p[[w]]; p[t] := y OD [ path_compression_postcondition p x y p0 ]" apply vcg_tc_simp apply (fact path_compression_1) apply (fact path_compression_2) using path_compression_3 by auto lemma path_compression_exists: "path_compression_precondition p x y \ \p' . path_compression_postcondition p' x y p" using tc_extract_function path_compression by blast definition "path_compression p x y \ (SOME p' . path_compression_postcondition p' x y p)" lemma path_compression_function: assumes "path_compression_precondition p x y" and "p' = path_compression p x y" shows "path_compression_postcondition p' x y p" by (metis assms path_compression_def path_compression_exists someI) subsection \Find-Set with Path Compression\ text \ We sequentially combine find-set and path compression. We consider implementations which use the previously derived functions and implementations which unfold their definitions. \ theorem find_set_path_compression: "VARS p y [ find_set_precondition p x \ p0 = p ] y := find_set p x; p := path_compression p x y [ path_compression_postcondition p x y p0 ]" apply vcg_tc_simp using find_set_function find_set_postcondition_def find_set_precondition_def path_compression_function path_compression_precondition_def by fastforce theorem find_set_path_compression_1: "VARS p t w y [ find_set_precondition p x \ p0 = p ] y := find_set p x; w := x; WHILE y \ p[[w]] INV { path_compression_invariant p x y p0 w } VAR { card { z . regular z \ z \ p\<^sup>T\<^sup>\ * w } } DO t := w; w := p[[w]]; p[t] := y OD [ path_compression_postcondition p x y p0 ]" apply vcg_tc_simp using find_set_function find_set_postcondition_def find_set_precondition_def path_compression_1 path_compression_precondition_def apply fastforce apply (fact path_compression_2) by (fact path_compression_3) theorem find_set_path_compression_2: "VARS p y [ find_set_precondition p x \ p0 = p ] y := x; WHILE y \ p[[y]] INV { find_set_invariant p x y \ p0 = p } VAR { card { z . regular z \ z \ p\<^sup>T\<^sup>\ * y } } DO y := p[[y]] OD; p := path_compression p x y [ path_compression_postcondition p x y p0 ]" apply vcg_tc_simp apply (fact find_set_1) apply (fact find_set_2) by (smt find_set_3 find_set_invariant_def find_set_postcondition_def find_set_precondition_def path_compression_function path_compression_precondition_def) theorem find_set_path_compression_3: "VARS p t w y [ find_set_precondition p x \ p0 = p ] y := x; WHILE y \ p[[y]] INV { find_set_invariant p x y \ p0 = p } VAR { card { z . regular z \ z \ p\<^sup>T\<^sup>\ * y } } DO y := p[[y]] OD; w := x; WHILE y \ p[[w]] INV { path_compression_invariant p x y p0 w } VAR { card { z . regular z \ z \ p\<^sup>T\<^sup>\ * w } } DO t := w; w := p[[w]]; p[t] := y OD [ path_compression_postcondition p x y p0 ]" apply vcg_tc_simp apply (simp add: find_set_1) apply (fact find_set_2) using find_set_3 find_set_invariant_def find_set_postcondition_def find_set_precondition_def path_compression_1 path_compression_precondition_def apply blast apply (fact path_compression_2) by (fact path_compression_3) text \ Find-set with path compression returns two results: the representative of the tree and the modified disjoint-set forest. \ lemma find_set_path_compression_exists: "find_set_precondition p x \ \p' y . path_compression_postcondition p' x y p" using tc_extract_function find_set_path_compression by blast definition "find_set_path_compression p x \ (SOME (p',y) . path_compression_postcondition p' x y p)" lemma find_set_path_compression_function: assumes "find_set_precondition p x" and "(p',y) = find_set_path_compression p x" shows "path_compression_postcondition p' x y p" proof - let ?P = "\(p',y) . path_compression_postcondition p' x y p" have "?P (SOME z . ?P z)" apply (unfold some_eq_ex) using assms(1) find_set_path_compression_exists by simp thus ?thesis using assms(2) find_set_path_compression_def by auto qed text \ We prove that \find_set_path_compression\ returns the same representative as \find_set\. \ lemma find_set_path_compression_find_set: assumes "find_set_precondition p x" shows "find_set p x = snd (find_set_path_compression p x)" proof - let ?r = "find_set p x" let ?p = "fst (find_set_path_compression p x)" let ?y = "snd (find_set_path_compression p x)" have 1: "find_set_postcondition p x ?r" by (simp add: assms find_set_function) have "path_compression_postcondition ?p x ?y p" using assms find_set_path_compression_function prod.collapse by blast thus "?r = ?y" using 1 by (smt assms same_root find_set_precondition_def find_set_postcondition_def path_compression_precondition_def path_compression_postcondition_def) qed text \ A weaker postcondition suffices to prove that the two forests have the same semantics; that is, they describe the same disjoint sets and have the same roots. \ lemma find_set_path_compression_path_compression_semantics: assumes "find_set_precondition p x" shows "fc (path_compression p x (find_set p x)) = fc (fst (find_set_path_compression p x))" and "path_compression p x (find_set p x) \ 1 = fst (find_set_path_compression p x) \ 1" proof - let ?r = "find_set p x" let ?q = "path_compression p x ?r" let ?p = "fst (find_set_path_compression p x)" let ?y = "snd (find_set_path_compression p x)" have 1: "path_compression_postcondition (path_compression p x ?r) x ?r p" using assms find_set_function find_set_postcondition_def find_set_precondition_def path_compression_function path_compression_precondition_def by auto have 2: "path_compression_postcondition ?p x ?y p" using assms find_set_path_compression_function prod.collapse by blast show "fc ?q = fc ?p" using 1 2 by (simp add: path_compression_postcondition_def) show "?q \ 1 = ?p \ 1" using 1 2 by (simp add: path_compression_postcondition_def) qed text \ With the current, stronger postcondition of path compression describing the precise effect of how links change, we can prove that the two forests are actually equal. \ lemma find_set_path_compression_find_set_pathcompression: assumes "find_set_precondition p x" shows "path_compression p x (find_set p x) = fst (find_set_path_compression p x)" proof - let ?r = "find_set p x" let ?q = "path_compression p x ?r" let ?p = "fst (find_set_path_compression p x)" let ?y = "snd (find_set_path_compression p x)" have 1: "path_compression_postcondition (path_compression p x ?r) x ?r p" using assms find_set_function find_set_postcondition_def find_set_precondition_def path_compression_function path_compression_precondition_def by auto have 2: "path_compression_postcondition ?p x ?y p" using assms find_set_path_compression_function prod.collapse by blast have "?r = ?y" by (simp add: assms find_set_path_compression_find_set) thus "?q = ?p" using 1 2 by (simp add: path_compression_postcondition_def) qed subsection \Union-Sets\ text \ We only consider a naive union-sets operation (without ranks). The semantics is the equivalence closure obtained after adding the link between the two given nodes, which requires those two elements to be in the same set. The implementation uses temporary variable \t\ to store the two results returned by find-set with path compression. The disjoint-set forest, which keeps being updated, is threaded through the sequence of operations. \ definition "union_sets_precondition p x y \ disjoint_set_forest p \ point x \ point y" definition "union_sets_postcondition p x y p0 \ union_sets_precondition p x y \ fc p = wcc (p0 \ x * y\<^sup>T)" lemma union_sets_1: assumes "union_sets_precondition p0 x y" and "path_compression_postcondition p1 x r p0" and "path_compression_postcondition p2 y s p1" shows "union_sets_postcondition (p2[r\s]) x y p0" proof (unfold union_sets_postcondition_def union_sets_precondition_def, intro conjI) let ?p = "p2[r\s]" have 1: "disjoint_set_forest p1 \ point r \ r = root p1 x \ p1 \ 1 = p0 \ 1 \ fc p1 = fc p0" using assms(2) path_compression_precondition_def path_compression_postcondition_def by auto have 2: "disjoint_set_forest p2 \ point s \ s = root p2 y \ p2 \ 1 = p1 \ 1 \ fc p2 = fc p1" using assms(3) path_compression_precondition_def path_compression_postcondition_def by auto hence 3: "fc p2 = fc p0" using 1 by simp show 4: "univalent ?p" using 1 2 update_univalent by blast show "total ?p" using 1 2 bijective_regular update_total by blast show "acyclic (?p \ -1)" proof (cases "r = s") case True thus ?thesis using 2 update_acyclic_5 by fastforce next case False hence "bot = r \ s" using 1 2 distinct_points by blast also have "... = r \ p2\<^sup>T\<^sup>\ * s" using 2 by (smt root_transitive_successor_loop) finally have "s \ p2\<^sup>\ * r = bot" using schroeder_1 conv_star_commute inf.sup_monoid.add_commute by fastforce thus ?thesis using 1 2 update_acyclic_4 by blast qed show "vector x" using assms(1) by (simp add: union_sets_precondition_def) show "injective x" using assms(1) by (simp add: union_sets_precondition_def) show "surjective x" using assms(1) by (simp add: union_sets_precondition_def) show "vector y" using assms(1) by (simp add: union_sets_precondition_def) show "injective y" using assms(1) by (simp add: union_sets_precondition_def) show "surjective y" using assms(1) by (simp add: union_sets_precondition_def) show "fc ?p = wcc (p0 \ x * y\<^sup>T)" proof (rule order.antisym) have "r = p1[[r]]" using 1 by (metis root_successor_loop) hence "r * r\<^sup>T \ p1\<^sup>T" using 1 eq_refl shunt_bijective by blast hence "r * r\<^sup>T \ p1" using 1 conv_order coreflexive_symmetric by fastforce hence "r * r\<^sup>T \ p1 \ 1" using 1 inf.boundedI by blast also have "... = p2 \ 1" using 2 by simp finally have "r * r\<^sup>T \ p2" by simp hence "r \ p2 * r" using 1 shunt_bijective by blast hence 5: "p2[[r]] \ r" using 2 shunt_mapping by blast have "r \ p2 \ r * (top \ r\<^sup>T * p2)" using 1 by (metis dedekind_1) also have "... = r * r\<^sup>T * p2" by (simp add: mult_assoc) also have "... \ r * r\<^sup>T" using 5 by (metis comp_associative conv_dist_comp conv_involutive conv_order mult_right_isotone) also have "... \ 1" using 1 by blast finally have 6: "r \ p2 \ 1" by simp have "p0 \ wcc p0" by (simp add: star.circ_sub_dist_1) also have "... = wcc p2" using 3 by (simp add: star_decompose_1) also have 7: "... \ wcc ?p" proof - have "wcc p2 = wcc ((-r \ p2) \ (r \ p2))" using 1 by (metis bijective_regular inf.sup_monoid.add_commute maddux_3_11_pp) also have "... \ wcc ((-r \ p2) \ 1)" using 6 wcc_isotone sup_right_isotone by simp also have "... = wcc (-r \ p2)" using wcc_with_loops by simp also have "... \ wcc ?p" using wcc_isotone sup_ge2 by blast finally show ?thesis by simp qed finally have 8: "p0 \ wcc ?p" by force have "r \ p1\<^sup>T\<^sup>\ * x" using 1 by (metis inf_le1) hence 9: "r * x\<^sup>T \ p1\<^sup>T\<^sup>\" using assms(1) shunt_bijective union_sets_precondition_def by blast hence "x * r\<^sup>T \ p1\<^sup>\" using conv_dist_comp conv_order conv_star_commute by force also have "... \ wcc p1" by (simp add: star.circ_sub_dist) also have "... = wcc p2" using 1 2 by (simp add: fc_wcc) also have "... \ wcc ?p" using 7 by simp finally have 10: "x * r\<^sup>T \ wcc ?p" by simp have 11: "r * s\<^sup>T \ wcc ?p" using 1 2 star.circ_sub_dist_1 sup_assoc vector_covector by auto have "s \ p2\<^sup>T\<^sup>\ * y" using 2 by (metis inf_le1) hence 12: "s * y\<^sup>T \ p2\<^sup>T\<^sup>\" using assms(1) shunt_bijective union_sets_precondition_def by blast also have "... \ wcc p2" using star_isotone sup_ge2 by blast also have "... \ wcc ?p" using 7 by simp finally have 13: "s * y\<^sup>T \ wcc ?p" by simp have "x \ x * r\<^sup>T * r \ y \ y * s\<^sup>T * s" using 1 2 shunt_bijective by blast hence "x * y\<^sup>T \ x * r\<^sup>T * r * (y * s\<^sup>T * s)\<^sup>T" using comp_isotone conv_isotone by blast also have "... = x * r\<^sup>T * r * s\<^sup>T * s * y\<^sup>T" by (simp add: comp_associative conv_dist_comp) also have "... \ wcc ?p * (r * s\<^sup>T) * (s * y\<^sup>T)" using 10 by (metis mult_left_isotone mult_assoc) also have "... \ wcc ?p * wcc ?p * (s * y\<^sup>T)" using 11 by (metis mult_left_isotone mult_right_isotone) also have "... \ wcc ?p * wcc ?p * wcc ?p" using 13 by (metis mult_right_isotone) also have "... = wcc ?p" by (simp add: star.circ_transitive_equal) finally have "p0 \ x * y\<^sup>T \ wcc ?p" using 8 by simp hence "wcc (p0 \ x * y\<^sup>T) \ wcc ?p" using wcc_below_wcc by simp thus "wcc (p0 \ x * y\<^sup>T) \ fc ?p" using 4 fc_wcc by simp have "-r \ p2 \ wcc p2" by (simp add: inf.coboundedI2 star.circ_sub_dist_1) also have "... = wcc p0" using 3 by (simp add: star_decompose_1) also have "... \ wcc (p0 \ x * y\<^sup>T)" by (simp add: wcc_isotone) finally have 14: "-r \ p2 \ wcc (p0 \ x * y\<^sup>T)" by simp have "r * x\<^sup>T \ wcc p1" using 9 inf.order_trans star.circ_sub_dist sup_commute by fastforce also have "... = wcc p0" using 1 by (simp add: star_decompose_1) also have "... \ wcc (p0 \ x * y\<^sup>T)" by (simp add: wcc_isotone) finally have 15: "r * x\<^sup>T \ wcc (p0 \ x * y\<^sup>T)" by simp have 16: "x * y\<^sup>T \ wcc (p0 \ x * y\<^sup>T)" using le_supE star.circ_sub_dist_1 by blast have "y * s\<^sup>T \ p2\<^sup>\" using 12 conv_dist_comp conv_order conv_star_commute by fastforce also have "... \ wcc p2" using star.circ_sub_dist sup_commute by fastforce also have "... = wcc p0" using 3 by (simp add: star_decompose_1) also have "... \ wcc (p0 \ x * y\<^sup>T)" by (simp add: wcc_isotone) finally have 17: "y * s\<^sup>T \ wcc (p0 \ x * y\<^sup>T)" by simp have "r \ r * x\<^sup>T * x \ s \ s * y\<^sup>T * y" using assms(1) shunt_bijective union_sets_precondition_def by blast hence "r * s\<^sup>T \ r * x\<^sup>T * x * (s * y\<^sup>T * y)\<^sup>T" using comp_isotone conv_isotone by blast also have "... = r * x\<^sup>T * x * y\<^sup>T * y * s\<^sup>T" by (simp add: comp_associative conv_dist_comp) also have "... \ wcc (p0 \ x * y\<^sup>T) * (x * y\<^sup>T) * (y * s\<^sup>T)" using 15 by (metis mult_left_isotone mult_assoc) also have "... \ wcc (p0 \ x * y\<^sup>T) * wcc (p0 \ x * y\<^sup>T) * (y * s\<^sup>T)" using 16 by (metis mult_left_isotone mult_right_isotone) also have "... \ wcc (p0 \ x * y\<^sup>T) * wcc (p0 \ x * y\<^sup>T) * wcc (p0 \ x * y\<^sup>T)" using 17 by (metis mult_right_isotone) also have "... = wcc (p0 \ x * y\<^sup>T)" by (simp add: star.circ_transitive_equal) finally have "?p \ wcc (p0 \ x * y\<^sup>T)" using 1 2 14 vector_covector by auto hence "wcc ?p \ wcc (p0 \ x * y\<^sup>T)" using wcc_below_wcc by blast thus "fc ?p \ wcc (p0 \ x * y\<^sup>T)" using 4 fc_wcc by simp qed qed theorem union_sets: "VARS p r s t [ union_sets_precondition p x y \ p0 = p ] t := find_set_path_compression p x; p := fst t; r := snd t; t := find_set_path_compression p y; p := fst t; s := snd t; p[r] := s [ union_sets_postcondition p x y p0 ]" proof vcg_tc_simp let ?t1 = "find_set_path_compression p0 x" let ?p1 = "fst ?t1" let ?r = "snd ?t1" let ?t2 = "find_set_path_compression ?p1 y" let ?p2 = "fst ?t2" let ?s = "snd ?t2" let ?p = "?p2[?r\?s]" assume 1: "union_sets_precondition p0 x y" hence 2: "path_compression_postcondition ?p1 x ?r p0" by (simp add: find_set_precondition_def union_sets_precondition_def find_set_path_compression_function) hence "path_compression_postcondition ?p2 y ?s ?p1" using 1 by (meson find_set_precondition_def union_sets_precondition_def find_set_path_compression_function path_compression_postcondition_def path_compression_precondition_def prod.collapse) thus "union_sets_postcondition (?p2[?r\?s]) x y p0" using 1 2 by (simp add: union_sets_1) qed lemma union_sets_exists: "union_sets_precondition p x y \ \p' . union_sets_postcondition p' x y p" using tc_extract_function union_sets by blast definition "union_sets p x y \ (SOME p' . union_sets_postcondition p' x y p)" lemma union_sets_function: assumes "union_sets_precondition p x y" and "p' = union_sets p x y" shows "union_sets_postcondition p' x y p" by (metis assms union_sets_def union_sets_exists someI) theorem union_sets_2: - "VARS p r s t + "VARS p r s [ union_sets_precondition p x y \ p0 = p ] r := find_set p x; p := path_compression p x r; s := find_set p y; p := path_compression p y s; p[r] := s [ union_sets_postcondition p x y p0 ]" proof vcg_tc_simp let ?r = "find_set p0 x" let ?p1 = "path_compression p0 x ?r" let ?s = "find_set ?p1 y" let ?p2 = "path_compression ?p1 y ?s" assume 1: "union_sets_precondition p0 x y" hence 2: "path_compression_postcondition ?p1 x ?r p0" using find_set_function find_set_postcondition_def find_set_precondition_def path_compression_function path_compression_precondition_def union_sets_precondition_def by auto hence "path_compression_postcondition ?p2 y ?s ?p1" using 1 find_set_function find_set_postcondition_def find_set_precondition_def path_compression_function path_compression_precondition_def union_sets_precondition_def path_compression_postcondition_def by meson thus "union_sets_postcondition (?p2[?r\?s]) x y p0" using 1 2 by (simp add: union_sets_1) qed end end diff --git a/thys/Relational_Disjoint_Set_Forests/More_Disjoint_Set_Forests.thy b/thys/Relational_Disjoint_Set_Forests/More_Disjoint_Set_Forests.thy new file mode 100644 --- /dev/null +++ b/thys/Relational_Disjoint_Set_Forests/More_Disjoint_Set_Forests.thy @@ -0,0 +1,3175 @@ +(* Title: More on Disjoint-Set Forests + Author: Walter Guttmann + Maintainer: Walter Guttmann +*) + +theory More_Disjoint_Set_Forests + +imports Disjoint_Set_Forests + +begin + +section \More on Array Access and Disjoint-Set Forests\ + +text \ +This section contains further results about directed acyclic graphs and relational array operations. +\ + +context stone_relation_algebra +begin + +text \Theorem 6.4\ + +lemma update_square: + assumes "point y" + shows "x[y\x[[x[[y]]]]] \ x * x \ x" +proof - + have "x[y\x[[x[[y]]]]] = (y \ y\<^sup>T * x * x) \ (-y \ x)" + by (simp add: conv_dist_comp) + also have "... \ (y \ y\<^sup>T) * x * x \ x" + by (smt assms inf.eq_refl inf.sup_monoid.add_commute inf_le1 sup_mono vector_inf_comp) + also have "... \ x * x \ x" + by (smt (z3) assms comp_associative conv_dist_comp coreflexive_comp_top_inf inf.cobounded2 sup_left_isotone symmetric_top_closed) + finally show ?thesis + . +qed + +text \Theorem 2.13\ + +lemma update_ub: + "x[y\z] \ x \ z\<^sup>T" + by (meson dual_order.trans inf.cobounded2 le_supI sup.cobounded1 sup_ge2) + +text \Theorem 6.7\ + +lemma update_square_ub: + "x[y\(x * x)\<^sup>T] \ x \ x * x" + by (metis conv_involutive update_ub) + +text \Theorem 2.14\ + +lemma update_same_sub: + assumes "u \ x = u \ z" + and "y \ u" + and "regular y" + shows "x[y\z\<^sup>T] = x" + by (smt (z3) assms conv_involutive inf.sup_monoid.add_commute inf.sup_relative_same_increasing maddux_3_11_pp) + +text \Theorem 2.15\ + +lemma update_point_get: + "point y \ x[y\z[[y]]] = x[y\z\<^sup>T]" + by (metis conv_involutive get_put inf_commute update_inf_same) + +text \Theorem 2.11\ + +lemma update_bot: + "x[bot\z] = x" + by simp + +text \Theorem 2.12\ + +lemma update_top: + "x[top\z] = z\<^sup>T" + by simp + +text \Theorem 2.6\ + +lemma update_same: + assumes "regular u" + shows "(x[y\z])[u\z] = x[y \ u\z]" +proof - + have "(x[y\z])[u\z] = (u \ z\<^sup>T) \ (-u \ y \ z\<^sup>T) \ (-u \ -y \ x)" + using inf.sup_monoid.add_assoc inf_sup_distrib1 sup_assoc by force + also have "... = (u \ z\<^sup>T) \ (y \ z\<^sup>T) \ (-(u \ y) \ x)" + by (metis assms inf_sup_distrib2 maddux_3_21_pp p_dist_sup) + also have "... = x[y \ u\z]" + using comp_inf.mult_right_dist_sup sup_commute by auto + finally show ?thesis + . +qed + +lemma update_same_3: + assumes "regular u" + and "regular v" + shows "((x[y\z])[u\z])[v\z] = x[y \ u \ v\z]" + by (metis assms update_same) + +text \Theorem 2.7\ + +lemma update_split: + assumes "regular w" + shows "x[y\z] = (x[y \ -w\z])[y \ w\z]" + by (smt (z3) assms comp_inf.semiring.distrib_left inf.left_commute inf.sup_monoid.add_commute inf_import_p maddux_3_11_pp maddux_3_12 p_dist_inf sup_assoc) + +text \Theorem 2.8\ + +lemma update_injective_swap: + assumes "injective x" + and "point y" + and "injective z" + and "vector z" + shows "injective ((x[y\x[[z]]])[z\x[[y]]])" +proof - + have 1: "(z \ y\<^sup>T * x) * (z \ y\<^sup>T * x)\<^sup>T \ 1" + using assms(3) injective_inf_closed by auto + have "(z \ y\<^sup>T * x) * (-z \ y \ z\<^sup>T * x)\<^sup>T \ (z \ y\<^sup>T * x) * (y\<^sup>T \ x\<^sup>T * z)" + by (metis conv_dist_comp conv_involutive conv_order inf.boundedE inf.boundedI inf.cobounded1 inf.cobounded2 mult_right_isotone) + also have "... = (z \ z\<^sup>T * x) * (y\<^sup>T \ x\<^sup>T * y)" + by (smt (z3) assms(2,4) covector_inf_comp_3 inf.left_commute inf.sup_monoid.add_commute comp_associative conv_dist_comp conv_involutive) + also have "... = (z \ z\<^sup>T) * x * x\<^sup>T * (y \ y\<^sup>T)" + by (smt (z3) assms(2,4) comp_associative inf.sup_monoid.add_commute vector_covector vector_inf_comp) + also have "... \ x * x\<^sup>T" + by (metis assms(2-4) comp_associative comp_right_one coreflexive_comp_top_inf inf.coboundedI2 mult_right_isotone vector_covector) + also have "... \ 1" + by (simp add: assms(1)) + finally have 2: "(z \ y\<^sup>T * x) * (-z \ y \ z\<^sup>T * x)\<^sup>T \ 1" + . + have "(z \ y\<^sup>T * x) * (-z \ -y \ x)\<^sup>T \ y\<^sup>T * x * (-y\<^sup>T \ x\<^sup>T)" + by (smt comp_isotone conv_complement conv_dist_inf inf.cobounded2 inf.sup_monoid.add_assoc) + also have "... = y\<^sup>T * x * x\<^sup>T \ -y\<^sup>T" + by (simp add: inf.commute assms(2) covector_comp_inf vector_conv_compl) + also have "... \ y\<^sup>T \ -y\<^sup>T" + by (metis assms(1) comp_associative comp_inf.mult_left_isotone comp_isotone comp_right_one mult_sub_right_one) + finally have 3: "(z \ y\<^sup>T * x) * (-z \ -y \ x)\<^sup>T \ 1" + using pseudo_complement by fastforce + have 4: "(-z \ y \ z\<^sup>T * x) * (z \ y\<^sup>T * x)\<^sup>T \ 1" + using 2 conv_dist_comp conv_order by force + have 5: "(-z \ y \ z\<^sup>T * x) * (-z \ y \ z\<^sup>T * x)\<^sup>T \ 1" + by (simp add: assms(2) inf_assoc inf_left_commute injective_inf_closed) + have "(-z \ y \ z\<^sup>T * x) * (-z \ -y \ x)\<^sup>T \ z\<^sup>T * x * (-z\<^sup>T \ x\<^sup>T)" + using comp_inf.mult_left_isotone comp_isotone conv_complement conv_dist_inf inf.cobounded1 inf.cobounded2 by auto + also have "... = z\<^sup>T * x * x\<^sup>T \ -z\<^sup>T" + by (metis assms(4) covector_comp_inf inf.sup_monoid.add_commute vector_conv_compl) + also have "... \ z\<^sup>T \ -z\<^sup>T" + by (metis assms(1) comp_associative comp_inf.mult_left_isotone comp_isotone comp_right_one mult_sub_right_one) + finally have 6: "(-z \ y \ z\<^sup>T * x) * (-z \ -y \ x)\<^sup>T \ 1" + using pseudo_complement by fastforce + have 7: "(-z \ -y \ x) * (z \ y\<^sup>T * x)\<^sup>T \ 1" + using 3 conv_dist_comp coreflexive_symmetric by fastforce + have 8: "(-z \ -y \ x) * (-z \ y \ z\<^sup>T * x)\<^sup>T \ 1" + using 6 conv_dist_comp coreflexive_symmetric by fastforce + have 9: "(-z \ -y \ x) * (-z \ -y \ x)\<^sup>T \ 1" + using assms(1) inf.sup_monoid.add_commute injective_inf_closed by auto + have "(x[y\x[[z]]])[z\x[[y]]] = (z \ y\<^sup>T * x) \ (-z \ y \ z\<^sup>T * x) \ (-z \ -y \ x)" + by (simp add: comp_inf.comp_left_dist_sup conv_dist_comp inf_assoc sup_monoid.add_assoc) + hence "((x[y\x[[z]]])[z\x[[y]]]) * ((x[y\x[[z]]])[z\x[[y]]])\<^sup>T = ((z \ y\<^sup>T * x) \ (-z \ y \ z\<^sup>T * x) \ (-z \ -y \ x)) * ((z \ y\<^sup>T * x)\<^sup>T \ (-z \ y \ z\<^sup>T * x)\<^sup>T \ (-z \ -y \ x)\<^sup>T)" + by (simp add: conv_dist_sup) + also have "... = (z \ y\<^sup>T * x) * ((z \ y\<^sup>T * x)\<^sup>T \ (-z \ y \ z\<^sup>T * x)\<^sup>T \ (-z \ -y \ x)\<^sup>T) \ + (-z \ y \ z\<^sup>T * x) * ((z \ y\<^sup>T * x)\<^sup>T \ (-z \ y \ z\<^sup>T * x)\<^sup>T \ (-z \ -y \ x)\<^sup>T) \ + (-z \ -y \ x) * ((z \ y\<^sup>T * x)\<^sup>T \ (-z \ y \ z\<^sup>T * x)\<^sup>T \ (-z \ -y \ x)\<^sup>T)" + using mult_right_dist_sup by auto + also have "... = (z \ y\<^sup>T * x) * (z \ y\<^sup>T * x)\<^sup>T \ (z \ y\<^sup>T * x) * (-z \ y \ z\<^sup>T * x)\<^sup>T \ (z \ y\<^sup>T * x) * (-z \ -y \ x)\<^sup>T \ + (-z \ y \ z\<^sup>T * x) * (z \ y\<^sup>T * x)\<^sup>T \ (-z \ y \ z\<^sup>T * x) * (-z \ y \ z\<^sup>T * x)\<^sup>T \ (-z \ y \ z\<^sup>T * x) * (-z \ -y \ x)\<^sup>T \ + (-z \ -y \ x) * (z \ y\<^sup>T * x)\<^sup>T \ (-z \ -y \ x) * (-z \ y \ z\<^sup>T * x)\<^sup>T \ (-z \ -y \ x) * (-z \ -y \ x)\<^sup>T" + using mult_left_dist_sup sup.left_commute sup_commute by auto + also have "... \ 1" + using 1 2 3 4 5 6 7 8 9 by simp_all + finally show ?thesis + . +qed + +lemma update_injective_swap_2: + assumes "injective x" + shows "injective ((x[y\x[[bot]]])[bot\x[[y]]])" + by (simp add: assms inf.sup_monoid.add_commute injective_inf_closed) + +text \Theorem 2.9\ + +lemma update_univalent_swap: + assumes "univalent x" + and "injective y" + and "vector y" + and "injective z" + and "vector z" + shows "univalent ((x[y\x[[z]]])[z\x[[y]]])" + by (simp add: assms read_injective update_univalent) + +text \Theorem 2.10\ + +lemma update_mapping_swap: + assumes "mapping x" + and "point y" + and "point z" + shows "mapping ((x[y\x[[z]]])[z\x[[y]]])" + by (simp add: assms bijective_regular read_injective read_surjective update_total update_univalent) + +text \Theorem 2.16 \mapping_inf_point_arc\ has been moved to theory \Relation_Algebras\ in entry \Stone_Relation_Algebras\\ + +end + +context stone_kleene_relation_algebra +begin + +lemma omit_redundant_points_2: + assumes "point p" + shows "p \ x\<^sup>\ = (p \ 1) \ (p \ x \ -p\<^sup>T) * (x \ -p\<^sup>T)\<^sup>\" +proof - + let ?p = "p \ 1" + let ?np = "-p \ 1" + have 1: "p \ x\<^sup>\ \ 1 = p \ 1" + by (metis inf.le_iff_sup inf.left_commute inf.sup_monoid.add_commute star.circ_reflexive) + have 2: "p \ 1 \ -p\<^sup>T = bot" + by (smt (z3) inf_bot_right inf_commute inf_left_commute one_inf_conv p_inf) + have "p \ x\<^sup>\ \ -1 = p \ x\<^sup>\ \ -p\<^sup>T" + by (metis assms antisymmetric_inf_diversity inf.cobounded1 inf.sup_relative_same_increasing vector_covector) + also have "... = (p \ 1 \ -p\<^sup>T) \ ((p \ x) * (-p \ x)\<^sup>\ \ -p\<^sup>T)" + by (simp add: assms omit_redundant_points comp_inf.semiring.distrib_right) + also have "... = (p \ x) * (-p \ x)\<^sup>\ \ -p\<^sup>T" + using 2 by simp + also have "... = ?p * x * (-p \ x)\<^sup>\ \ -p\<^sup>T" + by (metis assms vector_export_comp_unit) + also have "... = ?p * x * (?np * x)\<^sup>\ \ -p\<^sup>T" + by (metis assms vector_complement_closed vector_export_comp_unit) + also have "... = ?p * x * (?np * x)\<^sup>\ * ?np" + by (metis assms conv_complement covector_comp_inf inf.sup_monoid.add_commute mult_1_right one_inf_conv vector_conv_compl) + also have "... = ?p * x * ?np * (x * ?np)\<^sup>\" + using star_slide mult_assoc by auto + also have "... = (?p * x \ -p\<^sup>T) * (x * ?np)\<^sup>\" + by (metis assms conv_complement covector_comp_inf inf.sup_monoid.add_commute mult_1_right one_inf_conv vector_conv_compl) + also have "... = (?p * x \ -p\<^sup>T) * (x \ -p\<^sup>T)\<^sup>\" + by (metis assms conv_complement covector_comp_inf inf.sup_monoid.add_commute mult_1_right one_inf_conv vector_conv_compl) + also have "... = (p \ x \ -p\<^sup>T) * (x \ -p\<^sup>T)\<^sup>\" + by (metis assms vector_export_comp_unit) + finally show ?thesis + using 1 by (metis maddux_3_11_pp regular_one_closed) +qed + +text \Theorem 5.3\ + +lemma omit_redundant_points_3: + assumes "point p" + shows "p \ x\<^sup>\ = (p \ 1) \ (p \ (x \ -p\<^sup>T)\<^sup>+)" + by (simp add: assms inf_assoc vector_inf_comp omit_redundant_points_2) + +text \Theorem 6.1\ + +lemma even_odd_root: + assumes "acyclic (x \ -1)" + and "regular x" + and "univalent x" + shows "(x * x)\<^sup>T\<^sup>\ \ x\<^sup>T * (x * x)\<^sup>T\<^sup>\ = (1 \ x) * ((x * x)\<^sup>T\<^sup>\ \ x\<^sup>T * (x * x)\<^sup>T\<^sup>\)" +proof - + have 1: "univalent (x * x)" + by (simp add: assms(3) univalent_mult_closed) + have "x \ 1 \ top * (x \ 1)" + by (simp add: top_left_mult_increasing) + hence "x \ -(top * (x \ 1)) \ x \ -1" + using assms(2) p_shunting_swap pp_dist_comp by auto + hence "x\<^sup>\ * (x \ -(top * (x \ 1))) \ (x \ -1)\<^sup>\ * (x \ -1)" + using mult_right_isotone reachable_without_loops by auto + also have "... \ -1" + by (simp add: assms(1) star_plus) + finally have "(x \ -(top * (x \ 1)))\<^sup>T \ -x\<^sup>\" + using schroeder_4_p by force + hence "x\<^sup>T \ x\<^sup>\ \ (top * (x \ 1))\<^sup>T" + by (smt (z3) assms(2) conv_complement conv_dist_inf p_shunting_swap regular_closed_inf regular_closed_top regular_mult_closed regular_one_closed) + also have "... = (1 \ x) * top" + by (metis conv_dist_comp conv_dist_inf inf_commute one_inf_conv symmetric_one_closed symmetric_top_closed) + finally have 2: "(x\<^sup>T \ x\<^sup>\) * top \ (1 \ x) * top" + by (metis inf.orderE inf.orderI inf_commute inf_vector_comp) + have "1 \ x\<^sup>T\<^sup>+ \ (x\<^sup>T \ 1 * x\<^sup>\) * x\<^sup>T\<^sup>\" + by (metis conv_involutive conv_star_commute dedekind_2 inf_commute) + also have "... \ (x\<^sup>T \ x\<^sup>\) * top" + by (simp add: mult_right_isotone) + also have "... \ (1 \ x) * top" + using 2 by simp + finally have 3: "1 \ x\<^sup>T\<^sup>+ \ (1 \ x) * top" + . + have "x\<^sup>T \ (x\<^sup>T * x\<^sup>T)\<^sup>+ = 1 * x\<^sup>T \ (x\<^sup>T * x\<^sup>T)\<^sup>\ * x\<^sup>T * x\<^sup>T" + using star_plus mult_assoc by auto + also have "... = (1 \ (x\<^sup>T * x\<^sup>T)\<^sup>\ * x\<^sup>T) * x\<^sup>T" + using assms(3) injective_comp_right_dist_inf by force + also have "... \ (1 \ x\<^sup>T\<^sup>\ * x\<^sup>T) * x\<^sup>T" + by (meson comp_inf.mult_right_isotone comp_isotone inf.eq_refl star.circ_square) + also have "... \ (1 \ x) * top * x\<^sup>T" + using 3 by (simp add: mult_left_isotone star_plus) + also have "... \ (1 \ x) * top" + by (simp add: comp_associative mult_right_isotone) + finally have 4: "x\<^sup>T \ (x\<^sup>T * x\<^sup>T)\<^sup>+ \ (1 \ x) * top" + . + have "x\<^sup>T \ (x\<^sup>T * x\<^sup>T)\<^sup>\ = (x\<^sup>T \ 1) \ (x\<^sup>T \ (x\<^sup>T * x\<^sup>T)\<^sup>+)" + by (metis inf_sup_distrib1 star_left_unfold_equal) + also have "... \ (1 \ x) * top" + using 4 by (metis inf.sup_monoid.add_commute le_supI one_inf_conv top_right_mult_increasing) + finally have 4: "x\<^sup>T \ (x\<^sup>T * x\<^sup>T)\<^sup>\ \ (1 \ x) * top" + . + have "x\<^sup>T \ (x * x)\<^sup>\ \ -1 \ x\<^sup>T \ x\<^sup>\ \ -1" + by (simp add: inf.coboundedI2 inf.sup_monoid.add_commute star.circ_square) + also have "... = (x \ -1)\<^sup>\ \ (x \ -1)\<^sup>T" + using conv_complement conv_dist_inf inf_assoc inf_left_commute reachable_without_loops symmetric_one_closed by auto + also have "... = bot" + using assms(1) acyclic_star_below_complement_1 by auto + finally have 5: "x\<^sup>T \ (x * x)\<^sup>\ \ -1 = bot" + by (simp add: le_bot) + have "x\<^sup>T \ (x * x)\<^sup>\ = (x\<^sup>T \ (x * x)\<^sup>\ \ 1) \ (x\<^sup>T \ (x * x)\<^sup>\ \ -1)" + by (metis maddux_3_11_pp regular_one_closed) + also have "... = x\<^sup>T \ (x * x)\<^sup>\ \ 1" + using 5 by simp + also have "... = x\<^sup>T \ 1" + by (metis calculation comp_inf.semiring.distrib_left inf.sup_monoid.add_commute star.circ_transitive_equal star_involutive star_left_unfold_equal sup_inf_absorb) + finally have "(x\<^sup>T \ (x * x)\<^sup>\) \ (x\<^sup>T \ (x\<^sup>T * x\<^sup>T)\<^sup>\) \ (1 \ x) * top" + using 4 inf.sup_monoid.add_commute one_inf_conv top_right_mult_increasing by auto + hence "x\<^sup>T \ ((x * x)\<^sup>\ \ (x * x)\<^sup>T\<^sup>\) \ (1 \ x) * top" + by (simp add: comp_inf.semiring.distrib_left conv_dist_comp) + hence 6: "x\<^sup>T \ (x * x)\<^sup>T\<^sup>\ * (x * x)\<^sup>\ \ (1 \ x) * top" + using 1 by (simp add: cancel_separate_eq sup_commute) + have "(x * x)\<^sup>T\<^sup>\ \ x\<^sup>T * (x * x)\<^sup>T\<^sup>\ \ (x\<^sup>T \ (x * x)\<^sup>T\<^sup>\ * (x * x)\<^sup>\) * (x * x)\<^sup>T\<^sup>\" + by (metis conv_involutive conv_star_commute dedekind_2 inf_commute) + also have "... \ (1 \ x) * top * (x * x)\<^sup>T\<^sup>\" + using 6 by (simp add: mult_left_isotone) + also have "... = (1 \ x) * top" + by (simp add: comp_associative star.circ_left_top) + finally have "(x * x)\<^sup>T\<^sup>\ \ x\<^sup>T * (x * x)\<^sup>T\<^sup>\ = (x * x)\<^sup>T\<^sup>\ \ x\<^sup>T * (x * x)\<^sup>T\<^sup>\ \ (1 \ x) * top" + using inf.order_iff by auto + also have "... = (1 \ x) * ((x * x)\<^sup>T\<^sup>\ \ x\<^sup>T * (x * x)\<^sup>T\<^sup>\)" + by (metis coreflexive_comp_top_inf inf.cobounded1 inf.sup_monoid.add_commute) + finally show ?thesis + . +qed + +lemma update_square_plus: + "point y \ x[y\x[[x[[y]]]]] \ x\<^sup>+" + by (meson update_square comp_isotone dual_order.trans le_supI order_refl star.circ_increasing star.circ_mult_increasing) + +lemma update_square_ub_plus: + "x[y\(x * x)\<^sup>T] \ x\<^sup>+" + by (simp add: comp_isotone inf.coboundedI2 star.circ_increasing star.circ_mult_increasing) + +text \Theorem 6.2\ + +lemma acyclic_square: + assumes "acyclic (x \ -1)" + shows "x * x \ 1 = x \ 1" +proof (rule order.antisym) + have "1 \ x * x = 1 \ ((x \ -1) * x \ (x \ 1) * x)" + by (metis maddux_3_11_pp regular_one_closed semiring.distrib_right) + also have "... \ 1 \ ((x \ -1) * x \ x)" + by (metis inf.cobounded2 mult_1_left mult_left_isotone inf.sup_right_isotone semiring.add_left_mono) + also have "... = 1 \ ((x \ -1) * (x \ -1) \ (x \ -1) * (x \ 1) \ x)" + by (metis maddux_3_11_pp mult_left_dist_sup regular_one_closed) + also have "... \ (1 \ (x \ -1) * (x \ -1)) \ (x \ -1) * (x \ 1) \ x" + by (metis inf_le2 inf_sup_distrib1 semiring.add_left_mono sup_monoid.add_assoc) + also have "... \ (1 \ (x \ -1)\<^sup>+) \ (x \ -1) * (x \ 1) \ x" + by (metis comp_isotone inf.eq_refl inf.sup_right_isotone star.circ_increasing sup_monoid.add_commute sup_right_isotone) + also have "... = (x \ -1) * (x \ 1) \ x" + by (metis assms inf.le_iff_sup inf.sup_monoid.add_commute inf_import_p inf_p regular_one_closed sup_inf_absorb sup_monoid.add_commute) + also have "... = x" + by (metis comp_isotone inf.cobounded1 inf_le2 mult_1_right sup.absorb2) + finally show "x * x \ 1 \ x \ 1" + by (simp add: inf.sup_monoid.add_commute) + show "x \ 1 \ x * x \ 1" + by (metis coreflexive_idempotent inf_le1 inf_le2 le_infI mult_isotone) +qed + +lemma diagonal_update_square_aux: + assumes "acyclic (x \ -1)" + and "point y" + shows "1 \ y \ y\<^sup>T * x * x = 1 \ y \ x" +proof - + have 1: "1 \ y \ x \ y\<^sup>T * x * x" + by (metis comp_isotone coreflexive_idempotent inf.boundedE inf.cobounded1 inf.cobounded2 one_inf_conv) + have "1 \ y \ y\<^sup>T * x * x = 1 \ (y \ y\<^sup>T) * x * x" + by (simp add: assms(2) inf.sup_monoid.add_assoc vector_inf_comp) + also have "... = 1 \ (y \ 1) * x * x" + by (metis assms(2) inf.cobounded1 inf.sup_monoid.add_commute inf.sup_same_context one_inf_conv vector_covector) + also have "... \ 1 \ x * x" + by (metis comp_left_subdist_inf inf.sup_right_isotone le_infE mult_left_isotone mult_left_one) + also have "... \ x" + using assms(1) acyclic_square inf.sup_monoid.add_commute by auto + finally show ?thesis + using 1 by (metis inf.absorb2 inf.left_commute inf.sup_monoid.add_commute) +qed + +text \Theorem 6.5\ + +lemma diagonal_update_square: + assumes "acyclic (x \ -1)" + and "point y" + shows "(x[y\x[[x[[y]]]]]) \ 1 = x \ 1" +proof - + let ?xy = "x[[y]]" + let ?xxy = "x[[?xy]]" + let ?xyxxy = "x[y\?xxy]" + have "?xyxxy \ 1 = ((y \ y\<^sup>T * x * x) \ (-y \ x)) \ 1" + by (simp add: conv_dist_comp) + also have "... = (y \ y\<^sup>T * x * x \ 1) \ (-y \ x \ 1)" + by (simp add: inf_sup_distrib2) + also have "... = (y \ x \ 1) \ (-y \ x \ 1)" + using assms by (smt (verit, ccfv_threshold) diagonal_update_square_aux find_set_precondition_def inf_assoc inf_commute) + also have "... = x \ 1" + by (metis assms(2) bijective_regular comp_inf.mult_right_dist_sup inf.sup_monoid.add_commute maddux_3_11_pp) + finally show ?thesis + . +qed + +text \Theorem 6.6\ + +lemma fc_update_square: + assumes "mapping x" + and "point y" + shows "fc (x[y\x[[x[[y]]]]]) = fc x" +proof (rule order.antisym) + let ?xy = "x[[y]]" + let ?xxy = "x[[?xy]]" + let ?xyxxy = "x[y\?xxy]" + have 1: "y \ y\<^sup>T * x * x \ x * x" + by (smt (z3) assms(2) inf.cobounded2 inf.sup_monoid.add_commute inf.sup_same_context mult_1_left one_inf_conv vector_covector vector_inf_comp) + have 2: "?xyxxy = (y \ y\<^sup>T * x * x) \ (-y \ x)" + by (simp add: conv_dist_comp) + also have "... \ x * x \ x" + using 1 inf_le2 sup_mono by blast + also have "... \ x\<^sup>\" + by (simp add: star.circ_increasing star.circ_mult_upper_bound) + finally show "fc ?xyxxy \ fc x" + by (metis comp_isotone conv_order conv_star_commute star_involutive star_isotone) + have 3: "y \ x \ 1 \ fc ?xyxxy" + using inf.coboundedI1 inf.sup_monoid.add_commute reflexive_mult_closed star.circ_reflexive by auto + have 4: "y \ -1 \ -y\<^sup>T" + using assms(2) p_shunting_swap regular_one_closed vector_covector by auto + have "y \ x \ y\<^sup>T * x" + by (simp add: assms(2) vector_restrict_comp_conv) + also have "... \ y\<^sup>T * x * x * x\<^sup>T" + by (metis assms(1) comp_associative mult_1_right mult_right_isotone total_var) + finally have "y \ x \ -1 \ y \ -y\<^sup>T \ y\<^sup>T * x * x * x\<^sup>T" + using 4 by (smt (z3) inf.cobounded1 inf.coboundedI2 inf.sup_monoid.add_assoc inf.sup_monoid.add_commute inf_greatest) + also have "... = (y \ y\<^sup>T * x * x) * x\<^sup>T \ -y\<^sup>T" + by (metis assms(2) inf.sup_monoid.add_assoc inf.sup_monoid.add_commute vector_inf_comp) + also have "... = (y \ y\<^sup>T * x * x) * (x\<^sup>T \ -y\<^sup>T)" + using assms(2) covector_comp_inf vector_conv_compl by auto + also have "... = (y \ y\<^sup>T * x * x) * (-y \ x)\<^sup>T" + by (simp add: conv_complement conv_dist_inf inf_commute) + also have "... \ ?xyxxy * (-y \ x)\<^sup>T" + using 2 by (simp add: comp_left_increasing_sup) + also have "... \ ?xyxxy * ?xyxxy\<^sup>T" + by (simp add: conv_isotone mult_right_isotone) + also have "... \ fc ?xyxxy" + using comp_isotone star.circ_increasing by blast + finally have 5: "y \ x \ fc ?xyxxy" + using 3 by (smt (z3) comp_inf.semiring.distrib_left inf.le_iff_sup maddux_3_11_pp regular_one_closed) + have "x = (y \ x) \ (-y \ x)" + by (metis assms(2) bijective_regular inf.sup_monoid.add_commute maddux_3_11_pp) + also have "... \ fc ?xyxxy" + using 5 dual_order.trans fc_increasing sup.cobounded2 sup_least by blast + finally show "fc x \ fc ?xyxxy" + by (smt (z3) assms fc_equivalence fc_isotone fc_wcc read_injective star.circ_decompose_9 star_decompose_1 update_univalent) +qed + +text \Theorem 6.2\ + +lemma acyclic_plus_loop: + assumes "acyclic (x \ -1)" + shows "x\<^sup>+ \ 1 = x \ 1" +proof - + let ?r = "x \ 1" + let ?i = "x \ -1" + have "x\<^sup>+ \ 1 = (?i \ ?r)\<^sup>+ \ 1" + by (metis maddux_3_11_pp regular_one_closed) + also have "... = ((?i\<^sup>\ * ?r)\<^sup>\ * ?i\<^sup>+ \ (?i\<^sup>\ * ?r)\<^sup>+) \ 1" + using plus_sup by auto + also have "... \ (?i\<^sup>+ \ (?i\<^sup>\ * ?r)\<^sup>+) \ 1" + by (metis comp_associative dual_order.eq_iff maddux_3_11_pp reachable_without_loops regular_one_closed star.circ_plus_same star.circ_sup_9) + also have "... = (?i\<^sup>\ * ?r)\<^sup>+ \ 1" + by (smt (z3) assms comp_inf.mult_right_dist_sup inf.absorb2 inf.sup_monoid.add_commute inf_le2 maddux_3_11_pp pseudo_complement regular_one_closed) + also have "... \ ?i\<^sup>\ * ?r \ 1" + by (metis comp_associative dual_order.eq_iff maddux_3_11_pp reachable_without_loops regular_one_closed star.circ_sup_9 star_slide) + also have "... = (?r \ ?i\<^sup>+ * ?r) \ 1" + using comp_associative star.circ_loop_fixpoint sup_commute by force + also have "... \ x \ (?i\<^sup>+ * ?r \ 1)" + by (metis comp_inf.mult_right_dist_sup inf.absorb1 inf.cobounded1 inf.cobounded2) + also have "... \ x \ (-1 * ?r \ 1)" + by (meson assms comp_inf.comp_isotone mult_left_isotone order.refl semiring.add_left_mono) + also have "... = x" + by (metis comp_inf.semiring.mult_not_zero comp_right_one inf.cobounded2 inf_sup_absorb mult_right_isotone pseudo_complement sup.idem sup_inf_distrib1) + finally show ?thesis + by (meson inf.sup_same_context inf_le1 order_trans star.circ_mult_increasing) +qed + +lemma star_irreflexive_part_eq: + "x\<^sup>\ \ -1 = (x \ -1)\<^sup>+ \ -1" + by (metis reachable_without_loops star_plus_without_loops) + +text \Theorem 6.3\ + +lemma star_irreflexive_part: + "x\<^sup>\ \ -1 \ (x \ -1)\<^sup>+" + using star_irreflexive_part_eq by auto + +lemma square_irreflexive_part: + "x * x \ -1 \ (x \ -1)\<^sup>+" +proof - + have "x * x = (x \ 1) * x \ (x \ -1) * x" + by (metis maddux_3_11_pp mult_right_dist_sup regular_one_closed) + also have "... \ 1 * x \ (x \ -1) * x" + using comp_isotone inf.cobounded2 semiring.add_right_mono by blast + also have "... \ 1 \ (x \ -1) \ (x \ -1) * x" + by (metis inf.cobounded2 maddux_3_11_pp mult_1_left regular_one_closed sup_left_isotone) + also have "... = (x \ -1) * (x \ 1) \ 1" + by (simp add: mult_left_dist_sup sup_assoc sup_commute) + finally have "x * x \ -1 \ (x \ -1) * (x \ 1)" + using shunting_var_p by auto + also have "... = (x \ -1) * (x \ -1) \ (x \ -1)" + by (metis comp_right_one inf.sup_monoid.add_commute maddux_3_21_pp mult_left_dist_sup regular_one_closed sup_commute) + also have "... \ (x \ -1)\<^sup>+" + by (metis mult_left_isotone star.circ_increasing star.circ_mult_increasing star.circ_plus_same sup.bounded_iff) + finally show ?thesis + . +qed + +text \Theorem 6.3\ + +lemma square_irreflexive_part_2: + "x * x \ -1 \ x\<^sup>\ \ -1" + using comp_inf.mult_left_isotone star.circ_increasing star.circ_mult_upper_bound by blast + +text \Theorem 6.8\ + +lemma acyclic_update_square: + assumes "acyclic (x \ -1)" + shows "acyclic ((x[y\(x * x)\<^sup>T]) \ -1)" +proof - + have "((x[y\(x * x)\<^sup>T]) \ -1)\<^sup>+ \ ((x \ x * x) \ -1)\<^sup>+" + by (metis comp_inf.mult_right_isotone comp_isotone inf.sup_monoid.add_commute star_isotone update_square_ub) + also have "... = ((x \ -1) \ (x * x \ -1))\<^sup>+" + using comp_inf.semiring.distrib_right by auto + also have "... \ ((x \ -1)\<^sup>+)\<^sup>+" + by (smt (verit, del_insts) comp_isotone reachable_without_loops star.circ_mult_increasing star.circ_plus_same star.circ_right_slide star.circ_separate_5 star.circ_square star.circ_transitive_equal star.left_plus_circ sup.bounded_iff sup_ge1 square_irreflexive_part) + also have "... \ -1" + using assms by (simp add: acyclic_plus) + finally show ?thesis + . +qed + +text \Theorem 6.9\ + +lemma disjoint_set_forest_update_square: + assumes "disjoint_set_forest x" + and "vector y" + and "regular y" + shows "disjoint_set_forest (x[y\(x * x)\<^sup>T])" +proof (intro conjI) + show "univalent (x[y\(x * x)\<^sup>T])" + using assms update_univalent mapping_mult_closed univalent_conv_injective by blast + show "total (x[y\(x * x)\<^sup>T])" + using assms update_total total_conv_surjective total_mult_closed by blast + show "acyclic ((x[y\(x * x)\<^sup>T]) \ -1)" + using acyclic_update_square assms(1) by blast +qed + +lemma disjoint_set_forest_update_square_point: + assumes "disjoint_set_forest x" + and "point y" + shows "disjoint_set_forest (x[y\(x * x)\<^sup>T])" + using assms disjoint_set_forest_update_square bijective_regular by blast + +end + +section \Verifying Further Operations on Disjoint-Set Forests\ + +text \ +In this section we verify the init-sets, path-halving and path-splitting operations of disjoint-set forests. +\ + +class choose_point = + fixes choose_point :: "'a \ 'a" + +class stone_kleene_relation_algebra_choose_point_finite_regular = stone_kleene_relation_algebra + finite_regular_p_algebra + choose_point + + assumes choose_point_point: "vector x \ x \ bot \ point (choose_point x)" + assumes choose_point_decreasing: "choose_point x \ --x" +begin + +subclass stone_kleene_relation_algebra_tarski_finite_regular +proof unfold_locales + fix x + let ?p = "choose_point (x * top)" + let ?q = "choose_point ((?p \ x)\<^sup>T * top)" + let ?y = "?p \ ?q\<^sup>T" + assume 1: "regular x" "x \ bot" + hence 2: "x * top \ bot" + using le_bot top_right_mult_increasing by auto + hence 3: "point ?p" + by (simp add: choose_point_point comp_associative) + hence 4: "?p \ bot" + using 2 mult_right_zero by force + have "?p \ x \ bot" + proof + assume "?p \ x = bot" + hence 5: "x \ -?p" + using p_antitone_iff pseudo_complement by auto + have "?p \ --(x * top)" + by (simp add: choose_point_decreasing) + also have "... \ --(-?p * top)" + using 5 by (simp add: comp_isotone pp_isotone) + also have "... = -?p * top" + using regular_mult_closed by auto + also have "... = -?p" + using 3 vector_complement_closed by auto + finally have "?p = bot" + using inf_absorb2 by fastforce + thus False + using 4 by auto + qed + hence "(?p \ x)\<^sup>T * top \ bot" + by (metis comp_inf.semiring.mult_zero_left comp_right_one inf.sup_monoid.add_commute inf_top.left_neutral schroeder_1) + hence "point ?q" + using choose_point_point vector_top_closed mult_assoc by auto + hence 6: "arc ?y" + using 3 by (smt bijective_conv_mapping inf.sup_monoid.add_commute mapping_inf_point_arc) + have "?q \ --((?p \ x)\<^sup>T * top)" + by (simp add: choose_point_decreasing) + hence "?y \ ?p \ --((?p \ x)\<^sup>T * top)\<^sup>T" + by (metis conv_complement conv_isotone inf.sup_right_isotone) + also have "... = ?p \ --(top * (?p \ x))" + by (simp add: conv_dist_comp) + also have "... = ?p \ top * (?p \ x)" + using 1 3 bijective_regular pp_dist_comp by auto + also have "... = ?p \ ?p\<^sup>T * x" + using 3 by (metis comp_inf_vector conv_dist_comp inf.sup_monoid.add_commute inf_top_right symmetric_top_closed) + also have "... = (?p \ ?p\<^sup>T) * x" + using 3 by (simp add: vector_inf_comp) + also have "... \ 1 * x" + using 3 point_antisymmetric mult_left_isotone by blast + finally have "?y \ x" + by simp + thus "top * x * top = top" + using 6 by (smt (verit, ccfv_SIG) mult_assoc le_iff_sup mult_left_isotone semiring.distrib_left sup.orderE top.extremum) +qed + +subsection \Init-Sets\ + +text \ +A disjoint-set forest is initialised by applying \make_set\ to each node. +We prove that the resulting disjoint-set forest is the identity relation. +\ + +theorem init_sets: + "VARS h p x + [ True ] + h := top; + WHILE h \ bot + INV { regular h \ vector h \ p \ -h = 1 \ -h } + VAR { card { x . regular x \ x \ h } } + DO x := choose_point h; + p := make_set p x; + h[x] := bot + OD + [ p = 1 \ disjoint_set_forest p \ h = bot ]" +proof vcg_tc_simp + fix h p + let ?x = "choose_point h" + let ?m = "make_set p ?x" + assume 1: "regular h \ vector h \ p \ -h = 1 \ -h \ h \ bot" + show "vector (-?x \ h) \ + ?m \ (--?x \ -h) = 1 \ (--?x \ -h) \ + card { x . regular x \ x \ -?x \ x \ h } < card { x . regular x \ x \ h }" + proof (intro conjI) + show "vector (-?x \ h)" + using 1 choose_point_point vector_complement_closed vector_inf_closed by blast + have 2: "point ?x \ regular ?x" + using 1 bijective_regular choose_point_point by blast + have 3: "-h \ -?x" + using choose_point_decreasing p_antitone_iff by auto + have 4: "?x \ ?m = ?x * ?x\<^sup>T \ -?x \ ?m = -?x \ p" + using 1 choose_point_point make_set_function make_set_postcondition_def by auto + have "?m \ (--?x \ -h) = (?m \ ?x) \ (?m \ -h)" + using 2 comp_inf.comp_left_dist_sup by auto + also have "... = ?x * ?x\<^sup>T \ (?m \ -?x \ -h)" + using 3 4 by (smt (z3) inf_absorb2 inf_assoc inf_commute) + also have "... = ?x * ?x\<^sup>T \ (1 \ -h)" + using 1 3 4 inf.absorb2 inf.sup_monoid.add_assoc inf_commute by auto + also have "... = (1 \ ?x) \ (1 \ -h)" + using 2 by (metis inf.cobounded2 inf.sup_same_context one_inf_conv vector_covector) + also have "... = 1 \ (--?x \ -h)" + using 2 comp_inf.semiring.distrib_left by auto + finally show "?m \ (--?x \ -h) = 1 \ (--?x \ -h)" + . + have 5: "\ ?x \ -?x" + using 1 2 by (metis comp_commute_below_diversity conv_order inf.cobounded2 inf_absorb2 pseudo_complement strict_order_var top.extremum) + have 6: "?x \ h" + using 1 by (metis choose_point_decreasing) + show "card { x . regular x \ x \ -?x \ x \ h } < card { x . regular x \ x \ h }" + apply (rule psubset_card_mono) + using finite_regular apply simp + using 2 5 6 by auto + qed +qed + +end + +subsection \Path Halving\ + +text \ +Path halving is a variant of the path compression technique. +Similarly to path compression, we implement path halving independently of find-set, using a second while-loop which iterates over the same path to the root. +We prove that path halving preserves the equivalence-relational semantics of the disjoint-set forest and also preserves the roots of the component trees. +Additionally we prove the exact effect of path halving, which is to replace every other parent pointer with a pointer to the respective grandparent. +\ + +context stone_kleene_relation_algebra_tarski_finite_regular +begin + +definition "path_halving_invariant p x y p0 \ + find_set_precondition p x \ point y \ y \ p\<^sup>T\<^sup>\ * x \ y \ (p0 * p0)\<^sup>T\<^sup>\ * x \ + p0[(p0 * p0)\<^sup>T\<^sup>\ * x \ -(p0\<^sup>T\<^sup>\ * y)\(p0 * p0)\<^sup>T] = p \ + disjoint_set_forest p0" +definition "path_halving_postcondition p x y p0 \ + path_compression_precondition p x y \ p \ 1 = p0 \ 1 \ fc p = fc p0 \ + p0[(p0 * p0)\<^sup>T\<^sup>\ * x\(p0 * p0)\<^sup>T] = p" + +lemma path_halving_invariant_aux_1: + assumes "point x" + and "point y" + and "disjoint_set_forest p0" + shows "p0 \ wcc (p0[(p0 * p0)\<^sup>T\<^sup>\ * x \ -(p0\<^sup>T\<^sup>\ * y)\(p0 * p0)\<^sup>T])" +proof - + let ?p2 = "p0 * p0" + let ?p2t = "?p2\<^sup>T" + let ?p2ts = "?p2t\<^sup>\" + let ?px = "?p2ts * x" + let ?py = "-(p0\<^sup>T\<^sup>\ * y)" + let ?pxy = "?px \ ?py" + let ?p = "p0[?pxy\?p2t]" + have 1: "regular ?pxy" + using assms(1,3) bijective_regular find_set_precondition_def mapping_regular pp_dist_comp regular_closed_star regular_conv_closed path_halving_invariant_def by auto + have 2: "vector x \ vector ?px \ vector ?py" + using assms(1,2) find_set_precondition_def vector_complement_closed vector_mult_closed path_halving_invariant_def by auto + have 3: "?pxy \ p0 \ -?p2 \ -?px\<^sup>T" + proof - + have 4: "injective x \ univalent ?p2 \ regular p0" + using assms(1,3) find_set_precondition_def mapping_regular univalent_mult_closed path_halving_invariant_def by auto + have "?p2\<^sup>\ * p0 \ 1 \ p0\<^sup>+ \ 1" + using comp_inf.mult_left_isotone comp_isotone comp_right_one mult_sub_right_one star.circ_square star_slide by auto + also have "... \ p0" + using acyclic_plus_loop assms(3) path_halving_invariant_def by auto + finally have 5: "?p2\<^sup>\ * p0 \ 1 \ p0" + . + hence 6: "?p2ts * (1 \ -p0) \ -p0" + by (smt (verit, ccfv_SIG) conv_star_commute dual_order.trans inf.sup_monoid.add_assoc order.refl p_antitone_iff pseudo_complement schroeder_4_p schroeder_6_p) + have "?p2t\<^sup>+ * p0 \ 1 = ?p2ts * p0\<^sup>T * (p0\<^sup>T * p0) \ 1" + by (metis conv_dist_comp star_plus mult_assoc) + also have "... \ ?p2ts * p0\<^sup>T \ 1" + by (metis assms(3) comp_inf.mult_left_isotone comp_isotone comp_right_one mult_sub_right_one) + also have "... \ p0" + using 5 by (metis conv_dist_comp conv_star_commute inf_commute one_inf_conv star_slide) + finally have "?p2t\<^sup>+ * p0 \ -1 \ p0" + by (metis regular_one_closed shunting_var_p sup_commute) + hence 7: "?p2\<^sup>+ * (1 \ -p0) \ -p0" + by (smt (z3) conv_dist_comp conv_star_commute half_shunting inf.sup_monoid.add_assoc inf.sup_monoid.add_commute pseudo_complement schroeder_4_p schroeder_6_p star.circ_plus_same) + have "(1 \ ?px) * top * (1 \ ?px \ -p0) = ?px \ top * (1 \ ?px \ -p0)" + using 2 by (metis inf_commute vector_inf_one_comp mult_assoc) + also have "... = ?px \ ?px\<^sup>T * (1 \ -p0)" + using 2 by (smt (verit, ccfv_threshold) covector_inf_comp_3 inf.sup_monoid.add_assoc inf.sup_monoid.add_commute inf_top.left_neutral) + also have "... = ?px \ x\<^sup>T * ?p2\<^sup>\ * (1 \ -p0)" + by (simp add: conv_dist_comp conv_star_commute) + also have "... = (?px \ x\<^sup>T) * ?p2\<^sup>\ * (1 \ -p0)" + using 2 vector_inf_comp by auto + also have "... = ?p2ts * (x * x\<^sup>T) * ?p2\<^sup>\ * (1 \ -p0)" + using 2 vector_covector mult_assoc by auto + also have "... \ ?p2ts * ?p2\<^sup>\ * (1 \ -p0)" + using 4 by (metis inf.order_lesseq_imp mult_left_isotone star.circ_mult_upper_bound star.circ_reflexive) + also have "... = (?p2ts \ ?p2\<^sup>\) * (1 \ -p0)" + using 4 by (simp add: cancel_separate_eq) + also have "... = (?p2ts \ ?p2\<^sup>+) * (1 \ -p0)" + by (metis star.circ_plus_one star_plus_loops sup_assoc sup_commute) + also have "... \ -p0" + using 6 7 by (simp add: mult_right_dist_sup) + finally have "(1 \ ?px)\<^sup>T * p0 * (1 \ ?px \ -p0)\<^sup>T \ bot" + by (smt (z3) inf.boundedI inf_p top.extremum triple_schroeder_p) + hence 8: "(1 \ ?px) * p0 * (1 \ ?px \ -p0) = bot" + by (simp add: coreflexive_inf_closed coreflexive_symmetric le_bot) + have "?px \ p0 \ ?px\<^sup>T = (1 \ ?px) * p0 \ ?px\<^sup>T" + using 2 inf_commute vector_inf_one_comp by fastforce + also have "... = (1 \ ?px) * p0 * (1 \ ?px)" + using 2 by (metis comp_inf_vector mult_1_right vector_conv_covector) + also have "... = (1 \ ?px) * p0 * (1 \ ?px \ p0) \ (1 \ ?px) * p0 * (1 \ ?px \ -p0)" + using 4 by (metis maddux_3_11_pp mult_left_dist_sup) + also have "... = (1 \ ?px) * p0 * (1 \ ?px \ p0)" + using 8 by simp + also have "... \ ?p2" + by (metis comp_isotone coreflexive_comp_top_inf inf.cobounded1 inf.cobounded2) + finally have "?px \ p0 \ -?p2 \ -?px\<^sup>T" + using 4 p_shunting_swap regular_mult_closed by fastforce + thus ?thesis + by (meson comp_inf.mult_left_isotone dual_order.trans inf.cobounded1) + qed + have "p0 \ ?p2 * p0\<^sup>T" + by (metis assms(3) comp_associative comp_isotone comp_right_one eq_refl total_var) + hence "?pxy \ p0 \ -?p2 \ ?p2 * p0\<^sup>T" + by (metis inf.coboundedI1 inf.sup_monoid.add_commute) + hence "?pxy \ p0 \ -?p2 \ ?pxy \ ?p2 * p0\<^sup>T \ -?px\<^sup>T" + using 3 by (meson dual_order.trans inf.boundedI inf.cobounded1) + also have "... = (?pxy \ ?p2) * p0\<^sup>T \ -?px\<^sup>T" + using 2 vector_inf_comp by auto + also have "... = (?pxy \ ?p2) * (-?px \ p0)\<^sup>T" + using 2 by (simp add: covector_comp_inf inf.sup_monoid.add_commute vector_conv_compl conv_complement conv_dist_inf) + also have "... \ ?p * (-?px \ p0)\<^sup>T" + using comp_left_increasing_sup by auto + also have "... \ ?p * ?p\<^sup>T" + by (metis comp_inf.mult_right_isotone comp_isotone conv_isotone inf.eq_refl inf.sup_monoid.add_commute le_supI1 p_antitone_inf sup_commute) + also have "... \ wcc ?p" + using star.circ_sub_dist_2 by auto + finally have 9: "?pxy \ p0 \ -?p2 \ wcc ?p" + . + have "p0 = (?pxy \ p0) \ (-?pxy \ p0)" + using 1 by (metis inf.sup_monoid.add_commute maddux_3_11_pp) + also have "... \ (?pxy \ p0) \ ?p" + using sup_right_isotone by auto + also have "... = (?pxy \ p0 \ -?p2) \ (?pxy \ p0 \ ?p2) \ ?p" + by (smt (z3) assms(3) maddux_3_11_pp mapping_regular pp_dist_comp path_halving_invariant_def) + also have "... \ (?pxy \ p0 \ -?p2) \ (?pxy \ ?p2) \ ?p" + by (meson comp_inf.comp_left_subdist_inf inf.boundedE semiring.add_left_mono semiring.add_right_mono) + also have "... = (?pxy \ p0 \ -?p2) \ ?p" + using sup_assoc by auto + also have "... \ wcc ?p \ ?p" + using 9 sup_left_isotone by blast + also have "... \ wcc ?p" + by (simp add: star.circ_sub_dist_1) + finally show ?thesis + . +qed + +lemma path_halving_invariant_aux: + assumes "path_halving_invariant p x y p0" + shows "p[[y]] = p0[[y]]" + and "p[[p[[y]]]] = p0[[p0[[y]]]]" + and "p[[p[[p[[y]]]]]] = p0[[p0[[p0[[y]]]]]]" + and "p \ 1 = p0 \ 1" + and "fc p = fc p0" +proof - + let ?p2 = "p0 * p0" + let ?p2t = "?p2\<^sup>T" + let ?p2ts = "?p2t\<^sup>\" + let ?px = "?p2ts * x" + let ?py = "-(p0\<^sup>T\<^sup>\ * y)" + let ?pxy = "?px \ ?py" + let ?p = "p0[?pxy\?p2t]" + have "?p[[y]] = p0[[y]]" + apply (rule put_get_different_vector) + using assms find_set_precondition_def vector_complement_closed vector_inf_closed vector_mult_closed path_halving_invariant_def apply force + by (meson inf.cobounded2 order_lesseq_imp p_antitone_iff path_compression_1b) + thus 1: "p[[y]] = p0[[y]]" + using assms path_halving_invariant_def by auto + have "?p[[p0[[y]]]] = p0[[p0[[y]]]]" + apply (rule put_get_different_vector) + using assms find_set_precondition_def vector_complement_closed vector_inf_closed vector_mult_closed path_halving_invariant_def apply force + by (metis comp_isotone inf.boundedE inf.coboundedI2 inf.eq_refl p_antitone_iff selection_closed_id star.circ_increasing) + thus 2: "p[[p[[y]]]] = p0[[p0[[y]]]]" + using 1 assms path_halving_invariant_def by auto + have "?p[[p0[[p0[[y]]]]]] = p0[[p0[[p0[[y]]]]]]" + apply (rule put_get_different_vector) + using assms find_set_precondition_def vector_complement_closed vector_inf_closed vector_mult_closed path_halving_invariant_def apply force + by (metis comp_associative comp_isotone conv_dist_comp conv_involutive conv_order inf.coboundedI2 inf.le_iff_sup mult_left_isotone p_antitone_iff p_antitone_inf star.circ_increasing star.circ_transitive_equal) + thus "p[[p[[p[[y]]]]]] = p0[[p0[[p0[[y]]]]]]" + using 2 assms path_halving_invariant_def by auto + have 3: "regular ?pxy" + using assms bijective_regular find_set_precondition_def mapping_regular pp_dist_comp regular_closed_star regular_conv_closed path_halving_invariant_def by auto + have "p \ 1 = ?p \ 1" + using assms path_halving_invariant_def by auto + also have "... = (?pxy \ ?p2 \ 1) \ (-?pxy \ p0 \ 1)" + using comp_inf.semiring.distrib_right conv_involutive by auto + also have "... = (?pxy \ p0 \ 1) \ (-?pxy \ p0 \ 1)" + using assms acyclic_square path_halving_invariant_def inf.sup_monoid.add_assoc by auto + also have "... = (?pxy \ -?pxy) \ p0 \ 1" + using inf_sup_distrib2 by auto + also have "... = p0 \ 1" + using 3 by (metis inf.sup_monoid.add_commute inf_sup_distrib1 maddux_3_11_pp) + finally show "p \ 1 = p0 \ 1" + . + have "p \ p0\<^sup>+" + by (metis assms path_halving_invariant_def update_square_ub_plus) + hence 4: "fc p \ fc p0" + using conv_plus_commute fc_isotone star.left_plus_circ by fastforce + have "wcc p0 \ wcc ?p" + by (meson assms wcc_below_wcc path_halving_invariant_aux_1 path_halving_invariant_def find_set_precondition_def) + hence "fc p0 \ fc ?p" + using assms find_set_precondition_def path_halving_invariant_def fc_wcc by auto + thus "fc p = fc p0" + using 4 assms path_halving_invariant_def by auto +qed + +lemma path_halving_1: + "find_set_precondition p0 x \ path_halving_invariant p0 x x p0" +proof - + assume 1: "find_set_precondition p0 x" + show "path_halving_invariant p0 x x p0" + proof (unfold path_halving_invariant_def, intro conjI) + show "find_set_precondition p0 x" + using 1 by simp + show "vector x" "injective x" "surjective x" + using 1 find_set_precondition_def by auto + show "x \ p0\<^sup>T\<^sup>\ * x" + by (simp add: path_compression_1b) + show "x \ (p0 * p0)\<^sup>T\<^sup>\ * x" + by (simp add: path_compression_1b) + have "(p0 * p0)\<^sup>T\<^sup>\ * x \ p0\<^sup>T\<^sup>\ * x" + by (simp add: conv_dist_comp mult_left_isotone star.circ_square) + thus "p0[(p0 * p0)\<^sup>T\<^sup>\ * x \ - (p0\<^sup>T\<^sup>\ * x)\(p0 * p0)\<^sup>T] = p0" + by (smt (z3) inf.le_iff_sup inf_commute maddux_3_11_pp p_antitone_inf pseudo_complement) + show "univalent p0" "total p0" "acyclic (p0 \ -1)" + using 1 find_set_precondition_def by auto + qed +qed + +lemma path_halving_2: + "path_halving_invariant p x y p0 \ y \ p[[y]] \ path_halving_invariant (p[y\p[[p[[y]]]]]) x ((p[y\p[[p[[y]]]]])[[y]]) p0 \ card { z . regular z \ z \ (p[y\p[[p[[y]]]]])\<^sup>T\<^sup>\ * ((p[y\p[[p[[y]]]]])[[y]]) } < card { z . regular z \ z \ p\<^sup>T\<^sup>\ * y }" +proof - + let ?py = "p[[y]]" + let ?ppy = "p[[?py]]" + let ?pyppy = "p[y\?ppy]" + let ?p2 = "p0 * p0" + let ?p2t = "?p2\<^sup>T" + let ?p2ts = "?p2t\<^sup>\" + let ?px = "?p2ts * x" + let ?py2 = "-(p0\<^sup>T\<^sup>\ * y)" + let ?pxy = "?px \ ?py2" + let ?p = "p0[?pxy\?p2t]" + let ?pty = "p0\<^sup>T * y" + let ?pt2y = "p0\<^sup>T * p0\<^sup>T * y" + let ?pt2sy = "p0\<^sup>T\<^sup>\ * p0\<^sup>T * p0\<^sup>T * y" + assume 1: "path_halving_invariant p x y p0 \ y \ ?py" + have 2: "point ?pty \ point ?pt2y" + using 1 by (smt (verit) comp_associative read_injective read_surjective path_halving_invariant_def) + show "path_halving_invariant ?pyppy x (?pyppy[[y]]) p0 \ card { z . regular z \ z \ ?pyppy\<^sup>T\<^sup>\ * (?pyppy[[y]]) } < card { z . regular z \ z \ p\<^sup>T\<^sup>\ * y }" + proof + show "path_halving_invariant ?pyppy x (?pyppy[[y]]) p0" + proof (unfold path_halving_invariant_def, intro conjI) + show 3: "find_set_precondition ?pyppy x" + proof (unfold find_set_precondition_def, intro conjI) + show "univalent ?pyppy" + using 1 find_set_precondition_def read_injective update_univalent path_halving_invariant_def by auto + show "total ?pyppy" + using 1 bijective_regular find_set_precondition_def read_surjective update_total path_halving_invariant_def by force + show "acyclic (?pyppy \ -1)" + apply (rule update_acyclic_3) + using 1 find_set_precondition_def path_halving_invariant_def apply blast + using 1 2 comp_associative path_halving_invariant_aux(2) apply force + using 1 path_halving_invariant_def apply blast + by (metis inf.order_lesseq_imp mult_isotone star.circ_increasing star.circ_square mult_assoc) + show "vector x" "injective x" "surjective x" + using 1 find_set_precondition_def path_halving_invariant_def by auto + qed + show "vector (?pyppy[[y]])" + using 1 comp_associative path_halving_invariant_def by auto + show "injective (?pyppy[[y]])" + using 1 3 read_injective path_halving_invariant_def find_set_precondition_def by auto + show "surjective (?pyppy[[y]])" + using 1 3 read_surjective path_halving_invariant_def find_set_precondition_def by auto + show "?pyppy[[y]] \ ?pyppy\<^sup>T\<^sup>\ * x" + proof - + have "y = (y \ p\<^sup>T\<^sup>\) * x" + using 1 le_iff_inf vector_inf_comp path_halving_invariant_def by auto + also have "... = ((y \ 1) \ (y \ (p\<^sup>T \ -y\<^sup>T)\<^sup>+)) * x" + using 1 omit_redundant_points_3 path_halving_invariant_def by auto + also have "... \ (1 \ (y \ (p\<^sup>T \ -y\<^sup>T)\<^sup>+)) * x" + using 1 sup_inf_distrib2 vector_inf_comp path_halving_invariant_def by auto + also have "... \ (1 \ (p\<^sup>T \ -y\<^sup>T)\<^sup>+) * x" + by (simp add: inf.coboundedI2 mult_left_isotone) + also have "... = (p \ -y)\<^sup>T\<^sup>\ * x" + by (simp add: conv_complement conv_dist_inf star_left_unfold_equal) + also have "... \ ?pyppy\<^sup>T\<^sup>\ * x" + by (simp add: conv_isotone inf.sup_monoid.add_commute mult_left_isotone star_isotone) + finally show ?thesis + by (metis mult_isotone star.circ_increasing star.circ_transitive_equal mult_assoc) + qed + show "?pyppy[[y]] \ ?px" + proof - + have "?pyppy[[y]] = p[[?py]]" + using 1 put_get vector_mult_closed path_halving_invariant_def by force + also have "... = p0[[p0[[y]]]]" + using 1 path_halving_invariant_aux(2) by blast + also have "... = ?p2t * y" + by (simp add: conv_dist_comp mult_assoc) + also have "... \ ?p2t * ?px" + using 1 path_halving_invariant_def comp_associative mult_right_isotone by force + also have "... \ ?px" + by (metis comp_associative mult_left_isotone star.left_plus_below_circ) + finally show ?thesis + . + qed + show "p0[?px \ -(p0\<^sup>T\<^sup>\ * (?pyppy[[y]]))\?p2t] = ?pyppy" + proof - + have "?px \ ?pty = ?px \ p0\<^sup>T * ?px \ ?pty" + using 1 inf.absorb2 inf.sup_monoid.add_assoc mult_right_isotone path_halving_invariant_def by force + also have "... = (?p2ts \ p0\<^sup>T * ?p2ts) * x \ ?pty" + using 3 comp_associative find_set_precondition_def injective_comp_right_dist_inf by auto + also have "... = (1 \ p0) * (?p2ts \ p0\<^sup>T * ?p2ts) * x \ ?pty" + using 1 even_odd_root mapping_regular path_halving_invariant_def by auto + also have "... \ (1 \ p0) * top \ ?pty" + by (metis comp_associative comp_inf.mult_left_isotone comp_inf.star.circ_sub_dist_2 comp_left_subdist_inf dual_order.trans mult_right_isotone) + also have 4: "... = (1 \ p0\<^sup>T) * ?pty" + using coreflexive_comp_top_inf one_inf_conv by auto + also have "... \ ?pt2y" + by (simp add: mult_assoc mult_left_isotone) + finally have 5: "?px \ ?pty \ ?pt2y" + . + have 6: "p[?px \ -?pt2sy \ ?pty\?p2t] = p" + proof (cases "?pty \ ?px \ -?pt2sy") + case True + hence "?pty \ ?pt2y" + using 5 conv_dist_comp inf.absorb2 by auto + hence 7: "?pty = ?pt2y" + using 2 epm_3 by fastforce + have "p[?px \ -?pt2sy \ ?pty\?p2t] = p[?pty\?p2t]" + using True inf.absorb2 by auto + also have "... = p[?pty\?p2[[?pty]]]" + using 2 update_point_get by auto + also have "... = p[?pty\p0\<^sup>T * p0\<^sup>T * p0\<^sup>T * y]" + using comp_associative conv_dist_comp by auto + also have "... = p[?pty\?pt2y]" + using 7 mult_assoc by simp + also have "... = p[?pty\p[[?pty]]]" + using 1 path_halving_invariant_aux(1,2) mult_assoc by force + also have "... = p" + using 2 get_put by auto + finally show ?thesis + . + next + case False + have "mapping ?p2" + using 1 mapping_mult_closed path_halving_invariant_def by blast + hence 8: "regular (?px \ -?pt2sy)" + using 1 bijective_regular find_set_precondition_def mapping_regular pp_dist_comp regular_closed_star regular_conv_closed path_halving_invariant_def by auto + have "vector (?px \ -?pt2sy)" + using 1 find_set_precondition_def vector_complement_closed vector_inf_closed vector_mult_closed path_halving_invariant_def by force + hence "?pty \ -(?px \ -?pt2sy)" + using 2 8 point_in_vector_or_complement False by blast + hence "?px \ -?pt2sy \ ?pty = bot" + by (simp add: p_antitone_iff pseudo_complement) + thus ?thesis + by simp + qed + have 9: "p[?px \ -?pt2sy \ y\?p2t] = ?pyppy" + proof (cases "y \ -?pt2sy") + case True + hence "p[?px \ -?pt2sy \ y\?p2t] = p[y\?p2t]" + using 1 inf.absorb2 path_halving_invariant_def by auto + also have "... = ?pyppy" + using 1 by (metis comp_associative conv_dist_comp path_halving_invariant_aux(2) path_halving_invariant_def update_point_get) + finally show ?thesis + . + next + case False + have "vector (-?pt2sy)" + using 1 vector_complement_closed vector_mult_closed path_halving_invariant_def by blast + hence 10: "y \ ?pt2sy" + using 1 by (smt (verit, del_insts) False bijective_regular point_in_vector_or_complement regular_closed_star regular_mult_closed total_conv_surjective univalent_conv_injective path_halving_invariant_def) + hence "?px \ -?pt2sy \ y = bot" + by (simp add: inf.coboundedI2 p_antitone pseudo_complement) + hence 11: "p[?px \ -?pt2sy \ y\?p2t] = p" + by simp + have "y \ p0\<^sup>T\<^sup>+ * y" + using 10 by (metis mult_left_isotone order_lesseq_imp star.circ_plus_same star.left_plus_below_circ) + hence 12: "y = root p0 y" + using 1 loop_root_2 path_halving_invariant_def by blast + have "?pyppy = p[y\p0[[p0[[y]]]]]" + using 1 path_halving_invariant_aux(2) by force + also have "... = p[y\p0[[y]]]" + using 1 12 by (metis root_successor_loop path_halving_invariant_def) + also have "... = p[y\?py]" + using 1 path_halving_invariant_aux(1) by force + also have "... = p" + using 1 get_put path_halving_invariant_def by blast + finally show ?thesis + using 11 by simp + qed + have 13: "-?pt2sy = -(p0\<^sup>T\<^sup>\ * y) \ (-?pt2sy \ ?pty) \ (-?pt2sy \ y)" + proof (rule order.antisym) + have 14: "regular (p0\<^sup>T\<^sup>\ * y) \ regular ?pt2sy" + using 1 by (metis order.antisym conv_complement conv_dist_comp conv_involutive conv_star_commute forest_components_increasing mapping_regular pp_dist_star regular_mult_closed top.extremum path_halving_invariant_def) + have "p0\<^sup>T\<^sup>\ = p0\<^sup>T\<^sup>\ * p0\<^sup>T * p0\<^sup>T \ p0\<^sup>T \ 1" + using star.circ_back_loop_fixpoint star.circ_plus_same star_left_unfold_equal sup_commute by auto + hence "p0\<^sup>T\<^sup>\ * y \ ?pt2sy \ ?pty \ y" + by (metis inf.eq_refl mult_1_left mult_right_dist_sup) + also have "... = ?pt2sy \ (-?pt2sy \ ?pty) \ y" + using 14 by (metis maddux_3_21_pp) + also have "... = ?pt2sy \ (-?pt2sy \ ?pty) \ (-?pt2sy \ y)" + using 14 by (smt (z3) maddux_3_21_pp sup.left_commute sup_assoc) + hence "p0\<^sup>T\<^sup>\ * y \ -?pt2sy \ (-?pt2sy \ ?pty) \ (-?pt2sy \ y)" + using calculation half_shunting sup_assoc sup_commute by auto + thus "-?pt2sy \ -(p0\<^sup>T\<^sup>\ * y) \ (-?pt2sy \ ?pty) \ (-?pt2sy \ y)" + using 14 by (smt (z3) inf.sup_monoid.add_commute shunting_var_p sup.left_commute sup_commute) + have "-(p0\<^sup>T\<^sup>\ * y) \ -?pt2sy" + by (meson mult_left_isotone order.trans p_antitone star.right_plus_below_circ) + thus "-(p0\<^sup>T\<^sup>\ * y) \ (-?pt2sy \ ?pty) \ (-?pt2sy \ y) \ -?pt2sy" + by simp + qed + have "regular ?px" "regular ?pty" "regular y" + using 1 bijective_regular find_set_precondition_def mapping_regular pp_dist_comp regular_closed_star regular_conv_closed path_halving_invariant_def by auto + hence 15: "regular (?px \ -?pt2sy \ ?pty)" "regular (?px \ -?pt2sy \ y)" + by auto + have "p0[?px \ -(p0\<^sup>T\<^sup>\ * (?pyppy[[y]]))\?p2t] = p0[?px \ -(p0\<^sup>T\<^sup>\ * (p[[?py]]))\?p2t]" + using 1 put_get vector_mult_closed path_halving_invariant_def by auto + also have "... = p0[?px \ -?pt2sy\?p2t]" + using 1 comp_associative path_halving_invariant_aux(2) by force + also have "... = p0[?pxy \ (?px \ -?pt2sy \ ?pty) \ (?px \ -?pt2sy \ y)\?p2t]" + using 13 by (metis comp_inf.semiring.distrib_left inf.sup_monoid.add_assoc) + also have "... = (?p[?px \ -?pt2sy \ ?pty\?p2t])[?px \ -?pt2sy \ y\?p2t]" + using 15 by (smt (z3) update_same_3 comp_inf.semiring.mult_not_zero inf.sup_monoid.add_assoc inf.sup_monoid.add_commute) + also have "... = (p[?px \ -?pt2sy \ ?pty\?p2t])[?px \ -?pt2sy \ y\?p2t]" + using 1 path_halving_invariant_def by auto + also have "... = p[?px \ -?pt2sy \ y\?p2t]" + using 6 by simp + also have "... = ?pyppy" + using 9 by auto + finally show ?thesis + . + qed + show "univalent p0" "total p0" "acyclic (p0 \ -1)" + using 1 path_halving_invariant_def by auto + qed + let ?s = "{ z . regular z \ z \ p\<^sup>T\<^sup>\ * y }" + let ?t = "{ z . regular z \ z \ ?pyppy\<^sup>T\<^sup>\ * (?pyppy[[y]]) }" + have "?pyppy\<^sup>T\<^sup>\ * (?pyppy[[y]]) = ?pyppy\<^sup>T\<^sup>\ * (p[[?py]])" + using 1 put_get vector_mult_closed path_halving_invariant_def by force + also have "... \ p\<^sup>+\<^sup>T\<^sup>\ * (p[[?py]])" + using 1 path_halving_invariant_def update_square_plus conv_order mult_left_isotone star_isotone by force + also have "... = p\<^sup>T\<^sup>\ * p\<^sup>T * p\<^sup>T * y" + by (simp add: conv_plus_commute star.left_plus_circ mult_assoc) + also have "... \ p\<^sup>T\<^sup>+ * y" + by (metis mult_left_isotone star.left_plus_below_circ star_plus) + finally have 16: "?pyppy\<^sup>T\<^sup>\ * (?pyppy[[y]]) \ p\<^sup>T\<^sup>+ * y" + . + hence "?pyppy\<^sup>T\<^sup>\ * (?pyppy[[y]]) \ p\<^sup>T\<^sup>\ * y" + using mult_left_isotone order_lesseq_imp star.left_plus_below_circ by blast + hence 17: "?t \ ?s" + using order_trans by auto + have 18: "y \ ?s" + using 1 bijective_regular path_compression_1b path_halving_invariant_def by force + have 19: "\ y \ ?t" + proof + assume "y \ ?t" + hence "y \ ?pyppy\<^sup>T\<^sup>\ * (?pyppy[[y]])" + by simp + hence "y \ p\<^sup>T\<^sup>+ * y" + using 16 dual_order.trans by blast + hence "y = root p y" + using 1 find_set_precondition_def loop_root_2 path_halving_invariant_def by blast + hence "y = ?py" + using 1 by (metis find_set_precondition_def root_successor_loop path_halving_invariant_def) + thus False + using 1 by simp + qed + show "card ?t < card ?s" + apply (rule psubset_card_mono) + subgoal using finite_regular by simp + subgoal using 17 18 19 by auto + done + qed +qed + +lemma path_halving_3: + "path_halving_invariant p x y p0 \ y = p[[y]] \ path_halving_postcondition p x y p0" +proof - + assume 1: "path_halving_invariant p x y p0 \ y = p[[y]]" + show "path_halving_postcondition p x y p0" + proof (unfold path_halving_postcondition_def path_compression_precondition_def, intro conjI) + show "univalent p" "total p" "acyclic (p \ -1)" + using 1 find_set_precondition_def path_halving_invariant_def by blast+ + show "vector x" "injective x" "surjective x" + using 1 find_set_precondition_def path_halving_invariant_def by blast+ + show 2: "vector y" "injective y" "surjective y" + using 1 path_halving_invariant_def by blast+ + have "find_set_invariant p x y" + using 1 find_set_invariant_def path_halving_invariant_def by blast + thus "y = root p x" + using 1 find_set_3 find_set_postcondition_def by blast + show "p \ 1 = p0 \ 1" + using 1 path_halving_invariant_aux(4) by blast + show "fc p = fc p0" + using 1 path_halving_invariant_aux(5) by blast + have 3: "y = p0[[y]]" + using 1 path_halving_invariant_aux(1) by auto + hence "p0\<^sup>T\<^sup>\ * y = y" + using order.antisym path_compression_1b star_left_induct_mult_equal by auto + hence 4: "p0[(p0 * p0)\<^sup>T\<^sup>\ * x \ -y\(p0 * p0)\<^sup>T] = p" + using 1 path_halving_invariant_def by auto + have "(p0 * p0)\<^sup>T * y = y" + using 3 mult_assoc conv_dist_comp by auto + hence "y \ p0 * p0 = y \ p0" + using 2 3 by (metis update_postcondition) + hence 5: "y \ p = y \ p0 * p0" + using 1 2 3 by (smt update_postcondition) + have "p0[(p0 * p0)\<^sup>T\<^sup>\ * x\(p0 * p0)\<^sup>T] = (p0[(p0 * p0)\<^sup>T\<^sup>\ * x \ -y\(p0 * p0)\<^sup>T])[(p0 * p0)\<^sup>T\<^sup>\ * x \ y\(p0 * p0)\<^sup>T]" + using 1 bijective_regular path_halving_invariant_def update_split by blast + also have "... = p[(p0 * p0)\<^sup>T\<^sup>\ * x \ y\(p0 * p0)\<^sup>T]" + using 4 by simp + also have "... = p" + apply (rule update_same_sub) + using 5 apply simp + apply simp + using 1 bijective_regular inf.absorb2 path_halving_invariant_def by auto + finally show "p0[(p0 * p0)\<^sup>T\<^sup>\ * x\(p0 * p0)\<^sup>T] = p" + . + qed +qed + +theorem find_path_halving: + "VARS p y + [ find_set_precondition p x \ p0 = p ] + y := x; + WHILE y \ p[[y]] + INV { path_halving_invariant p x y p0 } + VAR { card { z . regular z \ z \ p\<^sup>T\<^sup>\ * y } } + DO p[y] := p[[p[[y]]]]; + y := p[[y]] + OD + [ path_halving_postcondition p x y p0 ]" + apply vcg_tc_simp + apply (fact path_halving_1) + apply (fact path_halving_2) + by (fact path_halving_3) + +subsection \Path Splitting\ + +text \ +Path splitting is another variant of the path compression technique. +We implement it again independently of find-set, using a second while-loop which iterates over the same path to the root. +We prove that path splitting preserves the equivalence-relational semantics of the disjoint-set forest and also preserves the roots of the component trees. +Additionally we prove the exact effect of path splitting, which is to replace every parent pointer with a pointer to the respective grandparent. +\ + +definition "path_splitting_invariant p x y p0 \ + find_set_precondition p x \ point y \ y \ p0\<^sup>T\<^sup>\ * x \ + p0[p0\<^sup>T\<^sup>\ * x \ -(p0\<^sup>T\<^sup>\ * y)\(p0 * p0)\<^sup>T] = p \ + disjoint_set_forest p0" +definition "path_splitting_postcondition p x y p0 \ + path_compression_precondition p x y \ p \ 1 = p0 \ 1 \ fc p = fc p0 \ + p0[p0\<^sup>T\<^sup>\ * x\(p0 * p0)\<^sup>T] = p" + +lemma path_splitting_invariant_aux_1: + assumes "point x" + and "point y" + and "disjoint_set_forest p0" + shows "(p0[p0\<^sup>T\<^sup>\ * x \ -(p0\<^sup>T\<^sup>\ * y)\(p0 * p0)\<^sup>T]) \ 1 = p0 \ 1" + and "fc (p0[p0\<^sup>T\<^sup>\ * x \ -(p0\<^sup>T\<^sup>\ * y)\(p0 * p0)\<^sup>T]) = fc p0" + and "p0\<^sup>T\<^sup>\ * x \ p0\<^sup>\ * root p0 x" +proof - + let ?p2 = "p0 * p0" + let ?p2t = "?p2\<^sup>T" + let ?px = "p0\<^sup>T\<^sup>\ * x" + let ?py = "-(p0\<^sup>T\<^sup>\ * y)" + let ?pxy = "?px \ ?py" + let ?q1 = "?pxy \ p0" + let ?q2 = "-?pxy \ p0" + let ?q3 = "?pxy \ ?p2" + let ?q4 = "-?pxy \ ?p2" + let ?p = "p0[?pxy\?p2t]" + let ?r0 = "root p0 x" + let ?rp = "root ?p x" + have 1: "regular ?px \ regular (p0\<^sup>T\<^sup>\ * y) \ regular ?pxy" + using assms bijective_regular find_set_precondition_def mapping_regular pp_dist_comp regular_closed_star regular_conv_closed path_halving_invariant_def regular_closed_inf by auto + have 2: "vector x \ vector ?px \ vector ?py \ vector ?pxy" + using assms(1,2) find_set_precondition_def vector_complement_closed vector_mult_closed path_halving_invariant_def vector_inf_closed by auto + have 3: "?r0 \ p0 * ?r0" + by (metis assms(3) dedekind_1 inf.le_iff_sup root_successor_loop top_greatest) + hence "?pxy \ p0 * ?r0 \ ?pxy \ ?p2 * ?r0" + by (metis comp_associative inf.eq_refl inf.sup_right_isotone mult_isotone) + hence 4: "?q1 * ?r0 \ ?q3 * ?r0" + using 2 by (simp add: vector_inf_comp) + have 5: "?q1 * ?q2 \ ?q3" + using 2 by (smt (z3) comp_isotone inf.cobounded1 inf.cobounded2 inf_greatest vector_export_comp) + have "?q1 * ?q2\<^sup>\ * ?r0 = ?q1 * ?r0 \ ?q1 * ?q2 * ?q2\<^sup>\ * ?r0" + by (metis comp_associative semiring.distrib_left star.circ_loop_fixpoint sup_commute) + also have "... \ ?q1 * ?r0 \ ?q3 * ?q2\<^sup>\ * ?r0" + using 5 by (meson mult_left_isotone sup_right_isotone) + also have "... \ ?q3 * ?r0 \ ?q3 * ?q2\<^sup>\ * ?r0" + using 4 sup_left_isotone by blast + also have "... = ?q3 * ?q2\<^sup>\ * ?r0" + by (smt (verit, del_insts) comp_associative semiring.distrib_left star.circ_loop_fixpoint star.circ_transitive_equal star_involutive sup_commute) + finally have 6: "?q1 * ?q2\<^sup>\ * ?r0 \ ?q3 * ?q2\<^sup>\ * ?r0" + . + have "?q1 * (-?pxy \ p0\<^sup>+) * ?pxy \ (?px \ p0) * (-?pxy \ p0\<^sup>+) * ?pxy" + by (meson comp_inf.comp_left_subdist_inf inf.boundedE mult_left_isotone) + also have "... \ (?px \ p0) * (-?pxy \ p0\<^sup>+) * ?py" + by (simp add: mult_right_isotone) + also have "... \ ?px\<^sup>T * (-?pxy \ p0\<^sup>+) * ?py" + proof - + have "?px \ p0 \ ?px\<^sup>T * p0" + using 2 by (simp add: vector_restrict_comp_conv) + also have "... \ ?px\<^sup>T" + by (metis comp_associative conv_dist_comp conv_involutive conv_star_commute mult_right_isotone star.circ_increasing star.circ_transitive_equal) + finally show ?thesis + using mult_left_isotone by auto + qed + also have "... = top * (?px \ -?pxy \ p0\<^sup>+) * ?py" + using 2 by (smt (z3) comp_inf.star_plus conv_dist_inf covector_inf_comp_3 inf_top.right_neutral vector_complement_closed vector_inf_closed) + also have "... \ top * (-?py \ p0\<^sup>+) * ?py" + by (metis comp_inf.comp_isotone comp_isotone inf.cobounded2 inf.eq_refl inf_import_p) + also have "... = top * (-?py \ p0\<^sup>+ \ ?py\<^sup>T) * top" + using 2 by (simp add: comp_associative covector_inf_comp_3) + also have "... = bot" + proof - + have "p0\<^sup>T\<^sup>\ * y \ -(y\<^sup>T * p0\<^sup>\) = p0\<^sup>T\<^sup>\ * y * y\<^sup>T * -p0\<^sup>\" + using 2 by (metis assms(2) bijective_conv_mapping comp_mapping_complement vector_covector vector_export_comp vector_mult_closed) + also have "... \ p0\<^sup>T\<^sup>\ * -p0\<^sup>\" + by (meson assms(2) mult_left_isotone order_refl shunt_bijective) + also have "... \ -p0\<^sup>\" + by (simp add: conv_complement conv_star_commute pp_increasing schroeder_6_p star.circ_transitive_equal) + also have "... \ -p0\<^sup>+" + by (simp add: p_antitone star.left_plus_below_circ) + finally have "-?py \ p0\<^sup>+ \ ?py\<^sup>T = bot" + by (metis comp_inf.p_pp_comp conv_complement conv_dist_comp conv_involutive conv_star_commute p_shunting_swap pp_isotone pseudo_complement_pp regular_closed_p) + thus ?thesis + by simp + qed + finally have 7: "?q1 * (-?pxy \ p0\<^sup>+) * ?pxy = bot" + using le_bot by blast + have "?q2\<^sup>+ \ -?pxy" + using 2 by (smt (z3) comp_isotone complement_conv_sub inf.order_trans inf.sup_right_divisibility inf_commute symmetric_top_closed top_greatest) + hence "?q2\<^sup>+ \ -?pxy \ p0\<^sup>+" + by (simp add: comp_isotone star_isotone) + hence 8: "?q1 * ?q2\<^sup>+ * ?pxy = bot" + using 7 mult_left_isotone mult_right_isotone le_bot by auto + have "?q1 * ?q2\<^sup>+ * ?q3\<^sup>\ = ?q1 * ?q2\<^sup>+ \ ?q1 * ?q2\<^sup>+ * ?q3\<^sup>+" + by (smt (z3) comp_associative star.circ_back_loop_fixpoint star.circ_plus_same sup_commute) + also have "... \ ?q1 * ?q2\<^sup>+ \ ?q1 * ?q2\<^sup>+ * ?pxy" + using 2 by (smt (z3) inf.cobounded1 mult_right_isotone sup_right_isotone vector_inf_comp) + finally have 9: "?q1 * ?q2\<^sup>+ * ?q3\<^sup>\ \ ?q1 * ?q2\<^sup>+" + using 8 by simp + have 10: "?q1 * ?q4 * ?pxy = bot" + proof - + have "?p2 \ p0\<^sup>+" + by (simp add: mult_right_isotone star.circ_increasing) + thus ?thesis + using 7 by (metis mult_left_isotone mult_right_isotone le_bot comp_inf.comp_isotone eq_refl) + qed + have 11: "?q1 * ?q2 * ?pxy = bot" + proof - + have "p0 \ p0\<^sup>+" + by (simp add: star.circ_mult_increasing) + thus ?thesis + using 7 by (metis mult_left_isotone mult_right_isotone le_bot comp_inf.comp_isotone eq_refl) + qed + have 12: "?q2 \ p0 * ?q3\<^sup>\ * ?q2\<^sup>\" + by (smt (verit, del_insts) conv_dist_comp conv_order conv_star_commute inf.coboundedI1 inf.orderE inf.sup_monoid.add_commute path_compression_1b) + have "?q3 * p0 * ?q3\<^sup>\ * ?q2\<^sup>\ = ?q1 * p0 * p0 * ?q3\<^sup>\ * ?q2\<^sup>\" + using 2 vector_inf_comp by auto + also have "... = ?q1 * (?q3 \ ?q4) * ?q3\<^sup>\ * ?q2\<^sup>\" + using 1 by (smt (z3) comp_associative comp_inf.mult_right_dist_sup comp_inf.star_slide inf_top.right_neutral regular_complement_top) + also have "... = ?q1 * ?q3 * ?q3\<^sup>\ * ?q2\<^sup>\ \ ?q1 * ?q4 * ?q3\<^sup>\ * ?q2\<^sup>\" + using mult_left_dist_sup mult_right_dist_sup by auto + also have "... \ ?q1 * ?q3\<^sup>\ * ?q2\<^sup>\ \ ?q1 * ?q4 * ?q3\<^sup>\ * ?q2\<^sup>\" + by (smt (z3) mult_left_isotone mult_left_sub_dist_sup_right sup_left_isotone sup_right_divisibility mult_assoc star.left_plus_below_circ) + also have "... = ?q1 * ?q3\<^sup>\ * ?q2\<^sup>\ \ ?q1 * ?q4 * ?q2\<^sup>\ \ ?q1 * ?q4 * ?q3\<^sup>+ * ?q2\<^sup>\" + by (smt (z3) semiring.combine_common_factor star.circ_back_loop_fixpoint star_plus sup_monoid.add_commute mult_assoc) + also have "... \ ?q1 * ?q3\<^sup>\ * ?q2\<^sup>\ \ ?q1 * ?q4 * ?q2\<^sup>\ \ ?q1 * ?q4 * ?pxy * ?q3\<^sup>\ * ?q2\<^sup>\" + by (smt (verit, ccfv_threshold) comp_isotone inf.sup_right_divisibility inf_commute order.refl semiring.add_left_mono mult_assoc) + also have "... = ?q1 * ?q3\<^sup>\ * ?q2\<^sup>\ \ ?q1 * ?q4 * ?q2\<^sup>\" + using 10 by simp + also have "... = ?q1 * ?q3\<^sup>\ * ?q2\<^sup>\ \ ?q1 * ?q2 * p0 * ?q2\<^sup>\" + using 2 by (smt vector_complement_closed vector_inf_comp mult_assoc) + also have "... = ?q1 * ?q3\<^sup>\ * ?q2\<^sup>\ \ ?q1 * ?q2 * (?q2 \ ?q1) * ?q2\<^sup>\" + using 1 by (smt (z3) comp_associative comp_inf.mult_right_dist_sup comp_inf.star_slide inf_top.right_neutral regular_complement_top) + also have "... = ?q1 * ?q3\<^sup>\ * ?q2\<^sup>\ \ ?q1 * ?q2 * ?q2 * ?q2\<^sup>\ \ ?q1 * ?q2 * ?q1 * ?q2\<^sup>\" + using mult_left_dist_sup mult_right_dist_sup sup_commute sup_left_commute by auto + also have "... \ ?q1 * ?q3\<^sup>\ * ?q2\<^sup>\ \ ?q1 * ?q2 * ?q2 * ?q2\<^sup>\ \ ?q1 * ?q2 * ?pxy * ?q2\<^sup>\" + by (smt (verit, ccfv_threshold) comp_isotone inf.sup_right_divisibility inf_commute order.refl semiring.add_left_mono mult_assoc) + also have "... = ?q1 * ?q3\<^sup>\ * ?q2\<^sup>\ \ ?q1 * ?q2 * ?q2 * ?q2\<^sup>\" + using 11 by simp + also have "... \ ?q1 * ?q3\<^sup>\ * ?q2\<^sup>\ \ ?q1 * ?q2\<^sup>\" + by (smt comp_associative comp_isotone mult_right_isotone star.circ_increasing star.circ_transitive_equal star.left_plus_below_circ sup_right_isotone) + also have "... = ?q1 * ?q3\<^sup>\ * ?q2\<^sup>\" + by (smt (verit, best) comp_associative semiring.distrib_left star.circ_loop_fixpoint star.circ_transitive_equal star_involutive) + finally have 13: "?q3 * p0 * ?q3\<^sup>\ * ?q2\<^sup>\ \ p0 * ?q3\<^sup>\ * ?q2\<^sup>\" + by (meson inf.cobounded2 mult_left_isotone order_lesseq_imp) + hence "?q3 * p0 * ?q3\<^sup>\ * ?q2\<^sup>\ \ ?q2 \ p0 * ?q3\<^sup>\ * ?q2\<^sup>\" + using 12 by simp + hence "?q3\<^sup>\ * ?q2 \ p0 * ?q3\<^sup>\ * ?q2\<^sup>\" + by (simp add: star_left_induct mult_assoc) + hence "?q1 * ?q3\<^sup>\ * ?q2 \ ?q1 * p0 * ?q3\<^sup>\ * ?q2\<^sup>\" + by (simp add: comp_associative mult_right_isotone) + hence "?q1 * ?q3\<^sup>\ * ?q2 \ ?q3\<^sup>+ * ?q2\<^sup>\" + using 2 by (simp add: vector_inf_comp) + hence 14: "?q1 * ?q3\<^sup>\ * ?q2 \ ?q3\<^sup>\ * ?q2\<^sup>\" + using mult_left_isotone order_lesseq_imp star.left_plus_below_circ by blast + have "p0 * ?r0 \ p0 * ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0" + by (metis comp_associative mult_1_right mult_left_isotone mult_right_isotone reflexive_mult_closed star.circ_reflexive) + hence 15: "?r0 \ p0 * ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0" + using 3 dual_order.trans by blast + have "?q3 * p0 * ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0 \ p0 * ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0" + using 13 mult_left_isotone by blast + hence "?q3 * p0 * ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0 \ ?r0 \ p0 * ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0" + using 15 by simp + hence "?q3\<^sup>\ * ?r0 \ p0 * ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0" + by (simp add: star_left_induct mult_assoc) + hence "?q1 * ?q3\<^sup>\ * ?r0 \ ?q1 * p0 * ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0" + by (simp add: comp_associative mult_right_isotone) + hence "?q1 * ?q3\<^sup>\ * ?r0 \ ?q3\<^sup>+ * ?q2\<^sup>\ * ?r0" + using 2 by (simp add: vector_inf_comp) + hence 16: "?q1 * ?q3\<^sup>\ * ?r0 \ ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0" + using mult_left_isotone order_lesseq_imp star.left_plus_below_circ by blast + have "?q1 * ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0 = ?q1 * ?q3\<^sup>\ * ?r0 \ ?q1 * ?q3\<^sup>\ * ?q2\<^sup>+ * ?r0" + by (smt (z3) comp_associative mult_right_dist_sup star.circ_back_loop_fixpoint star.circ_plus_same sup_commute) + also have "... \ ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0 \ ?q1 * ?q3\<^sup>\ * ?q2\<^sup>+ * ?r0" + using 16 sup_left_isotone by blast + also have "... \ ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0 \ ?q3\<^sup>\ * ?q2\<^sup>\ * ?q2\<^sup>\ * ?r0" + using 14 by (smt (z3) inf.eq_refl semiring.distrib_right star.circ_transitive_equal sup.absorb2 sup_monoid.add_commute mult_assoc) + also have "... = ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0" + by (simp add: comp_associative star.circ_transitive_equal) + finally have 17: "?q1 * ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0 \ ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0" + . + have "?r0 \ ?q2\<^sup>\ * ?r0" + using star.circ_loop_fixpoint sup_right_divisibility by auto + also have "... \ ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0" + using comp_associative star.circ_loop_fixpoint sup_right_divisibility by force + also have "... \ ?q2\<^sup>\ * ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0" + using comp_associative star.circ_loop_fixpoint sup_right_divisibility by force + finally have 18: "?r0 \ ?q2\<^sup>\ * ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0" + . + have "p0 * ?q2\<^sup>\ * ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0 = (?q2 \ ?q1) * ?q2\<^sup>\ * ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0" + using 1 by (smt (z3) comp_inf.mult_right_dist_sup comp_inf.star_plus inf_top.right_neutral regular_complement_top) + also have "... = ?q2 * ?q2\<^sup>\ * ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0 \ ?q1 * ?q2\<^sup>\ * ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0" + using mult_right_dist_sup by auto + also have "... \ ?q2\<^sup>\ * ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0 \ ?q1 * ?q2\<^sup>\ * ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0" + by (smt (z3) comp_left_increasing_sup star.circ_loop_fixpoint sup_left_isotone mult_assoc) + also have "... = ?q2\<^sup>\ * ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0 \ ?q1 * ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0 \ ?q1 * ?q2\<^sup>+ * ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0" + by (smt (z3) mult_left_dist_sup semiring.combine_common_factor star.circ_loop_fixpoint sup_monoid.add_commute mult_assoc) + also have "... \ ?q2\<^sup>\ * ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0 \ ?q1 * ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0 \ ?q1 * ?q2\<^sup>+ * ?q2\<^sup>\ * ?r0" + using 9 mult_left_isotone sup_right_isotone by auto + also have "... \ ?q2\<^sup>\ * ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0 \ ?q1 * ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0 \ ?q1 * ?q2\<^sup>\ * ?r0" + by (smt (z3) comp_associative comp_isotone inf.eq_refl semiring.add_right_mono star.circ_transitive_equal star.left_plus_below_circ sup_commute) + also have "... \ ?q2\<^sup>\ * ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0 \ ?q1 * ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0 \ ?q3 * ?q2\<^sup>\ * ?r0" + using 6 sup_right_isotone by blast + also have "... = ?q2\<^sup>\ * ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0 \ ?q3 * ?q2\<^sup>\ * ?r0" + using 17 by (smt (z3) le_iff_sup semiring.combine_common_factor semiring.distrib_right star.circ_loop_fixpoint sup_monoid.add_commute) + also have "... \ ?q2\<^sup>\ * ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0 \ ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0" + by (meson mult_left_isotone star.circ_increasing sup_right_isotone) + also have "... = ?q2\<^sup>\ * ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0" + by (smt (z3) comp_associative star.circ_loop_fixpoint star.circ_transitive_equal star_involutive) + finally have "p0 * ?q2\<^sup>\ * ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0 \ ?r0 \ ?q2\<^sup>\ * ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0" + using 18 sup.boundedI by blast + hence "p0\<^sup>\ * ?r0 \ ?q2\<^sup>\ * ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0" + by (simp add: comp_associative star_left_induct) + also have "... \ ?p\<^sup>\ * ?q3\<^sup>\ * ?q2\<^sup>\ * ?r0" + by (metis mult_left_isotone star.circ_sub_dist sup_commute) + also have "... \ ?p\<^sup>\ * ?p\<^sup>\ * ?q2\<^sup>\ * ?r0" + by (simp add: mult_left_isotone mult_right_isotone star_isotone) + also have "... \ ?p\<^sup>\ * ?p\<^sup>\ * ?p\<^sup>\ * ?r0" + by (metis mult_isotone order.refl star.circ_sub_dist sup_commute) + finally have 19: "p0\<^sup>\ * ?r0 \ ?p\<^sup>\ * ?r0" + by (simp add: star.circ_transitive_equal) + have 20: "?p\<^sup>\ \ p0\<^sup>\" + by (metis star.left_plus_circ star_isotone update_square_ub_plus) + hence 21: "p0\<^sup>\ * ?r0 = ?p\<^sup>\ * ?r0" + using 19 order.antisym mult_left_isotone by auto + have "?p \ 1 = (?q3 \ 1) \ (?q2 \ 1)" + using comp_inf.semiring.distrib_right conv_involutive by auto + also have "... = (?q1 \ 1) \ (?q2 \ 1)" + using assms(3) acyclic_square path_splitting_invariant_def inf.sup_monoid.add_assoc by auto + also have "... = (?pxy \ -?pxy) \ p0 \ 1" + using inf_sup_distrib2 by auto + also have "... = p0 \ 1" + using 1 by (metis inf.sup_monoid.add_commute inf_sup_distrib1 maddux_3_11_pp) + finally show 22: "?p \ 1 = p0 \ 1" + . + have "?p\<^sup>T\<^sup>\ * x \ p0\<^sup>T\<^sup>\ * x" + using 20 by (metis conv_isotone conv_star_commute mult_left_isotone) + hence 23: "?rp \ ?r0" + using 22 comp_inf.mult_left_isotone by auto + have 24: "disjoint_set_forest ?p" + using 1 2 assms(3) disjoint_set_forest_update_square by blast + hence 25: "point ?rp" + using root_point assms(1) by auto + have "?r0 * ?rp\<^sup>T = ?r0 * x\<^sup>T * ?p\<^sup>\ * (?p \ 1)" + by (smt (z3) comp_associative conv_dist_comp conv_dist_inf conv_involutive conv_star_commute inf.sup_monoid.add_commute one_inf_conv root_var star_one star_sup_one wcc_one) + also have "... \ (p0 \ 1) * p0\<^sup>T\<^sup>\ * 1 * ?p\<^sup>\ * (?p \ 1)" + by (smt (z3) assms(1) comp_associative mult_left_isotone mult_right_isotone root_var) + also have "... \ (p0 \ 1) * p0\<^sup>T\<^sup>\ * p0\<^sup>\ * (p0 \ 1)" + using 20 22 comp_isotone by force + also have "... = (p0 \ 1) * p0\<^sup>\ * (p0 \ 1) \ (p0 \ 1) * p0\<^sup>T\<^sup>\ * (p0 \ 1)" + by (simp add: assms(3) cancel_separate_eq sup_monoid.add_commute mult_assoc mult_left_dist_sup semiring.distrib_right) + also have "... = (p0 \ 1) * (p0 \ 1) \ (p0 \ 1) * p0\<^sup>T\<^sup>\ * (p0 \ 1)" + using univalent_root_successors assms(3) by simp + also have "... = (p0 \ 1) * (p0 \ 1) \ (p0 \ 1) * ((p0 \ 1) * p0\<^sup>\)\<^sup>T" + by (smt (z3) comp_associative conv_dist_comp conv_dist_inf conv_star_commute inf.sup_monoid.add_commute one_inf_conv star_one star_sup_one wcc_one) + also have "... = (p0 \ 1) * (p0 \ 1)" + by (metis univalent_root_successors assms(3) conv_dist_inf inf.sup_monoid.add_commute one_inf_conv sup_idem symmetric_one_closed) + also have "... \ 1" + by (simp add: coreflexive_mult_closed) + finally have "?r0 * ?rp\<^sup>T \ 1" + . + hence "?r0 \ 1 * ?rp" + using 25 shunt_bijective by blast + hence 26: "?r0 = ?rp" + using 23 order.antisym by simp + have "?px * ?r0\<^sup>T = ?px * x\<^sup>T * p0\<^sup>\ * (p0 \ 1)" + by (smt (z3) comp_associative conv_dist_comp conv_dist_inf conv_involutive conv_star_commute inf.sup_monoid.add_commute one_inf_conv root_var star_one star_sup_one wcc_one) + also have "... \ p0\<^sup>T\<^sup>\ * 1 * p0\<^sup>\ * (p0 \ 1)" + by (smt (z3) assms(1) comp_associative mult_left_isotone mult_right_isotone root_var) + also have "... = p0\<^sup>\ * (p0 \ 1) \ p0\<^sup>T\<^sup>\ * (p0 \ 1)" + by (simp add: assms(3) cancel_separate_eq sup_monoid.add_commute mult_right_dist_sup) + also have "... = p0\<^sup>\ * (p0 \ 1) \ ((p0 \ 1) * p0\<^sup>\)\<^sup>T" + by (smt (z3) conv_dist_comp conv_dist_inf conv_star_commute inf.sup_monoid.add_commute one_inf_conv star_one star_sup_one wcc_one) + also have "... = p0\<^sup>\ * (p0 \ 1) \ (p0 \ 1)" + by (metis univalent_root_successors assms(3) conv_dist_inf inf.sup_monoid.add_commute one_inf_conv symmetric_one_closed) + also have "... = p0\<^sup>\ * (p0 \ 1)" + by (metis conv_involutive path_compression_1b sup.absorb2 sup_commute) + also have "... \ p0\<^sup>\" + by (simp add: inf.coboundedI1 star.circ_increasing star.circ_mult_upper_bound) + finally have 27: "?px * ?r0\<^sup>T \ p0\<^sup>\" + . + thus 28: "?px \ p0\<^sup>\ * ?r0" + by (simp add: assms(1,3) root_point shunt_bijective) + have 29: "point ?r0" + using root_point assms(1,3) by auto + hence 30: "mapping (?r0\<^sup>T)" + using bijective_conv_mapping by blast + have "?r0 * (?px \ p0) = ?r0 * top * (?px \ p0)" + using 29 by force + also have "... = ?r0 * ?px\<^sup>T * p0" + using 29 by (metis assms(1) covector_inf_comp_3 vector_covector vector_mult_closed) + also have "... = ?r0 * x\<^sup>T * p0\<^sup>\ * p0" + using comp_associative conv_dist_comp conv_star_commute by auto + also have "... \ ?r0 * x\<^sup>T * p0\<^sup>\" + by (simp add: comp_associative mult_right_isotone star.circ_plus_same star.left_plus_below_circ) + also have "... = ?r0 * ?px\<^sup>T" + by (simp add: comp_associative conv_dist_comp conv_star_commute) + also have "... = (?px * ?r0\<^sup>T)\<^sup>T" + by (simp add: conv_dist_comp) + also have "... \ p0\<^sup>T\<^sup>\" + using 27 conv_isotone conv_star_commute by fastforce + finally have "?r0 * (?px \ p0) \ p0\<^sup>T\<^sup>\" + . + hence "?px \ p0 \ ?r0\<^sup>T * p0\<^sup>T\<^sup>\" + using 30 shunt_mapping by auto + hence "?px \ p0 \ p0\<^sup>\ * ?r0 \ ?r0\<^sup>T * p0\<^sup>T\<^sup>\" + using 28 inf.coboundedI2 inf.sup_monoid.add_commute by fastforce + also have "... = p0\<^sup>\ * ?r0 * ?r0\<^sup>T * p0\<^sup>T\<^sup>\" + using 29 by (smt (z3) vector_covector vector_inf_comp vector_mult_closed) + also have "... = ?p\<^sup>\ * ?r0 * ?r0\<^sup>T * ?p\<^sup>T\<^sup>\" + using 21 by (smt comp_associative conv_dist_comp conv_star_commute) + also have "... = ?p\<^sup>\ * ?rp * ?rp\<^sup>T * ?p\<^sup>T\<^sup>\" + using 26 by auto + also have "... \ ?p\<^sup>\ * 1 * ?p\<^sup>T\<^sup>\" + using 25 by (smt (z3) comp_associative mult_left_isotone mult_right_isotone) + finally have 31: "?px \ p0 \ fc ?p" + by auto + have "-?px \ p0 \ ?p" + by (simp add: inf.sup_monoid.add_commute le_supI1 sup_commute) + also have "... \ fc ?p" + using fc_increasing by auto + finally have "p0 \ fc ?p" + using 1 31 by (smt (z3) inf.sup_monoid.add_commute maddux_3_11_pp semiring.add_left_mono sup.orderE sup_commute) + also have "... \ wcc ?p" + using star.circ_sub_dist_3 by auto + finally have 32: "wcc p0 \ wcc ?p" + using wcc_below_wcc by blast + have "?p \ wcc p0" + by (simp add: inf.coboundedI1 inf.sup_monoid.add_commute star.circ_mult_upper_bound star.circ_sub_dist_1) + hence "wcc ?p \ wcc p0" + using wcc_below_wcc by blast + hence "wcc ?p = wcc p0" + using 32 order.antisym by blast + thus "fc ?p = fc p0" + using 24 assms(3) fc_wcc by auto +qed + +lemma path_splitting_invariant_aux: + assumes "path_splitting_invariant p x y p0" + shows "p[[y]] = p0[[y]]" + and "p[[p[[y]]]] = p0[[p0[[y]]]]" + and "p[[p[[p[[y]]]]]] = p0[[p0[[p0[[y]]]]]]" + and "p \ 1 = p0 \ 1" + and "fc p = fc p0" +proof - + let ?p2 = "p0 * p0" + let ?p2t = "?p2\<^sup>T" + let ?px = "p0\<^sup>T\<^sup>\ * x" + let ?py = "-(p0\<^sup>T\<^sup>\ * y)" + let ?pxy = "?px \ ?py" + let ?p = "p0[?pxy\?p2t]" + have "?p[[y]] = p0[[y]]" + apply (rule put_get_different_vector) + using assms find_set_precondition_def vector_complement_closed vector_inf_closed vector_mult_closed path_splitting_invariant_def apply force + by (meson inf.cobounded2 order_lesseq_imp p_antitone_iff path_compression_1b) + thus 1: "p[[y]] = p0[[y]]" + using assms path_splitting_invariant_def by auto + have "?p[[p0[[y]]]] = p0[[p0[[y]]]]" + apply (rule put_get_different_vector) + using assms find_set_precondition_def vector_complement_closed vector_inf_closed vector_mult_closed path_splitting_invariant_def apply force + by (metis comp_isotone inf.boundedE inf.coboundedI2 inf.eq_refl p_antitone_iff selection_closed_id star.circ_increasing) + thus 2: "p[[p[[y]]]] = p0[[p0[[y]]]]" + using 1 assms path_splitting_invariant_def by auto + have "?p[[p0[[p0[[y]]]]]] = p0[[p0[[p0[[y]]]]]]" + apply (rule put_get_different_vector) + using assms find_set_precondition_def vector_complement_closed vector_inf_closed vector_mult_closed path_splitting_invariant_def apply force + by (metis comp_associative comp_isotone conv_dist_comp conv_involutive conv_order inf.coboundedI2 inf.le_iff_sup mult_left_isotone p_antitone_iff p_antitone_inf star.circ_increasing star.circ_transitive_equal) + thus "p[[p[[p[[y]]]]]] = p0[[p0[[p0[[y]]]]]]" + using 2 assms path_splitting_invariant_def by auto + show "p \ 1 = p0 \ 1" + using assms path_splitting_invariant_aux_1(1) path_splitting_invariant_def find_set_precondition_def by auto + show "fc p = fc p0" + using assms path_splitting_invariant_aux_1(2) path_splitting_invariant_def find_set_precondition_def by auto +qed + +lemma path_splitting_1: + "find_set_precondition p0 x \ path_splitting_invariant p0 x x p0" +proof - + assume 1: "find_set_precondition p0 x" + show "path_splitting_invariant p0 x x p0" + proof (unfold path_splitting_invariant_def, intro conjI) + show "find_set_precondition p0 x" + using 1 by simp + show "vector x" "injective x" "surjective x" + using 1 find_set_precondition_def by auto + show "x \ p0\<^sup>T\<^sup>\ * x" + by (simp add: path_compression_1b) + have "(p0 * p0)\<^sup>T\<^sup>\ * x \ p0\<^sup>T\<^sup>\ * x" + by (simp add: conv_dist_comp mult_left_isotone star.circ_square) + thus "p0[p0\<^sup>T\<^sup>\ * x \ - (p0\<^sup>T\<^sup>\ * x)\(p0 * p0)\<^sup>T] = p0" + by (smt (z3) inf.le_iff_sup inf_commute maddux_3_11_pp p_antitone_inf pseudo_complement) + show "univalent p0" "total p0" "acyclic (p0 \ -1)" + using 1 find_set_precondition_def by auto + qed +qed + +lemma path_splitting_2: + "path_splitting_invariant p x y p0 \ y \ p[[y]] \ path_splitting_invariant (p[y\p[[p[[y]]]]]) x (p[[y]]) p0 \ card { z . regular z \ z \ (p[y\p[[p[[y]]]]])\<^sup>T\<^sup>\ * (p[[y]]) } < card { z . regular z \ z \ p\<^sup>T\<^sup>\ * y }" +proof - + let ?py = "p[[y]]" + let ?ppy = "p[[?py]]" + let ?pyppy = "p[y\?ppy]" + let ?p2 = "p0 * p0" + let ?p2t = "?p2\<^sup>T" + let ?p2ts = "?p2t\<^sup>\" + let ?px = "p0\<^sup>T\<^sup>\ * x" + let ?py2 = "-(p0\<^sup>T\<^sup>\ * y)" + let ?pxy = "?px \ ?py2" + let ?p = "p0[?pxy\?p2t]" + let ?pty = "p0\<^sup>T * y" + let ?pt2y = "p0\<^sup>T * p0\<^sup>T * y" + let ?pt2sy = "p0\<^sup>T\<^sup>\ * p0\<^sup>T * p0\<^sup>T * y" + let ?ptpy = "p0\<^sup>T\<^sup>+ * y" + assume 1: "path_splitting_invariant p x y p0 \ y \ ?py" + have 2: "point ?pty \ point ?pt2y" + using 1 by (smt (verit) comp_associative read_injective read_surjective path_splitting_invariant_def) + show "path_splitting_invariant ?pyppy x (p[[y]]) p0 \ card { z . regular z \ z \ ?pyppy\<^sup>T\<^sup>\ * (p[[y]]) } < card { z . regular z \ z \ p\<^sup>T\<^sup>\ * y }" + proof + show "path_splitting_invariant ?pyppy x (p[[y]]) p0" + proof (unfold path_splitting_invariant_def, intro conjI) + show 3: "find_set_precondition ?pyppy x" + proof (unfold find_set_precondition_def, intro conjI) + show "univalent ?pyppy" + using 1 find_set_precondition_def read_injective update_univalent path_splitting_invariant_def by auto + show "total ?pyppy" + using 1 bijective_regular find_set_precondition_def read_surjective update_total path_splitting_invariant_def by force + show "acyclic (?pyppy \ -1)" + apply (rule update_acyclic_3) + using 1 find_set_precondition_def path_splitting_invariant_def apply blast + using 1 2 comp_associative path_splitting_invariant_aux(2) apply force + using 1 path_splitting_invariant_def apply blast + by (metis inf.order_lesseq_imp mult_isotone star.circ_increasing star.circ_square mult_assoc) + show "vector x" "injective x" "surjective x" + using 1 find_set_precondition_def path_splitting_invariant_def by auto + qed + show "vector (p[[y]])" + using 1 comp_associative path_splitting_invariant_def by auto + show "injective (p[[y]])" + using 1 3 read_injective path_splitting_invariant_def find_set_precondition_def by auto + show "surjective (p[[y]])" + using 1 3 read_surjective path_splitting_invariant_def find_set_precondition_def by auto + show "p[[y]] \ ?px" + proof - + have "p[[y]] = p0[[y]]" + using 1 path_splitting_invariant_aux(1) by blast + also have "... \ p0\<^sup>T * ?px" + using 1 path_splitting_invariant_def mult_right_isotone by force + also have "... \ ?px" + by (metis comp_associative mult_left_isotone star.left_plus_below_circ) + finally show ?thesis + . + qed + show "p0[?px \ -(p0\<^sup>T\<^sup>\ * (p[[y]]))\?p2t] = ?pyppy" + proof - + have 4: "p[?px \ -?ptpy \ y\?p2t] = ?pyppy" + proof (cases "y \ -?ptpy") + case True + hence "p[?px \ -?ptpy \ y\?p2t] = p[y\?p2t]" + using 1 inf.absorb2 path_splitting_invariant_def by auto + also have "... = ?pyppy" + using 1 by (metis comp_associative conv_dist_comp path_splitting_invariant_aux(2) path_splitting_invariant_def update_point_get) + finally show ?thesis + . + next + case False + have "vector (-?ptpy)" + using 1 vector_complement_closed vector_mult_closed path_splitting_invariant_def by blast + hence 5: "y \ ?ptpy" + using 1 by (smt (verit, del_insts) False bijective_regular point_in_vector_or_complement regular_closed_star regular_mult_closed total_conv_surjective univalent_conv_injective path_splitting_invariant_def) + hence "?px \ -?ptpy \ y = bot" + by (simp add: inf.coboundedI2 p_antitone pseudo_complement) + hence 6: "p[?px \ -?ptpy \ y\?p2t] = p" + by simp + have 7: "y = root p0 y" + using 1 5 loop_root_2 path_splitting_invariant_def by blast + have "?pyppy = p[y\p0[[p0[[y]]]]]" + using 1 path_splitting_invariant_aux(2) by force + also have "... = p[y\p0[[y]]]" + using 1 7 by (metis root_successor_loop path_splitting_invariant_def) + also have "... = p[y\?py]" + using 1 path_splitting_invariant_aux(1) by force + also have "... = p" + using 1 get_put path_splitting_invariant_def by blast + finally show ?thesis + using 6 by simp + qed + have 8: "-?ptpy = ?py2 \ (-?ptpy \ y)" + proof (rule order.antisym) + have 9: "regular (p0\<^sup>T\<^sup>\ * y) \ regular ?ptpy" + using 1 bijective_regular mapping_conv_bijective pp_dist_star regular_mult_closed path_splitting_invariant_def by auto + have "p0\<^sup>T\<^sup>\ * y \ ?ptpy \ y" + by (simp add: star.circ_loop_fixpoint mult_assoc) + also have "... = ?ptpy \ (-?ptpy \ y)" + using 9 by (metis maddux_3_21_pp) + hence "p0\<^sup>T\<^sup>\ * y \ -?ptpy \ -?ptpy \ y" + using calculation half_shunting sup_commute by auto + thus "-?ptpy \ ?py2 \ (-?ptpy \ y)" + using 9 by (smt (z3) inf.sup_monoid.add_commute shunting_var_p sup.left_commute sup_commute) + have "-(p0\<^sup>T\<^sup>\ * y) \ -?ptpy" + by (simp add: comp_isotone p_antitone star.left_plus_below_circ) + thus "-(p0\<^sup>T\<^sup>\ * y) \ (-?ptpy \ y) \ -?ptpy" + by simp + qed + have "regular ?px" "regular y" + using 1 bijective_regular find_set_precondition_def mapping_regular pp_dist_comp regular_closed_star regular_conv_closed path_splitting_invariant_def by auto + hence 10: "regular (?px \ -?ptpy \ y)" + by auto + have "p0[?px \ -(p0\<^sup>T\<^sup>\ * (p[[y]]))\?p2t] = p0[?px \ -?ptpy\?p2t]" + using 1 by (smt comp_associative path_splitting_invariant_aux(1) star_plus) + also have "... = p0[?pxy \ (?px \ -?ptpy \ y)\?p2t]" + using 8 by (metis comp_inf.semiring.distrib_left inf.sup_monoid.add_assoc) + also have "... = ?p[?px \ -?ptpy \ y\?p2t]" + using 10 by (smt (z3) update_same comp_inf.semiring.mult_not_zero inf.sup_monoid.add_assoc inf.sup_monoid.add_commute) + also have "... = p[?px \ -?ptpy \ y\?p2t]" + using 1 path_splitting_invariant_def by auto + also have "... = ?pyppy" + using 4 by auto + finally show ?thesis + . + qed + show "univalent p0" "total p0" "acyclic (p0 \ -1)" + using 1 path_splitting_invariant_def by auto + qed + let ?s = "{ z . regular z \ z \ p\<^sup>T\<^sup>\ * y }" + let ?t = "{ z . regular z \ z \ ?pyppy\<^sup>T\<^sup>\ * (p[[y]]) }" + have "?pyppy\<^sup>T\<^sup>\ * (p[[y]]) \ p\<^sup>+\<^sup>T\<^sup>\ * (p[[y]])" + using 1 path_splitting_invariant_def update_square_plus conv_order mult_left_isotone star_isotone by force + also have "... = p\<^sup>T\<^sup>\ * p\<^sup>T * y" + by (simp add: conv_plus_commute star.left_plus_circ mult_assoc) + also have "... = p\<^sup>T\<^sup>+ * y" + by (simp add: star_plus) + finally have 11: "?pyppy\<^sup>T\<^sup>\ * (p[[y]]) \ p\<^sup>T\<^sup>+ * y" + . + hence "?pyppy\<^sup>T\<^sup>\ * (p[[y]]) \ p\<^sup>T\<^sup>\ * y" + using mult_left_isotone order_lesseq_imp star.left_plus_below_circ by blast + hence 12: "?t \ ?s" + using order_trans by auto + have 13: "y \ ?s" + using 1 bijective_regular path_compression_1b path_splitting_invariant_def by force + have 14: "\ y \ ?t" + proof + assume "y \ ?t" + hence "y \ ?pyppy\<^sup>T\<^sup>\ * (p[[y]])" + by simp + hence "y \ p\<^sup>T\<^sup>+ * y" + using 11 dual_order.trans by blast + hence "y = root p y" + using 1 find_set_precondition_def loop_root_2 path_splitting_invariant_def by blast + hence "y = ?py" + using 1 by (metis find_set_precondition_def root_successor_loop path_splitting_invariant_def) + thus False + using 1 by simp + qed + show "card ?t < card ?s" + apply (rule psubset_card_mono) + subgoal using finite_regular by simp + subgoal using 12 13 14 by auto + done + qed +qed + +lemma path_splitting_3: + "path_splitting_invariant p x y p0 \ y = p[[y]] \ path_splitting_postcondition p x y p0" +proof - + assume 1: "path_splitting_invariant p x y p0 \ y = p[[y]]" + show "path_splitting_postcondition p x y p0" + proof (unfold path_splitting_postcondition_def path_compression_precondition_def, intro conjI) + show "univalent p" "total p" "acyclic (p \ -1)" + using 1 find_set_precondition_def path_splitting_invariant_def by blast+ + show "vector x" "injective x" "surjective x" + using 1 find_set_precondition_def path_splitting_invariant_def by blast+ + show 2: "vector y" "injective y" "surjective y" + using 1 path_splitting_invariant_def by blast+ + show 3: "p \ 1 = p0 \ 1" + using 1 path_splitting_invariant_aux(4) by blast + show 4: "fc p = fc p0" + using 1 path_splitting_invariant_aux(5) by blast + have "y \ p0\<^sup>T\<^sup>\ * x" + using 1 path_splitting_invariant_def by simp + hence 5: "y * x\<^sup>T \ fc p0" + using 1 by (metis dual_order.trans fc_wcc find_set_precondition_def shunt_bijective star.circ_decompose_11 star_decompose_1 star_outer_increasing path_splitting_invariant_def) + have 6: "y = p0[[y]]" + using 1 path_splitting_invariant_aux(1) by auto + hence "y = root p0 y" + using 2 loop_root by auto + also have "... = root p0 x" + using 1 2 5 find_set_precondition_def path_splitting_invariant_def same_component_same_root by auto + also have "... = root p x" + using 1 3 4 by (metis find_set_precondition_def path_splitting_invariant_def same_root) + finally show "y = root p x" + . + have "p0\<^sup>T\<^sup>\ * y = y" + using 6 order.antisym path_compression_1b star_left_induct_mult_equal by auto + hence 7: "p0[p0\<^sup>T\<^sup>\ * x \ -y\(p0 * p0)\<^sup>T] = p" + using 1 path_splitting_invariant_def by auto + have "(p0 * p0)\<^sup>T * y = y" + using 6 mult_assoc conv_dist_comp by auto + hence "y \ p0 * p0 = y \ p0" + using 2 6 by (metis update_postcondition) + hence 8: "y \ p = y \ p0 * p0" + using 1 2 6 by (smt update_postcondition) + have "p0[p0\<^sup>T\<^sup>\ * x\(p0 * p0)\<^sup>T] = (p0[p0\<^sup>T\<^sup>\ * x \ -y\(p0 * p0)\<^sup>T])[p0\<^sup>T\<^sup>\ * x \ y\(p0 * p0)\<^sup>T]" + using 1 bijective_regular path_splitting_invariant_def update_split by blast + also have "... = p[p0\<^sup>T\<^sup>\ * x \ y\(p0 * p0)\<^sup>T]" + using 7 by simp + also have "... = p" + apply (rule update_same_sub) + using 8 apply simp + apply simp + using 1 bijective_regular inf.absorb2 path_splitting_invariant_def by auto + finally show "p0[p0\<^sup>T\<^sup>\ * x\(p0 * p0)\<^sup>T] = p" + . + qed +qed + +theorem find_path_splitting: + "VARS p t y + [ find_set_precondition p x \ p0 = p ] + y := x; + WHILE y \ p[[y]] + INV { path_splitting_invariant p x y p0 } + VAR { card { z . regular z \ z \ p\<^sup>T\<^sup>\ * y } } + DO t := p[[y]]; + p[y] := p[[p[[y]]]]; + y := t + OD + [ path_splitting_postcondition p x y p0 ]" + apply vcg_tc_simp + apply (fact path_splitting_1) + apply (fact path_splitting_2) + by (fact path_splitting_3) + +end + +section \Verifying Union by Rank\ + +text \ +In this section we verify the union-by-rank operation of disjoint-set forests. +The rank of a node is an upper bound of the height of the subtree rooted at that node. +The rank array of a disjoint-set forest maps each node to its rank. +This can be represented as a homogeneous relation since the possible rank values are $0, \dots, n-1$ where $n$ is the number of nodes of the disjoint-set forest. +\ + +subsection \Peano structures\ + +text \ +Since ranks are natural numbers we start by introducing basic Peano arithmetic. +Numbers are represented as (relational) points. +Constant \Z\ represents the number $0$. +Constant \S\ represents the successor function. +The successor of a number $x$ is obtained by the relational composition \S\<^sup>T * x\. +The composition \S * x\ results in the predecessor of $x$. +\ + +class peano_signature = + fixes Z :: "'a" + fixes S :: "'a" + +text \ +The numbers will be used in arrays, which are represented by homogeneous finite relations. +Such relations can only represent finitely many numbers. +This means that we weaken the Peano axioms, which are usually used to obtain (infinitely many) natural numbers. +Axiom \Z_point\ specifies that $0$ is a number. +Axiom \S_univalent\ specifies that every number has at most one `successor'. +Together with axiom \S_total\, which is added later, this means that every number has exactly one `successor'. +Axiom \S_injective\ specifies that numbers with the same successor are equal. +Axiom \S_star_Z_top\ specifies that every number can be obtained from $0$ by finitely many applications of the successor. +We omit the Peano axiom \S * Z = bot\ which would specify that $0$ is not the successor of any number. +Since only finitely many numbers will be represented, the remaining axioms will model successor modulo $m$ for some $m$ depending on the carrier of the algebra. +That is, the algebra will be able to represent numbers $0, \dots, m-1$ where the successor of $m-1$ is $0$. +\ + +class skra_peano_1 = stone_kleene_relation_algebra + stone_relation_algebra_tarski_consistent + peano_signature + + assumes Z_point: "point Z" + assumes S_univalent: "univalent S" + assumes S_injective: "injective S" + assumes S_star_Z_top: "S\<^sup>T\<^sup>\ * Z = top" +begin + +lemma conv_Z_Z: + "Z\<^sup>T * Z = top" + by (simp add: Z_point point_conv_comp) + +text \Theorem 9.2\ + +lemma Z_below_S_star: + "Z \ S\<^sup>\" +proof - + have "top * Z\<^sup>T \ S\<^sup>T\<^sup>\" + using S_star_Z_top Z_point shunt_bijective by blast + thus ?thesis + using Z_point conv_order conv_star_commute vector_conv_covector by force +qed + +text \Theorem 9.3\ + +lemma S_connected: + "S\<^sup>T\<^sup>\ * S\<^sup>\ = top" + by (metis Z_below_S_star S_star_Z_top mult_left_dist_sup sup.orderE sup_commute top.extremum) + +text \Theorem 9.4\ + +lemma S_star_connex: + "S\<^sup>\ \ S\<^sup>T\<^sup>\ = top" + using S_connected S_univalent cancel_separate_eq sup_commute by auto + +text \Theorem 9.5\ + +lemma Z_sup_conv_S_top: + "Z \ S\<^sup>T * top = top" + using S_star_Z_top star.circ_loop_fixpoint sup_commute by auto + +lemma top_S_sup_conv_Z: + "top * S \ Z\<^sup>T = top" + by (metis S_star_Z_top conv_dist_comp conv_involutive conv_star_commute star.circ_back_loop_fixpoint symmetric_top_closed) + +text \Theorem 9.1\ + +lemma S_inf_1_below_Z: + "S \ 1 \ Z" +proof - + have "(S \ 1) * S\<^sup>T \ S \ 1" + by (metis S_injective conv_dist_comp coreflexive_symmetric inf.boundedI inf.cobounded1 inf.cobounded2 injective_codomain) + hence "(S \ 1) * S\<^sup>T\<^sup>\ \ S \ 1" + using star_right_induct_mult by blast + hence "(S \ 1) * S\<^sup>T\<^sup>\ * Z \ (S \ 1) * Z" + by (simp add: mult_left_isotone) + also have "... \ Z" + by (metis comp_left_subdist_inf inf.boundedE mult_1_left) + finally show ?thesis + using S_star_Z_top inf.order_trans top_right_mult_increasing mult_assoc by auto +qed + +lemma S_inf_1_below_conv_Z: + "S \ 1 \ Z\<^sup>T" + using S_inf_1_below_Z conv_order coreflexive_symmetric by fastforce + +text \ +The successor operation provides a convenient way to compare two natural numbers. +Namely, $k < m$ if $m$ can be reached from $k$ by finitely many applications of the successor, formally \m \ S\<^sup>T\<^sup>\ * k\ or \k \ S\<^sup>\ * m\. +This does not work for numbers modulo $m$ since comparison depends on the chosen representative. +We therefore work with a modified successor relation \S'\, which is a partial function that computes the successor for all numbers except $m-1$. +If $S$ is surjective, the point \M\ representing the greatest number $m-1$ is the predecessor of $0$ under \S\. +If $S$ is not surjective (like for the set of all natural numbers), \M = bot\. +\ + +abbreviation "S' \ S \ -Z\<^sup>T" +abbreviation "M \ S * Z" + +text \Theorem 11.1\ + +lemma M_point_iff_S_surjective: + "point M \ surjective S" +proof + assume 1: "point M" + hence "1 \ Z\<^sup>T * S\<^sup>T * S * Z" + using comp_associative conv_dist_comp surjective_var by auto + hence "Z \ S\<^sup>T * S * Z" + using 1 Z_point bijective_reverse mult_assoc by auto + also have "... \ S\<^sup>T * top" + by (simp add: comp_isotone mult_assoc) + finally have "S\<^sup>T * S\<^sup>T * top \ Z \ S\<^sup>T * top" + using mult_isotone mult_assoc by force + hence "S\<^sup>T\<^sup>\ * Z \ S\<^sup>T * top" + by (simp add: star_left_induct mult_assoc) + thus "surjective S" + by (simp add: S_star_Z_top order.antisym surjective_conv_total) +next + assume "surjective S" + thus "point M" + by (metis S_injective Z_point comp_associative injective_mult_closed) +qed + +text \Theorem 10.1\ + +lemma S'_univalent: + "univalent S'" + by (simp add: S_univalent univalent_inf_closed) + +text \Theorem 10.2\ + +lemma S'_injective: + "injective S'" + by (simp add: S_injective injective_inf_closed) + +text \Theorem 10.9\ + +lemma S'_Z: + "S' * Z = bot" + by (simp add: Z_point covector_vector_comp injective_comp_right_dist_inf) + +text \Theorem 10.4\ + +lemma S'_irreflexive: + "irreflexive S'" + using S_inf_1_below_conv_Z order_lesseq_imp p_shunting_swap pp_increasing by blast + +end + +class skra_peano_2 = skra_peano_1 + + assumes S_total: "total S" +begin + +lemma S_mapping: + "mapping S" + by (simp add: S_total S_univalent) + +text \Theorem 11.2\ + +lemma M_bot_iff_S_not_surjective: + "M \ bot \ surjective S" +proof + assume "M \ bot" + hence "top * S * Z = top" + by (metis S_mapping Z_point bijective_regular comp_associative mapping_regular regular_mult_closed tarski) + hence "Z\<^sup>T \ top * S" + using M_point_iff_S_surjective S_injective Z_point comp_associative injective_mult_closed by auto + thus "surjective S" + using sup.orderE top_S_sup_conv_Z by fastforce +next + assume "surjective S" + thus "M \ bot" + using M_point_iff_S_surjective consistent covector_bot_closed by force +qed + +text \Theorem 11.3\ + +lemma M_point_or_bot: + "point M \ M = bot" + using M_bot_iff_S_not_surjective M_point_iff_S_surjective by blast + +text \Alternative way to express \S'\\ + +text \Theorem 12.1\ + +lemma S'_var: + "S' = S \ -M" +proof - + have "S' = S * (1 \ -Z\<^sup>T)" + by (simp add: Z_point covector_comp_inf vector_conv_compl) + also have "... = S * (1 \ -Z)" + by (metis conv_complement one_inf_conv) + also have "... = S * 1 \ S * -Z" + by (simp add: S_mapping univalent_comp_left_dist_inf) + also have "... = S \ -M" + by (simp add: comp_mapping_complement S_mapping) + finally show ?thesis + . +qed + +text \Special case of just $1$ number\ + +text \Theorem 12.2\ + +lemma M_is_Z_iff_1_is_top: + "M = Z \ 1 = top" +proof + assume "M = Z" + hence "Z = S\<^sup>T * Z" + by (metis S_mapping Z_point order.antisym bijective_reverse inf.eq_refl shunt_mapping) + thus "1 = top" + by (metis S_star_Z_top Z_point inf.eq_refl star_left_induct sup.absorb2 symmetric_top_closed top_le) +next + assume "1 = top" + thus "M = Z" + using S_mapping comp_right_one mult_1_left by auto +qed + +text \Theorem 12.3\ + +lemma S_irreflexive: + assumes "M \ Z" + shows "irreflexive S" +proof - + have "(S \ 1) * S\<^sup>T \ S \ 1" + by (smt (z3) S_injective S_mapping coreflexive_comp_top_inf dual_order.eq_iff inf.cobounded1 inf.sup_monoid.add_commute inf.sup_same_context mult_left_isotone one_inf_conv top_right_mult_increasing total_var) + hence "(S \ 1) * S\<^sup>T\<^sup>\ \ S \ 1" + using star_right_induct_mult by blast + hence "(S \ 1) * S\<^sup>T\<^sup>\ * Z \ (S \ 1) * Z" + by (simp add: mult_left_isotone) + also have "... = M \ Z" + by (simp add: Z_point injective_comp_right_dist_inf) + also have "... = bot" + by (smt (verit, ccfv_threshold) M_point_or_bot assms Z_point bijective_one_closed bijective_regular comp_associative conv_complement coreflexive_comp_top_inf epm_3 inf.sup_monoid.add_commute one_inf_conv regular_mult_closed star.circ_increasing star.circ_zero tarski vector_conv_covector vector_export_comp_unit) + finally have "S \ 1 \ bot" + using S_star_Z_top comp_associative le_bot top_right_mult_increasing by fastforce + thus ?thesis + using le_bot pseudo_complement by blast +qed + +text \ +We show that \S'\ satisfies most properties of \S\. +\ + +lemma M_regular: + "regular M" + using S_mapping Z_point bijective_regular mapping_regular regular_mult_closed by blast + +lemma S'_regular: + "regular S'" + using S_mapping mapping_regular by auto + +text \Theorem 10.3\ + +lemma S'_star_Z_top: + "S'\<^sup>T\<^sup>\ * Z = top" +proof - + have "S\<^sup>T\<^sup>\ * Z = (S' \ (S \ M))\<^sup>T\<^sup>\ * Z" + by (metis M_regular maddux_3_11_pp S'_var) + also have "... \ S'\<^sup>T\<^sup>\ * Z" + proof (cases "M = bot") + case True + thus ?thesis + by simp + next + case False + hence "point M" + using M_point_or_bot by auto + hence "arc (S \ M)" + using S_mapping mapping_inf_point_arc by blast + hence 1: "arc ((S \ M)\<^sup>T)" + using conv_involutive by auto + have 2: "S \ M \ Z\<^sup>T" + by (metis S'_var Z_point bijective_regular conv_complement inf.cobounded2 p_shunting_swap) + have "(S' \ (S \ M))\<^sup>T\<^sup>\ * Z = (S'\<^sup>T \ (S \ M)\<^sup>T)\<^sup>\ * Z" + by (simp add: S'_var conv_dist_sup) + also have "... = (S'\<^sup>T\<^sup>\ * (S \ M)\<^sup>T * S'\<^sup>T\<^sup>\ \ S'\<^sup>T\<^sup>\) * Z" + using 1 star_arc_decompose sup_commute by auto + also have "... = S'\<^sup>T\<^sup>\ * (S \ M)\<^sup>T * S'\<^sup>T\<^sup>\ * Z \ S'\<^sup>T\<^sup>\ * Z" + using mult_right_dist_sup by auto + also have "... \ S'\<^sup>T\<^sup>\ * Z\<^sup>T\<^sup>T * S'\<^sup>T\<^sup>\ * Z \ S'\<^sup>T\<^sup>\ * Z" + using 2 by (meson comp_isotone conv_isotone inf.eq_refl semiring.add_mono) + also have "... \ S'\<^sup>T\<^sup>\ * Z" + by (metis Z_point comp_associative conv_involutive le_supI mult_right_isotone top.extremum) + finally show ?thesis + . + qed + finally show ?thesis + using S_star_Z_top top_le by auto +qed + +text \Theorem 10.5\ + +lemma Z_below_S'_star: + "Z \ S'\<^sup>\" + by (metis S'_star_Z_top Z_point comp_associative comp_right_one conv_order conv_star_commute mult_right_isotone vector_conv_covector) + +text \Theorem 10.6\ + +lemma S'_connected: + "S'\<^sup>T\<^sup>\ * S'\<^sup>\ = top" + by (metis Z_below_S'_star S'_star_Z_top mult_left_dist_sup sup.orderE sup_commute top.extremum) + +text \Theorem 10.7\ + +lemma S'_star_connex: + "S'\<^sup>\ \ S'\<^sup>T\<^sup>\ = top" + using S'_connected S'_univalent cancel_separate_eq sup_commute by auto + +text \Theorem 10.8\ + +lemma Z_sup_conv_S'_top: + "Z \ S'\<^sup>T * top = top" + using S'_star_Z_top star.circ_loop_fixpoint sup_commute by auto + +lemma top_S'_sup_conv_Z: + "top * S' \ Z\<^sup>T = top" + by (metis S'_star_Z_top conv_dist_comp conv_involutive conv_star_commute star.circ_back_loop_fixpoint symmetric_top_closed) + +end + +subsection \Initialising Ranks\ + +text \ +We show that the rank array satisfies three properties which are established/preserved by the union-find operations. +First, every node has a rank, that is, the rank array is a mapping. +Second, the rank of a node is strictly smaller than the rank of its parent, except if the node is a root. +This implies that the rank of a node is an upper bound on the height of its subtree. +Third, the number of roots in the disjoint-set forest (the number of disjoint sets) is not larger than $m-k$ where $m$ is the total number of nodes and $k$ is the maximum rank of any node. +The third property is useful to show that ranks never overflow (exceed $m-1$). +To compare the number of roots and $m-k$ we use the existence of an injective univalent relation between the set of roots and the set of $m-k$ largest numbers, both represented as vectors. +The three properties are captured in \rank_property\. +\ + +class skra_peano_3 = stone_kleene_relation_algebra_tarski_finite_regular + skra_peano_2 +begin + +definition "card_less_eq v w \ \i . injective i \ univalent i \ regular i \ v \ i * w" +definition "rank_property p rank \ mapping rank \ (p \ -1) * rank \ rank * S'\<^sup>+ \ card_less_eq ((p \ 1) * top) (-(S'\<^sup>+ * rank\<^sup>T * top))" + +end + +class skra_peano_4 = stone_kleene_relation_algebra_choose_point_finite_regular + skra_peano_2 +begin + +subclass skra_peano_3 .. + +text \ +The initialisation loop is augmented by setting the rank of each node to $0$. +The resulting rank array satisfies the desired properties explained above. +\ + +theorem init_rank: + "VARS h p x rank + [ True ] + h := top; + WHILE h \ bot + INV { regular h \ vector h \ p \ -h = 1 \ -h \ rank \ -h = Z\<^sup>T \ -h } + VAR { card { x . regular x \ x \ h } } + DO x := choose_point h; + p := make_set p x; + rank[x] := Z; + h[x] := bot + OD + [ p = 1 \ disjoint_set_forest p \ rank = Z\<^sup>T \ rank_property p rank \ h = bot ]" +proof vcg_tc_simp + fix h p rank + let ?x = "choose_point h" + let ?m = "make_set p ?x" + let ?rank = "rank[?x\Z]" + assume 1: "regular h \ vector h \ p \ -h = 1 \ -h \ rank \ -h = Z\<^sup>T \ -h \ h \ bot" + show "vector (-?x \ h) \ + ?m \ (--?x \ -h) = 1 \ (--?x \ -h) \ + ?rank \ (--?x \ -h) = Z\<^sup>T \ (--?x \ -h) \ + card { x . regular x \ x \ -?x \ x \ h } < card { x . regular x \ x \ h }" + proof (intro conjI) + show "vector (-?x \ h)" + using 1 choose_point_point vector_complement_closed vector_inf_closed by blast + have 2: "point ?x \ regular ?x" + using 1 bijective_regular choose_point_point by blast + have 3: "-h \ -?x" + using choose_point_decreasing p_antitone_iff by auto + have 4: "?x \ ?m = ?x * ?x\<^sup>T \ -?x \ ?m = -?x \ p" + using 1 choose_point_point make_set_function make_set_postcondition_def by auto + have "?m \ (--?x \ -h) = (?m \ ?x) \ (?m \ -h)" + using 2 comp_inf.comp_left_dist_sup by auto + also have "... = ?x * ?x\<^sup>T \ (?m \ -?x \ -h)" + using 3 4 by (smt (z3) inf_absorb2 inf_assoc inf_commute) + also have "... = ?x * ?x\<^sup>T \ (1 \ -h)" + using 1 3 4 inf.absorb2 inf.sup_monoid.add_assoc inf_commute by auto + also have "... = (1 \ ?x) \ (1 \ -h)" + using 2 by (metis inf.cobounded2 inf.sup_same_context one_inf_conv vector_covector) + also have "... = 1 \ (--?x \ -h)" + using 2 comp_inf.semiring.distrib_left by auto + finally show "?m \ (--?x \ -h) = 1 \ (--?x \ -h)" + . + have 5: "?x \ ?rank = ?x \ Z\<^sup>T \ -?x \ ?rank = -?x \ rank" + by (smt (z3) inf_commute order_refl update_inf_different update_inf_same) + have "?rank \ (--?x \ -h) = (?rank \ ?x) \ (?rank \ -h)" + using 2 comp_inf.comp_left_dist_sup by auto + also have "... = (?x \ Z\<^sup>T) \ (?rank \ -?x \ -h)" + using 3 5 by (smt (z3) inf_absorb2 inf_assoc inf_commute) + also have "... = (Z\<^sup>T \ ?x) \ (Z\<^sup>T \ -h)" + using 1 3 5 inf.absorb2 inf.sup_monoid.add_assoc inf_commute by auto + also have "... = Z\<^sup>T \ (--?x \ -h)" + using 2 comp_inf.semiring.distrib_left by auto + finally show "?rank \ (--?x \ -h) = Z\<^sup>T \ (--?x \ -h)" + . + have 5: "\ ?x \ -?x" + using 1 2 by (metis comp_commute_below_diversity conv_order inf.cobounded2 inf_absorb2 pseudo_complement strict_order_var top.extremum) + have 6: "?x \ h" + using 1 by (metis choose_point_decreasing) + show "card { x . regular x \ x \ -?x \ x \ h } < card { x . regular x \ x \ h }" + apply (rule psubset_card_mono) + using finite_regular apply simp + using 2 5 6 by auto + qed +next + show "rank_property 1 (Z\<^sup>T)" + proof (unfold rank_property_def, intro conjI) + show "univalent (Z\<^sup>T)" "total (Z\<^sup>T)" + using Z_point surjective_conv_total by auto + show "(1 \ -1) * (Z\<^sup>T) \ (Z\<^sup>T) * S'\<^sup>+" + by simp + have "top \ 1 * - (S'\<^sup>+ * Z * top)" + by (simp add: S'_Z comp_associative star_simulation_right_equal) + thus "card_less_eq ((1 \ 1) * top) (-(S'\<^sup>+ * Z\<^sup>T\<^sup>T * top))" + by (metis conv_involutive inf.idem mapping_one_closed regular_one_closed card_less_eq_def bijective_one_closed) + qed +qed + +end + +subsection \Union by Rank\ + +text \ +We show that path compression and union-by-rank preserve the rank property. +\ + +context stone_kleene_relation_algebra_tarski_finite_regular +begin + +lemma union_sets_1_swap: + assumes "union_sets_precondition p0 x y" + and "path_compression_postcondition p1 x r p0" + and "path_compression_postcondition p2 y s p1" + shows "union_sets_postcondition (p2[s\r]) x y p0" +proof (unfold union_sets_postcondition_def union_sets_precondition_def, intro conjI) + let ?p = "p2[s\r]" + have 1: "disjoint_set_forest p1 \ point r \ r = root p1 x \ p1 \ 1 = p0 \ 1 \ fc p1 = fc p0" + using assms(2) path_compression_precondition_def path_compression_postcondition_def by auto + have 2: "disjoint_set_forest p2 \ point s \ s = root p2 y \ p2 \ 1 = p1 \ 1 \ fc p2 = fc p1" + using assms(3) path_compression_precondition_def path_compression_postcondition_def by auto + hence 3: "fc p2 = fc p0" + using 1 by simp + show 4: "univalent ?p" + using 1 2 update_univalent by blast + show "total ?p" + using 1 2 bijective_regular update_total by blast + show "acyclic (?p \ -1)" + proof (cases "r = s") + case True + thus ?thesis + using 2 update_acyclic_5 by fastforce + next + case False + hence "bot = s \ r" + using 1 2 distinct_points inf_commute by blast + also have "... = s \ p1\<^sup>T\<^sup>\ * r" + using 1 by (smt root_transitive_successor_loop) + also have "... = s \ p2\<^sup>T\<^sup>\ * r" + using 1 2 by (smt (z3) inf_assoc inf_commute same_root) + finally have "r \ p2\<^sup>\ * s = bot" + using schroeder_1 conv_star_commute inf.sup_monoid.add_commute by fastforce + thus ?thesis + using 1 2 update_acyclic_4 by blast + qed + show "vector x" + using assms(1) by (simp add: union_sets_precondition_def) + show "injective x" + using assms(1) by (simp add: union_sets_precondition_def) + show "surjective x" + using assms(1) by (simp add: union_sets_precondition_def) + show "vector y" + using assms(1) by (simp add: union_sets_precondition_def) + show "injective y" + using assms(1) by (simp add: union_sets_precondition_def) + show "surjective y" + using assms(1) by (simp add: union_sets_precondition_def) + show "fc ?p = wcc (p0 \ x * y\<^sup>T)" + proof (rule order.antisym) + have "s = p2[[s]]" + using 2 by (metis root_successor_loop) + hence "s * s\<^sup>T \ p2\<^sup>T" + using 2 eq_refl shunt_bijective by blast + hence "s * s\<^sup>T \ p2" + using 2 conv_order coreflexive_symmetric by fastforce + hence "s \ p2 * s" + using 2 shunt_bijective by blast + hence 5: "p2[[s]] \ s" + using 2 shunt_mapping by blast + have "s \ p2 \ s * (top \ s\<^sup>T * p2)" + using 2 by (metis dedekind_1) + also have "... = s * s\<^sup>T * p2" + by (simp add: mult_assoc) + also have "... \ s * s\<^sup>T" + using 5 by (metis comp_associative conv_dist_comp conv_involutive conv_order mult_right_isotone) + also have "... \ 1" + using 2 by blast + finally have 6: "s \ p2 \ 1" + by simp + have "p0 \ wcc p0" + by (simp add: star.circ_sub_dist_1) + also have "... = wcc p2" + using 3 by (simp add: star_decompose_1) + also have 7: "... \ wcc ?p" + proof - + have "wcc p2 = wcc ((-s \ p2) \ (s \ p2))" + using 2 by (metis bijective_regular inf.sup_monoid.add_commute maddux_3_11_pp) + also have "... \ wcc ((-s \ p2) \ 1)" + using 6 wcc_isotone sup_right_isotone by simp + also have "... = wcc (-s \ p2)" + using wcc_with_loops by simp + also have "... \ wcc ?p" + using wcc_isotone sup_ge2 by blast + finally show ?thesis + by simp + qed + finally have 8: "p0 \ wcc ?p" + by force + have "s \ p2\<^sup>T\<^sup>\ * y" + using 2 by (metis inf_le1) + hence 9: "s * y\<^sup>T \ p2\<^sup>T\<^sup>\" + using assms(1) shunt_bijective union_sets_precondition_def by blast + hence "y * s\<^sup>T \ p2\<^sup>\" + using conv_dist_comp conv_order conv_star_commute by force + also have "... \ wcc p2" + by (simp add: star.circ_sub_dist) + also have "... \ wcc ?p" + using 7 by simp + finally have 10: "y * s\<^sup>T \ wcc ?p" + by simp + have 11: "s * r\<^sup>T \ wcc ?p" + using 1 2 star.circ_sub_dist_1 sup_assoc vector_covector by auto + have "r \ p1\<^sup>T\<^sup>\ * x" + using 1 by (metis inf_le1) + hence 12: "r * x\<^sup>T \ p1\<^sup>T\<^sup>\" + using assms(1) shunt_bijective union_sets_precondition_def by blast + also have "... \ wcc p1" + using star_isotone sup_ge2 by blast + also have "... = wcc p2" + using 2 by (simp add: star_decompose_1) + also have "... \ wcc ?p" + using 7 by simp + finally have 13: "r * x\<^sup>T \ wcc ?p" + by simp + have "x \ x * r\<^sup>T * r \ y \ y * s\<^sup>T * s" + using 1 2 shunt_bijective by blast + hence "y * x\<^sup>T \ y * s\<^sup>T * s * (x * r\<^sup>T * r)\<^sup>T" + using comp_isotone conv_isotone by blast + also have "... = y * s\<^sup>T * s * r\<^sup>T * r * x\<^sup>T" + by (simp add: comp_associative conv_dist_comp) + also have "... \ wcc ?p * (s * r\<^sup>T) * (r * x\<^sup>T)" + using 10 by (metis mult_left_isotone mult_assoc) + also have "... \ wcc ?p * wcc ?p * (r * x\<^sup>T)" + using 11 by (metis mult_left_isotone mult_right_isotone) + also have "... \ wcc ?p * wcc ?p * wcc ?p" + using 13 by (metis mult_right_isotone) + also have "... = wcc ?p" + by (simp add: star.circ_transitive_equal) + finally have "x * y\<^sup>T \ wcc ?p" + by (metis conv_dist_comp conv_involutive conv_order wcc_equivalence) + hence "p0 \ x * y\<^sup>T \ wcc ?p" + using 8 by simp + hence "wcc (p0 \ x * y\<^sup>T) \ wcc ?p" + using wcc_below_wcc by simp + thus "wcc (p0 \ x * y\<^sup>T) \ fc ?p" + using 4 fc_wcc by simp + have "-s \ p2 \ wcc p2" + by (simp add: inf.coboundedI2 star.circ_sub_dist_1) + also have "... = wcc p0" + using 3 by (simp add: star_decompose_1) + also have "... \ wcc (p0 \ y * x\<^sup>T)" + by (simp add: wcc_isotone) + finally have 14: "-s \ p2 \ wcc (p0 \ y * x\<^sup>T)" + by simp + have "s * y\<^sup>T \ wcc p2" + using 9 inf.order_trans star.circ_sub_dist sup_commute by fastforce + also have "... = wcc p0" + using 1 2 by (simp add: star_decompose_1) + also have "... \ wcc (p0 \ y * x\<^sup>T)" + by (simp add: wcc_isotone) + finally have 15: "s * y\<^sup>T \ wcc (p0 \ y * x\<^sup>T)" + by simp + have 16: "y * x\<^sup>T \ wcc (p0 \ y * x\<^sup>T)" + using le_supE star.circ_sub_dist_1 by blast + have "x * r\<^sup>T \ p1\<^sup>\" + using 12 conv_dist_comp conv_order conv_star_commute by fastforce + also have "... \ wcc p1" + using star.circ_sub_dist sup_commute by fastforce + also have "... = wcc p0" + using 1 by (simp add: star_decompose_1) + also have "... \ wcc (p0 \ y * x\<^sup>T)" + by (simp add: wcc_isotone) + finally have 17: "x * r\<^sup>T \ wcc (p0 \ y * x\<^sup>T)" + by simp + have "r \ r * x\<^sup>T * x \ s \ s * y\<^sup>T * y" + using assms(1) shunt_bijective union_sets_precondition_def by blast + hence "s * r\<^sup>T \ s * y\<^sup>T * y * (r * x\<^sup>T * x)\<^sup>T" + using comp_isotone conv_isotone by blast + also have "... = s * y\<^sup>T * y * x\<^sup>T * x * r\<^sup>T" + by (simp add: comp_associative conv_dist_comp) + also have "... \ wcc (p0 \ y * x\<^sup>T) * (y * x\<^sup>T) * (x * r\<^sup>T)" + using 15 by (metis mult_left_isotone mult_assoc) + also have "... \ wcc (p0 \ y * x\<^sup>T) * wcc (p0 \ y * x\<^sup>T) * (x * r\<^sup>T)" + using 16 by (metis mult_left_isotone mult_right_isotone) + also have "... \ wcc (p0 \ y * x\<^sup>T) * wcc (p0 \ y * x\<^sup>T) * wcc (p0 \ y * x\<^sup>T)" + using 17 by (metis mult_right_isotone) + also have "... = wcc (p0 \ y * x\<^sup>T)" + by (simp add: star.circ_transitive_equal) + finally have "?p \ wcc (p0 \ y * x\<^sup>T)" + using 1 2 14 vector_covector by auto + hence "wcc ?p \ wcc (p0 \ y * x\<^sup>T)" + using wcc_below_wcc by blast + also have "... = wcc (p0 \ x * y\<^sup>T)" + using conv_dist_comp conv_dist_sup sup_assoc sup_commute by auto + finally show "fc ?p \ wcc (p0 \ x * y\<^sup>T)" + using 4 fc_wcc by simp + qed +qed + +lemma union_sets_1_skip: + assumes "union_sets_precondition p0 x y" + and "path_compression_postcondition p1 x r p0" + and "path_compression_postcondition p2 y r p1" + shows "union_sets_postcondition p2 x y p0" +proof (unfold union_sets_postcondition_def union_sets_precondition_def, intro conjI) + have 1: "point r \ r = root p1 x \ fc p1 = fc p0 \ disjoint_set_forest p2 \ r = root p2 y \ fc p2 = fc p1" + using assms(2,3) path_compression_precondition_def path_compression_postcondition_def by auto + thus "univalent p2" "total p2" "acyclic (p2 \ -1)" + by auto + show "vector x" "injective x" "surjective x" "vector y" "injective y" "surjective y" + using assms(1) union_sets_precondition_def by auto + have "r \ p1\<^sup>T\<^sup>\ * x" + using 1 by (metis inf_le1) + hence "r * x\<^sup>T \ p1\<^sup>T\<^sup>\" + using assms(1) shunt_bijective union_sets_precondition_def by blast + hence 2: "x * r\<^sup>T \ p1\<^sup>\" + using conv_dist_comp conv_order conv_star_commute by force + have "r \ p2\<^sup>T\<^sup>\ * y" + using 1 by (metis inf_le1) + hence 3: "r * y\<^sup>T \ p2\<^sup>T\<^sup>\" + using assms(1) shunt_bijective union_sets_precondition_def by blast + have "x * y\<^sup>T \ x * r\<^sup>T * r * y\<^sup>T" + using 1 mult_left_isotone shunt_bijective by blast + also have "... \ p1\<^sup>\ * p2\<^sup>T\<^sup>\" + using 2 3 by (metis comp_associative comp_isotone) + also have "... \ wcc p0" + using 1 by (metis star.circ_mult_upper_bound star_decompose_1 star_isotone sup_ge2 star.circ_sub_dist) + finally show "fc p2 = wcc (p0 \ x * y\<^sup>T)" + using 1 by (smt (z3) fc_star star_decompose_1 sup_absorb1 wcc_sup_wcc star.circ_sub_dist_3 sup_commute wcc_equivalence) +qed + +end + +context skra_peano_3 +begin + +lemma path_compression_preserves_rank_property: + assumes "path_compression_postcondition p x y p0" + and "disjoint_set_forest p0" + and "rank_property p0 rank" + shows "rank_property p rank" +proof (unfold rank_property_def, intro conjI) + let ?px = "p0\<^sup>T\<^sup>\ * x" + have 1: "point y" + using assms(1,2) path_compression_postcondition_def path_compression_precondition_def root_point by auto + have 2: "vector ?px" + using assms(1) comp_associative path_compression_postcondition_def path_compression_precondition_def by auto + have "root p0 x = root p x" + by (smt (verit) assms(1,2) path_compression_postcondition_def path_compression_precondition_def same_root) + hence "root p0 x = y" + using assms(1) path_compression_postcondition_def path_compression_precondition_def by auto + hence "?px \ p0\<^sup>\ * y" + by (meson assms(1,2) path_splitting_invariant_aux_1(3) path_compression_precondition_def path_compression_postcondition_def) + hence "?px * y\<^sup>T \ p0\<^sup>\" + using 1 shunt_bijective by blast + hence "?px \ y\<^sup>T \ p0\<^sup>\" + using 1 2 by (simp add: vector_covector) + also have "... = (p0 \ -1)\<^sup>+ \ 1" + using reachable_without_loops star_left_unfold_equal sup_commute by fastforce + finally have 3: "?px \ y\<^sup>T \ -1 \ (p0 \ -1)\<^sup>+" + using half_shunting by blast + have "p0[?px\y] = p" + using assms(1) path_compression_postcondition_def by simp + hence "(p \ -1) * rank = (?px \ y\<^sup>T \ -1) * rank \ (-?px \ p0 \ -1) * rank" + using inf_sup_distrib2 mult_right_dist_sup by force + also have "... \ (?px \ y\<^sup>T \ -1) * rank \ (p0 \ -1) * rank" + by (meson comp_inf.mult_semi_associative le_infE mult_left_isotone sup_right_isotone) + also have "... \ (?px \ y\<^sup>T \ -1) * rank \ rank * S'\<^sup>+" + using assms(3) rank_property_def sup_right_isotone by auto + also have "... \ (p0 \ -1)\<^sup>+ * rank \ rank * S'\<^sup>+" + using 3 mult_left_isotone sup_left_isotone by blast + also have "... \ rank * S'\<^sup>+" + proof - + have "(p0 \ -1)\<^sup>\ * rank \ rank * S'\<^sup>\" + using assms(3) rank_property_def star_simulation_left star.left_plus_circ by fastforce + hence "(p0 \ -1)\<^sup>+ * rank \ (p0 \ -1) * rank * S'\<^sup>\" + by (simp add: comp_associative mult_right_isotone) + also have "... \ rank * S'\<^sup>+" + by (smt (z3) assms(3) rank_property_def comp_associative comp_left_subdist_inf inf.boundedE inf.sup_right_divisibility star.circ_transitive_equal) + finally show ?thesis + by simp + qed + finally show "(p \ - 1) * rank \ rank * S'\<^sup>+" + . + show "univalent rank" "total rank" + using rank_property_def assms(3) by auto + show "card_less_eq ((p \ 1) * top) (- (S'\<^sup>+ * rank\<^sup>T * top))" + using assms(1,3) path_compression_postcondition_def rank_property_def by auto +qed + +theorem union_sets_by_rank: + "VARS p r s rank + [ union_sets_precondition p x y \ rank_property p rank \ p0 = p ] + r := find_set p x; + p := path_compression p x r; + s := find_set p y; + p := path_compression p y s; + IF r \ s THEN + IF rank[[r]] \ S'\<^sup>+ * (rank[[s]]) THEN + p[r] := s + ELSE + p[s] := r; + IF rank[[r]] = rank[[s]] THEN + rank[r] := S'\<^sup>T * (rank[[r]]) + ELSE + SKIP + FI + FI + ELSE + SKIP + FI + [ union_sets_postcondition p x y p0 \ rank_property p rank ]" +proof vcg_tc_simp + fix rank + let ?r = "find_set p0 x" + let ?p1 = "path_compression p0 x ?r" + let ?s = "find_set ?p1 y" + let ?p2 = "path_compression ?p1 y ?s" + let ?p5 = "path_compression ?p1 y ?r" + let ?rr = "rank[[?r]]" + let ?rs = "rank[[?s]]" + let ?rank = "rank[?r\S'\<^sup>T * ?rs]" + let ?p3 = "?p2[?r\?s]" + let ?p4 = "?p2[?s\?r]" + assume 1: "union_sets_precondition p0 x y \ rank_property p0 rank" + hence 2: "path_compression_postcondition ?p1 x ?r p0" + using find_set_function find_set_postcondition_def find_set_precondition_def path_compression_function path_compression_precondition_def union_sets_precondition_def by auto + hence 3: "path_compression_postcondition ?p2 y ?s ?p1" + using 1 find_set_function find_set_postcondition_def find_set_precondition_def path_compression_function path_compression_precondition_def union_sets_precondition_def path_compression_postcondition_def by meson + have "rank_property ?p1 rank" + using 1 2 path_compression_preserves_rank_property union_sets_precondition_def by blast + hence 4: "rank_property ?p2 rank" + using 1 2 3 by (meson path_compression_preserves_rank_property path_compression_postcondition_def path_compression_precondition_def) + have 5: "point ?r" "point ?s" + using 2 3 path_compression_postcondition_def path_compression_precondition_def by auto + hence 6: "point ?rr" "point ?rs" + using 1 comp_associative read_injective read_surjective rank_property_def by auto + have "top \ S'\<^sup>\ \ S'\<^sup>+\<^sup>T" + by (metis S'_star_connex conv_dist_comp conv_star_commute eq_refl star.circ_reflexive star_left_unfold_equal star_simulation_right_equal sup.orderE sup_monoid.add_assoc) + hence 7: "-S'\<^sup>+\<^sup>T \ S'\<^sup>\" + by (metis comp_inf.case_split_left comp_inf.star.circ_plus_one comp_inf.star.circ_sup_2 half_shunting) + show "(?r \ ?s \ (?rr \ S'\<^sup>+ * ?rs \ union_sets_postcondition ?p3 x y p0 \ rank_property ?p3 rank) \ + (\ ?rr \ S'\<^sup>+ * ?rs \ ((?rr = ?rs \ union_sets_postcondition ?p4 x y p0 \ rank_property ?p4 ?rank) \ + (?rr \ ?rs \ union_sets_postcondition ?p4 x y p0 \ rank_property ?p4 rank)))) \ + (?r = ?s \ union_sets_postcondition ?p5 x y p0 \ rank_property ?p5 rank)" + proof + show "?r \ ?s \ (?rr \ S'\<^sup>+ * ?rs \ union_sets_postcondition ?p3 x y p0 \ rank_property ?p3 rank) \ + (\ ?rr \ S'\<^sup>+ * ?rs \ ((?rr = ?rs \ union_sets_postcondition ?p4 x y p0 \ rank_property ?p4 ?rank) \ + (?rr \ ?rs \ union_sets_postcondition ?p4 x y p0 \ rank_property ?p4 rank)))" + proof + assume 8: "?r \ ?s" + show "(?rr \ S'\<^sup>+ * ?rs \ union_sets_postcondition ?p3 x y p0 \ rank_property ?p3 rank) \ + (\ ?rr \ S'\<^sup>+ * ?rs \ ((?rr = ?rs \ union_sets_postcondition ?p4 x y p0 \ rank_property ?p4 ?rank) \ + (?rr \ ?rs \ union_sets_postcondition ?p4 x y p0 \ rank_property ?p4 rank)))" + proof + show "?rr \ S'\<^sup>+ * ?rs \ union_sets_postcondition ?p3 x y p0 \ rank_property ?p3 rank" + proof + assume 9: "?rr \ S'\<^sup>+ * ?rs" + show "union_sets_postcondition ?p3 x y p0 \ rank_property ?p3 rank" + proof + show "union_sets_postcondition ?p3 x y p0" + using 1 2 3 by (simp add: union_sets_1) + show "rank_property ?p3 rank" + proof (unfold rank_property_def, intro conjI) + show "univalent rank" "total rank" + using 1 rank_property_def by auto + have "?s \ -?r" + using 5 8 by (meson order.antisym bijective_regular point_in_vector_or_complement point_in_vector_or_complement_2) + hence "?r \ ?s\<^sup>T \ 1 = bot" + by (metis (full_types) bot_least inf.left_commute inf.sup_monoid.add_commute one_inf_conv pseudo_complement) + hence "?p3 \ 1 \ ?p2" + by (smt half_shunting inf.cobounded2 pseudo_complement regular_one_closed semiring.add_mono sup_commute) + hence "(?p3 \ 1) * top \ (?p2 \ 1) * top" + by (simp add: mult_left_isotone) + thus "card_less_eq ((?p3 \ 1) * top) (-(S'\<^sup>+ * rank\<^sup>T * top))" + using 4 by (meson rank_property_def card_less_eq_def order_trans) + have "(?p3 \ -1) * rank = (?r \ ?s\<^sup>T \ -1) * rank \ (-?r \ ?p2 \ -1) * rank" + using comp_inf.semiring.distrib_right mult_right_dist_sup by auto + also have "... \ (?r \ ?s\<^sup>T \ -1) * rank \ (?p2 \ -1) * rank" + using comp_inf.mult_semi_associative mult_left_isotone sup_right_isotone by auto + also have "... \ (?r \ ?s\<^sup>T \ -1) * rank \ rank * S'\<^sup>+" + using 4 sup_right_isotone rank_property_def by blast + also have "... \ (?r \ ?s\<^sup>T) * rank \ rank * S'\<^sup>+" + using inf_le1 mult_left_isotone sup_left_isotone by blast + also have "... = ?r * ?s\<^sup>T * rank \ rank * S'\<^sup>+" + using 5 by (simp add: vector_covector) + also have "... = rank * S'\<^sup>+" + proof - + have "rank\<^sup>T * ?r \ S'\<^sup>+ * rank\<^sup>T * ?s" + using 9 comp_associative by auto + hence "?r \ rank * S'\<^sup>+ * rank\<^sup>T * ?s" + using 4 shunt_mapping comp_associative rank_property_def by auto + hence "?r * ?s\<^sup>T \ rank * S'\<^sup>+ * rank\<^sup>T" + using 5 shunt_bijective by blast + hence "?r * ?s\<^sup>T * rank \ rank * S'\<^sup>+" + using 4 shunt_bijective rank_property_def mapping_conv_bijective by auto + thus ?thesis + using sup_absorb2 by blast + qed + finally show "(?p3 \ -1) * rank \ rank * S'\<^sup>+" + . + qed + qed + qed + show "\ ?rr \ S'\<^sup>+ * ?rs \ ((?rr = ?rs \ union_sets_postcondition ?p4 x y p0 \ rank_property ?p4 ?rank) \ + (?rr \ ?rs \ union_sets_postcondition ?p4 x y p0 \ rank_property ?p4 rank))" + proof + assume "\ ?rr \ S'\<^sup>+ * ?rs" + hence "?rr \ -(S'\<^sup>+ * ?rs)" + using 6 by (meson point_in_vector_or_complement S'_regular bijective_regular regular_closed_star regular_mult_closed vector_mult_closed) + also have "... = -S'\<^sup>+ * ?rs" + using 6 comp_bijective_complement by simp + finally have "?rs \ -S'\<^sup>+\<^sup>T * ?rr" + using 6 by (metis bijective_reverse conv_complement) + also have "... \ S'\<^sup>\ * ?rr" + using 7 by (simp add: mult_left_isotone) + also have "... = S'\<^sup>+ * ?rr \ ?rr" + using star.circ_loop_fixpoint mult_assoc by auto + finally have 10: "?rs \ -?rr \ S'\<^sup>+ * ?rr" + using half_shunting by blast + show "((?rr = ?rs \ union_sets_postcondition ?p4 x y p0 \ rank_property ?p4 ?rank) \ + (?rr \ ?rs \ union_sets_postcondition ?p4 x y p0 \ rank_property ?p4 rank))" + proof + show "?rr = ?rs \ union_sets_postcondition ?p4 x y p0 \ rank_property ?p4 ?rank" + proof + assume 11: "?rr = ?rs" + show "union_sets_postcondition ?p4 x y p0 \ rank_property ?p4 ?rank" + proof + show "union_sets_postcondition ?p4 x y p0" + using 1 2 3 by (simp add: union_sets_1_swap) + show "rank_property ?p4 ?rank" + proof (unfold rank_property_def, intro conjI) + show "univalent ?rank" + using 4 5 6 by (meson S'_univalent read_injective update_univalent rank_property_def) + have "card_less_eq ((?p2 \ 1) * top) (-(S'\<^sup>+ * rank\<^sup>T * top))" + using 4 rank_property_def by blast + from this obtain i where 12: "injective i \ univalent i \ regular i \ (?p2 \ 1) * top \ i * -(S'\<^sup>+ * rank\<^sup>T * top)" + using card_less_eq_def by blast + let ?i = "(i[?s\i[[i * ?rr]]])[i * ?rr\i[[?s]]]" + have 13: "?i = (i * ?rr \ ?s\<^sup>T * i) \ (-(i * ?rr) \ ?s \ ?rr\<^sup>T * i\<^sup>T * i) \ (-(i * ?rr) \ -?s \ i)" + by (smt (z3) conv_dist_comp conv_involutive inf.sup_monoid.add_assoc inf_sup_distrib1 sup_assoc) + have 14: "injective ?i" + apply (rule update_injective_swap) + subgoal using 12 by simp + subgoal using 5 by simp + subgoal using 6 12 injective_mult_closed by simp + subgoal using 5 comp_associative by simp + done + have 15: "univalent ?i" + apply (rule update_univalent_swap) + subgoal using 12 by simp + subgoal using 5 by simp + subgoal using 5 by simp + subgoal using 6 12 injective_mult_closed by simp + subgoal using 5 comp_associative by simp + done + have 16: "regular ?i" + using 5 6 12 by (smt (z3) bijective_regular p_dist_inf p_dist_sup pp_dist_comp regular_closed_inf regular_conv_closed) + have 17: "regular (i * ?rr)" + using 6 12 bijective_regular regular_mult_closed by blast + have 18: "find_set_precondition ?p1 y" + using 2 3 find_set_precondition_def path_compression_postcondition_def path_compression_precondition_def by blast + hence "?s = root ?p1 y" + by (meson find_set_function find_set_postcondition_def) + also have "... = root ?p2 y" + using 3 18 by (smt (z3) find_set_precondition_def path_compression_postcondition_def path_compression_precondition_def same_root) + also have "... \ (?p2 \ 1) * top" + by simp + also have "... \ i * -(S'\<^sup>+ * rank\<^sup>T * top)" + using 12 by simp + finally have 19: "?s \ i * -(S'\<^sup>+ * rank\<^sup>T * top)" + . + have "(?p4 \ 1) * top \ ?i * -(S'\<^sup>+ * ?rank\<^sup>T * top)" + proof - + have "?r \ -?s" + using 5 8 by (meson order.antisym bijective_regular point_in_vector_or_complement point_in_vector_or_complement_2) + hence "?s \ ?r\<^sup>T \ 1 = bot" + by (metis (full_types) bot_least inf.left_commute inf.sup_monoid.add_commute one_inf_conv pseudo_complement) + hence "?p4 \ 1 \ -?s \ ?p2" + by (smt (z3) bot_least comp_inf.semiring.distrib_left inf.cobounded2 inf.sup_monoid.add_commute le_supI) + hence "(?p4 \ 1) * top \ (-?s \ ?p2 \ 1) * top" + by (simp add: mult_left_isotone) + also have "... = -?s \ (?p2 \ 1) * top" + using 5 inf_assoc vector_complement_closed vector_inf_comp by auto + also have "... = (i * ?rr \ -?s \ (?p2 \ 1) * top) \ (-(i * ?rr) \ -?s \ (?p2 \ 1) * top)" + using 17 by (smt (z3) comp_inf.star_plus inf_sup_distrib2 inf_top.right_neutral regular_complement_top) + also have "... \ ?i * (-(S'\<^sup>+ * ?rank\<^sup>T * top))" + proof (rule sup_least) + have "?rank\<^sup>T * top = (?r \ (S'\<^sup>T * ?rs)\<^sup>T)\<^sup>T * top \ (-?r \ rank)\<^sup>T * top" + using conv_dist_sup mult_right_dist_sup by auto + also have "... = (?r\<^sup>T \ S'\<^sup>T * ?rs) * top \ (-?r\<^sup>T \ rank\<^sup>T) * top" + using conv_complement conv_dist_inf conv_involutive by auto + also have "... = S'\<^sup>T * ?rs * (?r \ top) \ (-?r\<^sup>T \ rank\<^sup>T) * top" + using 5 by (smt (z3) covector_inf_comp_3 inf_commute) + also have "... = S'\<^sup>T * ?rs * (?r \ top) \ rank\<^sup>T * (-?r \ top)" + using 5 by (smt (z3) conv_complement vector_complement_closed covector_inf_comp_3 inf_commute) + also have "... = S'\<^sup>T * ?rs * ?r \ rank\<^sup>T * -?r" + by simp + also have "... \ S'\<^sup>T * ?rs * ?r \ rank\<^sup>T * top" + using mult_right_isotone sup_right_isotone by force + also have "... \ S'\<^sup>T * ?rs \ rank\<^sup>T * top" + using 5 6 by (metis inf.eq_refl mult_assoc) + finally have "S'\<^sup>+ * ?rank\<^sup>T * top \ S'\<^sup>+ * S'\<^sup>T * ?rs \ S'\<^sup>+ * rank\<^sup>T * top" + by (smt comp_associative mult_left_dist_sup mult_right_isotone) + also have "... = S'\<^sup>\ * (S' * S'\<^sup>T) * ?rs \ S'\<^sup>+ * rank\<^sup>T * top" + by (smt star_plus mult_assoc) + also have "... \ S'\<^sup>\ * ?rs \ S'\<^sup>+ * rank\<^sup>T * top" + by (metis S'_injective comp_right_one mult_left_isotone mult_right_isotone sup_left_isotone) + also have "... = ?rs \ S'\<^sup>+ * ?rs \ S'\<^sup>+ * rank\<^sup>T * top" + using comp_associative star.circ_loop_fixpoint sup_commute by fastforce + also have "... = ?rs \ S'\<^sup>+ * rank\<^sup>T * top" + by (smt (verit, del_insts) comp_associative mult_left_dist_sup sup.orderE sup_assoc sup_commute top.extremum) + finally have 20: "S'\<^sup>+ * ?rank\<^sup>T * top \ ?rs \ S'\<^sup>+ * rank\<^sup>T * top" + . + have "?s * ?s\<^sup>T = (?s \ i * -(S'\<^sup>+ * rank\<^sup>T * top)) * ?s\<^sup>T" + using 19 inf.orderE by fastforce + also have "... = (?s \ i * -(S'\<^sup>+ * rank\<^sup>T * top)) * top \ ?s\<^sup>T" + using 5 by (smt (z3) covector_comp_inf vector_conv_covector vector_covector vector_top_closed) + also have "... = ?s \ i * -(S'\<^sup>+ * rank\<^sup>T * top) * top \ ?s\<^sup>T" + using 5 vector_inf_comp by auto + also have "... \ 1 \ i * -(S'\<^sup>+ * rank\<^sup>T * top) * top" + using 5 by (smt (verit, ccfv_SIG) inf.cobounded1 inf.cobounded2 inf_greatest order_trans vector_covector) + also have "... = 1 \ i * -(S'\<^sup>+ * rank\<^sup>T * top)" + using comp_associative vector_complement_closed vector_top_closed by auto + also have "... \ 1 \ i * -(S'\<^sup>+ * rank\<^sup>T)" + by (meson comp_inf.mult_right_isotone mult_right_isotone p_antitone top_right_mult_increasing) + also have "... \ 1 \ i * S'\<^sup>\\<^sup>T * rank\<^sup>T" + proof - + have "S'\<^sup>\\<^sup>T * rank\<^sup>T \ S'\<^sup>+ * rank\<^sup>T = (S'\<^sup>T\<^sup>\ \ S'\<^sup>+) * rank\<^sup>T" + by (simp add: conv_star_commute mult_right_dist_sup) + also have "... = (S'\<^sup>T\<^sup>\ \ S'\<^sup>\) * rank\<^sup>T" + by (smt (z3) comp_associative semiring.distrib_right star.circ_loop_fixpoint sup.left_commute sup_commute sup_idem) + also have "... = top * rank\<^sup>T" + by (simp add: S'_star_connex sup_commute) + also have "... = top" + using 4 rank_property_def total_conv_surjective by blast + finally have "-(S'\<^sup>+ * rank\<^sup>T) \ S'\<^sup>\\<^sup>T * rank\<^sup>T" + by (metis half_shunting inf.idem top_greatest) + thus ?thesis + using comp_associative inf.sup_right_isotone mult_right_isotone by auto + qed + also have "... = 1 \ rank * S'\<^sup>\ * i\<^sup>T" + by (metis comp_associative conv_dist_comp conv_involutive one_inf_conv) + also have "... \ rank * S'\<^sup>\ * i\<^sup>T" + by simp + finally have "?s \ rank * S'\<^sup>\ * i\<^sup>T * ?s" + using 5 shunt_bijective by auto + hence "?rs \ S'\<^sup>\ * i\<^sup>T * ?s" + using 4 shunt_mapping comp_associative rank_property_def by auto + hence "?s * (i * ?rr \ -?s \ (?p2 \ 1) * top) \ ?s * (i * S'\<^sup>\ * i\<^sup>T * ?s \ -?s \ (?p2 \ 1) * top)" + using 11 comp_associative comp_inf.mult_left_isotone comp_isotone inf.eq_refl by auto + also have "... = ?s * ((i * S'\<^sup>+ * i\<^sup>T * ?s \ i * i\<^sup>T * ?s) \ -?s \ (?p2 \ 1) * top)" + by (metis comp_associative mult_left_dist_sup star.circ_loop_fixpoint) + also have "... \ ?s * ((i * S'\<^sup>+ * i\<^sup>T * ?s \ 1 * ?s) \ -?s \ (?p2 \ 1) * top)" + using 12 by (metis mult_left_isotone sup_right_isotone comp_inf.mult_left_isotone mult_right_isotone) + also have "... = ?s * (i * S'\<^sup>+ * i\<^sup>T * ?s \ -?s \ (?p2 \ 1) * top)" + using comp_inf.comp_right_dist_sup by simp + also have "... \ ?s * (i * S'\<^sup>+ * i\<^sup>T * ?s \ (?p2 \ 1) * top)" + using comp_inf.mult_left_isotone inf.cobounded1 mult_right_isotone by blast + also have "... \ ?s * (i * S'\<^sup>+ * i\<^sup>T * ?s \ i * -(S'\<^sup>+ * rank\<^sup>T * top))" + using 12 comp_inf.mult_right_isotone mult_right_isotone by auto + also have "... = ?s * (i * S'\<^sup>+ * i\<^sup>T * ?s \ i) * -(S'\<^sup>+ * rank\<^sup>T * top)" + using 5 by (simp add: comp_associative vector_inf_comp) + also have "... = (?s \ (i * S'\<^sup>+ * i\<^sup>T * ?s)\<^sup>T) * i * -(S'\<^sup>+ * rank\<^sup>T * top)" + using 5 covector_inf_comp_3 mult_assoc by auto + also have "... = (?s \ ?s\<^sup>T * i * S'\<^sup>+\<^sup>T * i\<^sup>T) * i * -(S'\<^sup>+ * rank\<^sup>T * top)" + using conv_dist_comp conv_involutive mult_assoc by auto + also have "... = (?s \ ?s\<^sup>T) * i * S'\<^sup>+\<^sup>T * i\<^sup>T * i * -(S'\<^sup>+ * rank\<^sup>T * top)" + using 5 vector_inf_comp by auto + also have "... \ i * S'\<^sup>+\<^sup>T * i\<^sup>T * i * -(S'\<^sup>+ * rank\<^sup>T * top)" + using 5 by (metis point_antisymmetric mult_left_isotone mult_left_one) + also have "... \ i * S'\<^sup>+\<^sup>T * -(S'\<^sup>+ * rank\<^sup>T * top)" + using 12 by (smt mult_left_isotone mult_right_isotone mult_assoc comp_right_one) + also have "... \ i * -(S'\<^sup>\ * rank\<^sup>T * top)" + proof - + have "S'\<^sup>+ * S'\<^sup>\ * rank\<^sup>T * top \ S'\<^sup>+ * rank\<^sup>T * top" + by (simp add: comp_associative star.circ_transitive_equal) + hence "S'\<^sup>+\<^sup>T * -(S'\<^sup>+ * rank\<^sup>T * top) \ -(S'\<^sup>\ * rank\<^sup>T * top)" + by (smt (verit, ccfv_SIG) comp_associative conv_complement_sub_leq mult_right_isotone order.trans p_antitone) + thus ?thesis + by (simp add: comp_associative mult_right_isotone) + qed + also have "... \ i * -(S'\<^sup>+ * ?rank\<^sup>T * top)" + proof - + have "S'\<^sup>+ * ?rank\<^sup>T * top \ ?rs \ S'\<^sup>+ * rank\<^sup>T * top" + using 20 by simp + also have "... \ rank\<^sup>T * top \ S'\<^sup>+ * rank\<^sup>T * top" + using mult_right_isotone sup_left_isotone top.extremum by blast + also have "... = S'\<^sup>\ * rank\<^sup>T * top" + using comp_associative star.circ_loop_fixpoint sup_commute by auto + finally show ?thesis + using mult_right_isotone p_antitone by blast + qed + finally have "?s * (i * ?rr \ -?s \ (?p2 \ 1) * top) \ i * -(S'\<^sup>+ * ?rank\<^sup>T * top)" + . + hence "i * ?rr \ -?s \ (?p2 \ 1) * top \ ?s\<^sup>T * i * -(S'\<^sup>+ * ?rank\<^sup>T * top)" + using 5 shunt_mapping bijective_conv_mapping mult_assoc by auto + hence "i * ?rr \ -?s \ (?p2 \ 1) * top \ i * ?rr \ ?s\<^sup>T * i * -(S'\<^sup>+ * ?rank\<^sup>T * top)" + by (simp add: inf.sup_monoid.add_assoc) + also have "... = (i * ?rr \ ?s\<^sup>T * i) * -(S'\<^sup>+ * ?rank\<^sup>T * top)" + using 6 vector_inf_comp vector_mult_closed by simp + also have "... \ ?i * -(S'\<^sup>+ * ?rank\<^sup>T * top)" + using 13 comp_left_increasing_sup sup_assoc by auto + finally show "i * ?rr \ -?s \ (?p2 \ 1) * top \ ?i * -(S'\<^sup>+ * ?rank\<^sup>T * top)" + . + have "-(i * ?rr) \ (?p2 \ 1) * top \ -(i * ?rr) \ i * -(S'\<^sup>+ * rank\<^sup>T * top)" + using 12 inf.sup_right_isotone by auto + also have "... \ -(i * ?rr) \ i * -(?rs \ S'\<^sup>+ * rank\<^sup>T * top)" + proof - + have 21: "regular (?rs \ S'\<^sup>+ * rank\<^sup>T * top)" + using 4 6 rank_property_def mapping_regular S'_regular pp_dist_star regular_conv_closed regular_mult_closed bijective_regular regular_closed_sup by auto + have "?rs \ S'\<^sup>+ * rank\<^sup>T * top \ S'\<^sup>+ * rank\<^sup>T * top \ ?rr" + using 11 by simp + hence "(?rs \ S'\<^sup>+ * rank\<^sup>T * top) \ -(S'\<^sup>+ * rank\<^sup>T * top) \ ?rr" + using half_shunting sup_commute by auto + hence "-(S'\<^sup>+ * rank\<^sup>T * top) \ -(?rs \ S'\<^sup>+ * rank\<^sup>T * top) \ ?rr" + using 21 by (metis inf.sup_monoid.add_commute shunting_var_p sup_commute) + hence "i * -(S'\<^sup>+ * rank\<^sup>T * top) \ i * -(?rs \ S'\<^sup>+ * rank\<^sup>T * top) \ i * ?rr" + by (metis mult_left_dist_sup mult_right_isotone) + hence "-(i * ?rr) \ i * -(S'\<^sup>+ * rank\<^sup>T * top) \ i * -(?rs \ S'\<^sup>+ * rank\<^sup>T * top)" + using half_shunting inf.sup_monoid.add_commute by fastforce + thus ?thesis + using inf.le_sup_iff by blast + qed + also have "... \ -(i * ?rr) \ i * -(S'\<^sup>+ * ?rank\<^sup>T * top)" + using 20 by (meson comp_inf.mult_right_isotone mult_right_isotone p_antitone) + finally have "-(i * ?rr) \ -?s \ (?p2 \ 1) * top \ -(i * ?rr) \ -?s \ i * -(S'\<^sup>+ * ?rank\<^sup>T * top)" + by (smt (z3) inf.boundedI inf.cobounded1 inf.coboundedI2 inf.sup_monoid.add_assoc inf.sup_monoid.add_commute) + also have "... \ ?i * (-(S'\<^sup>+ * ?rank\<^sup>T * top))" + using 5 6 13 by (smt (z3) sup_commute vector_complement_closed vector_inf_comp vector_mult_closed comp_left_increasing_sup) + finally show "-(i * ?rr) \ -?s \ (?p2 \ 1) * top \ ?i * -(S'\<^sup>+ * ?rank\<^sup>T * top)" + . + qed + finally show ?thesis + . + qed + thus "card_less_eq ((?p4 \ 1) * top) (-(S'\<^sup>+ * ?rank\<^sup>T * top))" + using 14 15 16 card_less_eq_def by auto + have "?s \ i * -(S'\<^sup>+ * rank\<^sup>T * top)" + using 19 by simp + also have "... \ i * -(S'\<^sup>+ * ?rr)" + using mult_right_isotone p_antitone top.extremum mult_assoc by auto + also have "... = i * -S'\<^sup>+ * ?rr" + using 6 comp_bijective_complement mult_assoc by fastforce + finally have "?rr \ -S'\<^sup>T\<^sup>+ * i\<^sup>T * ?s" + using 5 6 by (metis conv_complement conv_dist_comp conv_plus_commute bijective_reverse) + also have "... \ S'\<^sup>\ * i\<^sup>T * ?s" + using 7 conv_plus_commute mult_left_isotone by auto + finally have 22: "?rr \ S'\<^sup>\ * i\<^sup>T * ?s" + . + have "?r = root ?p1 x" + using 2 path_compression_postcondition_def path_compression_precondition_def by blast + also have "... = root ?p2 x" + using 3 18 by (smt (z3) find_set_precondition_def path_compression_postcondition_def path_compression_precondition_def same_root) + also have "... \ (?p2 \ 1) * top" + by simp + also have "... \ i * -(S'\<^sup>+ * rank\<^sup>T * top)" + using 12 by simp + also have "... \ i * -(S'\<^sup>+ * ?rr)" + using mult_right_isotone p_antitone top.extremum mult_assoc by auto + also have "... = i * -S'\<^sup>+ * ?rr" + using 6 comp_bijective_complement mult_assoc by fastforce + finally have "?rr \ -S'\<^sup>T\<^sup>+ * i\<^sup>T * ?r" + using 5 6 by (metis conv_complement conv_dist_comp conv_plus_commute bijective_reverse) + also have "... \ S'\<^sup>\ * i\<^sup>T * ?r" + using 7 conv_plus_commute mult_left_isotone by auto + finally have "?rr \ S'\<^sup>\ * i\<^sup>T * ?r" + . + hence "?rr \ S'\<^sup>\ * i\<^sup>T * ?r \ S'\<^sup>\ * i\<^sup>T * ?s" + using 22 inf.boundedI by blast + also have "... = (S'\<^sup>+ * i\<^sup>T * ?r \ i\<^sup>T * ?r) \ S'\<^sup>\ * i\<^sup>T * ?s" + by (simp add: star.circ_loop_fixpoint mult_assoc) + also have "... \ S'\<^sup>+ * i\<^sup>T * ?r \ (i\<^sup>T * ?r \ S'\<^sup>\ * i\<^sup>T * ?s)" + by (metis comp_inf.mult_right_dist_sup eq_refl inf.cobounded1 semiring.add_mono) + also have "... \ S' * top \ (i\<^sup>T * ?r \ S'\<^sup>\ * i\<^sup>T * ?s)" + using comp_associative mult_right_isotone sup_left_isotone top.extremum by auto + also have "... = S' * top \ (i\<^sup>T * ?r \ (S'\<^sup>+ * i\<^sup>T * ?s \ i\<^sup>T * ?s))" + by (simp add: star.circ_loop_fixpoint mult_assoc) + also have "... \ S' * top \ S'\<^sup>+ * i\<^sup>T * ?s \ (i\<^sup>T * ?r \ i\<^sup>T * ?s)" + by (smt (z3) comp_inf.semiring.distrib_left inf.sup_right_divisibility star.circ_loop_fixpoint sup_assoc sup_commute sup_inf_distrib1) + also have "... \ S' * top \ (i\<^sup>T * ?r \ i\<^sup>T * ?s)" + by (metis comp_associative mult_right_isotone order.refl sup.orderE top.extremum) + also have "... = S' * top \ i\<^sup>T * (?r \ ?s)" + using 12 conv_involutive univalent_comp_left_dist_inf by auto + also have "... = S' * top" + using 5 8 distinct_points by auto + finally have "top \ ?rr\<^sup>T * S' * top" + using 6 by (smt conv_involutive shunt_mapping bijective_conv_mapping mult_assoc) + hence "surjective (S'\<^sup>T * ?rs)" + using 6 11 by (smt conv_dist_comp conv_involutive point_conv_comp top_le) + thus "total ?rank" + using 4 5 bijective_regular update_total rank_property_def by blast + show "(?p4 \ -1) * ?rank \ ?rank * S'\<^sup>+" + proof - + have 23: "univalent ?p2" + using 3 path_compression_postcondition_def path_compression_precondition_def by blast + have 24: "?r \ (?p4 \ -1) * ?rank \ ?s\<^sup>T * rank * S' * S'\<^sup>+" + proof - + have "?r \ (?p4 \ -1) * ?rank = (?r \ ?p4 \ -1) * ?rank" + using 5 vector_complement_closed vector_inf_comp inf_assoc by auto + also have "... = (?r \ -?s \ ?p4 \ -1) * ?rank" + using 5 8 by (smt (z3) order.antisym bijective_regular point_in_vector_or_complement point_in_vector_or_complement_2 inf_absorb1) + also have "... = (?r \ -?s \ ?p2 \ -1) * ?rank" + by (simp add: inf.left_commute inf.sup_monoid.add_commute inf_sup_distrib1 inf_assoc) + also have "... \ (?r \ ?p2 \ -1) * ?rank" + using inf.sup_left_isotone inf_le1 mult_left_isotone by blast + also have "... = bot" + proof - + have "?r = root ?p1 x" + using 2 path_compression_postcondition_def path_compression_precondition_def by blast + also have "... = root ?p2 x" + using 3 18 by (smt (z3) find_set_precondition_def path_compression_postcondition_def path_compression_precondition_def same_root) + also have "... \ (?p2 \ 1) * top" + by simp + finally have "?r \ ?p2 \ (?p2 \ 1) * top \ ?p2" + using inf.sup_left_isotone by blast + also have "... \ (?p2 \ 1) * (?p2 \ 1)\<^sup>T * ?p2" + by (smt (z3) comp_associative comp_inf.star_plus dedekind_1 inf_top_right order_lesseq_imp) + also have "... = (?p2 \ 1) * (?p2 \ 1) * ?p2" + using coreflexive_symmetric by force + also have "... \ (?p2 \ 1) * ?p2" + by (metis coreflexive_comp_top_inf inf.cobounded2 mult_left_isotone) + also have "... \ ?p2 \ 1" + by (smt 23 comp_inf.mult_semi_associative conv_dist_comp conv_dist_inf conv_order equivalence_one_closed inf.absorb1 inf.sup_monoid.add_assoc injective_codomain) + also have "... \ 1" + by simp + finally have "?r \ ?p2 \ 1" + . + thus ?thesis + by (metis pseudo_complement regular_one_closed semiring.mult_not_zero) + qed + finally show ?thesis + using bot_least le_bot by blast + qed + have 25: "-?r \ (?p4 \ -1) * ?rank \ rank * S'\<^sup>+" + proof - + have "-?r \ (?p4 \ -1) * ?rank = (-?r \ ?p4 \ -1) * ?rank" + using 5 vector_complement_closed vector_inf_comp inf_assoc by auto + also have "... = (-?r \ (?s \ -?s) \ ?p4 \ -1) * ?rank" + using 5 bijective_regular inf_top_right regular_complement_top by auto + also have "... = (-?r \ ?s \ ?p4 \ -1) * ?rank \ (-?r \ -?s \ ?p4 \ -1) * ?rank" + by (smt (z3) inf_sup_distrib1 inf_sup_distrib2 mult_right_dist_sup) + also have "... = (-?r \ ?s \ ?r\<^sup>T \ -1) * ?rank \ (-?r \ -?s \ ?p2 \ -1) * ?rank" + using 5 by (smt (z3) bijective_regular comp_inf.comp_left_dist_sup inf_assoc inf_commute inf_top_right mult_right_dist_sup regular_complement_top) + also have "... \ (?s \ ?r\<^sup>T \ -1) * ?rank \ (-?s \ ?p2 \ -1) * ?rank" + by (smt (z3) comp_inf.semiring.distrib_left inf.cobounded2 inf.sup_monoid.add_assoc mult_left_isotone mult_right_dist_sup) + also have "... \ (?s \ ?r\<^sup>T) * ?rank \ (?p2 \ -1) * ?rank" + by (smt (z3) inf.cobounded1 inf.cobounded2 inf.sup_monoid.add_assoc mult_left_isotone semiring.add_mono) + also have "... = ?s * (?r \ ?rank) \ (?p2 \ -1) * ?rank" + using 5 by (simp add: covector_inf_comp_3) + also have "... = ?s * (?r \ (S'\<^sup>T * ?rs)\<^sup>T) \ (?p2 \ -1) * ?rank" + using inf_commute update_inf_same mult_assoc by force + also have "... = ?s * (?r \ ?s\<^sup>T * rank * S') \ (?p2 \ -1) * ?rank" + using comp_associative conv_dist_comp conv_involutive by auto + also have "... \ ?s * ?s\<^sup>T * rank * S' \ (?p2 \ -1) * ?rank" + using comp_associative inf.cobounded2 mult_right_isotone semiring.add_right_mono by auto + also have "... \ 1 * rank * S' \ (?p2 \ -1) * ?rank" + using 5 by (meson mult_left_isotone order.refl semiring.add_mono) + also have "... = rank * S' \ (?p2 \ -1) * (?r \ (S'\<^sup>T * ?rr)\<^sup>T) \ (?p2 \ -1) * (-?r \ rank)" + using 11 comp_associative mult_1_left mult_left_dist_sup sup_assoc by auto + also have "... \ rank * S' \ (?p2 \ -1) * (?r \ ?r\<^sup>T * rank * S') \ (?p2 \ -1) * rank" + using comp_associative conv_dist_comp conv_involutive inf.cobounded1 inf.sup_monoid.add_commute mult_right_isotone semiring.add_left_mono by auto + also have "... = rank * S' \ (?p2 \ -1) * (?r \ ?r\<^sup>T) * rank * S' \ (?p2 \ -1) * rank" + using 5 comp_associative vector_inf_comp by auto + also have "... \ rank * S' \ (?p2 \ -1) * rank * S' \ (?p2 \ -1) * rank" + using 5 by (metis point_antisymmetric mult_left_isotone mult_right_isotone sup_left_isotone sup_right_isotone comp_right_one) + also have "... \ rank * S' \ rank * S'\<^sup>+ * S' \ (?p2 \ -1) * rank" + using 4 by (metis rank_property_def mult_left_isotone sup_left_isotone sup_right_isotone) + also have "... \ rank * S' \ rank * S'\<^sup>+ * S' \ rank * S'\<^sup>+" + using 4 by (metis rank_property_def sup_right_isotone) + also have "... \ rank * S'\<^sup>+" + using comp_associative eq_refl le_sup_iff mult_right_isotone star.circ_mult_increasing star.circ_plus_same star.left_plus_below_circ by auto + finally show ?thesis + . + qed + have "(?p4 \ -1) * ?rank = (?r \ (?p4 \ -1) * ?rank) \ (-?r \ (?p4 \ -1) * ?rank)" + using 5 by (smt (verit, ccfv_threshold) bijective_regular inf_commute inf_sup_distrib2 inf_top_right regular_complement_top) + also have "... \ (?r \ ?s\<^sup>T * rank * S' * S'\<^sup>+) \ (-?r \ rank * S'\<^sup>+)" + using 24 25 by (meson inf.boundedI inf.cobounded1 semiring.add_mono) + also have "... = (?r \ ?s\<^sup>T * rank * S') * S'\<^sup>+ \ (-?r \ rank) * S'\<^sup>+" + using 5 vector_complement_closed vector_inf_comp by auto + also have "... = ?rank * S'\<^sup>+" + using conv_dist_comp mult_right_dist_sup by auto + finally show ?thesis + . + qed + qed + qed + qed + show "?rr \ ?rs \ union_sets_postcondition ?p4 x y p0 \ rank_property ?p4 rank" + proof + assume "?rr \ ?rs" + hence "?rs \ ?rr = bot" + using 6 by (meson bijective_regular dual_order.eq_iff point_in_vector_or_complement point_in_vector_or_complement_2 pseudo_complement) + hence 26: "?rs \ S'\<^sup>+ * ?rr" + using 10 le_iff_inf pseudo_complement by auto + show "union_sets_postcondition ?p4 x y p0 \ rank_property ?p4 rank" + proof + show "union_sets_postcondition ?p4 x y p0" + using 1 2 3 by (simp add: union_sets_1_swap) + show "rank_property ?p4 rank" + proof (unfold rank_property_def, intro conjI) + show "univalent rank" "total rank" + using 1 rank_property_def by auto + have "?r \ -?s" + using 5 8 by (meson order.antisym bijective_regular point_in_vector_or_complement point_in_vector_or_complement_2) + hence "?s \ ?r\<^sup>T \ 1 = bot" + by (metis (full_types) bot_least inf.left_commute inf.sup_monoid.add_commute one_inf_conv pseudo_complement) + hence "?p4 \ 1 \ ?p2" + by (smt half_shunting inf.cobounded2 pseudo_complement regular_one_closed semiring.add_mono sup_commute) + hence "(?p4 \ 1) * top \ (?p2 \ 1) * top" + by (simp add: mult_left_isotone) + thus "card_less_eq ((?p4 \ 1) * top) (-(S'\<^sup>+ * rank\<^sup>T * top))" + using 4 by (meson rank_property_def card_less_eq_def order_trans) + have "(?p4 \ -1) * rank = (?s \ ?r\<^sup>T \ -1) * rank \ (-?s \ ?p2 \ -1) * rank" + using comp_inf.semiring.distrib_right mult_right_dist_sup by auto + also have "... \ (?s \ ?r\<^sup>T \ -1) * rank \ (?p2 \ -1) * rank" + using comp_inf.mult_semi_associative mult_left_isotone sup_right_isotone by auto + also have "... \ (?s \ ?r\<^sup>T \ -1) * rank \ rank * S'\<^sup>+" + using 4 sup_right_isotone rank_property_def by blast + also have "... \ (?s \ ?r\<^sup>T) * rank \ rank * S'\<^sup>+" + using inf_le1 mult_left_isotone sup_left_isotone by blast + also have "... = ?s * ?r\<^sup>T * rank \ rank * S'\<^sup>+" + using 5 by (simp add: vector_covector) + also have "... = rank * S'\<^sup>+" + proof - + have "rank\<^sup>T * ?s \ S'\<^sup>+ * rank\<^sup>T * ?r" + using 26 comp_associative by auto + hence "?s \ rank * S'\<^sup>+ * rank\<^sup>T * ?r" + using 4 shunt_mapping comp_associative rank_property_def by auto + hence "?s * ?r\<^sup>T \ rank * S'\<^sup>+ * rank\<^sup>T" + using 5 shunt_bijective by blast + hence "?s * ?r\<^sup>T * rank \ rank * S'\<^sup>+" + using 4 shunt_bijective rank_property_def mapping_conv_bijective by auto + thus ?thesis + using sup_absorb2 by blast + qed + finally show "(?p4 \ -1) * rank \ rank * S'\<^sup>+" + . + qed + qed + qed + qed + qed + qed + qed + show "?r = ?s \ union_sets_postcondition ?p5 x y p0 \ rank_property ?p5 rank" + proof + assume 27: "?r = ?s" + show "union_sets_postcondition ?p5 x y p0 \ rank_property ?p5 rank" + proof + show "union_sets_postcondition ?p5 x y p0" + using 1 2 3 27 by (simp add: union_sets_1_skip) + show "rank_property ?p5 rank" + using 4 27 by simp + qed + qed + qed +qed + +end + +end + diff --git a/thys/Relational_Disjoint_Set_Forests/ROOT b/thys/Relational_Disjoint_Set_Forests/ROOT --- a/thys/Relational_Disjoint_Set_Forests/ROOT +++ b/thys/Relational_Disjoint_Set_Forests/ROOT @@ -1,16 +1,17 @@ chapter AFP session Relational_Disjoint_Set_Forests (AFP) = Stone_Kleene_Relation_Algebras + options [timeout = 600] sessions "HOL-Hoare" theories Disjoint_Set_Forests + More_Disjoint_Set_Forests document_files "root.tex" "root.bib" diff --git a/thys/Relational_Disjoint_Set_Forests/document/root.tex b/thys/Relational_Disjoint_Set_Forests/document/root.tex --- a/thys/Relational_Disjoint_Set_Forests/document/root.tex +++ b/thys/Relational_Disjoint_Set_Forests/document/root.tex @@ -1,61 +1,64 @@ \documentclass[11pt,a4paper]{article} \usepackage[T1]{fontenc} \usepackage{isabelle,isabellesym} \usepackage{amssymb,ragged2e} \usepackage{pdfsetup} \isabellestyle{it} \renewenvironment{isamarkuptext}{\par\isastyletext\begin{isapar}\justifying\color{blue}}{\end{isapar}} \renewcommand\labelitemi{$*$} %\urlstyle{rm} \begin{document} \title{Relational Disjoint-Set Forests} \author{Walter Guttmann} \maketitle \begin{abstract} We give a simple relation-algebraic semantics of read and write operations on associative arrays. The array operations seamlessly integrate with assignments in the Hoare-logic library. - Using relation algebras and Kleene algebras we verify the correctness of an array-based implementation of disjoint-set forests with a naive union operation and a find operation with path compression. + Using relation algebras and Kleene algebras we verify the correctness of an array-based implementation of disjoint-set forests using the union-by-rank strategy and find operations with path compression, path splitting and path halving. \end{abstract} \tableofcontents \section{Overview} Relation algebras and Kleene algebras have previously been used to reason about graphs and graph algorithms \cite{BackhouseCarre1975,Berghammer1999,BerghammerStruth2010,BerghammerKargerWolf1998,GondranMinoux2008,HoefnerMoeller2012,Moeller1993}. The operations of these algebras manipulate entire graphs, which is useful for specification but not directly intended for implementation. Low-level array access is a key ingredient for efficient algorithms \cite{CormenLeisersonRivest1990}. We give a relation-algebraic semantics for such read/write access to associative arrays. This allows us to extend relation-algebraic verification methods to a lower level of more efficient implementations. In this theory we focus on arrays with the same index and value sets, which can be modelled as homogeneous relations and therefore as elements of relation algebras and Kleene algebras \cite{Kozen1994,Tarski1941}. -We implement and verify the correctness of disjoint-set forests with path compression and naive union \cite{CormenLeisersonRivest1990,GallerFisher1964,Tarjan1975}. +We implement and verify the correctness of disjoint-set forests with path compression strategies and union-by-rank \cite{CormenLeisersonRivest1990,GallerFisher1964,Tarjan1975}. In order to prepare this theory for future applications with weighted graphs, the verification uses Stone relation algebras, which have weaker axioms than relation algebras \cite{Guttmann2018c}. Section 2 contains the simple relation-algebraic semantics of associative array read and write and basic properties of these access operations. In Section 3 we give a Kleene-relation-algebraic semantics of disjoint-set forests. -The make-set, find-set and union-sets operations are implemented and verified in Section 4. +The make-set operation, find-set with path compression and the naive union-sets operation are implemented and verified in Section 4. +Section 5 presents further results on disjoint-set forests and relational array access. +The initialisation of disjoint-set forests, path halving and path splitting are implemented and verified in Section 6. +In Section 7 we study relational Peano structures and implement and verify union-by-rank. -This Isabelle/HOL theory formally verifies results in \cite{Guttmann2020b}. -Theorem numbers from this paper are mentioned in the theory for reference. +This Isabelle/HOL theory formally verifies results in \cite{Guttmann2020b} and in an extended version of that paper. +Theorem numbers from the extended version of the paper are mentioned in the theories for reference. See the paper for further details and related work. Several Isabelle/HOL theories are related to disjoint sets. The theory \texttt{HOL/Library/Disjoint\_Sets.thy} contains results about partitions and sets of disjoint sets and does not consider their implementation. An implementation of disjoint-set forests with path compression and a size-based heuristic in the Imperative/HOL framework is verified in Archive of Formal Proofs entry \cite{LammichMeis2012}. Improved automation of this proof is considered in Archive of Formal Proofs entry \cite{Zhan2018}. These approaches are based on logical specifications whereas the present theory uses relation algebras and Kleene algebras. \begin{flushleft} \input{session} \end{flushleft} \bibliographystyle{abbrv} \bibliography{root} \end{document} diff --git a/thys/Stone_Algebras/P_Algebras.thy b/thys/Stone_Algebras/P_Algebras.thy --- a/thys/Stone_Algebras/P_Algebras.thy +++ b/thys/Stone_Algebras/P_Algebras.thy @@ -1,1324 +1,1328 @@ (* Title: Pseudocomplemented Algebras Author: Walter Guttmann Maintainer: Walter Guttmann *) section \Pseudocomplemented Algebras\ text \ This theory expands lattices with a pseudocomplement operation. In particular, we consider the following algebraic structures: \begin{itemize} \item pseudocomplemented lattices (p-algebras) \item pseudocomplemented distributive lattices (distributive p-algebras) \item Stone algebras \item Heyting semilattices \item Heyting lattices \item Heyting algebras \item Heyting-Stone algebras \item Brouwer algebras \item Boolean algebras \end{itemize} Most of these structures and many results in this theory are discussed in \cite{BalbesDwinger1974,Birkhoff1967,Blyth2005,Curry1977,Graetzer1971,Maddux1996}. \ theory P_Algebras imports Lattice_Basics begin subsection \P-Algebras\ text \ In this section we add a pseudocomplement operation to lattices and to distributive lattices. \ subsubsection \Pseudocomplemented Lattices\ text \ The pseudocomplement of an element \y\ is the greatest element whose meet with \y\ is the least element of the lattice. \ class p_algebra = bounded_lattice + uminus + assumes pseudo_complement: "x \ y = bot \ x \ -y" begin subclass sup_inf_top_bot_uminus_ord . text \ Regular elements and dense elements are frequently used in pseudocomplemented algebras. \ abbreviation "regular x \ x = --x" abbreviation "dense x \ -x = bot" abbreviation "complemented x \ \y . x \ y = bot \ x \ y = top" abbreviation "in_p_image x \ \y . x = -y" abbreviation "selection s x \ s = --s \ x" abbreviation "dense_elements \ { x . dense x }" abbreviation "regular_elements \ { x . in_p_image x }" lemma p_bot [simp]: "-bot = top" using inf_top.left_neutral pseudo_complement top_unique by blast lemma p_top [simp]: "-top = bot" by (metis eq_refl inf_top.comm_neutral pseudo_complement) text \ The pseudocomplement satisfies the following half of the requirements of a complement. \ lemma inf_p [simp]: "x \ -x = bot" using inf.commute pseudo_complement by fastforce lemma p_inf [simp]: "-x \ x = bot" by (simp add: inf_commute) lemma pp_inf_p: "--x \ -x = bot" by simp text \ The double complement is a closure operation. \ lemma pp_increasing: "x \ --x" using inf_p pseudo_complement by blast lemma ppp [simp]: "---x = -x" by (metis order.antisym inf.commute order_trans pseudo_complement pp_increasing) lemma pp_idempotent: "----x = --x" by simp lemma regular_in_p_image_iff: "regular x \ in_p_image x" by auto lemma pseudo_complement_pp: "x \ y = bot \ --x \ -y" by (metis inf_commute pseudo_complement ppp) lemma p_antitone: "x \ y \ -y \ -x" by (metis inf_commute order_trans pseudo_complement pp_increasing) lemma p_antitone_sup: "-(x \ y) \ -x" by (simp add: p_antitone) lemma p_antitone_inf: "-x \ -(x \ y)" by (simp add: p_antitone) lemma p_antitone_iff: "x \ -y \ y \ -x" using order_lesseq_imp p_antitone pp_increasing by blast lemma pp_isotone: "x \ y \ --x \ --y" by (simp add: p_antitone) lemma pp_isotone_sup: "--x \ --(x \ y)" by (simp add: p_antitone) lemma pp_isotone_inf: "--(x \ y) \ --x" by (simp add: p_antitone) text \ One of De Morgan's laws holds in pseudocomplemented lattices. \ lemma p_dist_sup [simp]: "-(x \ y) = -x \ -y" apply (rule order.antisym) apply (simp add: p_antitone) using inf_le1 inf_le2 le_sup_iff p_antitone_iff by blast lemma p_supdist_inf: "-x \ -y \ -(x \ y)" by (simp add: p_antitone) lemma pp_dist_pp_sup [simp]: "--(--x \ --y) = --(x \ y)" by simp lemma p_sup_p [simp]: "-(x \ -x) = bot" by simp lemma pp_sup_p [simp]: "--(x \ -x) = top" by simp lemma dense_pp: "dense x \ --x = top" by (metis p_bot p_top ppp) lemma dense_sup_p: "dense (x \ -x)" by simp lemma regular_char: "regular x \ (\y . x = -y)" by auto lemma pp_inf_bot_iff: "x \ y = bot \ --x \ y = bot" by (simp add: pseudo_complement_pp) text \ Weak forms of the shunting property hold. Most require a pseudocomplemented element on the right-hand side. \ lemma p_shunting_swap: "x \ y \ -z \ x \ z \ -y" by (metis inf_assoc inf_commute pseudo_complement) lemma pp_inf_below_iff: "x \ y \ -z \ --x \ y \ -z" by (simp add: inf_commute p_shunting_swap) lemma p_inf_pp [simp]: "-(x \ --y) = -(x \ y)" apply (rule order.antisym) apply (simp add: inf.coboundedI2 p_antitone pp_increasing) using inf_commute p_antitone_iff pp_inf_below_iff by auto lemma p_inf_pp_pp [simp]: "-(--x \ --y) = -(x \ y)" by (simp add: inf_commute) lemma regular_closed_inf: "regular x \ regular y \ regular (x \ y)" by (metis p_dist_sup ppp) lemma regular_closed_p: "regular (-x)" by simp lemma regular_closed_pp: "regular (--x)" by simp lemma regular_closed_bot: "regular bot" by simp lemma regular_closed_top: "regular top" by simp lemma pp_dist_inf [simp]: "--(x \ y) = --x \ --y" by (metis p_dist_sup p_inf_pp_pp ppp) lemma inf_import_p [simp]: "x \ -(x \ y) = x \ -y" apply (rule order.antisym) using p_shunting_swap apply fastforce using inf.sup_right_isotone p_antitone by auto text \ Pseudocomplements are unique. \ lemma p_unique: "(\x . x \ y = bot \ x \ z) \ z = -y" using inf.order_eq_iff pseudo_complement by auto lemma maddux_3_5: "x \ x = x \ -(y \ -y)" by simp lemma shunting_1_pp: "x \ --y \ x \ -y = bot" by (simp add: pseudo_complement) lemma pp_pp_inf_bot_iff: "x \ y = bot \ --x \ --y = bot" by (simp add: pseudo_complement_pp) lemma inf_pp_semi_commute: "x \ --y \ --(x \ y)" using inf.eq_refl p_antitone_iff p_inf_pp by presburger lemma inf_pp_commute: "--(--x \ y) = --x \ --y" by simp lemma sup_pp_semi_commute: "x \ --y \ --(x \ y)" by (simp add: p_antitone_iff) lemma regular_sup: "regular z \ (x \ z \ y \ z \ --(x \ y) \ z)" apply (rule iffI) apply (metis le_supI pp_isotone) using dual_order.trans sup_ge2 pp_increasing pp_isotone_sup by blast lemma dense_closed_inf: "dense x \ dense y \ dense (x \ y)" by (simp add: dense_pp) lemma dense_closed_sup: "dense x \ dense y \ dense (x \ y)" by simp lemma dense_closed_pp: "dense x \ dense (--x)" by simp lemma dense_closed_top: "dense top" by simp lemma dense_up_closed: "dense x \ x \ y \ dense y" using dense_pp top_le pp_isotone by auto lemma regular_dense_top: "regular x \ dense x \ x = top" using p_bot by blast lemma selection_char: "selection s x \ (\y . s = -y \ x)" by (metis inf_import_p inf_commute regular_closed_p) lemma selection_closed_inf: "selection s x \ selection t x \ selection (s \ t) x" by (metis inf_assoc inf_commute inf_idem pp_dist_inf) lemma selection_closed_pp: "regular x \ selection s x \ selection (--s) x" by (metis pp_dist_inf) lemma selection_closed_bot: "selection bot x" by simp lemma selection_closed_id: "selection x x" using inf.le_iff_sup pp_increasing by auto text \ Conjugates are usually studied for Boolean algebras, however, some of their properties generalise to pseudocomplemented algebras. \ lemma conjugate_unique_p: assumes "conjugate f g" and "conjugate f h" shows "uminus \ g = uminus \ h" proof - have "\x y . x \ g y = bot \ x \ h y = bot" using assms conjugate_def inf.commute by simp hence "\x y . x \ -(g y) \ x \ -(h y)" using inf.commute pseudo_complement by simp hence "\y . -(g y) = -(h y)" using order.eq_iff by blast thus ?thesis by auto qed lemma conjugate_symmetric: "conjugate f g \ conjugate g f" by (simp add: conjugate_def inf_commute) lemma additive_isotone: "additive f \ isotone f" by (metis additive_def isotone_def le_iff_sup) lemma dual_additive_antitone: assumes "dual_additive f" shows "isotone (uminus \ f)" proof - have "\x y . f (x \ y) \ f x" using assms dual_additive_def by simp hence "\x y . x \ y \ f y \ f x" by (metis sup_absorb2) hence "\x y . x \ y \ -(f x) \ -(f y)" by (simp add: p_antitone) thus ?thesis by (simp add: isotone_def) qed lemma conjugate_dual_additive: assumes "conjugate f g" shows "dual_additive (uminus \ f)" proof - have 1: "\x y z . -z \ -(f (x \ y)) \ -z \ -(f x) \ -z \ -(f y)" proof (intro allI) fix x y z have "(-z \ -(f (x \ y))) = (f (x \ y) \ -z = bot)" by (simp add: p_antitone_iff pseudo_complement) also have "... = ((x \ y) \ g(-z) = bot)" using assms conjugate_def by auto also have "... = (x \ y \ -(g(-z)))" by (simp add: pseudo_complement) also have "... = (x \ -(g(-z)) \ y \ -(g(-z)))" by (simp add: le_sup_iff) also have "... = (x \ g(-z) = bot \ y \ g(-z) = bot)" by (simp add: pseudo_complement) also have "... = (f x \ -z = bot \ f y \ -z = bot)" using assms conjugate_def by auto also have "... = (-z \ -(f x) \ -z \ -(f y))" by (simp add: p_antitone_iff pseudo_complement) finally show "-z \ -(f (x \ y)) \ -z \ -(f x) \ -z \ -(f y)" by simp qed have "\x y . -(f (x \ y)) = -(f x) \ -(f y)" proof (intro allI) fix x y have "-(f x) \ -(f y) = --(-(f x) \ -(f y))" by simp hence "-(f x) \ -(f y) \ -(f (x \ y))" using 1 by (metis inf_le1 inf_le2) thus "-(f (x \ y)) = -(f x) \ -(f y)" using 1 order.antisym by fastforce qed thus ?thesis using dual_additive_def by simp qed lemma conjugate_isotone_pp: "conjugate f g \ isotone (uminus \ uminus \ f)" by (simp add: comp_assoc conjugate_dual_additive dual_additive_antitone) lemma conjugate_char_1_pp: "conjugate f g \ (\x y . f(x \ -(g y)) \ --f x \ -y \ g(y \ -(f x)) \ --g y \ -x)" proof assume 1: "conjugate f g" show "\x y . f(x \ -(g y)) \ --f x \ -y \ g(y \ -(f x)) \ --g y \ -x" proof (intro allI) fix x y have 2: "f(x \ -(g y)) \ -y" using 1 by (simp add: conjugate_def pseudo_complement) have "f(x \ -(g y)) \ --f(x \ -(g y))" by (simp add: pp_increasing) also have "... \ --f x" using 1 conjugate_isotone_pp isotone_def by simp finally have 3: "f(x \ -(g y)) \ --f x \ -y" using 2 by simp have 4: "isotone (uminus \ uminus \ g)" using 1 conjugate_isotone_pp conjugate_symmetric by auto have 5: "g(y \ -(f x)) \ -x" using 1 by (metis conjugate_def inf.cobounded2 inf_commute pseudo_complement) have "g(y \ -(f x)) \ --g(y \ -(f x))" by (simp add: pp_increasing) also have "... \ --g y" using 4 isotone_def by auto finally have "g(y \ -(f x)) \ --g y \ -x" using 5 by simp thus "f(x \ -(g y)) \ --f x \ -y \ g(y \ -(f x)) \ --g y \ -x" using 3 by simp qed next assume 6: "\x y . f(x \ -(g y)) \ --f x \ -y \ g(y \ -(f x)) \ --g y \ -x" hence 7: "\x y . f x \ y = bot \ x \ g y = bot" by (metis inf.le_iff_sup inf.le_sup_iff inf_commute pseudo_complement) have "\x y . x \ g y = bot \ f x \ y = bot" using 6 by (metis inf.le_iff_sup inf.le_sup_iff inf_commute pseudo_complement) thus "conjugate f g" using 7 conjugate_def by auto qed lemma conjugate_char_1_isotone: "conjugate f g \ isotone f \ isotone g \ f(x \ -(g y)) \ f x \ -y \ g(y \ -(f x)) \ g y \ -x" by (simp add: conjugate_char_1_pp ord.isotone_def) lemma dense_lattice_char_1: "(\x y . x \ y = bot \ x = bot \ y = bot) \ (\x . x \ bot \ dense x)" by (metis inf_top.left_neutral p_bot p_inf pp_inf_bot_iff) lemma dense_lattice_char_2: "(\x y . x \ y = bot \ x = bot \ y = bot) \ (\x . regular x \ x = bot \ x = top)" by (metis dense_lattice_char_1 inf_top.left_neutral p_inf regular_closed_p regular_closed_top) lemma restrict_below_Rep_eq: "x \ --y \ z \ x \ y = x \ z \ y" by (metis inf.absorb2 inf.commute inf.left_commute pp_increasing) (* lemma p_inf_sup_below: "-x \ (x \ y) \ y" nitpick [expect=genuine] oops lemma complement_p: "x \ y = bot \ x \ y = top \ -x = y" nitpick [expect=genuine] oops lemma complemented_regular: "complemented x \ regular x" nitpick [expect=genuine] oops *) end text \ The following class gives equational axioms for the pseudocomplement operation. \ class p_algebra_eq = bounded_lattice + uminus + assumes p_bot_eq: "-bot = top" and p_top_eq: "-top = bot" and inf_import_p_eq: "x \ -(x \ y) = x \ -y" begin lemma inf_p_eq: "x \ -x = bot" by (metis inf_bot_right inf_import_p_eq inf_top_right p_top_eq) subclass p_algebra apply unfold_locales apply (rule iffI) apply (metis inf.orderI inf_import_p_eq inf_top.right_neutral p_bot_eq) by (metis (full_types) inf.left_commute inf.orderE inf_bot_right inf_commute inf_p_eq) end subsubsection \Pseudocomplemented Distributive Lattices\ text \ We obtain further properties if we assume that the lattice operations are distributive. \ class pd_algebra = p_algebra + bounded_distrib_lattice begin lemma p_inf_sup_below: "-x \ (x \ y) \ y" by (simp add: inf_sup_distrib1) lemma pp_inf_sup_p [simp]: "--x \ (x \ -x) = x" using inf.absorb2 inf_sup_distrib1 pp_increasing by auto lemma complement_p: "x \ y = bot \ x \ y = top \ -x = y" by (metis pseudo_complement inf.commute inf_top.left_neutral sup.absorb_iff1 sup.commute sup_bot.right_neutral sup_inf_distrib2 p_inf) lemma complemented_regular: "complemented x \ regular x" using complement_p inf.commute sup.commute by fastforce lemma regular_inf_dense: "\y z . regular y \ dense z \ x = y \ z" by (metis pp_inf_sup_p dense_sup_p ppp) lemma maddux_3_12 [simp]: "(x \ -y) \ (x \ y) = x" by (metis p_inf sup_bot_right sup_inf_distrib1) lemma maddux_3_13 [simp]: "(x \ y) \ -x = y \ -x" by (simp add: inf_sup_distrib2) lemma maddux_3_20: "((v \ w) \ (-v \ x)) \ -((v \ y) \ (-v \ z)) = (v \ w \ -y) \ (-v \ x \ -z)" proof - have "v \ w \ -(v \ y) \ -(-v \ z) = v \ w \ -(v \ y)" by (meson inf.cobounded1 inf_absorb1 le_infI1 p_antitone_iff) also have "... = v \ w \ -y" using inf.sup_relative_same_increasing inf_import_p inf_le1 by blast finally have 1: "v \ w \ -(v \ y) \ -(-v \ z) = v \ w \ -y" . have "-v \ x \ -(v \ y) \ -(-v \ z) = -v \ x \ -(-v \ z)" by (simp add: inf.absorb1 le_infI1 p_antitone_inf) also have "... = -v \ x \ -z" by (simp add: inf.assoc inf_left_commute) finally have 2: "-v \ x \ -(v \ y) \ -(-v \ z) = -v \ x \ -z" . have "((v \ w) \ (-v \ x)) \ -((v \ y) \ (-v \ z)) = (v \ w \ -(v \ y) \ -(-v \ z)) \ (-v \ x \ -(v \ y) \ -(-v \ z))" by (simp add: inf_assoc inf_sup_distrib2) also have "... = (v \ w \ -y) \ (-v \ x \ -z)" using 1 2 by simp finally show ?thesis . qed lemma order_char_1: "x \ y \ x \ y \ -x" by (metis inf.sup_left_isotone inf_sup_absorb le_supI1 maddux_3_12 sup_commute) lemma order_char_2: "x \ y \ x \ -x \ y \ -x" using order_char_1 by auto +lemma half_shunting: + "x \ y \ z \ x \ -z \ y" + by (metis inf.sup_right_isotone inf_commute inf_sup_distrib1 sup.boundedE maddux_3_12) + (* lemma pp_dist_sup [simp]: "--(x \ y) = --x \ --y" nitpick [expect=genuine] oops lemma regular_closed_sup: "regular x \ regular y \ regular (x \ y)" nitpick [expect=genuine] oops lemma regular_complemented_iff: "regular x \ complemented x" nitpick [expect=genuine] oops lemma selection_closed_sup: "selection s x \ selection t x \ selection (s \ t) x" nitpick [expect=genuine] oops lemma stone [simp]: "-x \ --x = top" nitpick [expect=genuine] oops *) end subsection \Stone Algebras\ text \ A Stone algebra is a distributive lattice with a pseudocomplement that satisfies the following equation. We thus obtain the other half of the requirements of a complement at least for the regular elements. \ class stone_algebra = pd_algebra + assumes stone [simp]: "-x \ --x = top" begin text \ As a consequence, we obtain both De Morgan's laws for all elements. \ lemma p_dist_inf [simp]: "-(x \ y) = -x \ -y" proof (rule p_unique[THEN sym], rule allI, rule iffI) fix w assume "w \ (x \ y) = bot" hence "w \ --x \ y = bot" using inf_commute inf_left_commute pseudo_complement by auto hence 1: "w \ --x \ -y" by (simp add: pseudo_complement) have "w = (w \ -x) \ (w \ --x)" using distrib_imp2 sup_inf_distrib1 by auto thus "w \ -x \ -y" using 1 by (metis inf_le2 sup.mono) next fix w assume "w \ -x \ -y" thus "w \ (x \ y) = bot" using order_trans p_supdist_inf pseudo_complement by blast qed lemma pp_dist_sup [simp]: "--(x \ y) = --x \ --y" by simp lemma regular_closed_sup: "regular x \ regular y \ regular (x \ y)" by simp text \ The regular elements are precisely the ones having a complement. \ lemma regular_complemented_iff: "regular x \ complemented x" by (metis inf_p stone complemented_regular) lemma selection_closed_sup: "selection s x \ selection t x \ selection (s \ t) x" by (simp add: inf_sup_distrib2) lemma huntington_3_pp [simp]: "-(-x \ -y) \ -(-x \ y) = --x" by (metis p_dist_inf p_inf sup.commute sup_bot_left sup_inf_distrib1) lemma maddux_3_3 [simp]: "-(x \ y) \ -(x \ -y) = -x" by (simp add: sup_commute sup_inf_distrib1) lemma maddux_3_11_pp: "(x \ -y) \ (x \ --y) = x" by (metis inf_sup_distrib1 inf_top_right stone) lemma maddux_3_19_pp: "(-x \ y) \ (--x \ z) = (--x \ y) \ (-x \ z)" proof - have "(--x \ y) \ (-x \ z) = (--x \ z) \ (y \ -x) \ (y \ z)" by (simp add: inf.commute inf_sup_distrib1 sup.assoc) also have "... = (--x \ z) \ (y \ -x) \ (y \ z \ (-x \ --x))" by simp also have "... = (--x \ z) \ ((y \ -x) \ (y \ -x \ z)) \ (y \ z \ --x)" using inf_sup_distrib1 sup_assoc inf_commute inf_assoc by presburger also have "... = (--x \ z) \ (y \ -x) \ (y \ z \ --x)" by simp also have "... = ((--x \ z) \ (--x \ z \ y)) \ (y \ -x)" by (simp add: inf_assoc inf_commute sup.left_commute sup_commute) also have "... = (--x \ z) \ (y \ -x)" by simp finally show ?thesis by (simp add: inf_commute sup_commute) qed lemma compl_inter_eq_pp: "--x \ y = --x \ z \ -x \ y = -x \ z \ y = z" by (metis inf_commute inf_p inf_sup_distrib1 inf_top.right_neutral p_bot p_dist_inf) lemma maddux_3_21_pp [simp]: "--x \ (-x \ y) = --x \ y" by (simp add: sup.commute sup_inf_distrib1) lemma shunting_2_pp: "x \ --y \ -x \ --y = top" by (metis inf_top_left p_bot p_dist_inf pseudo_complement) lemma shunting_p: "x \ y \ -z \ x \ -z \ -y" by (metis inf.assoc p_dist_inf p_shunting_swap pseudo_complement) text \ The following weak shunting property is interesting as it does not require the element \z\ on the right-hand side to be regular. \ lemma shunting_var_p: "x \ -y \ z \ x \ z \ --y" proof assume "x \ -y \ z" hence "z \ --y = --y \ (z \ x \ -y)" by (simp add: sup.absorb1 sup.commute) thus "x \ z \ --y" by (metis inf_commute maddux_3_21_pp sup.commute sup.left_commute sup_left_divisibility) next assume "x \ z \ --y" thus "x \ -y \ z" by (metis inf.mono maddux_3_12 sup_ge2) qed (* Whether conjugate_char_2_pp can be proved in pd_algebra or in p_algebra is unknown. *) lemma conjugate_char_2_pp: "conjugate f g \ f bot = bot \ g bot = bot \ (\x y . f x \ y \ --(f(x \ --(g y))) \ g y \ x \ --(g(y \ --(f x))))" proof assume 1: "conjugate f g" hence 2: "dual_additive (uminus \ g)" using conjugate_symmetric conjugate_dual_additive by auto show "f bot = bot \ g bot = bot \ (\x y . f x \ y \ --(f(x \ --(g y))) \ g y \ x \ --(g(y \ --(f x))))" proof (intro conjI) show "f bot = bot" using 1 by (metis conjugate_def inf_idem inf_bot_left) next show "g bot = bot" using 1 by (metis conjugate_def inf_idem inf_bot_right) next show "\x y . f x \ y \ --(f(x \ --(g y))) \ g y \ x \ --(g(y \ --(f x)))" proof (intro allI) fix x y have 3: "y \ -(f(x \ -(g y)))" using 1 by (simp add: conjugate_def pseudo_complement inf_commute) have 4: "x \ -(g(y \ -(f x)))" using 1 conjugate_def inf.commute pseudo_complement by fastforce have "y \ -(f(x \ --(g y))) = y \ -(f(x \ -(g y))) \ -(f(x \ --(g y)))" using 3 by (simp add: inf.le_iff_sup inf_commute) also have "... = y \ -(f((x \ -(g y)) \ (x \ --(g y))))" using 1 conjugate_dual_additive dual_additive_def inf_assoc by auto also have "... = y \ -(f x)" by (simp add: maddux_3_11_pp) also have "... \ -(f x)" by simp finally have 5: "f x \ y \ --(f(x \ --(g y)))" by (simp add: inf_commute p_shunting_swap) have "x \ -(g(y \ --(f x))) = x \ -(g(y \ -(f x))) \ -(g(y \ --(f x)))" using 4 by (simp add: inf.le_iff_sup inf_commute) also have "... = x \ -(g((y \ -(f x)) \ (y \ --(f x))))" using 2 by (simp add: dual_additive_def inf_assoc) also have "... = x \ -(g y)" by (simp add: maddux_3_11_pp) also have "... \ -(g y)" by simp finally have "g y \ x \ --(g(y \ --(f x)))" by (simp add: inf_commute p_shunting_swap) thus "f x \ y \ --(f(x \ --(g y))) \ g y \ x \ --(g(y \ --(f x)))" using 5 by simp qed qed next assume "f bot = bot \ g bot = bot \ (\x y . f x \ y \ --(f(x \ --(g y))) \ g y \ x \ --(g(y \ --(f x))))" thus "conjugate f g" by (unfold conjugate_def, metis inf_commute le_bot pp_inf_bot_iff regular_closed_bot) qed lemma conjugate_char_2_pp_additive: assumes "conjugate f g" and "additive f" and "additive g" shows "f x \ y \ f(x \ --(g y)) \ g y \ x \ g(y \ --(f x))" proof - have "f x \ y = f ((x \ --g y) \ (x \ -g y)) \ y" by (simp add: sup.commute sup_inf_distrib1) also have "... = (f (x \ --g y) \ y) \ (f (x \ -g y) \ y)" using assms(2) additive_def inf_sup_distrib2 by auto also have "... = f (x \ --g y) \ y" by (metis assms(1) conjugate_def inf_le2 pseudo_complement sup_bot.right_neutral) finally have 2: "f x \ y \ f (x \ --g y)" by simp have "g y \ x = g ((y \ --f x) \ (y \ -f x)) \ x" by (simp add: sup.commute sup_inf_distrib1) also have "... = (g (y \ --f x) \ x) \ (g (y \ -f x) \ x)" using assms(3) additive_def inf_sup_distrib2 by auto also have "... = g (y \ --f x) \ x" by (metis assms(1) conjugate_def inf.cobounded2 pseudo_complement sup_bot.right_neutral inf_commute) finally have "g y \ x \ g (y \ --f x)" by simp thus ?thesis using 2 by simp qed (* lemma compl_le_swap2_iff: "-x \ y \ -y \ x" nitpick [expect=genuine] oops lemma huntington_3: "x = -(-x \ -y) \ -(-x \ y)" nitpick [expect=genuine] oops lemma maddux_3_1: "x \ -x = y \ -y" nitpick [expect=genuine] oops lemma maddux_3_4: "x \ (y \ -y) = z \ -z" nitpick [expect=genuine] oops lemma maddux_3_11: "x = (x \ y) \ (x \ -y)" nitpick [expect=genuine] oops lemma maddux_3_19: "(-x \ y) \ (x \ z) = (x \ y) \ (-x \ z)" nitpick [expect=genuine] oops lemma compl_inter_eq: "x \ y = x \ z \ -x \ y = -x \ z \ y = z" nitpick [expect=genuine] oops lemma maddux_3_21: "x \ y = x \ (-x \ y)" nitpick [expect=genuine] oops lemma shunting_1: "x \ y \ x \ -y = bot" nitpick [expect=genuine] oops lemma shunting_2: "x \ y \ -x \ y = top" nitpick [expect=genuine] oops lemma conjugate_unique: "conjugate f g \ conjugate f h \ g = h" nitpick [expect=genuine] oops lemma conjugate_isotone_pp: "conjugate f g \ isotone f" nitpick [expect=genuine] oops lemma conjugate_char_1: "conjugate f g \ (\x y . f(x \ -(g y)) \ f x \ -y \ g(y \ -(f x)) \ g y \ -x)" nitpick [expect=genuine] oops lemma conjugate_char_2: "conjugate f g \ f bot = bot \ g bot = bot \ (\x y . f x \ y \ f(x \ g y) \ g y \ x \ g(y \ f x))" nitpick [expect=genuine] oops lemma shunting: "x \ y \ z \ x \ z \ -y" nitpick [expect=genuine] oops lemma shunting_var: "x \ -y \ z \ x \ z \ y" nitpick [expect=genuine] oops lemma sup_compl_top: "x \ -x = top" nitpick [expect=genuine] oops lemma selection_closed_p: "selection s x \ selection (-s) x" nitpick [expect=genuine] oops lemma selection_closed_pp: "selection s x \ selection (--s) x" nitpick [expect=genuine] oops *) end abbreviation stone_algebra_isomorphism :: "('a::stone_algebra \ 'b::stone_algebra) \ bool" where "stone_algebra_isomorphism f \ sup_inf_top_bot_uminus_isomorphism f" text \ Every bounded linear order can be expanded to a Stone algebra. The pseudocomplement takes \bot\ to the \top\ and every other element to \bot\. \ class linorder_stone_algebra_expansion = linorder_lattice_expansion + uminus + assumes uminus_def [simp]: "-x = (if x = bot then top else bot)" begin subclass stone_algebra apply unfold_locales using bot_unique min_def top_le by auto text \ The regular elements are the least and greatest elements. All elements except the least element are dense. \ lemma regular_bot_top: "regular x \ x = bot \ x = top" by simp lemma not_bot_dense: "x \ bot \ --x = top" by simp end subsection \Heyting Algebras\ text \ In this section we add a relative pseudocomplement operation to semilattices and to lattices. \ subsubsection \Heyting Semilattices\ text \ The pseudocomplement of an element \y\ relative to an element \z\ is the least element whose meet with \y\ is below \z\. This can be stated as a Galois connection. Specialising \z = bot\ gives (non-relative) pseudocomplements. Many properties can already be shown if the underlying structure is just a semilattice. \ class implies = fixes implies :: "'a \ 'a \ 'a" (infixl "\" 65) class heyting_semilattice = semilattice_inf + implies + assumes implies_galois: "x \ y \ z \ x \ y \ z" begin lemma implies_below_eq [simp]: "y \ (x \ y) = y" using implies_galois inf.absorb_iff1 inf.cobounded1 by blast lemma implies_increasing: "x \ y \ x" by (simp add: inf.orderI) lemma implies_galois_swap: "x \ y \ z \ y \ x \ z" by (metis implies_galois inf_commute) lemma implies_galois_var: "x \ y \ z \ y \ x \ z" by (simp add: implies_galois_swap implies_galois) lemma implies_galois_increasing: "x \ y \ (x \ y)" using implies_galois by blast lemma implies_galois_decreasing: "(y \ x) \ y \ x" using implies_galois by blast lemma implies_mp_below: "x \ (x \ y) \ y" using implies_galois_decreasing inf_commute by auto lemma implies_isotone: "x \ y \ z \ x \ z \ y" using implies_galois order_trans by blast lemma implies_antitone: "x \ y \ y \ z \ x \ z" by (meson implies_galois_swap order_lesseq_imp) lemma implies_isotone_inf: "x \ (y \ z) \ x \ y" by (simp add: implies_isotone) lemma implies_antitone_inf: "x \ z \ (x \ y) \ z" by (simp add: implies_antitone) lemma implies_curry: "x \ (y \ z) = (x \ y) \ z" by (metis implies_galois_decreasing implies_galois inf_assoc order.antisym) lemma implies_curry_flip: "x \ (y \ z) = y \ (x \ z)" by (simp add: implies_curry inf_commute) lemma triple_implies [simp]: "((x \ y) \ y) \ y = x \ y" using implies_antitone implies_galois_swap order.eq_iff by auto lemma implies_mp_eq [simp]: "x \ (x \ y) = x \ y" by (metis implies_below_eq implies_mp_below inf_left_commute inf.absorb2) lemma implies_dist_implies: "x \ (y \ z) \ (x \ y) \ (x \ z)" using implies_curry implies_curry_flip by auto lemma implies_import_inf [simp]: "x \ ((x \ y) \ (x \ z)) = x \ (y \ z)" by (metis implies_curry implies_mp_eq inf_commute) lemma implies_dist_inf: "x \ (y \ z) = (x \ y) \ (x \ z)" proof - have "(x \ y) \ (x \ z) \ x \ y \ z" by (simp add: implies_galois) hence "(x \ y) \ (x \ z) \ x \ (y \ z)" using implies_galois by blast thus ?thesis by (simp add: implies_isotone order.eq_iff) qed lemma implies_itself_top: "y \ x \ x" by (simp add: implies_galois_swap implies_increasing) lemma inf_implies_top: "z \ (x \ y) \ x" using implies_galois_var le_infI1 by blast lemma inf_inf_implies [simp]: "z \ ((x \ y) \ x) = z" by (simp add: inf_implies_top inf_absorb1) lemma le_implies_top: "x \ y \ z \ x \ y" using implies_antitone implies_itself_top order.trans by blast lemma le_iff_le_implies: "x \ y \ x \ x \ y" using implies_galois inf_idem by force lemma implies_inf_isotone: "x \ y \ (x \ z) \ (y \ z)" by (metis implies_curry implies_galois_increasing implies_isotone) lemma implies_transitive: "(x \ y) \ (y \ z) \ x \ z" using implies_dist_implies implies_galois_var implies_increasing order_lesseq_imp by blast lemma implies_inf_absorb [simp]: "x \ (x \ y) = x \ y" using implies_dist_inf implies_itself_top inf.absorb_iff2 by auto lemma implies_implies_absorb [simp]: "x \ (x \ y) = x \ y" by (simp add: implies_curry) lemma implies_inf_identity: "(x \ y) \ y = y" by (simp add: inf_commute) lemma implies_itself_same: "x \ x = y \ y" by (simp add: le_implies_top order.eq_iff) end text \ The following class gives equational axioms for the relative pseudocomplement operation (inequalities can be written as equations). \ class heyting_semilattice_eq = semilattice_inf + implies + assumes implies_mp_below: "x \ (x \ y) \ y" and implies_galois_increasing: "x \ y \ (x \ y)" and implies_isotone_inf: "x \ (y \ z) \ x \ y" begin subclass heyting_semilattice apply unfold_locales apply (rule iffI) apply (metis implies_galois_increasing implies_isotone_inf inf.absorb2 order_lesseq_imp) by (metis implies_mp_below inf_commute order_trans inf_mono order_refl) end text \ The following class allows us to explicitly give the pseudocomplement of an element relative to itself. \ class bounded_heyting_semilattice = bounded_semilattice_inf_top + heyting_semilattice begin lemma implies_itself [simp]: "x \ x = top" using implies_galois inf_le2 top_le by blast lemma implies_order: "x \ y \ x \ y = top" by (metis implies_galois inf_top.left_neutral top_unique) lemma inf_implies [simp]: "(x \ y) \ x = top" using implies_order inf_le1 by blast lemma top_implies [simp]: "top \ x = x" by (metis implies_mp_eq inf_top.left_neutral) end subsubsection \Heyting Lattices\ text \ We obtain further properties if the underlying structure is a lattice. In particular, the lattice operations are automatically distributive in this case. \ class heyting_lattice = lattice + heyting_semilattice begin lemma sup_distrib_inf_le: "(x \ y) \ (x \ z) \ x \ (y \ z)" proof - have "x \ z \ y \ (x \ (y \ z))" using implies_galois_var implies_increasing sup.bounded_iff sup.cobounded2 by blast hence "x \ y \ (x \ z) \ (x \ (y \ z))" using implies_galois_swap implies_increasing le_sup_iff by blast thus ?thesis by (simp add: implies_galois) qed subclass distrib_lattice apply unfold_locales using distrib_sup_le order.eq_iff sup_distrib_inf_le by auto lemma implies_isotone_sup: "x \ y \ x \ (y \ z)" by (simp add: implies_isotone) lemma implies_antitone_sup: "(x \ y) \ z \ x \ z" by (simp add: implies_antitone) lemma implies_sup: "x \ z \ (y \ z) \ ((x \ y) \ z)" proof - have "(x \ z) \ (y \ z) \ y \ z" by (simp add: implies_galois) hence "(x \ z) \ (y \ z) \ (x \ y) \ z" using implies_galois_swap implies_galois_var by fastforce thus ?thesis by (simp add: implies_galois) qed lemma implies_dist_sup: "(x \ y) \ z = (x \ z) \ (y \ z)" apply (rule order.antisym) apply (simp add: implies_antitone) by (simp add: implies_sup implies_galois) lemma implies_antitone_isotone: "(x \ y) \ (x \ y) \ x \ y" by (simp add: implies_antitone_sup implies_dist_inf le_infI2) lemma implies_antisymmetry: "(x \ y) \ (y \ x) = (x \ y) \ (x \ y)" by (metis implies_dist_sup implies_inf_absorb inf.commute) lemma sup_inf_implies [simp]: "(x \ y) \ (x \ y) = y" by (simp add: inf_sup_distrib2 sup.absorb2) lemma implies_subdist_sup: "(x \ y) \ (x \ z) \ x \ (y \ z)" by (simp add: implies_isotone) lemma implies_subdist_inf: "(x \ z) \ (y \ z) \ (x \ y) \ z" by (simp add: implies_antitone) lemma implies_sup_absorb: "(x \ y) \ z \ (x \ z) \ (y \ z)" by (metis implies_dist_sup implies_isotone_sup implies_increasing inf_inf_implies le_sup_iff sup_inf_implies) lemma sup_below_implies_implies: "x \ y \ (x \ y) \ y" by (simp add: implies_dist_sup implies_galois_swap implies_increasing) end class bounded_heyting_lattice = bounded_lattice + heyting_lattice begin subclass bounded_heyting_semilattice .. lemma implies_bot [simp]: "bot \ x = top" using implies_galois top_unique by fastforce end subsubsection \Heyting Algebras\ text \ The pseudocomplement operation can be defined in Heyting algebras, but it is typically not part of their signature. We add the definition as an axiom so that we can use the class hierarchy, for example, to inherit results from the class \pd_algebra\. \ class heyting_algebra = bounded_heyting_lattice + uminus + assumes uminus_eq: "-x = x \ bot" begin subclass pd_algebra apply unfold_locales using bot_unique implies_galois uminus_eq by auto lemma boolean_implies_below: "-x \ y \ x \ y" by (simp add: implies_increasing implies_isotone uminus_eq) lemma negation_implies: "-(x \ y) = --x \ -y" proof (rule order.antisym) show "-(x \ y) \ --x \ -y" using boolean_implies_below p_antitone by auto next have "x \ -y \ (x \ y) = bot" by (metis implies_mp_eq inf_p inf_bot_left inf_commute inf_left_commute) hence "--x \ -y \ (x \ y) = bot" using pp_inf_bot_iff inf_assoc by auto thus "--x \ -y \ -(x \ y)" by (simp add: pseudo_complement) qed lemma double_negation_dist_implies: "--(x \ y) = --x \ --y" apply (rule order.antisym) apply (metis pp_inf_below_iff implies_galois_decreasing implies_galois negation_implies ppp) by (simp add: p_antitone_iff negation_implies) (* lemma stone: "-x \ --x = top" nitpick [expect=genuine] oops *) end text \ The following class gives equational axioms for Heyting algebras. \ class heyting_algebra_eq = bounded_lattice + implies + uminus + assumes implies_mp_eq: "x \ (x \ y) = x \ y" and implies_import_inf: "x \ ((x \ y) \ (x \ z)) = x \ (y \ z)" and inf_inf_implies: "z \ ((x \ y) \ x) = z" and uminus_eq_eq: "-x = x \ bot" begin subclass heyting_algebra apply unfold_locales apply (rule iffI) apply (metis implies_import_inf inf.sup_left_divisibility inf_inf_implies le_iff_inf) apply (metis implies_mp_eq inf.commute inf.le_sup_iff inf.sup_right_isotone) by (simp add: uminus_eq_eq) end text \ A relative pseudocomplement is not enough to obtain the Stone equation, so we add it in the following class. \ class heyting_stone_algebra = heyting_algebra + assumes heyting_stone: "-x \ --x = top" begin subclass stone_algebra by unfold_locales (simp add: heyting_stone) (* lemma pre_linear: "(x \ y) \ (y \ x) = top" nitpick [expect=genuine] oops *) end subsubsection \Brouwer Algebras\ text \ Brouwer algebras are dual to Heyting algebras. The dual pseudocomplement of an element \y\ relative to an element \x\ is the least element whose join with \y\ is above \x\. We can now use the binary operation provided by Boolean algebras in Isabelle/HOL because it is compatible with dual relative pseudocomplements (not relative pseudocomplements). \ class brouwer_algebra = bounded_lattice + minus + uminus + assumes minus_galois: "x \ y \ z \ x - y \ z" and uminus_eq_minus: "-x = top - x" begin sublocale brouwer: heyting_algebra where inf = sup and less_eq = greater_eq and less = greater and sup = inf and bot = top and top = bot and implies = "\x y . y - x" apply unfold_locales apply simp apply simp apply simp apply simp apply (metis minus_galois sup_commute) by (simp add: uminus_eq_minus) lemma curry_minus: "x - (y \ z) = (x - y) - z" by (simp add: brouwer.implies_curry sup_commute) lemma minus_subdist_sup: "(x - z) \ (y - z) \ (x \ y) - z" by (simp add: brouwer.implies_dist_inf) lemma inf_sup_minus: "(x \ y) \ (x - y) = x" by (simp add: inf.absorb1 brouwer.inf_sup_distrib2) end subsection \Boolean Algebras\ text \ This section integrates Boolean algebras in the above hierarchy. In particular, we strengthen several results shown above. \ context boolean_algebra begin text \ Every Boolean algebra is a Stone algebra, a Heyting algebra and a Brouwer algebra. \ subclass stone_algebra apply unfold_locales apply (rule iffI) apply (metis compl_sup_top inf.orderI inf_bot_right inf_sup_distrib1 inf_top_right sup_inf_absorb) using inf.commute inf.sup_right_divisibility apply fastforce by simp sublocale heyting: heyting_algebra where implies = "\x y . -x \ y" apply unfold_locales apply (rule iffI) using shunting_var_p sup_commute apply fastforce using shunting_var_p sup_commute apply force by simp subclass brouwer_algebra apply unfold_locales apply (simp add: diff_eq shunting_var_p sup.commute) by (simp add: diff_eq) lemma huntington_3 [simp]: "-(-x \ -y) \ -(-x \ y) = x" using huntington_3_pp by auto lemma maddux_3_1: "x \ -x = y \ -y" by simp lemma maddux_3_4: "x \ (y \ -y) = z \ -z" by simp lemma maddux_3_11 [simp]: "(x \ y) \ (x \ -y) = x" using brouwer.maddux_3_12 sup.commute by auto lemma maddux_3_19: "(-x \ y) \ (x \ z) = (x \ y) \ (-x \ z)" using maddux_3_19_pp by auto lemma compl_inter_eq: "x \ y = x \ z \ -x \ y = -x \ z \ y = z" by (metis inf_commute maddux_3_11) lemma maddux_3_21 [simp]: "x \ (-x \ y) = x \ y" by (simp add: sup_inf_distrib1) lemma shunting_1: "x \ y \ x \ -y = bot" by (simp add: pseudo_complement) lemma uminus_involutive: "uminus \ uminus = id" by auto lemma uminus_injective: "uminus \ f = uminus \ g \ f = g" by (metis comp_assoc id_o minus_comp_minus) lemma conjugate_unique: "conjugate f g \ conjugate f h \ g = h" using conjugate_unique_p uminus_injective by blast lemma dual_additive_additive: "dual_additive (uminus \ f) \ additive f" by (metis additive_def compl_eq_compl_iff dual_additive_def p_dist_sup o_def) lemma conjugate_additive: "conjugate f g \ additive f" by (simp add: conjugate_dual_additive dual_additive_additive) lemma conjugate_isotone: "conjugate f g \ isotone f" by (simp add: conjugate_additive additive_isotone) lemma conjugate_char_1: "conjugate f g \ (\x y . f(x \ -(g y)) \ f x \ -y \ g(y \ -(f x)) \ g y \ -x)" by (simp add: conjugate_char_1_pp) lemma conjugate_char_2: "conjugate f g \ f bot = bot \ g bot = bot \ (\x y . f x \ y \ f(x \ g y) \ g y \ x \ g(y \ f x))" by (simp add: conjugate_char_2_pp) lemma shunting: "x \ y \ z \ x \ z \ -y" by (simp add: heyting.implies_galois sup.commute) lemma shunting_var: "x \ -y \ z \ x \ z \ y" by (simp add: shunting) end class non_trivial_stone_algebra = non_trivial_bounded_order + stone_algebra class non_trivial_boolean_algebra = non_trivial_stone_algebra + boolean_algebra end diff --git a/thys/Stone_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,3269 +1,3300 @@ (* 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 order.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) text \Lemma \reflexive_inf_star\ was contributed by Nicolas Robinson-O'Brien.\ 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) lemma equivalence_star_closed: "equivalence x \ equivalence (x\<^sup>\)" by (simp add: conv_star_commute star.circ_reflexive star.circ_transitive_equal) lemma equivalence_plus_closed: "equivalence x \ equivalence (x\<^sup>+)" by (simp add: conv_star_commute star.circ_reflexive star.circ_sup_one_left_unfold star.circ_transitive_equal) 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 order.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: order.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_down_closed: "x \ y \ acyclic y \ acyclic x" using comp_isotone star_isotone by fastforce lemma forest_down_closed: "x \ y \ forest y \ forest x" using conv_isotone mult_isotone star_isotone by fastforce 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 order.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: order.antisym star.circ_reflexive) qed lemma acyclic_asymmetric: "acyclic w \ asymmetric w" 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: order.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) 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: order.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 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 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 text \Lemmas \forests_bot_1\, \forests_bot_2\, \forests_bot_3\ and \fc_comp_eq_fc\ were contributed by Nicolas Robinson-O'Brien.\ 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 +lemma acyclic_plus: + "acyclic x \ acyclic (x\<^sup>+)" + by (simp add: star.circ_transitive_equal star.left_plus_circ mult_assoc) + 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 reachable_without_loops: "x\<^sup>\ = (x \ -1)\<^sup>\" proof (rule order.antisym) have "x * (x \ -1)\<^sup>\ = (x \ 1) * (x \ -1)\<^sup>\ \ (x \ -1) * (x \ -1)\<^sup>\" by (metis maddux_3_11_pp mult_right_dist_sup regular_one_closed) also have "... \ (x \ -1)\<^sup>\" by (metis inf.cobounded2 le_supI mult_left_isotone star.circ_circ_mult star.left_plus_below_circ star_involutive star_one) finally show "x\<^sup>\ \ (x \ -1)\<^sup>\" by (metis inf.cobounded2 maddux_3_11_pp regular_one_closed star.circ_circ_mult star.circ_sup_2 star_involutive star_sub_one) next show "(x \ -1)\<^sup>\ \ x\<^sup>\" by (simp add: star_isotone) qed lemma plus_reachable_without_loops: "x\<^sup>+ = (x \ -1)\<^sup>+ \ (x \ 1)" by (metis comp_associative maddux_3_11_pp regular_one_closed star.circ_back_loop_fixpoint star.circ_loop_fixpoint sup_assoc reachable_without_loops) lemma star_plus_without_loops: "x\<^sup>\ \ -1 = x\<^sup>+ \ -1" by (metis maddux_3_13 star_left_unfold_equal) 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 order.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: order.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 order.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 order.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: order.antisym mult_right_isotone star_isotone) also have "... = r\<^sup>T * t\<^sup>\" using assms(5,6) order.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) 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) order.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: order.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: order.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) order.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: order.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: order.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 +class stone_kleene_relation_algebra_tarski = stone_kleene_relation_algebra + stone_relation_algebra_tarski + +class kleene_relation_algebra_tarski = kleene_relation_algebra + stone_kleene_relation_algebra_tarski +begin + +subclass relation_algebra_tarski .. + end +class stone_kleene_relation_algebra_consistent = stone_kleene_relation_algebra + stone_relation_algebra_consistent + +class kleene_relation_algebra_consistent = kleene_relation_algebra + stone_kleene_relation_algebra_consistent +begin + +subclass relation_algebra_consistent .. + +end + +class stone_kleene_relation_algebra_tarski_consistent = stone_kleene_relation_algebra + stone_relation_algebra_tarski_consistent + +class kleene_relation_algebra_tarski_consistent = kleene_relation_algebra + stone_kleene_relation_algebra_tarski_consistent +begin + +subclass relation_algebra_tarski_consistent .. + +end + +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,1973 +1,2020 @@ (* 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 order.eq_iff top_greatest) lemma conv_dist_inf: "(x \ y)\<^sup>T = x\<^sup>T \ y\<^sup>T" apply (rule order.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 order.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 order.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 order.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 order.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 order.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 order.antisym comp_right_subdist_inf dedekind_1 order.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 order.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 test_comp_test_inf: "(x \ 1) * y * (z \ 1) = (x \ 1) * y \ y * (z \ 1)" by (smt comp_right_one comp_right_subdist_inf coreflexive_comp_top_inf inf.left_commute inf.orderE inf_le2 mult_assoc) lemma test_comp_test_top: "y \ (x \ 1) * top * (z \ 1) = (x \ 1) * y * (z \ 1)" proof - have "\u v . (v \ u\<^sup>T)\<^sup>T = v\<^sup>T \ u" using conv_dist_inf by auto thus ?thesis by (smt conv_dist_comp conv_involutive coreflexive_comp_top_inf inf.cobounded2 inf.left_commute inf.sup_monoid.add_commute symmetric_one_closed mult_assoc symmetric_top_closed) 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 order.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 order.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 lemma injective_codomain: assumes "injective x" shows "x * (x \ 1) = x \ 1" proof (rule order.antisym) show "x * (x \ 1) \ x \ 1" by (metis assms comp_right_one dual_order.trans inf.boundedI inf.cobounded1 inf.sup_monoid.add_commute mult_right_isotone one_inf_conv) next show "x \ 1 \ x * (x \ 1)" by (metis coreflexive_idempotent inf.cobounded1 inf.cobounded2 mult_left_isotone) qed 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 order.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 order.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 text \Lemmas \arc_eq_1\ and \arc_eq_2\ were contributed by Nicolas Robinson-O'Brien.\ 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: order.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 lemma points_arc: "point x \ point y \ arc (x * y\<^sup>T)" by (metis comp_associative conv_dist_comp conv_involutive equivalence_top_closed) lemma point_arc: "point x \ arc (x * x\<^sup>T)" by (simp add: points_arc) +lemma arc_expanded_1: + "arc e \ e * x * e\<^sup>T \ 1" + by (meson arc_expanded order_trans top_greatest mult_left_isotone mult_right_isotone) + +lemma arc_expanded_2: + "arc e \ e\<^sup>T * x * e \ 1" + by (meson arc_expanded order_trans top_greatest mult_left_isotone mult_right_isotone) + +lemma point_conv_comp: + "point x \ x\<^sup>T * x = top" + using order_eq_iff shunt_bijective top_greatest vector_conv_covector by blast + +lemma point_antisymmetric: + "point x \ antisymmetric x" + by (simp add: vector_covector) + +lemma mapping_inf_point_arc: + assumes "mapping x" + and "point y" + shows "arc (x \ y)" +proof (unfold arc_expanded, intro conjI) + show "(x \ y) * top * (x \ y)\<^sup>T \ 1" + by (metis assms conv_dist_comp covector_conv_vector inf.orderE inf.sup_monoid.add_commute surjective_conv_total top.extremum top_right_mult_increasing vector_export_comp) + have "(x \ y)\<^sup>T * top * (x \ y) = x\<^sup>T * y * (x \ y)" + by (simp add: assms(2) conv_dist_inf covector_inf_comp_3) + also have "... = x\<^sup>T * (y \ y\<^sup>T) * x" + by (simp add: assms(2) comp_associative covector_inf_comp_3 inf.sup_monoid.add_commute) + also have "... \ x\<^sup>T * x" + by (metis assms(2) comp_right_one mult_left_isotone mult_right_isotone vector_covector) + also have "... \ 1" + by (simp add: assms(1)) + finally show "(x \ y)\<^sup>T * top * (x \ y) \ 1" + . + show "top * (x \ y) * top = top" + by (metis assms inf_top_right inf_vector_comp mult_assoc) +qed + 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 order.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 order.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 order.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 order.eq_iff top_right_mult_increasing) lemma covector_complement_closed: "covector x \ covector (-x)" by (metis conv_complement_sub_leq conv_top order.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 order.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 order.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 order.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: order.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 order.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 order.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 +lemma antisymmetric_inf_diversity: + "antisymmetric x \ x \ -1 = x \ -x\<^sup>T" + by (smt (verit, del_insts) inf.orderE inf.sup_monoid.add_assoc inf.sup_monoid.add_commute inf_import_p one_inf_conv) + 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: order.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 distinct_points: assumes "point x" and "point y" and "x \ y" shows "x \ y = bot" by (metis assms order.antisym comp_bijective_complement inf.sup_monoid.add_commute mult_left_one pseudo_complement regular_one_closed point_in_vector_or_pseudo_complement) 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 point_atomic_vector: assumes "point x" and "vector y" and "regular y" and "y \ x" shows "y = x \ y = bot" proof (cases "x \ -y") case True thus ?thesis using assms(4) inf.absorb2 pseudo_complement by force next case False thus ?thesis using assms point_in_vector_or_pseudo_complement by fastforce qed lemma point_in_vector_or_complement_2: assumes "point x" and "vector y" and "regular y" and "\ y \ -x" shows "x \ y" using assms point_in_vector_or_pseudo_complement p_antitone_iff by fastforce text \The next three lemmas \arc_in_arc_or_complement\, \arc_in_sup_arc\ and \different_arc_in_sup_arc\ were contributed by Nicolas Robinson-O'Brien.\ 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) order.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 +lemma point_not_bot: + "point p \ p \ bot" + using consistent by force + 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 +class relation_algebra_tarski_consistent = relation_algebra + stone_relation_algebra_tarski_consistent + end