diff --git a/thys/Aggregation_Algebras/M_Choose_Component.thy b/thys/Aggregation_Algebras/M_Choose_Component.thy new file mode 100644 --- /dev/null +++ b/thys/Aggregation_Algebras/M_Choose_Component.thy @@ -0,0 +1,149 @@ +(* Title: An Operation to Select Components in Algebras with Minimisation + Author: Nicolas Robinson-O'Brien + Maintainer: Walter Guttmann +*) + +section \An Operation to Select Components in Algebras with Minimisation\ + +text \ +In this theory we show that an operation to select components of a graph can be defined in m-Kleene Algebras. +This work is by Nicolas Robinson-O'Brien. +\ + +theory M_Choose_Component + +imports + Stone_Relation_Algebras.Choose_Component + Matrix_Aggregation_Algebras + +begin + +text \ +Every \m_kleene_algebra\ is an instance of \choose_component_algebra\ when the \choose_component\ operation is defined as follows: +\ + +context m_kleene_algebra +begin + +definition "m_choose_component e v \ + if vector_classes e v then + e * minarc(v) * top + else + bot" + +sublocale m_choose_component_algebra: choose_component_algebra where choose_component = m_choose_component +proof + fix e v + show "m_choose_component e v \ -- v" + proof (cases "vector_classes e v") + case True + hence "m_choose_component e v = e * minarc(v) * top" + by (simp add: m_choose_component_def) + also have "... \ e * --v * top" + by (simp add: comp_isotone minarc_below) + also have "... = e * v * top" + using True vector_classes_def by auto + also have "... \ v * top" + using True vector_classes_def mult_assoc by auto + finally show ?thesis + using True vector_classes_def by auto + next + case False + hence "m_choose_component e v = bot" + using False m_choose_component_def by auto + thus ?thesis + by simp + qed +next + fix e v + show "vector (m_choose_component e v)" + proof (cases "vector_classes e v") + case True + thus ?thesis + by (simp add: mult_assoc m_choose_component_def) + next + case False + thus ?thesis + by (simp add: m_choose_component_def) + qed +next + fix e v + show "regular (m_choose_component e v)" + using minarc_regular regular_mult_closed vector_classes_def m_choose_component_def by auto +next + fix e v + show "m_choose_component e v * (m_choose_component e v)\<^sup>T \ e" + proof (cases "vector_classes e v") + case True + assume 1: "vector_classes e v" + hence "m_choose_component e v * (m_choose_component e v)\<^sup>T = e * minarc(v) * top * (e * minarc(v) * top)\<^sup>T" + by (simp add: m_choose_component_def) + also have "... = e * minarc(v) * top * top\<^sup>T * minarc(v)\<^sup>T * e\<^sup>T" + by (metis comp_associative conv_dist_comp) + also have "... = e * minarc(v) * top * top * minarc(v)\<^sup>T * e" + using True vector_classes_def by auto + also have "... = e * minarc(v) * top * minarc(v)\<^sup>T * e" + by (simp add: comp_associative) + also have "... \ e" + proof (cases "v = bot") + case True + thus ?thesis + by (simp add: True minarc_bot) + next + case False + assume 3: "v \ bot" + hence "e * minarc(v) * top * minarc(v)\<^sup>T \ e * 1" + using 3 minarc_arc arc_expanded comp_associative mult_right_isotone by fastforce + hence "e * minarc(v) * top * minarc(v)\<^sup>T * e \ e * 1 * e" + using mult_left_isotone by auto + also have "... = e" + using True preorder_idempotent vector_classes_def by auto + thus ?thesis + using calculation by auto + qed + thus ?thesis + by (simp add: calculation) + next + case False + thus ?thesis + by (simp add: m_choose_component_def) + qed +next + fix e v + show "e * m_choose_component e v \ m_choose_component e v" + proof (cases "vector_classes e v") + case True + thus ?thesis + using comp_right_one dual_order.eq_iff mult_isotone vector_classes_def m_choose_component_def mult_assoc by metis + next + case False + thus ?thesis + by (simp add: m_choose_component_def) + qed +next + fix e v + show "vector_classes e v \ m_choose_component e v \ bot" + proof (cases "vector_classes e v") + case True + hence "m_choose_component e v \ minarc(v) * top" + using vector_classes_def m_choose_component_def comp_associative minarc_arc shunt_bijective by fastforce + also have "... \ minarc(v)" + using calculation dual_order.trans top_right_mult_increasing by blast + thus ?thesis + using le_bot minarc_bot_iff vector_classes_def by fastforce + next + case False + thus ?thesis + by blast + qed +qed + +sublocale m_choose_component_algebra_tarski: choose_component_algebra_tarski where choose_component = m_choose_component + .. + +end + +class m_kleene_algebra_choose_component = m_kleene_algebra + choose_component_algebra + +end + diff --git a/thys/Aggregation_Algebras/Matrix_Aggregation_Algebras.thy b/thys/Aggregation_Algebras/Matrix_Aggregation_Algebras.thy --- a/thys/Aggregation_Algebras/Matrix_Aggregation_Algebras.thy +++ b/thys/Aggregation_Algebras/Matrix_Aggregation_Algebras.thy @@ -1,1017 +1,1046 @@ (* Title: Matrix Algebras for Aggregation and Minimisation Author: Walter Guttmann Maintainer: Walter Guttmann *) section \Matrix Algebras for Aggregation and Minimisation\ text \ This theory formalises aggregation orders and lattices as introduced in \<^cite>\"Guttmann2018a"\. Aggregation in these algebras is an associative and commutative operation satisfying additional properties related to the partial order and its least element. We apply the aggregation operation to finite matrices over the aggregation algebras, which shows that they form an s-algebra. By requiring the aggregation algebras to be linearly ordered, we also obtain that the matrices form an m-algebra. This is an intermediate step in demonstrating that weighted graphs form an s-algebra and an m-algebra. The present theory specifies abstract properties for the edge weights and shows that matrices over such weights form an instance of s-algebras and m-algebras. A second step taken in a separate theory gives concrete instances of edge weights satisfying the abstract properties introduced here. Lifting the aggregation to matrices requires finite sums over elements taken from commutative semigroups with an element that is a unit on the image of the semigroup operation. Because standard sums assume a commutative monoid we have to derive a number of properties of these generalised sums as their statements or proofs are different. \ theory Matrix_Aggregation_Algebras imports Stone_Kleene_Relation_Algebras.Matrix_Kleene_Algebras Aggregation_Algebras Semigroups_Big begin no_notation inf (infixl "\" 70) and uminus ("- _" [81] 80) subsection \Aggregation Orders and Finite Sums\ text \ An aggregation order is a partial order with a least element and an associative commutative operation satisfying certain properties. Axiom \add_add_bot\ introduces almost a commutative monoid; we require that \bot\ is a unit only on the image of the aggregation operation. This is necessary since it is not a unit of a number of aggregation operations we are interested in. Axiom \add_right_isotone\ states that aggregation is $\leq$-isotone on the image of the aggregation operation. Its assumption $x \neq bot$ is necessary because the introduction of new edges can decrease the aggregated value. Axiom \add_bot\ expresses that aggregation is zero-sum-free. \ class aggregation_order = order_bot + ab_semigroup_add + assumes add_right_isotone: "x \ bot \ x + bot \ y + bot \ x + z \ y + z" assumes add_add_bot [simp]: "x + y + bot = x + y" assumes add_bot: "x + y = bot \ x = bot" begin abbreviation "zero \ bot + bot" sublocale aggregation: ab_semigroup_add_0 where plus = plus and zero = zero apply unfold_locales using add_assoc add_add_bot by auto lemma add_bot_bot_bot: "x + bot + bot + bot = x + bot" by simp end text \ We introduce notation for finite sums over aggregation orders. The index variable of the summation ranges over the finite universe of its type. Finite sums are defined recursively using the binary aggregation and $\bot + \bot$ for the empty sum. \ syntax (xsymbols) "_sum_ab_semigroup_add_0" :: "idt \ 'a::bounded_semilattice_sup_bot \ 'a" ("(\\<^sub>_ _)" [0,10] 10) translations "\\<^sub>x t" => "XCONST ab_semigroup_add_0.sum_0 XCONST plus (XCONST plus XCONST bot XCONST bot) (\x . t) { x . CONST True }" text \ The following are basic properties of such sums. \ lemma agg_sum_bot: "(\\<^sub>k bot::'a::aggregation_order) = bot + bot" proof (induct rule: infinite_finite_induct) case (infinite A) thus ?case by simp next case empty thus ?case by simp next case (insert x F) thus ?case by (metis add.commute add_add_bot aggregation.sum_0.insert) qed lemma agg_sum_bot_bot: "(\\<^sub>k bot + bot::'a::aggregation_order) = bot + bot" by (rule aggregation.sum_0.neutral_const) lemma agg_sum_not_bot_1: fixes f :: "'a::finite \ 'b::aggregation_order" assumes "f i \ bot" shows "(\\<^sub>k f k) \ bot" by (metis assms add_bot aggregation.sum_0.remove finite_code mem_Collect_eq) lemma agg_sum_not_bot: fixes f :: "('a::finite,'b::aggregation_order) square" assumes "f (i,j) \ bot" shows "(\\<^sub>k \\<^sub>l f (k,l)) \ bot" by (metis assms agg_sum_not_bot_1) lemma agg_sum_distrib: fixes f g :: "'a \ 'b::aggregation_order" shows "(\\<^sub>k f k + g k) = (\\<^sub>k f k) + (\\<^sub>k g k)" by (rule aggregation.sum_0.distrib) lemma agg_sum_distrib_2: fixes f g :: "('a,'b::aggregation_order) square" shows "(\\<^sub>k \\<^sub>l f (k,l) + g (k,l)) = (\\<^sub>k \\<^sub>l f (k,l)) + (\\<^sub>k \\<^sub>l g (k,l))" proof - have "(\\<^sub>k \\<^sub>l f (k,l) + g (k,l)) = (\\<^sub>k (\\<^sub>l f (k,l)) + (\\<^sub>l g (k,l)))" by (metis (no_types) aggregation.sum_0.distrib) also have "... = (\\<^sub>k \\<^sub>l f (k,l)) + (\\<^sub>k \\<^sub>l g (k,l))" by (metis (no_types) aggregation.sum_0.distrib) finally show ?thesis . qed lemma agg_sum_add_bot: fixes f :: "'a \ 'b::aggregation_order" shows "(\\<^sub>k f k) = (\\<^sub>k f k) + bot" by (metis (no_types) add_add_bot aggregation.sum_0.F_one) lemma agg_sum_add_bot_2: fixes f :: "'a \ 'b::aggregation_order" shows "(\\<^sub>k f k + bot) = (\\<^sub>k f k)" proof - have "(\\<^sub>k f k + bot) = (\\<^sub>k f k) + (\\<^sub>k::'a bot::'b)" using agg_sum_distrib by simp also have "... = (\\<^sub>k f k) + (bot + bot)" by (metis agg_sum_bot) also have "... = (\\<^sub>k f k)" by simp finally show ?thesis by simp qed lemma agg_sum_commute: fixes f :: "('a,'b::aggregation_order) square" shows "(\\<^sub>k \\<^sub>l f (k,l)) = (\\<^sub>l \\<^sub>k f (k,l))" by (rule aggregation.sum_0.swap) lemma agg_delta: fixes f :: "'a::finite \ 'b::aggregation_order" shows "(\\<^sub>l if l = j then f l else zero) = f j + bot" apply (subst aggregation.sum_0.delta) apply simp by (metis add.commute add.left_commute add_add_bot mem_Collect_eq) lemma agg_delta_1: fixes f :: "'a::finite \ 'b::aggregation_order" shows "(\\<^sub>l if l = j then f l else bot) = f j + bot" proof - let ?f = "(\l . if l = j then f l else bot)" let ?S = "{l::'a . True}" show ?thesis proof (cases "j \ ?S") case False thus ?thesis by simp next case True let ?A = "?S - {j}" let ?B = "{j}" from True have eq: "?S = ?A \ ?B" by blast have dj: "?A \ ?B = {}" by simp have fAB: "finite ?A" "finite ?B" by auto have "aggregation.sum_0 ?f ?S = aggregation.sum_0 ?f ?A + aggregation.sum_0 ?f ?B" using aggregation.sum_0.union_disjoint[OF fAB dj, of ?f, unfolded eq [symmetric]] by simp also have "... = aggregation.sum_0 (\l . bot) ?A + aggregation.sum_0 ?f ?B" by (subst aggregation.sum_0.cong[where ?B="?A"]) simp_all also have "... = zero + aggregation.sum_0 ?f ?B" by (metis (no_types, lifting) add.commute add_add_bot aggregation.sum_0.F_g_one aggregation.sum_0.neutral) also have "... = zero + (f j + zero)" by simp also have "... = f j + bot" by (metis add.commute add.left_commute add_add_bot) finally show ?thesis . qed qed lemma agg_delta_2: fixes f :: "('a::finite,'b::aggregation_order) square" shows "(\\<^sub>k \\<^sub>l if k = i \ l = j then f (k,l) else bot) = f (i,j) + bot" proof - have "\k . (\\<^sub>l if k = i \ l = j then f (k,l) else bot) = (if k = i then f (k,j) + bot else zero)" proof fix k have "(\\<^sub>l if k = i \ l = j then f (k,l) else bot) = (\\<^sub>l if l = j then if k = i then f (k,l) else bot else bot)" by meson also have "... = (if k = i then f (k,j) else bot) + bot" by (rule agg_delta_1) finally show "(\\<^sub>l if k = i \ l = j then f (k,l) else bot) = (if k = i then f (k,j) + bot else zero)" by simp qed hence "(\\<^sub>k \\<^sub>l if k = i \ l = j then f (k,l) else bot) = (\\<^sub>k if k = i then f (k,j) + bot else zero)" using aggregation.sum_0.cong by auto also have "... = f (i,j) + bot" apply (subst agg_delta) by simp finally show ?thesis . qed subsection \Matrix Aggregation\ text \ The following definitions introduce the matrix of unit elements, componentwise aggregation and aggregation on matrices. The aggregation of a matrix is a single value, but because s-algebras are single-sorted the result has to be encoded as a matrix of the same type (size) as the input. We store the aggregated matrix value in the `first' entry of a matrix, setting all other entries to the unit value. The first entry is determined by requiring an enumeration of indices. It does not have to be the first entry; any fixed location in the matrix would work as well. \ definition zero_matrix :: "('a,'b::{plus,bot}) square" ("mzero") where "mzero = (\e . bot + bot)" definition plus_matrix :: "('a,'b::plus) square \ ('a,'b) square \ ('a,'b) square" (infixl "\\<^sub>M" 65) where "plus_matrix f g = (\e . f e + g e)" definition sum_matrix :: "('a::enum,'b::{plus,bot}) square \ ('a,'b) square" ("sum\<^sub>M _" [80] 80) where "sum_matrix f = (\(i,j) . if i = hd enum_class.enum \ j = i then \\<^sub>k \\<^sub>l f (k,l) else bot + bot)" text \ Basic properties of these operations are given in the following. \ lemma bot_plus_bot: "mbot \\<^sub>M mbot = mzero" by (simp add: plus_matrix_def bot_matrix_def zero_matrix_def) lemma sum_bot: "sum\<^sub>M (mbot :: ('a::enum,'b::aggregation_order) square) = mzero" proof (rule ext, rule prod_cases) fix i j :: "'a" have "(sum\<^sub>M mbot :: ('a,'b) square) (i,j) = (if i = hd enum_class.enum \ j = i then \\<^sub>(k::'a) \\<^sub>(l::'a) bot else bot + bot)" by (unfold sum_matrix_def bot_matrix_def) simp also have "... = bot + bot" using agg_sum_bot aggregation.sum_0.neutral by fastforce also have "... = mzero (i,j)" by (simp add: zero_matrix_def) finally show "(sum\<^sub>M mbot :: ('a,'b) square) (i,j) = mzero (i,j)" . qed lemma sum_plus_bot: fixes f :: "('a::enum,'b::aggregation_order) square" shows "sum\<^sub>M f \\<^sub>M mbot = sum\<^sub>M f" proof (rule ext, rule prod_cases) let ?h = "hd enum_class.enum" fix i j have "(sum\<^sub>M f \\<^sub>M mbot) (i,j) = (if i = ?h \ j = i then (\\<^sub>k \\<^sub>l f (k,l)) + bot else zero + bot)" by (simp add: plus_matrix_def bot_matrix_def sum_matrix_def) also have "... = (if i = ?h \ j = i then \\<^sub>k \\<^sub>l f (k,l) else zero)" by (metis (no_types, lifting) add_add_bot aggregation.sum_0.F_one) also have "... = (sum\<^sub>M f) (i,j)" by (simp add: sum_matrix_def) finally show "(sum\<^sub>M f \\<^sub>M mbot) (i,j) = (sum\<^sub>M f) (i,j)" by simp qed lemma sum_plus_zero: fixes f :: "('a::enum,'b::aggregation_order) square" shows "sum\<^sub>M f \\<^sub>M mzero = sum\<^sub>M f" by (rule ext, rule prod_cases) (simp add: plus_matrix_def zero_matrix_def sum_matrix_def) lemma agg_matrix_bot: fixes f :: "('a,'b::aggregation_order) square" assumes "\i j . f (i,j) = bot" shows "f = mbot" apply (unfold bot_matrix_def) using assms by auto text \ We consider a different implementation of matrix aggregation which stores the aggregated value in all entries of the matrix instead of a particular one. This does not require an enumeration of the indices. All results continue to hold using this alternative implementation. \ definition sum_matrix_2 :: "('a,'b::{plus,bot}) square \ ('a,'b) square" ("sum2\<^sub>M _" [80] 80) where "sum_matrix_2 f = (\e . \\<^sub>k \\<^sub>l f (k,l))" lemma sum_bot_2: "sum2\<^sub>M (mbot :: ('a,'b::aggregation_order) square) = mzero" proof fix e have "(sum2\<^sub>M mbot :: ('a,'b) square) e = (\\<^sub>(k::'a) \\<^sub>(l::'a) bot)" by (unfold sum_matrix_2_def bot_matrix_def) simp also have "... = bot + bot" using agg_sum_bot aggregation.sum_0.neutral by fastforce also have "... = mzero e" by (simp add: zero_matrix_def) finally show "(sum2\<^sub>M mbot :: ('a,'b) square) e = mzero e" . qed lemma sum_plus_bot_2: fixes f :: "('a,'b::aggregation_order) square" shows "sum2\<^sub>M f \\<^sub>M mbot = sum2\<^sub>M f" proof fix e have "(sum2\<^sub>M f \\<^sub>M mbot) e = (\\<^sub>k \\<^sub>l f (k,l)) + bot" by (simp add: plus_matrix_def bot_matrix_def sum_matrix_2_def) also have "... = (\\<^sub>k \\<^sub>l f (k,l))" by (metis (no_types, lifting) add_add_bot aggregation.sum_0.F_one) also have "... = (sum2\<^sub>M f) e" by (simp add: sum_matrix_2_def) finally show "(sum2\<^sub>M f \\<^sub>M mbot) e = (sum2\<^sub>M f) e" by simp qed lemma sum_plus_zero_2: fixes f :: "('a,'b::aggregation_order) square" shows "sum2\<^sub>M f \\<^sub>M mzero = sum2\<^sub>M f" by (simp add: plus_matrix_def zero_matrix_def sum_matrix_2_def) subsection \Aggregation Lattices\ text \ We extend aggregation orders to dense bounded distributive lattices. Axiom \add_lattice\ implements the inclusion-exclusion principle at the level of edge weights. \ class aggregation_lattice = bounded_distrib_lattice + dense_lattice + aggregation_order + assumes add_lattice: "x + y = (x \ y) + (x \ y)" text \ Aggregation lattices form a Stone relation algebra by reusing the meet operation as composition, using identity as converse and a standard implementation of pseudocomplement. \ class aggregation_algebra = aggregation_lattice + uminus + one + times + conv + assumes uminus_def [simp]: "-x = (if x = bot then top else bot)" assumes one_def [simp]: "1 = top" assumes times_def [simp]: "x * y = x \ y" assumes conv_def [simp]: "x\<^sup>T = x" begin subclass stone_algebra apply unfold_locales using bot_meet_irreducible bot_unique by auto subclass stone_relation_algebra apply unfold_locales prefer 9 using bot_meet_irreducible apply auto[1] by (simp_all add: inf.assoc le_infI2 inf_sup_distrib1 inf_sup_distrib2 inf.commute inf.left_commute) end text \ We show that matrices over aggregation lattices form an s-algebra using the above operations. \ interpretation agg_square_s_algebra: s_algebra where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix::('a::enum,'b::aggregation_algebra) square" and top = top_matrix and uminus = uminus_matrix and one = one_matrix and times = times_matrix and conv = conv_matrix and plus = plus_matrix and sum = sum_matrix proof fix f g h :: "('a,'b) square" show "f \ mbot \ sum\<^sub>M f \ sum\<^sub>M g \ h \\<^sub>M sum\<^sub>M f \ h \\<^sub>M sum\<^sub>M g" proof let ?h = "hd enum_class.enum" assume 1: "f \ mbot \ sum\<^sub>M f \ sum\<^sub>M g" hence "\k l . f (k,l) \ bot" by (meson agg_matrix_bot) hence 2: "(\\<^sub>k \\<^sub>l f (k,l)) \ bot" using agg_sum_not_bot by blast have "(\\<^sub>k \\<^sub>l f (k,l)) = (sum\<^sub>M f) (?h,?h)" by (simp add: sum_matrix_def) also have "... \ (sum\<^sub>M g) (?h,?h)" using 1 by (simp add: less_eq_matrix_def) also have "... = (\\<^sub>k \\<^sub>l g (k,l))" by (simp add: sum_matrix_def) finally have "(\\<^sub>k \\<^sub>l f (k,l)) \ (\\<^sub>k \\<^sub>l g (k,l))" by simp hence 3: "(\\<^sub>k \\<^sub>l f (k,l)) + bot \ (\\<^sub>k \\<^sub>l g (k,l)) + bot" by (metis (no_types, lifting) add_add_bot aggregation.sum_0.F_one) show "h \\<^sub>M sum\<^sub>M f \ h \\<^sub>M sum\<^sub>M g" proof (unfold less_eq_matrix_def, rule allI, rule prod_cases, unfold plus_matrix_def) fix i j have 4: "h (i,j) + (\\<^sub>k \\<^sub>l f (k,l)) \ h (i,j) + (\\<^sub>k \\<^sub>l g (k,l))" using 2 3 by (metis (no_types, lifting) add_right_isotone add.commute) have "h (i,j) + (sum\<^sub>M f) (i,j) = h (i,j) + (if i = ?h \ j = i then \\<^sub>k \\<^sub>l f (k,l) else zero)" by (simp add: sum_matrix_def) also have "... = (if i = ?h \ j = i then h (i,j) + (\\<^sub>k \\<^sub>l f (k,l)) else h (i,j) + zero)" by simp also have "... \ (if i = ?h \ j = i then h (i,j) + (\\<^sub>k \\<^sub>l g (k,l)) else h (i,j) + zero)" using 4 order.eq_iff by auto also have "... = h (i,j) + (if i = ?h \ j = i then \\<^sub>k \\<^sub>l g (k,l) else zero)" by simp finally show "h (i,j) + (sum\<^sub>M f) (i,j) \ h (i,j) + (sum\<^sub>M g) (i,j)" by (simp add: sum_matrix_def) qed qed next fix f :: "('a,'b) square" show "sum\<^sub>M f \\<^sub>M sum\<^sub>M mbot = sum\<^sub>M f" by (simp add: sum_bot sum_plus_zero) next fix f g :: "('a,'b) square" show "sum\<^sub>M f \\<^sub>M sum\<^sub>M g = sum\<^sub>M (f \ g) \\<^sub>M sum\<^sub>M (f \ g)" proof (rule ext, rule prod_cases) fix i j let ?h = "hd enum_class.enum" have "(sum\<^sub>M f \\<^sub>M sum\<^sub>M g) (i,j) = (sum\<^sub>M f) (i,j) + (sum\<^sub>M g) (i,j)" by (simp add: plus_matrix_def) also have "... = (if i = ?h \ j = i then \\<^sub>k \\<^sub>l f (k,l) else zero) + (if i = ?h \ j = i then \\<^sub>k \\<^sub>l g (k,l) else zero)" by (simp add: sum_matrix_def) also have "... = (if i = ?h \ j = i then (\\<^sub>k \\<^sub>l f (k,l)) + (\\<^sub>k \\<^sub>l g (k,l)) else zero)" by simp also have "... = (if i = ?h \ j = i then \\<^sub>k \\<^sub>l f (k,l) + g (k,l) else zero)" using agg_sum_distrib_2 by (metis (no_types)) also have "... = (if i = ?h \ j = i then \\<^sub>k \\<^sub>l (f (k,l) \ g (k,l)) + (f (k,l) \ g (k,l)) else zero)" using add_lattice aggregation.sum_0.cong by (metis (no_types, lifting)) also have "... = (if i = ?h \ j = i then \\<^sub>k \\<^sub>l (f \ g) (k,l) + (f \ g) (k,l) else zero)" by (simp add: sup_matrix_def inf_matrix_def) also have "... = (if i = ?h \ j = i then (\\<^sub>k \\<^sub>l (f \ g) (k,l)) + (\\<^sub>k \\<^sub>l (f \ g) (k,l)) else zero)" using agg_sum_distrib_2 by (metis (no_types)) also have "... = (if i = ?h \ j = i then \\<^sub>k \\<^sub>l (f \ g) (k,l) else zero) + (if i = ?h \ j = i then \\<^sub>k \\<^sub>l (f \ g) (k,l) else zero)" by simp also have "... = (sum\<^sub>M (f \ g)) (i,j) + (sum\<^sub>M (f \ g)) (i,j)" by (simp add: sum_matrix_def) also have "... = (sum\<^sub>M (f \ g) \\<^sub>M sum\<^sub>M (f \ g)) (i,j)" by (simp add: plus_matrix_def) finally show "(sum\<^sub>M f \\<^sub>M sum\<^sub>M g) (i,j) = (sum\<^sub>M (f \ g) \\<^sub>M sum\<^sub>M (f \ g)) (i,j)" . qed next fix f :: "('a,'b) square" show "sum\<^sub>M (f\<^sup>t) = sum\<^sub>M f" proof (rule ext, rule prod_cases) fix i j let ?h = "hd enum_class.enum" have "(sum\<^sub>M (f\<^sup>t)) (i,j) = (if i = ?h \ j = i then \\<^sub>k \\<^sub>l (f\<^sup>t) (k,l) else zero)" by (simp add: sum_matrix_def) also have "... = (if i = ?h \ j = i then \\<^sub>k \\<^sub>l (f (l,k))\<^sup>T else zero)" by (simp add: conv_matrix_def) also have "... = (if i = ?h \ j = i then \\<^sub>k \\<^sub>l f (l,k) else zero)" by simp also have "... = (if i = ?h \ j = i then \\<^sub>l \\<^sub>k f (l,k) else zero)" by (metis agg_sum_commute) also have "... = (sum\<^sub>M f) (i,j)" by (simp add: sum_matrix_def) finally show "(sum\<^sub>M (f\<^sup>t)) (i,j) = (sum\<^sub>M f) (i,j)" . qed qed text \ We show the same for the alternative implementation that stores the result of aggregation in all elements of the matrix. \ interpretation agg_square_s_algebra_2: s_algebra where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix::('a::finite,'b::aggregation_algebra) square" and top = top_matrix and uminus = uminus_matrix and one = one_matrix and times = times_matrix and conv = conv_matrix and plus = plus_matrix and sum = sum_matrix_2 proof fix f g h :: "('a,'b) square" show "f \ mbot \ sum2\<^sub>M f \ sum2\<^sub>M g \ h \\<^sub>M sum2\<^sub>M f \ h \\<^sub>M sum2\<^sub>M g" proof assume 1: "f \ mbot \ sum2\<^sub>M f \ sum2\<^sub>M g" hence "\k l . f (k,l) \ bot" by (meson agg_matrix_bot) hence 2: "(\\<^sub>k \\<^sub>l f (k,l)) \ bot" using agg_sum_not_bot by blast obtain c :: 'a where True by simp have "(\\<^sub>k \\<^sub>l f (k,l)) = (sum2\<^sub>M f) (c,c)" by (simp add: sum_matrix_2_def) also have "... \ (sum2\<^sub>M g) (c,c)" using 1 by (simp add: less_eq_matrix_def) also have "... = (\\<^sub>k \\<^sub>l g (k,l))" by (simp add: sum_matrix_2_def) finally have "(\\<^sub>k \\<^sub>l f (k,l)) \ (\\<^sub>k \\<^sub>l g (k,l))" by simp hence 3: "(\\<^sub>k \\<^sub>l f (k,l)) + bot \ (\\<^sub>k \\<^sub>l g (k,l)) + bot" by (metis (no_types, lifting) add_add_bot aggregation.sum_0.F_one) show "h \\<^sub>M sum2\<^sub>M f \ h \\<^sub>M sum2\<^sub>M g" proof (unfold less_eq_matrix_def, rule allI, unfold plus_matrix_def) fix e have "h e + (sum2\<^sub>M f) e = h e + (\\<^sub>k \\<^sub>l f (k,l))" by (simp add: sum_matrix_2_def) also have "... \ h e + (\\<^sub>k \\<^sub>l g (k,l))" using 2 3 by (metis (no_types, lifting) add_right_isotone add.commute) finally show "h e + (sum2\<^sub>M f) e \ h e + (sum2\<^sub>M g) e" by (simp add: sum_matrix_2_def) qed qed next fix f :: "('a,'b) square" show "sum2\<^sub>M f \\<^sub>M sum2\<^sub>M mbot = sum2\<^sub>M f" by (simp add: sum_bot_2 sum_plus_zero_2) next fix f g :: "('a,'b) square" show "sum2\<^sub>M f \\<^sub>M sum2\<^sub>M g = sum2\<^sub>M (f \ g) \\<^sub>M sum2\<^sub>M (f \ g)" proof fix e have "(sum2\<^sub>M f \\<^sub>M sum2\<^sub>M g) e = (sum2\<^sub>M f) e + (sum2\<^sub>M g) e" by (simp add: plus_matrix_def) also have "... = (\\<^sub>k \\<^sub>l f (k,l)) + (\\<^sub>k \\<^sub>l g (k,l))" by (simp add: sum_matrix_2_def) also have "... = (\\<^sub>k \\<^sub>l f (k,l) + g (k,l))" using agg_sum_distrib_2 by (metis (no_types)) also have "... = (\\<^sub>k \\<^sub>l (f (k,l) \ g (k,l)) + (f (k,l) \ g (k,l)))" using add_lattice aggregation.sum_0.cong by (metis (no_types, lifting)) also have "... = (\\<^sub>k \\<^sub>l (f \ g) (k,l) + (f \ g) (k,l))" by (simp add: sup_matrix_def inf_matrix_def) also have "... = (\\<^sub>k \\<^sub>l (f \ g) (k,l)) + (\\<^sub>k \\<^sub>l (f \ g) (k,l))" using agg_sum_distrib_2 by (metis (no_types)) also have "... = (sum2\<^sub>M (f \ g)) e + (sum2\<^sub>M (f \ g)) e" by (simp add: sum_matrix_2_def) also have "... = (sum2\<^sub>M (f \ g) \\<^sub>M sum2\<^sub>M (f \ g)) e" by (simp add: plus_matrix_def) finally show "(sum2\<^sub>M f \\<^sub>M sum2\<^sub>M g) e = (sum2\<^sub>M (f \ g) \\<^sub>M sum2\<^sub>M (f \ g)) e" . qed next fix f :: "('a,'b) square" show "sum2\<^sub>M (f\<^sup>t) = sum2\<^sub>M f" proof fix e have "(sum2\<^sub>M (f\<^sup>t)) e = (\\<^sub>k \\<^sub>l (f\<^sup>t) (k,l))" by (simp add: sum_matrix_2_def) also have "... = (\\<^sub>k \\<^sub>l (f (l,k))\<^sup>T)" by (simp add: conv_matrix_def) also have "... = (\\<^sub>k \\<^sub>l f (l,k))" by simp also have "... = (\\<^sub>l \\<^sub>k f (l,k))" by (metis agg_sum_commute) also have "... = (sum2\<^sub>M f) e" by (simp add: sum_matrix_2_def) finally show "(sum2\<^sub>M (f\<^sup>t)) e = (sum2\<^sub>M f) e" . qed qed subsection \Matrix Minimisation\ text \ We construct an operation that finds the minimum entry of a matrix. Because a matrix can have several entries with the same minimum value, we introduce a lexicographic order on the indices to make the operation deterministic. The order is obtained by enumerating the universe of the index. \ primrec enum_pos' :: "'a list \ 'a::enum \ nat" where "enum_pos' Nil x = 0" | "enum_pos' (y#ys) x = (if x = y then 0 else 1 + enum_pos' ys x)" lemma enum_pos'_inverse: "List.member xs x \ xs!(enum_pos' xs x) = x" apply (induct xs) apply (simp add: member_rec(2)) by (metis diff_add_inverse enum_pos'.simps(2) less_one member_rec(1) not_add_less1 nth_Cons') text \ The following function finds the position of an index in the enumerated universe. \ fun enum_pos :: "'a::enum \ nat" where "enum_pos x = enum_pos' (enum_class.enum::'a list) x" lemma enum_pos_inverse [simp]: "enum_class.enum!(enum_pos x) = x" apply (unfold enum_pos.simps) apply (rule enum_pos'_inverse) by (metis in_enum List.member_def) lemma enum_pos_injective [simp]: "enum_pos x = enum_pos y \ x = y" by (metis enum_pos_inverse) text \ The position in the enumerated universe determines the order. \ abbreviation enum_pos_less_eq :: "'a::enum \ 'a \ bool" where "enum_pos_less_eq x y \ (enum_pos x \ enum_pos y)" abbreviation enum_pos_less :: "'a::enum \ 'a \ bool" where "enum_pos_less x y \ (enum_pos x < enum_pos y)" sublocale enum < enum_order: order where less_eq = "\x y . enum_pos_less_eq x y" and less = "\x y . enum_pos x < enum_pos y" apply unfold_locales by auto text \ Based on this, a lexicographic order is defined on pairs, which represent locations in a matrix. \ abbreviation enum_lex_less :: "'a::enum \ 'a \ 'a \ 'a \ bool" where "enum_lex_less \ (\(i,j) (k,l) . enum_pos_less i k \ (i = k \ enum_pos_less j l))" abbreviation enum_lex_less_eq :: "'a::enum \ 'a \ 'a \ 'a \ bool" where "enum_lex_less_eq \ (\(i,j) (k,l) . enum_pos_less i k \ (i = k \ enum_pos_less_eq j l))" text \ The $m$-operation determines the location of the non-$\bot$ minimum element which is first in the lexicographic order. The result is returned as a regular matrix with $\top$ at that location and $\bot$ everywhere else. In the weighted-graph model, this represents a single unweighted edge of the graph. \ definition minarc_matrix :: "('a::enum,'b::{bot,ord,plus,top}) square \ ('a,'b) square" ("minarc\<^sub>M _" [80] 80) where "minarc_matrix f = (\e . if f e \ bot \ (\d . (f d \ bot \ (f e + bot \ f d + bot \ (enum_lex_less d e \ f e + bot \ f d + bot)))) then top else bot)" lemma minarc_at_most_one: fixes f :: "('a::enum,'b::{aggregation_order,top}) square" assumes "(minarc\<^sub>M f) e \ bot" and "(minarc\<^sub>M f) d \ bot" shows "e = d" proof - have 1: "f e + bot \ f d + bot" by (metis assms minarc_matrix_def) have "f d + bot \ f e + bot" by (metis assms minarc_matrix_def) hence "f e + bot = f d + bot" using 1 by simp hence "\ enum_lex_less d e \ \ enum_lex_less e d" using assms by (unfold minarc_matrix_def) (metis (lifting)) thus ?thesis using enum_pos_injective less_linear by auto qed subsection \Linear Aggregation Lattices\ text \ We now assume that the aggregation order is linear and forms a bounded lattice. It follows that these structures are aggregation lattices. A linear order on matrix entries is necessary to obtain a unique minimum entry. \ class linear_aggregation_lattice = linear_bounded_lattice + aggregation_order begin subclass aggregation_lattice apply unfold_locales by (metis add_commute sup_inf_selective) sublocale heyting: bounded_heyting_lattice where implies = "\x y . if x \ y then top else y" apply unfold_locales by (simp add: inf_less_eq) end text \ Every non-empty set with a transitive total relation has a least element with respect to this relation. \ lemma least_order: assumes transitive: "\x y z . le x y \ le y z \ le x z" and total: "\x y . le x y \ le y x" shows "finite A \ A \ {} \ \x . x \ A \ (\y . y \ A \ le x y)" proof (induct A rule: finite_ne_induct) case singleton thus ?case using total by auto next case insert thus ?case by (metis insert_iff transitive total) qed lemma minarc_at_least_one: fixes f :: "('a::enum,'b::linear_aggregation_lattice) square" assumes "f \ mbot" shows "\e . (minarc\<^sub>M f) e = top" proof - let ?nbe = "{ (e,f e) | e . f e \ bot }" have 1: "finite ?nbe" using finite_code finite_image_set by blast have 2: "?nbe \ {}" using assms agg_matrix_bot by fastforce let ?le = "\(e::'a \ 'a,fe::'b) (d::'a \ 'a,fd) . fe + bot \ fd + bot" have 3: "\x y z . ?le x y \ ?le y z \ ?le x z" by auto have 4: "\x y . ?le x y \ ?le y x" by (simp add: linear) have "\x . x \ ?nbe \ (\y . y \ ?nbe \ ?le x y)" by (rule least_order, rule 3, rule 4, rule 1, rule 2) then obtain e fe where 5: "(e,fe) \ ?nbe \ (\y . y \ ?nbe \ ?le (e,fe) y)" by auto let ?me = "{ e . f e \ bot \ f e + bot = fe + bot }" have 6: "finite ?me" using finite_code finite_image_set by blast have 7: "?me \ {}" using 5 by auto have 8: "\x y z . enum_lex_less_eq x y \ enum_lex_less_eq y z \ enum_lex_less_eq x z" by auto have 9: "\x y . enum_lex_less_eq x y \ enum_lex_less_eq y x" by auto have "\x . x \ ?me \ (\y . y \ ?me \ enum_lex_less_eq x y)" by (rule least_order, rule 8, rule 9, rule 6, rule 7) then obtain m where 10: "m \ ?me \ (\y . y \ ?me \ enum_lex_less_eq m y)" by auto have 11: "f m \ bot" using 10 5 by auto have 12: "\d. f d \ bot \ f m + bot \ f d + bot" using 10 5 by simp have "\d. f d \ bot \ enum_lex_less d m \ f m + bot \ f d + bot" using 10 by fastforce hence "(minarc\<^sub>M f) m = top" using 11 12 by (simp add: minarc_matrix_def) thus ?thesis by blast qed text \ Linear aggregation lattices form a Stone relation algebra by reusing the meet operation as composition, using identity as converse and a standard implementation of pseudocomplement. \ class linear_aggregation_algebra = linear_aggregation_lattice + uminus + one + times + conv + assumes uminus_def_2 [simp]: "-x = (if x = bot then top else bot)" assumes one_def_2 [simp]: "1 = top" assumes times_def_2 [simp]: "x * y = x \ y" assumes conv_def_2 [simp]: "x\<^sup>T = x" begin subclass aggregation_algebra apply unfold_locales using inf_dense by auto lemma regular_bot_top_2: "regular x \ x = bot \ x = top" by simp sublocale heyting: heyting_stone_algebra where implies = "\x y . if x \ y then top else y" apply unfold_locales apply (simp add: order.antisym) by auto end text \ We show that matrices over linear aggregation lattices form an m-algebra using the above operations. \ interpretation agg_square_m_algebra: m_algebra where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix::('a::enum,'b::linear_aggregation_algebra) square" and top = top_matrix and uminus = uminus_matrix and one = one_matrix and times = times_matrix and conv = conv_matrix and plus = plus_matrix and sum = sum_matrix and minarc = minarc_matrix proof fix f :: "('a,'b) square" show "minarc\<^sub>M f \ \\f" proof (unfold less_eq_matrix_def, rule allI) fix e :: "'a \ 'a" have "(minarc\<^sub>M f) e \ (if f e \ bot then top else --(f e))" by (simp add: minarc_matrix_def) also have "... = --(f e)" by simp also have "... = (\\f) e" by (simp add: uminus_matrix_def) finally show "(minarc\<^sub>M f) e \ (\\f) e" . qed next fix f :: "('a,'b) square" let ?at = "bounded_distrib_allegory_signature.arc mone times_matrix less_eq_matrix mtop conv_matrix" show "f \ mbot \ ?at (minarc\<^sub>M f)" proof assume 1: "f \ mbot" have "minarc\<^sub>M f \ mtop \ (minarc\<^sub>M f \ mtop)\<^sup>t = minarc\<^sub>M f \ mtop \ (minarc\<^sub>M f)\<^sup>t" by (metis matrix_bounded_idempotent_semiring.surjective_top_closed matrix_monoid.mult_assoc matrix_stone_relation_algebra.conv_dist_comp matrix_stone_relation_algebra.conv_top) also have "... \ mone" proof (unfold less_eq_matrix_def, rule allI, rule prod_cases) fix i j have "(minarc\<^sub>M f \ mtop \ (minarc\<^sub>M f)\<^sup>t) (i,j) = (\\<^sub>l (\\<^sub>k (minarc\<^sub>M f) (i,k) * mtop (k,l)) * ((minarc\<^sub>M f)\<^sup>t) (l,j))" by (simp add: times_matrix_def) also have "... = (\\<^sub>l (\\<^sub>k (minarc\<^sub>M f) (i,k) * top) * ((minarc\<^sub>M f) (j,l))\<^sup>T)" by (simp add: top_matrix_def conv_matrix_def) also have "... = (\\<^sub>l \\<^sub>k (minarc\<^sub>M f) (i,k) * top * ((minarc\<^sub>M f) (j,l))\<^sup>T)" by (metis comp_right_dist_sum) also have "... = (\\<^sub>l \\<^sub>k if i = j \ l = k then (minarc\<^sub>M f) (i,k) * top * ((minarc\<^sub>M f) (j,l))\<^sup>T else bot)" apply (rule sup_monoid.sum.cong) apply simp by (metis (no_types, lifting) comp_left_zero comp_right_zero conv_bot prod.inject minarc_at_most_one) also have "... = (if i = j then (\\<^sub>l \\<^sub>k if l = k then (minarc\<^sub>M f) (i,k) * top * ((minarc\<^sub>M f) (j,l))\<^sup>T else bot) else bot)" by auto also have "... \ (if i = j then top else bot)" by simp also have "... = mone (i,j)" by (simp add: one_matrix_def) finally show "(minarc\<^sub>M f \ mtop \ (minarc\<^sub>M f)\<^sup>t) (i,j) \ mone (i,j)" . qed finally have 2: "minarc\<^sub>M f \ mtop \ (minarc\<^sub>M f \ mtop)\<^sup>t \ mone" . have 3: "mtop \ (minarc\<^sub>M f \ mtop) = mtop" proof (rule ext, rule prod_cases) fix i j from minarc_at_least_one obtain ei ej where "(minarc\<^sub>M f) (ei,ej) = top" using 1 by force hence 4: "top * top \ (\\<^sub>l (minarc\<^sub>M f) (ei,l) * top)" by (metis comp_inf.ub_sum) have "top * (\\<^sub>l (minarc\<^sub>M f) (ei,l) * top) \ (\\<^sub>k top * (\\<^sub>l (minarc\<^sub>M f) (k,l) * top))" by (rule comp_inf.ub_sum) hence "top \ (\\<^sub>k top * (\\<^sub>l (minarc\<^sub>M f) (k,l) * top))" using 4 by auto also have "... = (\\<^sub>k mtop (i,k) * (\\<^sub>l (minarc\<^sub>M f) (k,l) * mtop (l,j)))" by (simp add: top_matrix_def) also have "... = (mtop \ (minarc\<^sub>M f \ mtop)) (i,j)" by (simp add: times_matrix_def) finally show "(mtop \ (minarc\<^sub>M f \ mtop)) (i,j) = mtop (i,j)" by (simp add: eq_iff top_matrix_def) qed have "(minarc\<^sub>M f)\<^sup>t \ mtop \ ((minarc\<^sub>M f)\<^sup>t \ mtop)\<^sup>t = (minarc\<^sub>M f)\<^sup>t \ mtop \ (minarc\<^sub>M f)" by (metis matrix_stone_relation_algebra.comp_associative matrix_stone_relation_algebra.conv_dist_comp matrix_stone_relation_algebra.conv_involutive matrix_stone_relation_algebra.conv_top matrix_bounded_idempotent_semiring.surjective_top_closed) also have "... \ mone" proof (unfold less_eq_matrix_def, rule allI, rule prod_cases) fix i j have "((minarc\<^sub>M f)\<^sup>t \ mtop \ minarc\<^sub>M f) (i,j) = (\\<^sub>l (\\<^sub>k ((minarc\<^sub>M f)\<^sup>t) (i,k) * mtop (k,l)) * (minarc\<^sub>M f) (l,j))" by (simp add: times_matrix_def) also have "... = (\\<^sub>l (\\<^sub>k ((minarc\<^sub>M f) (k,i))\<^sup>T * top) * (minarc\<^sub>M f) (l,j))" by (simp add: top_matrix_def conv_matrix_def) also have "... = (\\<^sub>l \\<^sub>k ((minarc\<^sub>M f) (k,i))\<^sup>T * top * (minarc\<^sub>M f) (l,j))" by (metis comp_right_dist_sum) also have "... = (\\<^sub>l \\<^sub>k if i = j \ l = k then ((minarc\<^sub>M f) (k,i))\<^sup>T * top * (minarc\<^sub>M f) (l,j) else bot)" apply (rule sup_monoid.sum.cong) apply simp by (metis (no_types, lifting) comp_left_zero comp_right_zero conv_bot prod.inject minarc_at_most_one) also have "... = (if i = j then (\\<^sub>l \\<^sub>k if l = k then ((minarc\<^sub>M f) (k,i))\<^sup>T * top * (minarc\<^sub>M f) (l,j) else bot) else bot)" by auto also have "... \ (if i = j then top else bot)" by simp also have "... = mone (i,j)" by (simp add: one_matrix_def) finally show "((minarc\<^sub>M f)\<^sup>t \ mtop \ (minarc\<^sub>M f)) (i,j) \ mone (i,j)" . qed finally have 5: "(minarc\<^sub>M f)\<^sup>t \ mtop \ ((minarc\<^sub>M f)\<^sup>t \ mtop)\<^sup>t \ mone" . have "mtop \ ((minarc\<^sub>M f)\<^sup>t \ mtop) = mtop" using 3 by (metis matrix_monoid.mult_assoc matrix_stone_relation_algebra.conv_dist_comp matrix_stone_relation_algebra.conv_top) thus "?at (minarc\<^sub>M f)" using 2 3 5 by blast qed next fix f g :: "('a,'b) square" let ?at = "bounded_distrib_allegory_signature.arc mone times_matrix less_eq_matrix mtop conv_matrix" show "?at g \ g \ f \ mbot \ sum\<^sub>M (minarc\<^sub>M f \ f) \ sum\<^sub>M (g \ f)" proof assume 1: "?at g \ g \ f \ mbot" hence 2: "g = \\g" using matrix_stone_relation_algebra.arc_regular by blast show "sum\<^sub>M (minarc\<^sub>M f \ f) \ sum\<^sub>M (g \ f)" proof (unfold less_eq_matrix_def, rule allI, rule prod_cases) fix i j from minarc_at_least_one obtain ei ej where 3: "(minarc\<^sub>M f) (ei,ej) = top" using 1 by force hence 4: "\k l . \(k = ei \ l = ej) \ (minarc\<^sub>M f) (k,l) = bot" by (metis (mono_tags, opaque_lifting) bot.extremum inf.bot_unique prod.inject minarc_at_most_one) from agg_matrix_bot obtain di dj where 5: "(g \ f) (di,dj) \ bot" using 1 by force hence 6: "g (di,dj) \ bot" by (metis inf_bot_left inf_matrix_def) hence 7: "g (di,dj) = top" using 2 by (metis uminus_matrix_def uminus_def) hence 8: "(g \ f) (di,dj) = f (di,dj)" by (metis inf_matrix_def inf_top.left_neutral) have 9: "\k l . k \ di \ g (k,l) = bot" proof (intro allI, rule impI) fix k l assume 10: "k \ di" have "top * (g (k,l))\<^sup>T = g (di,dj) * top * (g\<^sup>t) (l,k)" using 7 by (simp add: conv_matrix_def) also have "... \ (\\<^sub>n g (di,n) * top) * (g\<^sup>t) (l,k)" by (metis comp_inf.ub_sum comp_right_dist_sum) also have "... \ (\\<^sub>m (\\<^sub>n g (di,n) * top) * (g\<^sup>t) (m,k))" by (metis comp_inf.ub_sum) also have "... = (g \ mtop \ g\<^sup>t) (di,k)" by (simp add: times_matrix_def top_matrix_def) also have "... \ mone (di,k)" using 1 by (metis matrix_stone_relation_algebra.arc_expanded less_eq_matrix_def) also have "... = bot" apply (unfold one_matrix_def) using 10 by auto finally have "g (k,l) \ top" using 5 by (metis bot.extremum conv_def inf.bot_unique mult.left_neutral one_def) thus "g (k,l) = bot" using 2 by (metis uminus_def uminus_matrix_def) qed have "\k l . l \ dj \ g (k,l) = bot" proof (intro allI, rule impI) fix k l assume 11: "l \ dj" have "(g (k,l))\<^sup>T * top = (g\<^sup>t) (l,k) * top * g (di,dj)" using 7 by (simp add: comp_associative conv_matrix_def) also have "... \ (\\<^sub>n (g\<^sup>t) (l,n) * top) * g (di,dj)" by (metis comp_inf.ub_sum comp_right_dist_sum) also have "... \ (\\<^sub>m (\\<^sub>n (g\<^sup>t) (l,n) * top) * g (m,dj))" by (metis comp_inf.ub_sum) also have "... = (g\<^sup>t \ mtop \ g) (l,dj)" by (simp add: times_matrix_def top_matrix_def) also have "... \ mone (l,dj)" using 1 by (metis matrix_stone_relation_algebra.arc_expanded less_eq_matrix_def) also have "... = bot" apply (unfold one_matrix_def) using 11 by auto finally have "g (k,l) \ top" using 5 by (metis bot.extremum comp_right_one conv_def one_def top.extremum_unique) thus "g (k,l) = bot" using 2 by (metis uminus_def uminus_matrix_def) qed hence 12: "\k l . \(k = di \ l = dj) \ (g \ f) (k,l) = bot" using 9 by (metis inf_bot_left inf_matrix_def) have "(\\<^sub>k \\<^sub>l (minarc\<^sub>M f \ f) (k,l)) = (\\<^sub>k \\<^sub>l if k = ei \ l = ej then (minarc\<^sub>M f \ f) (k,l) else (minarc\<^sub>M f \ f) (k,l))" by simp also have "... = (\\<^sub>k \\<^sub>l if k = ei \ l = ej then (minarc\<^sub>M f \ f) (k,l) else (minarc\<^sub>M f) (k,l) \ f (k,l))" by (unfold inf_matrix_def) simp also have "... = (\\<^sub>k \\<^sub>l if k = ei \ l = ej then (minarc\<^sub>M f \ f) (k,l) else bot)" apply (rule aggregation.sum_0.cong) apply simp using 4 by (metis inf_bot_left) also have "... = (minarc\<^sub>M f \ f) (ei,ej) + bot" by (unfold agg_delta_2) simp also have "... = f (ei,ej) + bot" using 3 by (simp add: inf_matrix_def) also have "... \ (g \ f) (di,dj) + bot" using 3 5 6 7 8 by (metis minarc_matrix_def) also have "... = (\\<^sub>k \\<^sub>l if k = di \ l = dj then (g \ f) (k,l) else bot)" by (unfold agg_delta_2) simp also have "... = (\\<^sub>k \\<^sub>l if k = di \ l = dj then (g \ f) (k,l) else (g \ f) (k,l))" apply (rule aggregation.sum_0.cong) apply simp using 12 by metis also have "... = (\\<^sub>k \\<^sub>l (g \ f) (k,l))" by simp finally show "(sum\<^sub>M (minarc\<^sub>M f \ f)) (i,j) \ (sum\<^sub>M (g \ f)) (i,j)" by (simp add: sum_matrix_def) qed qed next fix f g :: "('a,'b) square" let ?h = "hd enum_class.enum" show "sum\<^sub>M f \ sum\<^sub>M g \ sum\<^sub>M g \ sum\<^sub>M f" proof (cases "(sum\<^sub>M f) (?h,?h) \ (sum\<^sub>M g) (?h,?h)") case 1: True have "sum\<^sub>M f \ sum\<^sub>M g" apply (unfold less_eq_matrix_def, rule allI, rule prod_cases) using 1 by (unfold sum_matrix_def) auto thus ?thesis by simp next case False hence 2: "(sum\<^sub>M g) (?h,?h) \ (sum\<^sub>M f) (?h,?h)" by (simp add: linear) have "sum\<^sub>M g \ sum\<^sub>M f" apply (unfold less_eq_matrix_def, rule allI, rule prod_cases) using 2 by (unfold sum_matrix_def) auto thus ?thesis by simp qed next have "finite { f :: ('a,'b) square . (\e . regular (f e)) }" by (unfold regular_bot_top_2, rule finite_set_of_finite_funs_pred) auto thus "finite { f :: ('a,'b) square . matrix_p_algebra.regular f }" by (unfold uminus_matrix_def) meson qed text \ We show the same for the alternative implementation that stores the result of aggregation in all elements of the matrix. \ interpretation agg_square_m_algebra_2: m_algebra where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix::('a::enum,'b::linear_aggregation_algebra) square" and top = top_matrix and uminus = uminus_matrix and one = one_matrix and times = times_matrix and conv = conv_matrix and plus = plus_matrix and sum = sum_matrix_2 and minarc = minarc_matrix proof fix f :: "('a,'b) square" show "minarc\<^sub>M f \ \\f" by (simp add: agg_square_m_algebra.minarc_below) next fix f :: "('a,'b) square" let ?at = "bounded_distrib_allegory_signature.arc mone times_matrix less_eq_matrix mtop conv_matrix" show "f \ mbot \ ?at (minarc\<^sub>M f)" by (simp add: agg_square_m_algebra.minarc_arc) next fix f g :: "('a,'b) square" let ?at = "bounded_distrib_allegory_signature.arc mone times_matrix less_eq_matrix mtop conv_matrix" show "?at g \ g \ f \ mbot \ sum2\<^sub>M (minarc\<^sub>M f \ f) \ sum2\<^sub>M (g \ f)" proof let ?h = "hd enum_class.enum" assume "?at g \ g \ f \ mbot" hence "sum\<^sub>M (minarc\<^sub>M f \ f) \ sum\<^sub>M (g \ f)" by (simp add: agg_square_m_algebra.minarc_min) hence "(sum\<^sub>M (minarc\<^sub>M f \ f)) (?h,?h) \ (sum\<^sub>M (g \ f)) (?h,?h)" by (simp add: less_eq_matrix_def) thus "sum2\<^sub>M (minarc\<^sub>M f \ f) \ sum2\<^sub>M (g \ f)" by (simp add: sum_matrix_def sum_matrix_2_def less_eq_matrix_def) qed next fix f g :: "('a,'b) square" let ?h = "hd enum_class.enum" have "sum\<^sub>M f \ sum\<^sub>M g \ sum\<^sub>M g \ sum\<^sub>M f" by (simp add: agg_square_m_algebra.sum_linear) hence "(sum\<^sub>M f) (?h,?h) \ (sum\<^sub>M g) (?h,?h) \ (sum\<^sub>M g) (?h,?h) \ (sum\<^sub>M f) (?h,?h)" using less_eq_matrix_def by auto thus "sum2\<^sub>M f \ sum2\<^sub>M g \ sum2\<^sub>M g \ sum2\<^sub>M f" by (simp add: sum_matrix_def sum_matrix_2_def less_eq_matrix_def) next show "finite { f :: ('a,'b) square . matrix_p_algebra.regular f }" by (simp add: agg_square_m_algebra.finite_regular) qed text \ By defining the Kleene star as $\top$ aggregation lattices form a Kleene algebra. \ class aggregation_kleene_algebra = aggregation_algebra + star + assumes star_def [simp]: "x\<^sup>\ = top" begin subclass stone_kleene_relation_algebra apply unfold_locales by (simp_all add: inf.assoc le_infI2 inf_sup_distrib1 inf_sup_distrib2) end class linear_aggregation_kleene_algebra = linear_aggregation_algebra + star + assumes star_def_2 [simp]: "x\<^sup>\ = top" begin subclass aggregation_kleene_algebra apply unfold_locales by simp end interpretation agg_square_m_kleene_algebra: m_kleene_algebra where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix::('a::enum,'b::linear_aggregation_kleene_algebra) square" and top = top_matrix and uminus = uminus_matrix and one = one_matrix and times = times_matrix and conv = conv_matrix and star = star_matrix and plus = plus_matrix and sum = sum_matrix and minarc = minarc_matrix .. interpretation agg_square_m_kleene_algebra_2: m_kleene_algebra where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix::('a::enum,'b::linear_aggregation_kleene_algebra) square" and top = top_matrix and uminus = uminus_matrix and one = one_matrix and times = times_matrix and conv = conv_matrix and star = star_matrix and plus = plus_matrix and sum = sum_matrix_2 and minarc = minarc_matrix .. +class linorder_stone_relation_algebra_plus_expansion = linorder_stone_relation_algebra_expansion + plus + + assumes plus_def: "plus = sup" +begin + +subclass linear_aggregation_algebra + apply unfold_locales + using plus_def sup_monoid.add_assoc apply blast + using plus_def sup_monoid.add_commute apply blast + using comp_inf.semiring.add_mono plus_def apply auto[1] + using plus_def apply force + using bot_eq_sup_iff plus_def apply blast + apply simp + apply simp + using times_inf apply presburger + by simp + end +class linorder_stone_kleene_relation_algebra_plus_expansion = linorder_stone_kleene_relation_algebra_expansion + linorder_stone_relation_algebra_plus_expansion +begin + +subclass linear_aggregation_kleene_algebra + apply unfold_locales + by simp + +end + +class linorder_stone_kleene_relation_algebra_tarski_consistent_plus_expansion = linorder_stone_kleene_relation_algebra_tarski_consistent_expansion + linorder_stone_kleene_relation_algebra_plus_expansion + +end + diff --git a/thys/Aggregation_Algebras/ROOT b/thys/Aggregation_Algebras/ROOT --- a/thys/Aggregation_Algebras/ROOT +++ b/thys/Aggregation_Algebras/ROOT @@ -1,12 +1,13 @@ chapter AFP session Aggregation_Algebras = Stone_Kleene_Relation_Algebras + options [timeout = 600] theories Semigroups_Big Aggregation_Algebras Matrix_Aggregation_Algebras Linear_Aggregation_Algebras + M_Choose_Component document_files "root.tex" "root.bib" diff --git a/thys/Aggregation_Algebras/document/root.tex b/thys/Aggregation_Algebras/document/root.tex --- a/thys/Aggregation_Algebras/document/root.tex +++ b/thys/Aggregation_Algebras/document/root.tex @@ -1,53 +1,55 @@ \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{$*$} \begin{document} \title{Aggregation Algebras} \author{Walter Guttmann} \maketitle \begin{abstract} We develop algebras for aggregation and minimisation for weight matrices and for edge weights in graphs. We show numerous instances of these algebras based on linearly ordered commutative semigroups. \end{abstract} \tableofcontents \section{Overview} This document describes the following four theory files: \begin{itemize} \item Big sums over semigroups generalises parts of Isabelle/HOL's theory of finite summation \texttt{Groups\_Big.thy} from commutative monoids to commutative semigroups with a unit element only on the image of the semigroup operation. \item Aggregation Algebras introduces s-algebras, m-algebras and m-Kleene-algebras with operations for aggregating the elements of a weight matrix and finding the edge with minimal weight. \item Matrix Aggregation Algebras introduces aggregation orders, aggregation lattices and linear aggregation lattices. Matrices over these structures form s-algebras and m-algebras. \item Linear Aggregation Algebras shows numerous instances based on linearly ordered commutative semigroups. They include aggregations used for the minimum weight spanning tree problem and for the minimum bottleneck spanning tree problem, as well as arbitrary t-norms and t-conorms. \end{itemize} Three theory files, which were originally part of this entry, have been moved elsewhere: \begin{itemize} \item A theory for total-correctness proofs in Hoare logic became part of Isabelle/HOL's theory \texttt{Hoare/Hoare\_Logic.thy}. \item A theory with simple total-correctness proof examples became Isabelle/HOL's theory \texttt{Hoare/ExamplesTC.thy}. \item A theory proving total correctness of Kruskal's and Prim's minimum spanning tree algorithms based on m-Kleene-algebras using Hoare logic was split into two theories that became part of AFP entry \cite{GuttmannRobinsonOBrien2020}. \end{itemize} +Following a refactoring, the selection of components of graphs in m-Kleene-algebras, which was originally part of Nicolas Robinson-O'Brien's theory \texttt{Relational\_Minimum\_Spanning\_Trees/Boruvka.thy}, has been moved into a new theory in this entry. + The development is based on Stone-Kleene relation algebras \cite{Guttmann2017b,Guttmann2017c}. The algebras for aggregation and minimisation, their application to weighted graphs and the verification of Prim's and Kruskal's minimum spanning tree algorithms, and various instances of aggregation are described in \cite{Guttmann2016c,Guttmann2018a,Guttmann2018b}. Related work is discussed in these papers. \begin{flushleft} \input{session} \end{flushleft} \bibliographystyle{abbrv} \bibliography{root} \end{document} diff --git a/thys/Relational_Disjoint_Set_Forests/Matrix_Peano_Algebras.thy b/thys/Relational_Disjoint_Set_Forests/Matrix_Peano_Algebras.thy new file mode 100644 --- /dev/null +++ b/thys/Relational_Disjoint_Set_Forests/Matrix_Peano_Algebras.thy @@ -0,0 +1,498 @@ +(* Title: Matrix Peano Algebras + Author: Walter Guttmann + Maintainer: Walter Guttmann +*) + +section \Matrix Peano Algebras\ + +text \ +We define a Boolean matrix representation of natural numbers up to \n\, where \n\ the size of an enumeration type \'a::enum\. +Numbers (obtained by \Z_matrix\ for \0\ and \N_matrix n\ for \n\) are represented as relational vectors. +The total successor function (\S_matrix\, modulo \n\) and the partial successor function (\S'_matrix\, for numbers up to \n-1\) are relations that are (partial) functions. + +We show that this representation satisfies the Peano axioms. +We also implement a function \CP_matrix\ that chooses a number in a non-empty set. +\ + +theory Matrix_Peano_Algebras + +imports Aggregation_Algebras.M_Choose_Component Relational_Disjoint_Set_Forests.More_Disjoint_Set_Forests + +begin + +definition Z_matrix :: "('a::enum,'b::{bot,top}) square" ("mZero") where "mZero = (\(i,j) . if i = hd enum_class.enum then top else bot)" +definition S_matrix :: "('a::enum,'b::{bot,top}) square" ("msuccmod") where "msuccmod = (\(i,j) . let e = (enum_class.enum :: 'a list) in if (\k . Suc k i = e ! k \ j = e ! Suc k) \ (i = e ! minus_class.minus (length e) 1 \ j = hd e) then top else bot)" +definition S'_matrix :: "('a::enum,'b::{bot,top}) square" ("msucc") where "msucc = (\(i,j) . let e = (enum_class.enum :: 'a list) in if \k . Suc k i = e ! k \ j = e ! Suc k then top else bot)" +definition N_matrix :: "nat \ ('a::enum,'b::{bot,top}) square" ("mnat") where "mnat n = (\(i,j) . if i = enum_class.enum ! n then top else bot)" +definition CP_matrix :: "('a::enum,'b::{bot,uminus}) square \ ('a,'b) square" ("mcp") where "mcp f = (\(i,j) . if Some i = find (\x . f (x,x) \ bot) enum_class.enum then uminus_class.uminus (uminus_class.uminus (f (i,j))) else bot)" + +lemma N_matrix_power_S: + "n < length (enum_class.enum :: 'a list) \ mnat n = (matrix_monoid.power (msuccmod\<^sup>t) n) \ (mZero :: ('a::enum,'b::stone_relation_algebra) square)" +proof (induct n) + let ?z = "mZero :: ('a,'b) square" + let ?s = "msuccmod :: ('a,'b) square" + let ?e = "enum_class.enum :: 'a list" + let ?h = "hd ?e" + let ?l = "length ?e" + let ?g = "?e ! minus_class.minus ?l 1" + let ?p = "matrix_monoid.power (?s\<^sup>t)" + case 0 + have "(UNIV :: 'a set) \ {}" + by simp + hence "?h = ?e ! 0" + by (simp add: hd_conv_nth UNIV_enum) + thus ?case + by (simp add: N_matrix_def Z_matrix_def) + case (Suc n) + assume 1: "n < ?l \ mnat n = ?p n \ ?z" + show "Suc n < ?l \ mnat (Suc n) = ?p (Suc n) \ ?z" + proof + assume 2: "Suc n < ?l" + hence "n < ?l" + by simp + hence "\l l = n)" + using nth_eq_iff_index_eq enum_distinct by auto + hence 3: "\i . (\l i = ?e ! Suc l) \ (i = ?e ! Suc n)" + by auto + have 4: "\i . (\l . Suc l ?e ! n = ?e ! l \ i = ?e ! Suc l) \ (i = ?e ! Suc n)" + apply (rule iffI) + using 3 apply (metis Suc_lessD) + using 2 by auto + show "mnat (Suc n) = ?p (Suc n) \ ?z" + proof (rule ext, rule prod_cases) + fix i j :: 'a + have "(?p (Suc n) \ ?z) (i,j) = (?s\<^sup>t \ mnat n) (i,j)" + using 1 2 by (simp add: matrix_monoid.mult_assoc) + also have "... = (\\<^sub>k ((?s (k,i))\<^sup>T * mnat n (k,j)))" + by (simp add: times_matrix_def conv_matrix_def) + also have "... = (\\<^sub>k ((if (\l . Suc l k = ?e ! l \ i = ?e ! Suc l) \ (k = ?g \ i = ?h) then top else bot)\<^sup>T * (if k = ?e ! n then top else bot)))" + by (simp add: S_matrix_def N_matrix_def) + also have "... = (\\<^sub>k ((if (\l . Suc l k = ?e ! l \ i = ?e ! Suc l) \ (k = ?g \ i = ?h) then top else bot) * (if k = ?e ! n then top else bot)))" + by (smt (verit, best) sup_monoid.sum.cong symmetric_bot_closed symmetric_top_closed) + also have "... = (\\<^sub>k (if (\l . Suc l k = ?e ! l \ i = ?e ! Suc l \ k = ?e ! n) \ (k = ?g \ i = ?h \ k = ?e ! n) then top else bot))" + by (smt (verit, best) covector_bot_closed idempotent_bot_closed sup_monoid.sum.cong surjective_top_closed vector_bot_closed) + also have "... = (if \l . Suc l ?e ! n = ?e ! l \ i = ?e ! Suc l then top else bot)" + proof - + have "\k . \(k = ?g \ i = ?h \ k = ?e ! n)" + using 2 distinct_conv_nth[of ?e] enum_distinct by auto + thus ?thesis + by (smt (verit, del_insts) comp_inf.ub_sum sup.order_iff sup_monoid.sum.neutral sup_top_right) + qed + also have "... = (if i = ?e ! Suc n then top else bot)" + using 4 by simp + also have "... = mnat (Suc n) (i,j)" + by (simp add: N_matrix_def) + finally show "mnat (Suc n) (i,j) = (?p (Suc n) \ ?z) (i,j)" + by simp + qed + qed +qed + +lemma N_matrix_power_S': + "n < length (enum_class.enum :: 'a list) \ mnat n = (matrix_monoid.power (msucc\<^sup>t) n) \ (mZero :: ('a::enum,'b::stone_relation_algebra) square)" +proof (induct n) + let ?z = "mZero :: ('a,'b) square" + let ?s = "msucc :: ('a,'b) square" + let ?e = "enum_class.enum :: 'a list" + let ?h = "hd ?e" + let ?l = "length ?e" + let ?p = "matrix_monoid.power (?s\<^sup>t)" + case 0 + have "(UNIV :: 'a set) \ {}" + by simp + hence "?h = ?e ! 0" + by (simp add: hd_conv_nth UNIV_enum) + thus ?case + by (simp add: N_matrix_def Z_matrix_def) + case (Suc n) + assume 1: "n < ?l \ mnat n = ?p n \ ?z" + show "Suc n < ?l \ mnat (Suc n) = ?p (Suc n) \ ?z" + proof + assume 2: "Suc n < ?l" + hence "n < ?l" + by simp + hence "\l l = n)" + using nth_eq_iff_index_eq enum_distinct by auto + hence 3: "\i . (\l i = ?e ! Suc l) \ (i = ?e ! Suc n)" + by auto + have 4: "\i . (\l . Suc l ?e ! n = ?e ! l \ i = ?e ! Suc l) \ (i = ?e ! Suc n)" + apply (rule iffI) + using 3 apply (metis Suc_lessD) + using 2 by auto + show "mnat (Suc n) = ?p (Suc n) \ ?z" + proof (rule ext, rule prod_cases) + fix i j :: 'a + have "(?p (Suc n) \ ?z) (i,j) = (?s\<^sup>t \ mnat n) (i,j)" + using 1 2 by (simp add: matrix_monoid.mult_assoc) + also have "... = (\\<^sub>k ((?s (k,i))\<^sup>T * mnat n (k,j)))" + by (simp add: times_matrix_def conv_matrix_def) + also have "... = (\\<^sub>k ((if \l . Suc l k = ?e ! l \ i = ?e ! Suc l then top else bot)\<^sup>T * (if k = ?e ! n then top else bot)))" + by (simp add: S'_matrix_def N_matrix_def) + also have "... = (\\<^sub>k ((if \l . Suc l k = ?e ! l \ i = ?e ! Suc l then top else bot) * (if k = ?e ! n then top else bot)))" + by (smt (verit, best) sup_monoid.sum.cong symmetric_bot_closed symmetric_top_closed) + also have "... = (\\<^sub>k (if \l . Suc l k = ?e ! l \ i = ?e ! Suc l \ k = ?e ! n then top else bot))" + by (smt (verit, best) covector_bot_closed idempotent_bot_closed sup_monoid.sum.cong surjective_top_closed vector_bot_closed) + also have "... = (if \l . Suc l ?e ! n = ?e ! l \ i = ?e ! Suc l then top else bot)" + by (smt (verit, del_insts) comp_inf.ub_sum sup.order_iff sup_monoid.sum.neutral sup_top_right) + also have "... = (if i = ?e ! Suc n then top else bot)" + using 4 by simp + also have "... = mnat (Suc n) (i,j)" + by (simp add: N_matrix_def) + finally show "mnat (Suc n) (i,j) = (?p (Suc n) \ ?z) (i,j)" + by simp + qed + qed +qed + +syntax + "_sum_sup_monoid" :: "idt \ nat \ 'a::bounded_semilattice_sup_bot \ 'a" ("(\_<_ . _)" [0,51,10] 10) +translations + "\x "XCONST sup_monoid.sum (\x . t) { x . x < y }" + +context bounded_semilattice_sup_bot +begin + +lemma ub_sum_nat: + fixes f :: "nat \ 'a" + assumes "i < l" + shows "f i \ (\k 'a" + assumes "\k x" + shows "(\k x" + apply (rule finite_subset_induct[where A="{k . k < l}"]) + by (simp_all add: assms) + +end + +lemma ext_sum_nat: + fixes l :: nat + shows "(\kk mtop = ?z" + proof (rule ext, rule prod_cases) + fix i j :: 'a + have "(?z \ mtop) (i,j) = (\\<^sub>k (?z (i,k) * top))" + by (simp add: times_matrix_def top_matrix_def) + also have "... = (\\<^sub>k::'a (if i = ?h then top else bot) * top)" + by (simp add: Z_matrix_def) + also have "... = (if i = ?h then top else bot) * (top :: 'b)" + using sum_const by blast + also have "... = ?z (i,j)" + by (simp add: Z_matrix_def) + finally show "(?z \ mtop) (i,j) = ?z (i,j)" + . + qed + have 2: "?z \ ?z\<^sup>t \ mone" + proof (unfold less_eq_matrix_def, rule allI, rule prod_cases) + fix i j :: 'a + have "(?z \ ?z\<^sup>t) (i,j) = (\\<^sub>k (?z (i,k) * (?z (j,k))\<^sup>T))" + by (simp add: times_matrix_def conv_matrix_def) + also have "... = (\\<^sub>k::'a (if i = ?h then top else bot) * (if j = ?h then top else bot))" + by (simp add: Z_matrix_def) + also have "... = (if i = ?h then top else bot) * (if j = ?h then top else bot)" + using sum_const by blast + also have "... \ mone (i,j)" + by (simp add: one_matrix_def) + finally show "(?z \ ?z\<^sup>t) (i,j) \ mone (i,j)" + . + qed + have 3: "mtop \ ?z = mtop" + proof (rule ext, rule prod_cases) + fix i j :: 'a + have "mtop (i,j) = (top::'b) * (if ?h = ?h then top else bot)" + by (simp add: top_matrix_def) + also have "... \ (\\<^sub>k::'a (top * (if k = ?h then top else bot)))" + by (rule ub_sum) + also have "... = (\\<^sub>k (top * ?z (k,j)))" + by (simp add: Z_matrix_def) + also have "... = (mtop \ ?z) (i,j)" + by (simp add: times_matrix_def top_matrix_def) + finally show "(mtop \ ?z) (i,j) = mtop (i,j)" + by (simp add: inf.le_bot top_matrix_def) + qed + show "matrix_stone_relation_algebra.point ?z" + using 1 2 3 by simp + have "\i (j::'a) (k::'a) . (\lm i = ?e ! Suc l \ k = ?e ! m \ j = ?e ! Suc m) \ i = j" + using distinct_conv_nth enum_distinct by auto + hence 4: "\i (j::'a) (k::'a) . (\l m . Suc l Suc m k = ?e ! l \ i = ?e ! Suc l \ k = ?e ! m \ j = ?e ! Suc m) \ i = j" + by (metis Suc_lessD) + show "?s\<^sup>t \ ?s \ mone" + proof (unfold less_eq_matrix_def, rule allI, rule prod_cases) + fix i j :: 'a + have "(?s\<^sup>t \ ?s) (i,j) = (\\<^sub>k (?s (k,i) * ?s (k,j)))" + by (simp add: times_matrix_def conv_matrix_def) + also have "... = (\\<^sub>k::'a ((if (\l . Suc l k = ?e ! l \ i = ?e ! Suc l) \ (k = ?g \ i = ?h) then top else bot) * (if (\m . Suc m k = ?e ! m \ j = ?e ! Suc m) \ (k = ?g \ j = ?h) then top else bot)))" + by (simp add: S_matrix_def) + also have "... = (\\<^sub>k::'a (if (\l m . Suc l Suc m k = ?e ! l \ i = ?e ! Suc l \ k = ?e ! m \ j = ?e ! Suc m) \ (k = ?g \ i = ?h \ j = ?h) then top else bot))" + proof - + have 5: "\k . \((\l . Suc l k = ?e ! l \ i = ?e ! Suc l) \ (k = ?g \ j = ?h))" + using distinct_conv_nth[of ?e] enum_distinct by auto + have "\k . \((k = ?g \ i = ?h) \ (\m . Suc m k = ?e ! m \ j = ?e ! Suc m))" + using distinct_conv_nth[of ?e] enum_distinct by auto + thus ?thesis + using 5 by (smt (verit) covector_bot_closed idempotent_bot_closed sup_monoid.sum.cong surjective_top_closed vector_bot_closed) + qed + also have "... \ (\\<^sub>k::'a (if i = j then top else bot))" + using 4 by (smt (verit, best) comp_inf.ub_sum order_top_class.top_greatest sup_monoid.sum.not_neutral_contains_not_neutral top.extremum_uniqueI) + also have "... \ (if i = j then top else bot)" + by (simp add: sum_const) + also have "... = mone (i,j)" + by (simp add: one_matrix_def) + finally show "(?s\<^sup>t \ ?s) (i,j) \ mone (i,j)" + . + qed + have 6: "\i (j::'a) (k::'a) . (\l m . Suc l Suc m i = ?e ! l \ k = ?e ! Suc l \ j = ?e ! m \ k = ?e ! Suc m) \ i = j" + using distinct_conv_nth enum_distinct by auto + show "?s \ ?s\<^sup>t \ mone" + proof (unfold less_eq_matrix_def, rule allI, rule prod_cases) + fix i j :: 'a + have "(?s \ ?s\<^sup>t) (i,j) = (\\<^sub>k (?s (i,k) * ?s (j,k)))" + by (simp add: times_matrix_def conv_matrix_def) + also have "... = (\\<^sub>k::'a ((if (\l . Suc l i = ?e ! l \ k = ?e ! Suc l) \ (i = ?g \ k = ?h) then top else bot) * (if (\m . Suc m j = ?e ! m \ k = ?e ! Suc m) \ (j = ?g \ k = ?h) then top else bot)))" + by (simp add: S_matrix_def) + also have "... = (\\<^sub>k::'a (if (\l m . Suc l Suc m i = ?e ! l \ k = ?e ! Suc l \ j = ?e ! m \ k = ?e ! Suc m) \ (i = ?g \ k = ?h \ j = ?g) then top else bot))" + proof - + have 7: "\l . Suc l 0k . \((\l . Suc l i = ?e ! l \ k = ?e ! Suc l) \ (j = ?g \ k = ?h))" + using 7 8 distinct_conv_nth[of ?e] enum_distinct by auto + have "\k . \((i = ?g \ k = ?h) \ (\m . Suc m j = ?e ! m \ k = ?e ! Suc m))" + using 7 8 distinct_conv_nth[of ?e] enum_distinct by auto + thus ?thesis + using 9 by (smt (verit) covector_bot_closed idempotent_bot_closed sup_monoid.sum.cong surjective_top_closed vector_bot_closed) + qed + also have "... \ (\\<^sub>k::'a (if i = j then top else bot))" + using 6 by (smt (verit, best) comp_inf.ub_sum order_top_class.top_greatest sup_monoid.sum.not_neutral_contains_not_neutral top.extremum_uniqueI) + also have "... \ (if i = j then top else bot)" + by simp + also have "... = mone (i,j)" + by (simp add: one_matrix_def) + finally show "(?s \ ?s\<^sup>t) (i,j) \ mone (i,j)" + . + qed + have "(mtop :: ('a,'b) square) = (\kk set ?e" + using UNIV_enum by auto + from this obtain k where 6: "k < ?l \ i = ?e ! k" + by (metis in_set_conv_nth) + hence "(\k . if i = ?e ! k then top else bot) k \ (\k (\kkkkk ?z)" + proof - + have "\k . k mnat k = ?p k \ ?z" + using N_matrix_power_S by auto + thus ?thesis + by (metis (no_types, lifting) mem_Collect_eq sup_monoid.sum.cong) + qed + also have "... \ ?s\<^sup>t\<^sup>\ \ ?z" + proof (unfold less_eq_matrix_def, rule allI, rule prod_cases) + fix i j :: 'a + have "(\k ?z) (i,j) = (\k ?z) (i,j))" + by (metis ext_sum_nat) + also have "... \ (?s\<^sup>t\<^sup>\ \ ?z) (i,j)" + apply (rule lub_sum_nat) + by (metis less_eq_matrix_def matrix_idempotent_semiring.mult_left_isotone matrix_kleene_algebra.star.power_below_circ) + finally show "(\k ?z) (i,j) \ (?s\<^sup>t\<^sup>\ \ ?z) (i,j)" + . + qed + finally show "?s\<^sup>t\<^sup>\ \ ?z = mtop" + by (simp add: matrix_order.antisym_conv) +qed + +interpretation matrix_skra_peano_2: skra_peano_2 where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix :: ('a::enum,'b::linorder_stone_kleene_relation_algebra_tarski_consistent_expansion) square" and top = top_matrix and uminus = uminus_matrix and one = one_matrix and times = times_matrix and conv = conv_matrix and star = star_matrix and Z = Z_matrix and S = S_matrix +proof + let ?s = "msuccmod :: ('a,'b) square" + let ?e = "enum_class.enum :: 'a list" + let ?h = "hd ?e" + let ?l = "length ?e" + let ?g = "?e ! minus_class.minus ?l 1" + show "matrix_bounded_idempotent_semiring.total ?s" + proof (rule ext, rule prod_cases) + fix i j :: 'a + have "(?s \ mtop) (i,j) = (\\<^sub>k (?s (i,k) * top))" + by (simp add: times_matrix_def top_matrix_def) + also have "... = (\\<^sub>k::'a if (\l . Suc l i = ?e ! l \ k = ?e ! Suc l) \ (i = ?g \ k = ?h) then top else bot)" + by (simp add: S_matrix_def) + also have "... = top" + proof - + have "\i . (\l . Suc l i = ?e ! l) \ i = ?g" + by (metis in_set_conv_nth in_enum Suc_lessI diff_Suc_1) + hence "\i . \k . (\l . Suc l i = ?e ! l \ k = ?e ! Suc l) \ (i = ?g \ k = ?h)" + by blast + thus ?thesis + by (smt (verit, ccfv_threshold) comp_inf.ub_sum top.extremum_uniqueI) + qed + finally show "(?s \ mtop) (i,j) = mtop (i,j)" + by (simp add: top_matrix_def) + qed +qed + +interpretation matrix_skra_peano_3: skra_peano_3 where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix :: ('a::enum,'b::linorder_stone_kleene_relation_algebra_tarski_consistent_expansion) square" and top = top_matrix and uminus = uminus_matrix and one = one_matrix and times = times_matrix and conv = conv_matrix and star = star_matrix and Z = Z_matrix and S = S_matrix +proof (unfold_locales, rule finite_surj) + show "finite (UNIV :: 'a rel set)" + by simp + let ?f = "\R p . if p \ R then top else bot" + show "{ f :: ('a,'b) square . matrix_p_algebra.regular f } \ range ?f" + proof + fix f :: "('a,'b) square" + let ?R = "{ (x,y) . f (x,y) = top }" + assume "f \ { f . matrix_p_algebra.regular f }" + hence "matrix_p_algebra.regular f" + by simp + hence "\p . f p \ top \ f p = bot" + by (metis linorder_stone_algebra_expansion_class.uminus_def uminus_matrix_def) + hence "f = ?f ?R" + by fastforce + thus "f \ range ?f" + by blast + qed +qed + +interpretation matrix_skra_peano_4: skra_peano_4 where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix :: ('a::enum,'b::linorder_stone_kleene_relation_algebra_tarski_consistent_plus_expansion) square" and top = top_matrix and uminus = uminus_matrix and one = one_matrix and times = times_matrix and conv = conv_matrix and star = star_matrix and Z = Z_matrix and S = S_matrix and choose_point = agg_square_m_kleene_algebra_2.m_choose_component_algebra_tarski.choose_component_point + apply unfold_locales + apply (simp add: agg_square_m_kleene_algebra_2.m_choose_component_algebra_tarski.choose_component_point_point) + by (simp add: agg_square_m_kleene_algebra_2.m_choose_component_algebra_tarski.choose_component_point_decreasing) + +interpretation matrix'_skra_peano_1: skra_peano_1 where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix :: ('a::enum,'b::linorder_stone_kleene_relation_algebra_tarski_consistent_expansion) square" and top = top_matrix and uminus = uminus_matrix and one = one_matrix and times = times_matrix and conv = conv_matrix and star = star_matrix and Z = Z_matrix and S = S'_matrix +proof + let ?z = "mZero :: ('a,'b) square" + let ?s = "msucc :: ('a,'b) square" + let ?e = "enum_class.enum :: 'a list" + let ?l = "length ?e" + let ?p = "matrix_monoid.power (?s\<^sup>t)" + show "matrix_stone_relation_algebra.point ?z" + using matrix_skra_peano_1.Z_point by auto + have "\i (j::'a) (k::'a) . (\lm i = ?e ! Suc l \ k = ?e ! m \ j = ?e ! Suc m) \ i = j" + using distinct_conv_nth enum_distinct by auto + hence 4: "\i (j::'a) (k::'a) . (\l m . Suc l Suc m k = ?e ! l \ i = ?e ! Suc l \ k = ?e ! m \ j = ?e ! Suc m) \ i = j" + by (metis Suc_lessD) + show "?s\<^sup>t \ ?s \ mone" + proof (unfold less_eq_matrix_def, rule allI, rule prod_cases) + fix i j :: 'a + have "(?s\<^sup>t \ ?s) (i,j) = (\\<^sub>k (?s (k,i) * ?s (k,j)))" + by (simp add: times_matrix_def conv_matrix_def) + also have "... = (\\<^sub>k::'a ((if \l . Suc l k = ?e ! l \ i = ?e ! Suc l then top else bot) * (if \m . Suc m k = ?e ! m \ j = ?e ! Suc m then top else bot)))" + by (simp add: S'_matrix_def) + also have "... = (\\<^sub>k::'a (if (\l m . Suc l Suc m k = ?e ! l \ i = ?e ! Suc l \ k = ?e ! m \ j = ?e ! Suc m) then top else bot))" + by (smt (verit) covector_bot_closed idempotent_bot_closed sup_monoid.sum.cong surjective_top_closed vector_bot_closed) + also have "... \ (\\<^sub>k::'a (if i = j then top else bot))" + using 4 by (smt (verit, best) comp_inf.ub_sum order_top_class.top_greatest sup_monoid.sum.not_neutral_contains_not_neutral top.extremum_uniqueI) + also have "... \ (if i = j then top else bot)" + by (simp add: sum_const) + also have "... = mone (i,j)" + by (simp add: one_matrix_def) + finally show "(?s\<^sup>t \ ?s) (i,j) \ mone (i,j)" + . + qed + have 5: "\i (j::'a) (k::'a) . (\l m . Suc l Suc m i = ?e ! l \ k = ?e ! Suc l \ j = ?e ! m \ k = ?e ! Suc m) \ i = j" + using distinct_conv_nth enum_distinct by auto + show "?s \ ?s\<^sup>t \ mone" + proof (unfold less_eq_matrix_def, rule allI, rule prod_cases) + fix i j :: 'a + have "(?s \ ?s\<^sup>t) (i,j) = (\\<^sub>k (?s (i,k) * ?s (j,k)))" + by (simp add: times_matrix_def conv_matrix_def) + also have "... = (\\<^sub>k::'a ((if \l . Suc l i = ?e ! l \ k = ?e ! Suc l then top else bot) * (if \m . Suc m j = ?e ! m \ k = ?e ! Suc m then top else bot)))" + by (simp add: S'_matrix_def) + also have "... = (\\<^sub>k::'a (if (\l m . Suc l Suc m i = ?e ! l \ k = ?e ! Suc l \ j = ?e ! m \ k = ?e ! Suc m) then top else bot))" + by (smt (verit) covector_bot_closed idempotent_bot_closed sup_monoid.sum.cong surjective_top_closed vector_bot_closed) + also have "... \ (\\<^sub>k::'a (if i = j then top else bot))" + using 5 by (smt (verit, best) comp_inf.ub_sum order_top_class.top_greatest sup_monoid.sum.not_neutral_contains_not_neutral top.extremum_uniqueI) + also have "... \ (if i = j then top else bot)" + by simp + also have "... = mone (i,j)" + by (simp add: one_matrix_def) + finally show "(?s \ ?s\<^sup>t) (i,j) \ mone (i,j)" + . + qed + have "(mtop :: ('a,'b) square) = (\kk set ?e" + using UNIV_enum by auto + from this obtain k where 6: "k < ?l \ i = ?e ! k" + by (metis in_set_conv_nth) + hence "(\k . if i = ?e ! k then top else bot) k \ (\k (\kkkkk ?z)" + proof - + have "\k . k mnat k = ?p k \ ?z" + using N_matrix_power_S' by auto + thus ?thesis + by (metis (no_types, lifting) mem_Collect_eq sup_monoid.sum.cong) + qed + also have "... \ ?s\<^sup>t\<^sup>\ \ ?z" + proof (unfold less_eq_matrix_def, rule allI, rule prod_cases) + fix i j :: 'a + have "(\k ?z) (i,j) = (\k ?z) (i,j))" + by (metis ext_sum_nat) + also have "... \ (?s\<^sup>t\<^sup>\ \ ?z) (i,j)" + apply (rule lub_sum_nat) + by (metis less_eq_matrix_def matrix_idempotent_semiring.mult_left_isotone matrix_kleene_algebra.star.power_below_circ) + finally show "(\k ?z) (i,j) \ (?s\<^sup>t\<^sup>\ \ ?z) (i,j)" + . + qed + finally show "?s\<^sup>t\<^sup>\ \ ?z = mtop" + by (simp add: matrix_order.antisym_conv) +qed + +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 --- a/thys/Relational_Disjoint_Set_Forests/More_Disjoint_Set_Forests.thy +++ b/thys/Relational_Disjoint_Set_Forests/More_Disjoint_Set_Forests.thy @@ -1,3187 +1,3187 @@ (* 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" text \ Using the \choose_point\ operation we define a simple for-each-loop abstraction as syntactic sugar translated to a while-loop. Regular vector \h\ describes the set of all elements that are yet to be processed. It is made explicit so that the invariant can refer to it. \ syntax "_Foreach" :: "idt \ idt \ 'assn \ 'com \ 'com" ("(1FOREACH _/ USING _/ INV {_} //DO _ /OD)" [0,0,0,0] 61) translations "FOREACH x USING h INV { i } DO c OD" => "h := CONST top; WHILE h \ CONST bot INV { CONST regular h \ CONST vector h \ i } VAR { h\ } DO x := CONST choose_point h; c; h[x] := CONST bot OD" 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 ] FOREACH x USING h INV { p - h = 1 - h } DO p := make_set p x 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 } < 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 } < 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 \ ((p[y\p[[p[[y]]]]])\<^sup>T\<^sup>\ * ((p[y\p[[p[[y]]]]])[[y]]))\ < (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 \ (?pyppy\<^sup>T\<^sup>\ * (?pyppy[[y]]))\ < (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 { (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 \ ((p[y\p[[p[[y]]]]])\<^sup>T\<^sup>\ * (p[[y]]))\ < (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 \ (?pyppy\<^sup>T\<^sup>\ * (p[[y]]))\ < (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 { (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 + +class skra_peano_1 = stone_kleene_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 (roots p) (-(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_ranks: "VARS h p x rank [ True ] FOREACH x USING h INV { p - h = 1 - h \ rank - h = Z\<^sup>T - h } DO p := make_set p x; rank[x] := Z 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 } < 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 } < 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 (roots 1) (-(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 syntax "_Cond1" :: "'bexp \ 'com \ 'com" ("(1IF _/ THEN _/ FI)" [0,0] 61) translations "IF b THEN c FI" == "IF b THEN c ELSE SKIP FI" 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 (roots p) (-(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]]) FI FI 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 "roots ?p3 \ roots ?p2" by (simp add: mult_left_isotone) thus "card_less_eq (roots ?p3) (-(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 (roots ?p2) (-(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 \ roots ?p2 \ 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 "... \ roots ?p2" 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 "roots ?p4 \ ?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 "roots ?p4 \ roots (-?s \ ?p2)" by (simp add: mult_left_isotone) also have "... = -?s \ roots ?p2" using 5 inf_assoc vector_complement_closed vector_inf_comp by auto also have "... = (i * ?rr \ -?s \ roots ?p2) \ (-(i * ?rr) \ -?s \ roots ?p2)" 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 \ roots ?p2) \ ?s * (i * S'\<^sup>\ * i\<^sup>T * ?s \ -?s \ roots ?p2)" 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 \ roots ?p2)" 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 \ roots ?p2)" 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 \ roots ?p2)" using comp_inf.comp_right_dist_sup by simp also have "... \ ?s * (i * S'\<^sup>+ * i\<^sup>T * ?s \ roots ?p2)" 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 \ roots ?p2) \ i * -(S'\<^sup>+ * ?rank\<^sup>T * top)" . hence "i * ?rr \ -?s \ roots ?p2 \ ?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 \ roots ?p2 \ 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 \ roots ?p2 \ ?i * -(S'\<^sup>+ * ?rank\<^sup>T * top)" . have "-(i * ?rr) \ roots ?p2 \ -(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 \ roots ?p2 \ -(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 \ roots ?p2 \ ?i * -(S'\<^sup>+ * ?rank\<^sup>T * top)" . qed finally show ?thesis . qed thus "card_less_eq (roots ?p4) (-(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 "... \ roots ?p2" 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 "... \ roots ?p2" by simp finally have "?r \ ?p2 \ roots ?p2 \ ?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 "roots ?p4 \ roots ?p2" by (simp add: mult_left_isotone) thus "card_less_eq (roots ?p4) (-(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,17 +1,18 @@ chapter AFP -session Relational_Disjoint_Set_Forests = Stone_Kleene_Relation_Algebras + +session Relational_Disjoint_Set_Forests = Aggregation_Algebras + options [timeout = 600] sessions "HOL-Hoare" theories Disjoint_Set_Forests More_Disjoint_Set_Forests + Matrix_Peano_Algebras document_files "root.tex" "root.bib" diff --git a/thys/Relational_Disjoint_Set_Forests/document/root.bib b/thys/Relational_Disjoint_Set_Forests/document/root.bib --- a/thys/Relational_Disjoint_Set_Forests/document/root.bib +++ b/thys/Relational_Disjoint_Set_Forests/document/root.bib @@ -1,199 +1,209 @@ @STRING{afp = {Archive of Formal Proofs}} +@STRING{arxiv = {arXiv}} @STRING{fac = {Formal Aspects of Computing}} @STRING{is = {Information Sciences}} @STRING{jlamp = {Journal of Logical and Algebraic Methods in Programming}} @STRING{lncs = {Lecture Notes in Computer Science}} @STRING{mitp = {MIT Press}} @STRING{sv = {Springer}} @Book{BackWright1998, author = {Back, R.-J. and von Wright, J.}, title = {Refinement Calculus}, publisher = sv, address = {New York}, year = 1998, note = {} } @Article{BackhouseCarre1975, author = {Backhouse, R. C. and Carr{\'e}, B. A.}, title = {Regular Algebra Applied to Path-finding Problems}, journal = {Journal of the Institute of Mathematics and its Applications}, volume = 15, number = 2, pages = {161--186}, year = 1975, note = {} } @Article{Berghammer1999, author = {Berghammer, R.}, title = {Combining relational calculus and the {Dijkstra--Gries} method for deriving relational programs}, journal = is, volume = 119, number = {3--4}, pages = {155--171}, year = 1999, note = {} } @InProceedings{BerghammerKargerWolf1998, author = {Berghammer, R. and von Karger, B. and Wolf, A.}, title = {Relation-Algebraic Derivation of Spanning Tree Algorithms}, editor = {Jeuring, J.}, booktitle = {Mathematics of Program Construction (MPC 1998)}, publisher = sv, series = lncs, volume = 1422, pages = {23--43}, year = 1998, note = {} } @InProceedings{BerghammerStruth2010, author = {Berghammer, R. and Struth, G.}, title = {On Automated Program Construction and Verification}, editor = {Bolduc, C. and Desharnais, J. and Ktari, B.}, booktitle = {Mathematics of Program Construction (MPC 2010)}, publisher = sv, series = lncs, volume = 6120, pages = {22--41}, year = 2010, note = {} } @Book{CormenLeisersonRivest1990, author = {Cormen, T. H. and Leiserson, C. E. and Rivest, R. L.}, title = {Introduction to Algorithms}, publisher = mitp, year = 1990, note = {} } @Article{FosterGreenwaldMoorePierceSchmitt2007, author = {Foster, J. N. and Greenwald, M. B. and Moore, J. T. and Pierce, B. C. and Schmitt, A.}, title = {Combinators for Bidirectional Tree Transformations: A Linguistic Approach to the View-Update Problem}, journal = toplas, volume = 29, number = {3:17}, pages = {1--65}, year = 2007, note = {} } @Article{GallerFisher1964, author = {Galler, B. A. and Fisher, M. J.}, title = {An Improved Equivalence Algorithm}, journal = cacm, volume = 7, number = 5, pages = {301--303}, year = 1964, note = {} } @Book{GondranMinoux2008, author = {Gondran, M. and Minoux, M.}, title = {Graphs, Dioids and Semirings}, publisher = sv, year = 2008, note = {} } @Article{Guttmann2018c, author = {Guttmann, W.}, title = {Verifying Minimum Spanning Tree Algorithms with {Stone} Relation Algebras}, journal = jlamp, volume = 101, pages = {132--150}, year = 2018, note = {} } @InProceedings{Guttmann2020b, author = {Guttmann, W.}, title = {Verifying the Correctness of Disjoint-Set Forests with {Kleene} Relation Algebras}, editor = {Fahrenberg, U. and Jipsen, P. and Winter, M.}, booktitle = {Relational and Algebraic Methods in Computer Science (RAMiCS 2020)}, publisher = sv, series = lncs, volume = 12062, pages = {134--151}, year = 2020, note = {} } +@Article{Guttmann2023, + author = {Guttmann, W.}, + title = {Relation-Algebraic Verification of Disjoint-Set Forests}, + journal = arxiv, + volume = {2301.10311}, + year = 2023, + note = {\url{https://arxiv.org/abs/2301.10311}} +} + @Article{HoefnerMoeller2012, author = {H{\"o}fner, P. and M{\"o}ller, B.}, title = {Dijkstra, {Floyd} and {Warshall} meet {Kleene}}, journal = fac, volume = 24, number = 4, pages = {459--476}, year = 2012, note = {} } @Article{Kozen1994, author = {Kozen, D.}, title = {A completeness theorem for {Kleene} algebras and the algebra of regular events}, journal = {Information and Computation}, volume = 110, number = 2, pages = {366--390}, year = 1994, note = {} } @Article{LammichMeis2012, author = {Lammich, P. and Meis, R.}, title = {A Separation Logic Framework for {I}mperative {HOL}}, journal = afp, year = 2012, note = {} } @InProceedings{Moeller1993, author = {M{\"o}ller, B.}, title = {Derivation of Graph and Pointer Algorithms}, editor = {M{\"o}ller, B. and Partsch, H. A. and Schuman, S. A.}, booktitle = {Formal Program Development}, publisher = sv, series = lncs, volume = 755, pages = {123--160}, year = 1993, note = {} } @Article{Tarjan1975, author = {Tarjan, R. E.}, title = {Efficiency of a Good But Not Linear Set Union Algorithm}, journal = jacm, volume = 22, number = 2, pages = {215--225}, year = 1975, note = {} } @Article{Tarski1941, author = {Tarski, A.}, title = {On the calculus of relations}, journal = {The Journal of Symbolic Logic}, volume = 6, number = 3, pages = {73--89}, year = 1941, note = {} } @Article{Zhan2018, author = {Zhan, B.}, title = {Verifying Imperative Programs using {Auto2}}, journal = afp, year = 2018, note = {} } 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,64 +1,65 @@ \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} +\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 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 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 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. +Section 8 instantiates the Peano axioms by Boolean matrices. -This Isabelle/HOL theory formally verifies results in \cite{Guttmann2020b} and in an extended version of that paper. +This Isabelle/HOL theory formally verifies results in \cite{Guttmann2020b} and an extended version of that paper \cite{Guttmann2023}. 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/Relational_Minimum_Spanning_Trees/Boruvka.thy b/thys/Relational_Minimum_Spanning_Trees/Boruvka.thy --- a/thys/Relational_Minimum_Spanning_Trees/Boruvka.thy +++ b/thys/Relational_Minimum_Spanning_Trees/Boruvka.thy @@ -1,3819 +1,3655 @@ (* Title: Borůvka's Minimum Spanning Tree Algorithm Author: Nicolas Robinson-O'Brien Maintainer: Walter Guttmann *) section \Bor\r{u}vka's Minimum Spanning Tree Algorithm\ text \ In this theory we prove partial correctness of Bor\r{u}vka's minimum spanning tree algorithm. \ theory Boruvka imports + Aggregation_Algebras.M_Choose_Component Relational_Disjoint_Set_Forests.Disjoint_Set_Forests Kruskal begin subsection \General results\ text \ The proof is carried out in $m$-$k$-Stone-Kleene relation algebras. In this section we give results that hold more generally. \ context stone_kleene_relation_algebra begin lemma He_eq_He_THe_star: assumes "arc e" and "equivalence H" shows "H * e = H * e * (top * H * e)\<^sup>\" proof - let ?x = "H * e" have 1: "H * e \ H * e * (top * H * e)\<^sup>\" using comp_isotone star.circ_reflexive by fastforce have "H * e * (top * H * e)\<^sup>\ = H * e * (top * e)\<^sup>\" by (metis assms(2) preorder_idempotent surjective_var) also have "... \ H * e * (1 \ top * (e * top)\<^sup>\ * e)" by (metis eq_refl star.circ_mult_1) also have "... \ H * e * (1 \ top * top * e)" by (simp add: star.circ_left_top) also have "... = H * e \ H * e * top * e" by (simp add: mult.semigroup_axioms semiring.distrib_left semigroup.assoc) finally have 2: "H * e * (top * H * e)\<^sup>\ \ H * e" using assms arc_top_arc mult_assoc by auto thus ?thesis using 1 2 by simp qed lemma path_through_components: assumes "equivalence H" and "arc e" shows "(H * (d \ e))\<^sup>\ = (H * d)\<^sup>\ \ (H * d)\<^sup>\ * H * e * (H * d)\<^sup>\" proof - have "H * e * (H * d)\<^sup>\ * H * e \ H * e * top * H * e" by (simp add: comp_isotone) also have "... = H * e * top * e" by (metis assms(1) preorder_idempotent surjective_var mult_assoc) also have "... = H * e" using assms(2) arc_top_arc mult_assoc by auto finally have 1: "H * e * (H * d)\<^sup>\ * H * e \ H * e" by simp have "(H * (d \ e))\<^sup>\ = (H * d \ H * e)\<^sup>\" by (simp add: comp_left_dist_sup) also have "... = (H * d)\<^sup>\ \ (H * d)\<^sup>\ * H * e * (H * d)\<^sup>\" using 1 star_separate_3 by (simp add: mult_assoc) finally show ?thesis by simp qed lemma simplify_f: assumes "regular p" and "regular e" shows "(f \ - e\<^sup>T \ - p) \ (f \ - e\<^sup>T \ p) \ (f \ - e\<^sup>T \ p)\<^sup>T \ (f \ - e\<^sup>T \ - p)\<^sup>T \ e\<^sup>T \ e = f \ f\<^sup>T \ e \ e\<^sup>T" proof - have "(f \ - e\<^sup>T \ - p) \ (f \ - e\<^sup>T \ p) \ (f \ - e\<^sup>T \ p)\<^sup>T \ (f \ - e\<^sup>T \ - p)\<^sup>T \ e\<^sup>T \ e = (f \ - e\<^sup>T \ - p) \ (f \ - e\<^sup>T \ p) \ (f\<^sup>T \ - e \ p\<^sup>T) \ (f\<^sup>T \ - e \ - p\<^sup>T) \ e\<^sup>T \ e" by (simp add: conv_complement conv_dist_inf) also have "... = ((f \ (f \ - e\<^sup>T \ p)) \ (- e\<^sup>T \ (f \ - e\<^sup>T \ p)) \ (- p \ (f \ - e\<^sup>T \ p))) \ ((f\<^sup>T \ (f\<^sup>T \ - e \ - p\<^sup>T)) \ (- e \ (f\<^sup>T \ - e \ - p\<^sup>T)) \ (p\<^sup>T \ (f\<^sup>T \ - e \ - p\<^sup>T))) \ e\<^sup>T \ e" by (metis sup_inf_distrib2 sup_assoc) also have "... = ((f \ f) \ (f \ - e\<^sup>T) \ (f \ p) \ (- e\<^sup>T \ f) \ (- e\<^sup>T \ - e\<^sup>T) \ (- e\<^sup>T \ p) \ (- p \ f) \ (- p \ - e\<^sup>T) \ (- p \ p)) \ ((f\<^sup>T \ f\<^sup>T) \ (f\<^sup>T \ - e) \ (f\<^sup>T \ - p\<^sup>T) \ (- e \ f\<^sup>T) \ (- e \ - e) \ (- e \ - p\<^sup>T) \ (p\<^sup>T \ f\<^sup>T) \ (p\<^sup>T \ - e) \ (p\<^sup>T \ - p\<^sup>T)) \ e\<^sup>T \ e" using sup_inf_distrib1 sup_assoc inf_assoc sup_inf_distrib1 by simp also have "... = ((f \ f) \ (f \ - e\<^sup>T) \ (f \ p) \ (f \ - p) \ (- e\<^sup>T \ f) \ (- e\<^sup>T \ - e\<^sup>T) \ (- e\<^sup>T \ p) \ (- e\<^sup>T \ - p) \ (- p \ p)) \ ((f\<^sup>T \ f\<^sup>T) \ (f\<^sup>T \ - e) \ (f\<^sup>T \ - p\<^sup>T) \ (- e \ f\<^sup>T) \ (f\<^sup>T \ p\<^sup>T) \ (- e \ - e) \ (- e \ - p\<^sup>T) \ (- e \ p\<^sup>T) \ (p\<^sup>T \ - p\<^sup>T)) \ e\<^sup>T \ e" by (smt abel_semigroup.commute inf.abel_semigroup_axioms inf.left_commute sup.abel_semigroup_axioms) also have "... = (f \ - e\<^sup>T \ (- p \ p)) \ (f\<^sup>T \ - e \ (p\<^sup>T \ - p\<^sup>T)) \ e\<^sup>T \ e" by (smt inf.sup_monoid.add_assoc inf.sup_monoid.add_commute inf_sup_absorb sup.idem) also have "... = (f \ - e\<^sup>T) \ (f\<^sup>T \ - e) \ e\<^sup>T \ e" by (metis assms(1) conv_complement inf_top_right stone) also have "... = (f \ e\<^sup>T) \ (- e\<^sup>T \ e\<^sup>T) \ (f\<^sup>T \ e) \ (- e \ e)" by (metis sup.left_commute sup_assoc sup_inf_distrib2) finally show ?thesis by (metis abel_semigroup.commute assms(2) conv_complement inf_top_right stone sup.abel_semigroup_axioms sup_assoc) qed lemma simplify_forest_components_f: assumes "regular p" and "regular e" and "injective (f \ - e\<^sup>T \ - p \ (f \ - e\<^sup>T \ p)\<^sup>T \ e)" and "injective f" shows "forest_components ((f \ - e\<^sup>T \ - p) \ (f \ -e\<^sup>T \ p)\<^sup>T \ e) = (f \ f\<^sup>T \ e \ e\<^sup>T)\<^sup>\" proof - have "forest_components ((f \ - e\<^sup>T \ - p) \ (f \ -e\<^sup>T \ p)\<^sup>T \ e) = wcc ((f \ - e\<^sup>T \ - p) \ (f \ - e\<^sup>T \ p)\<^sup>T \ e)" by (simp add: assms(3) forest_components_wcc) also have "... = ((f \ - e\<^sup>T \ - p) \ (f \ - e\<^sup>T \ p)\<^sup>T \ e \ (f \ - e\<^sup>T \ - p)\<^sup>T \ (f \ - e\<^sup>T \ p) \ e\<^sup>T)\<^sup>\" using conv_dist_sup sup_assoc by auto also have "... = ((f \ - e\<^sup>T \ - p) \ (f \ - e\<^sup>T \ p) \ (f \ - e\<^sup>T \ p)\<^sup>T \ (f \ - e\<^sup>T \ - p)\<^sup>T \ e\<^sup>T \ e)\<^sup>\" using sup_assoc sup_commute by auto also have "... = (f \ f\<^sup>T \ e \ e\<^sup>T)\<^sup>\" using assms(1, 2, 3, 4) simplify_f by auto finally show ?thesis by simp qed lemma components_disj_increasing: assumes "regular p" and "regular e" and "injective (f \ - e\<^sup>T \ - p \ (f \ - e\<^sup>T \ p)\<^sup>T \ e)" and "injective f" shows "forest_components f \ forest_components (f \ - e\<^sup>T \ - p \ (f \ - e\<^sup>T \ p)\<^sup>T \ e)" proof - have 1: "forest_components ((f \ - e\<^sup>T \ - p) \ (f \ -e\<^sup>T \ p)\<^sup>T \ e) = (f \ f\<^sup>T \ e \ e\<^sup>T)\<^sup>\" using simplify_forest_components_f assms(1, 2, 3, 4) by blast have "forest_components f = wcc f" by (simp add: assms(4) forest_components_wcc) also have "... \ (f \ f\<^sup>T \ e\<^sup>T \ e)\<^sup>\" by (simp add: le_supI2 star_isotone sup_commute) finally show ?thesis using 1 sup.left_commute sup_commute by simp qed lemma fch_equivalence: assumes "forest h" shows "equivalence (forest_components h)" by (simp add: assms forest_components_equivalence) lemma forest_modulo_equivalence_path_split_1: assumes "arc a" and "equivalence H" shows "(H * d)\<^sup>\ * H * a * top = (H * (d \ - a))\<^sup>\ * H * a * top" proof - let ?H = "H" let ?x = "?H * (d \ -a)" let ?y = "?H * a" let ?a = "?H * a * top" let ?d = "?H * d" have 1: "?d\<^sup>\ * ?a \ ?x\<^sup>\ * ?a" proof - have "?x\<^sup>\ *?y * ?x\<^sup>\ * ?a \ ?x\<^sup>\ * ?a * ?a" by (smt mult_left_isotone star.circ_right_top top_right_mult_increasing mult_assoc) also have "... = ?x\<^sup>\ * ?a * a * top" by (metis ex231e mult_assoc) also have "... = ?x\<^sup>\ * ?a" by (simp add: assms(1) mult_assoc) finally have 11: "?x\<^sup>\ *?y * ?x\<^sup>\ * ?a \ ?x\<^sup>\ * ?a" by simp have "?d\<^sup>\ * ?a = (?H * (d \ a) \ ?H * (d \ -a))\<^sup>\ * ?a" proof - have 12: "regular a" using assms(1) arc_regular by simp have "?H * ((d \ a) \ (d \ - a)) = ?H * (d \ top)" using 12 by (metis inf_top_right maddux_3_11_pp) thus ?thesis using mult_left_dist_sup by auto qed also have "... \ (?y \ ?x)\<^sup>\ * ?a" by (metis comp_inf.coreflexive_idempotent comp_isotone inf.cobounded1 inf.sup_monoid.add_commute semiring.add_mono star_isotone top.extremum) also have "... = (?x \ ?y)\<^sup>\ * ?a" by (simp add: sup_commute mult_assoc) also have "... = ?x\<^sup>\ * ?a \ (?x\<^sup>\ * ?y * (?x\<^sup>\ * ?y)\<^sup>\ * ?x\<^sup>\) * ?a" by (smt mult_right_dist_sup star.circ_sup_9 star.circ_unfold_sum mult_assoc) also have "... \ ?x\<^sup>\ * ?a \ (?x\<^sup>\ * ?y * (top * ?y)\<^sup>\ * ?x\<^sup>\) * ?a" proof - have "(?x\<^sup>\ * ?y)\<^sup>\ \ (top * ?y)\<^sup>\" by (simp add: mult_left_isotone star_isotone) thus ?thesis by (metis comp_inf.coreflexive_idempotent comp_inf.transitive_star eq_refl mult_left_dist_sup top.extremum mult_assoc) qed also have "... = ?x\<^sup>\ * ?a \ (?x\<^sup>\ * ?y * ?x\<^sup>\) * ?a" using assms(1, 2) He_eq_He_THe_star arc_regular mult_assoc by auto finally have 13: "(?H * d)\<^sup>\ * ?a \ ?x\<^sup>\ * ?a \ ?x\<^sup>\ * ?y * ?x\<^sup>\ * ?a" by (simp add: mult_assoc) have 14: "?x\<^sup>\ * ?y * ?x\<^sup>\ * ?a \ ?x\<^sup>\ * ?a" using 11 mult_assoc by auto thus ?thesis using 13 14 sup.absorb1 by auto qed have 2: "?d\<^sup>\ * ?a \ ?x\<^sup>\ *?a" by (simp add: comp_isotone star_isotone) thus ?thesis using 1 2 order.antisym mult_assoc by simp qed lemma dTransHd_le_1: assumes "equivalence H" and "univalent (H * d)" shows "d\<^sup>T * H * d \ 1" proof - have "d\<^sup>T * H\<^sup>T * H * d \ 1" using assms(2) conv_dist_comp mult_assoc by auto thus ?thesis using assms(1) mult_assoc by (simp add: preorder_idempotent) qed lemma HcompaT_le_compHaT: assumes "equivalence H" and "injective (a * top)" shows "-H * a * top \ - (H * a * top)" proof - have "a * top * a\<^sup>T \ 1" by (metis assms(2) conv_dist_comp symmetric_top_closed vector_top_closed mult_assoc) hence "a * top * a\<^sup>T * H \ H" using assms(1) comp_isotone order_trans by blast hence "a * top * top * a\<^sup>T * H \ H" by (simp add: vector_mult_closed) hence "a * top * (H * a * top)\<^sup>T \ H" by (metis assms(1) conv_dist_comp symmetric_top_closed vector_top_closed mult_assoc) thus ?thesis using assms(2) comp_injective_below_complement mult_assoc by auto qed subsection \Forests modulo an equivalence\ text \ In the graphical interpretation, the arcs of d are directed towards the root(s) of the \forest_modulo_equivalence\. \ definition "forest_modulo_equivalence x d \ equivalence x \ univalent (x * d) \ x \ d * d\<^sup>T \ 1 \ (x * d)\<^sup>+ \ x \ bot" definition "forest_modulo_equivalence_path a b H d \ arc a \ arc b \ a\<^sup>T * top \ (H * d)\<^sup>\ * H * b * top" lemma d_separates_forest_modulo_equivalence_1: assumes "forest_modulo_equivalence x d" shows "x * d \ -x" proof - have "x * d \ (x * d)\<^sup>+" using star.circ_mult_increasing by simp also have "... \ -x" using assms(1) forest_modulo_equivalence_def le_bot pseudo_complement by blast finally show ?thesis by simp qed lemma d_separates_forest_modulo_equivalence_2: shows "forest_modulo_equivalence x d \ d * x \ -x" using forest_modulo_equivalence_def schroeder_6_p d_separates_forest_modulo_equivalence_1 by metis lemma d_separates_forest_modulo_equivalence_3: assumes "forest_modulo_equivalence x d" shows "d \ -x" proof - have "1 \ x" using assms(1) forest_modulo_equivalence_def by auto then have "d \ x * d" using mult_left_isotone by fastforce also have "... \ (x * d)\<^sup>+" using star.circ_mult_increasing by simp also have "... \ -x" using assms(1) forest_modulo_equivalence_def le_bot pseudo_complement by blast finally show ?thesis by simp qed lemma d_separates_forest_modulo_equivalence_4: shows "forest_modulo_equivalence x d \ d\<^sup>T \ -x" using d_separates_forest_modulo_equivalence_3 forest_modulo_equivalence_def conv_isotone symmetric_complement_closed by metis lemma d_separates_forest_modulo_equivalence_5: shows "forest_modulo_equivalence x d \ d \ d\<^sup>T \ -x" using d_separates_forest_modulo_equivalence_3 d_separates_forest_modulo_equivalence_4 sup_least by blast lemma d_separates_forest_modulo_equivalence_6: shows "forest_modulo_equivalence x d \ d * x \ x * d \ -x" using d_separates_forest_modulo_equivalence_1 d_separates_forest_modulo_equivalence_2 sup_least by blast lemma d_separates_forest_modulo_equivalence_7: shows "forest_modulo_equivalence x d \ x * (d \ d\<^sup>T) * x \ -x" using d_separates_forest_modulo_equivalence_5 forest_modulo_equivalence_def inf.sup_monoid.add_commute preorder_idempotent pseudo_complement triple_schroeder_p by metis lemma d_separates_forest_modulo_equivalence_8: shows "forest_modulo_equivalence x d \ (d * x)\<^sup>T \ -x" using d_separates_forest_modulo_equivalence_2 forest_modulo_equivalence_def conv_isotone symmetric_complement_closed by metis lemma d_separates_forest_modulo_equivalence_9: shows "forest_modulo_equivalence x d \ (x * d)\<^sup>T \ -x" by (metis d_separates_forest_modulo_equivalence_1 forest_modulo_equivalence_def conv_isotone symmetric_complement_closed) lemma d_separates_forest_modulo_equivalence_10: shows "forest_modulo_equivalence x d \ (d * x)\<^sup>+ \ -x" using forest_modulo_equivalence_def le_bot pseudo_complement schroeder_5_p star_slide mult_assoc by metis lemma d_separates_forest_modulo_equivalence_11: shows "forest_modulo_equivalence x d \ (x * d)\<^sup>+ \ -x" using forest_modulo_equivalence_def le_bot pseudo_complement by blast lemma d_separates_forest_modulo_equivalence_12: shows "forest_modulo_equivalence x d \ (d * x)\<^sup>T\<^sup>+ \ -x" using d_separates_forest_modulo_equivalence_10 forest_modulo_equivalence_def conv_isotone conv_plus_commute symmetric_complement_closed by metis lemma d_separates_x_in_forest_13: shows "forest_modulo_equivalence x d \ (x * d)\<^sup>T\<^sup>+ \ -x" using d_separates_forest_modulo_equivalence_11 forest_modulo_equivalence_def conv_isotone conv_plus_commute symmetric_complement_closed by metis lemma irreflexive_d_in_forest_modulo_equivalence: shows "forest_modulo_equivalence x d \ irreflexive d" by (metis d_separates_forest_modulo_equivalence_3 forest_modulo_equivalence_def inf.cobounded2 inf.left_commute inf.orderE pseudo_complement) lemma univalent_d_in_forest_modulo_equivalence: assumes "forest_modulo_equivalence x d" shows "univalent d" proof - have "d\<^sup>T * d \ d\<^sup>T * x\<^sup>T * x * d" using assms(1) forest_modulo_equivalence_def comp_isotone comp_right_one mult_sub_right_one by metis also have "... \ 1" using assms(1) forest_modulo_equivalence_def comp_associative conv_dist_comp by auto finally show ?thesis by simp qed lemma acyclic_d_in_forest_modulo_equivalence: assumes "forest_modulo_equivalence x d" shows "acyclic d" proof - have "d\<^sup>\ \ (x * d)\<^sup>\" using assms(1) forest_modulo_equivalence_def mult_left_isotone star.circ_circ_mult star.circ_circ_mult_1 star.circ_extra_circ star.left_plus_circ star_involutive star_isotone star_one star_slide mult_assoc by metis then have "d * d\<^sup>\ \ d * (x * d)\<^sup>\" using mult_right_isotone by blast also have "... \ x * d * (x * d)\<^sup>\" using assms(1) forest_modulo_equivalence_def eq_refl inf.order_trans mult_isotone star.circ_circ_mult_1 star_involutive star_one star_outer_increasing mult_assoc by metis also have "... \ -x" using assms d_separates_forest_modulo_equivalence_11 by blast also have "... \ -1" using assms(1) forest_modulo_equivalence_def p_antitone by blast finally show ?thesis by simp qed lemma acyclic_dt_d_in_forest_modulo_equivalence: shows "forest_modulo_equivalence x d \ acyclic (d\<^sup>T)" using acyclic_d_in_forest_modulo_equivalence conv_plus_commute irreflexive_conv_closed by fastforce lemma dt_forest_modulo_equivalence_forest: shows "forest_modulo_equivalence x d \ forest (d\<^sup>T)" using acyclic_dt_d_in_forest_modulo_equivalence univalent_d_in_forest_modulo_equivalence by simp lemma var_forest_modulo_equivalence_axiom: shows "forest_modulo_equivalence x d \ d\<^sup>T * x * d \ 1" using forest_modulo_equivalence_def comp_associative conv_dist_comp preorder_idempotent by metis lemma forest_modulo_equivalence_wcc: assumes "forest_modulo_equivalence x d" shows "(x * d)\<^sup>\ * (x * d)\<^sup>T\<^sup>\ = ((x * d) \ (x * d)\<^sup>T)\<^sup>\" using assms(1) forest_modulo_equivalence_def fc_wcc by force lemma forest_modulo_equivalence_direction_1: assumes "forest_modulo_equivalence x d" shows "(x * d)\<^sup>\ \ (x * d)\<^sup>T = bot" using assms(1) d_separates_forest_modulo_equivalence_11 forest_modulo_equivalence_def acyclic_star_below_complement_1 order_lesseq_imp p_antitone_iff by meson lemma forest_modulo_equivalence_direction_2: assumes "forest_modulo_equivalence x d" shows "(x * d)\<^sup>T\<^sup>\ \ (x * d) \ bot" using assms(1) forest_modulo_equivalence_direction_1 comp_inf.idempotent_bot_closed conv_inf_bot_iff conv_star_commute inf.sup_left_divisibility by metis lemma forest_modulo_equivalence_separate: assumes "forest_modulo_equivalence x d" shows "(x * d)\<^sup>\ * (x * d)\<^sup>T\<^sup>\ \ (x * d)\<^sup>T * (x * d) \ 1" proof - have "(x * d)\<^sup>\ \ (x * d)\<^sup>T * (x * d) = (1 \ (x * d)\<^sup>+) \ (x * d)\<^sup>T * (x * d)" using star_left_unfold_equal by simp also have "... = (1 \ (x * d)\<^sup>T * (x * d)) \ ((x * d)\<^sup>+ \ (x * d)\<^sup>T * (x * d))" using comp_inf.semiring.distrib_right by simp also have "... \ 1 \ ((x * d)\<^sup>+ \ (x * d)\<^sup>T * (x * d))" using inf.cobounded1 semiring.add_right_mono by blast also have "... = 1 \ ((x * d)\<^sup>\ \ (x * d)\<^sup>T) * (x * d)" using assms(1) forest_modulo_equivalence_def forest_modulo_equivalence_direction_1 comp_inf.semiring.mult_zero_right inf.sup_left_divisibility le_infI2 semiring.mult_not_zero sup.orderE by metis also have "... \ 1 \ bot" using assms(1) forest_modulo_equivalence_direction_1 by simp finally have 2: "(x * d)\<^sup>\ \ (x * d)\<^sup>T * (x * d) \ 1" by simp then have 3: "(x * d)\<^sup>T\<^sup>\ \ (x * d)\<^sup>T * (x * d) \ 1" using assms(1) forest_modulo_equivalence_def conv_dist_comp conv_dist_inf conv_involutive conv_star_commute coreflexive_symmetric by metis have "((x * d)\<^sup>\ \ (x * d)\<^sup>T\<^sup>\) \ ((x * d)\<^sup>T * (x * d)) \ 1" using 2 3 inf_sup_distrib2 by simp thus ?thesis using assms(1) le_infI2 forest_modulo_equivalence_def by blast qed lemma forest_modulo_equivalence_path_trans_closure: assumes "forest_modulo_equivalence x d" shows "(x * d\<^sup>T)\<^sup>+ * x * d * x * d\<^sup>T \ (x * d\<^sup>T)\<^sup>+" proof - have "(x * d\<^sup>T)\<^sup>+ * x * d * x * d\<^sup>T = (x * d\<^sup>T)\<^sup>\ * x * d\<^sup>T * x * d * x * d\<^sup>T" using comp_associative star.circ_plus_same by metis also have "... \ (x * d\<^sup>T)\<^sup>\ * x * 1 * x * d\<^sup>T" using assms(1) forest_modulo_equivalence_def var_forest_modulo_equivalence_axiom comp_associative mult_left_isotone mult_right_isotone by metis also have "... \ (x * d\<^sup>T)\<^sup>\ * x * d\<^sup>T" using assms(1) forest_modulo_equivalence_def by (simp add: preorder_idempotent mult_assoc) finally show ?thesis using star.circ_plus_same mult_assoc by simp qed text \ The \forest_modulo_equivalence\ structure is preserved if d is decreased. \ lemma forest_modulo_equivalence_decrease_d: assumes "forest_modulo_equivalence x d" shows "forest_modulo_equivalence x (d \ c)" proof (unfold forest_modulo_equivalence_def, intro conjI) show "reflexive x" using assms(1) forest_modulo_equivalence_def by blast next show "transitive x" using assms(1) forest_modulo_equivalence_def by blast next show "symmetric x" using assms(1) forest_modulo_equivalence_def by blast next show "univalent (x * (d \ c))" proof - have "(x * (d \ c))\<^sup>T * x * (d \ c) \ (x * d)\<^sup>T * x * d" using conv_order mult_isotone by auto also have "... \ 1" using assms(1) forest_modulo_equivalence_def mult_assoc by auto finally show ?thesis using mult_assoc by auto qed next show "coreflexive (x \ ((d \ c) * (d \ c)\<^sup>T))" proof - have "x \ (d \ c) * (d \ c)\<^sup>T \ x \ d * d\<^sup>T" using conv_dist_inf inf.sup_right_isotone mult_isotone by auto thus ?thesis using assms(1) forest_modulo_equivalence_def order_lesseq_imp by blast qed next show "(x * (d \ c))\<^sup>+ \ x \ bot" proof - have "(x * (d \ c))\<^sup>+ \ (x * d)\<^sup>+" using comp_isotone star_isotone by simp thus ?thesis using assms d_separates_forest_modulo_equivalence_11 dual_order.eq_iff dual_order.trans pseudo_complement by blast qed qed lemma expand_forest_modulo_equivalence: assumes "forest_modulo_equivalence H d" shows "(d\<^sup>T * H)\<^sup>\ * (H * d)\<^sup>\ = (d\<^sup>T * H)\<^sup>\ \ (H * d)\<^sup>\" proof - have "(H * d)\<^sup>T * H * d \ 1" using assms forest_modulo_equivalence_def mult_assoc by auto hence "d\<^sup>T * H * H * d \ 1" using assms forest_modulo_equivalence_def conv_dist_comp by auto thus ?thesis by (simp add: cancel_separate_eq comp_associative) qed lemma forest_modulo_equivalence_path_bot: assumes "arc a" and "a \ d" and "forest_modulo_equivalence H d" shows "(d \ - a)\<^sup>T * (H * a * top) \ bot" proof - have 1: "d\<^sup>T * H * d \ 1" using assms(3) forest_modulo_equivalence_def dTransHd_le_1 by blast hence "d * - 1 * d\<^sup>T \ - H" using triple_schroeder_p by force hence "d * - 1 * d\<^sup>T \ 1 \ - H" by (simp add: le_supI2) hence "d * d\<^sup>T \ d * - 1 * d\<^sup>T \ 1 \ - H" by (metis assms(3) forest_modulo_equivalence_def inf_commute regular_one_closed shunting_p le_supI) hence "d * 1 * d\<^sup>T \ d * - 1 * d\<^sup>T \ 1 \ - H" by simp hence "d * (1 \ - 1) * d\<^sup>T \ 1 \ - H" using comp_associative mult_right_dist_sup by (simp add: mult_left_dist_sup) hence "d * top * d\<^sup>T \ 1 \ - H" using regular_complement_top by auto hence "d * top * a\<^sup>T \ 1 \ - H" using assms(2) conv_isotone dual_order.trans mult_right_isotone by blast hence "d * (a * top)\<^sup>T \ 1 \ - H" by (simp add: comp_associative conv_dist_comp) hence "d \ (1 \ - H) * (a * top)" by (simp add: assms(1) shunt_bijective) hence "d \ a * top \ - H * a * top" by (simp add: comp_associative mult_right_dist_sup) also have "... \ a * top \ - (H * a * top)" using assms(1, 3) HcompaT_le_compHaT forest_modulo_equivalence_def sup_right_isotone by auto finally have "d \ a * top \ - (H * a * top)" by simp hence "d \ --( H * a * top) \ a * top" using shunting_var_p by auto hence 2:"d \ H * a * top \ a * top" using inf.sup_right_isotone order.trans pp_increasing by blast have 3:"d \ H * a * top \ top * a" proof - have "d \ H * a * top \ (H * a \ d * top\<^sup>T) * (top \ (H * a)\<^sup>T * d)" by (metis dedekind inf_commute) also have "... = d * top \ H * a * a\<^sup>T * H\<^sup>T * d" by (simp add: conv_dist_comp inf_vector_comp mult_assoc) also have "... \ d * top \ H * a * d\<^sup>T * H\<^sup>T * d" using assms(2) mult_right_isotone mult_left_isotone conv_isotone inf.sup_right_isotone by auto also have "... = d * top \ H * a * d\<^sup>T * H * d" using assms(3) forest_modulo_equivalence_def by auto also have "... \ d * top \ H * a * 1" using 1 by (metis inf.sup_right_isotone mult_right_isotone mult_assoc) also have "... \ H * a" by simp also have "... \ top * a" by (simp add: mult_left_isotone) finally have "d \ H * a * top \ top * a" by simp thus ?thesis by simp qed have "d \ H * a * top \ a * top \ top * a" using 2 3 by simp also have "... = a * top * top * a" by (metis comp_associative comp_inf.star.circ_decompose_9 comp_inf.star_star_absorb comp_inf_covector vector_inf_comp vector_top_closed) also have "... = a * top * a" by (simp add: vector_mult_closed) finally have 4:"d \ H * a * top \ a" by (simp add: assms(1) arc_top_arc) hence "d \ - a \ -(H * a * top)" using assms(1) arc_regular p_shunting_swap by fastforce hence "(d \ - a) * top \ -(H * a * top)" using mult.semigroup_axioms p_antitone_iff schroeder_4_p semigroup.assoc by fastforce thus ?thesis by (simp add: schroeder_3_p) qed lemma forest_modulo_equivalence_path_split_2: assumes "arc a" and "a \ d" and "forest_modulo_equivalence H d" shows "(H * (d \ - a))\<^sup>\ * H * a * top = (H * ((d \ - a) \ (d \ - a)\<^sup>T))\<^sup>\ * H * a * top" proof - let ?lhs = "(H * (d \ - a))\<^sup>\ * H * a * top" have 1: "d\<^sup>T * H * d \ 1" using assms(3) forest_modulo_equivalence_def dTransHd_le_1 by blast have 2: "H * a * top \ ?lhs" by (metis le_iff_sup star.circ_loop_fixpoint star.circ_transitive_equal star_involutive sup_commute mult_assoc) have "(d \ - a)\<^sup>T * (H * (d \ - a))\<^sup>\ * (H * a * top) = (d \ - a)\<^sup>T * H * (d \ - a) * (H * (d \ - a))\<^sup>\ * (H * a * top)" proof - have "(d \ - a)\<^sup>T * (H * (d \ - a))\<^sup>\ * (H * a * top) = (d \ - a)\<^sup>T * (1 \ H * (d \ - a) * (H * (d \ - a))\<^sup>\) * (H * a * top)" by (simp add: star_left_unfold_equal) also have "... = (d \ - a)\<^sup>T * H * a * top \ (d \ - a)\<^sup>T * H * (d \ - a) * (H * (d \ - a))\<^sup>\ * (H * a * top)" by (smt mult_left_dist_sup star.circ_loop_fixpoint star.circ_mult_1 star_slide sup_commute mult_assoc) also have "... = bot \ (d \ - a)\<^sup>T * H * (d \ - a) * (H * (d \ - a))\<^sup>\ * (H * a * top)" by (metis assms(1, 2, 3) forest_modulo_equivalence_path_bot mult_assoc le_bot) thus ?thesis by (simp add: calculation) qed also have "... \ d\<^sup>T * H * d * (H * (d \ - a))\<^sup>\ * (H * a * top)" using conv_isotone inf.cobounded1 mult_isotone by auto also have "... \ 1 * (H * (d \ - a))\<^sup>\ * (H * a * top)" using 1 by (metis le_iff_sup mult_right_dist_sup) finally have 3: "(d \ - a)\<^sup>T * (H * (d \ - a))\<^sup>\ * (H * a * top) \ ?lhs" using mult_assoc by auto hence 4: "H * (d \ - a)\<^sup>T * (H * (d \ - a))\<^sup>\ * (H * a * top) \ ?lhs" proof - have "H * (d \ - a)\<^sup>T * (H * (d \ - a))\<^sup>\ * (H * a * top) \ H * (H * (d \ - a))\<^sup>\ * H * a * top" using 3 mult_right_isotone mult_assoc by auto also have "... = H * H * ((d \ - a) * H)\<^sup>\ * H * a * top" by (metis assms(3) forest_modulo_equivalence_def star_slide mult_assoc preorder_idempotent) also have "... = H * ((d \ - a) * H)\<^sup>\ * H * a * top" using assms(3) forest_modulo_equivalence_def preorder_idempotent by fastforce finally show ?thesis by (metis assms(3) forest_modulo_equivalence_def preorder_idempotent star_slide mult_assoc) qed have 5: "(H * (d \ - a) \ H * (d \ - a)\<^sup>T) * (H * (d \ - a))\<^sup>\ * H * a * top \ ?lhs" proof - have 51: "H * (d \ - a) * (H * (d \ - a))\<^sup>\ * H * a * top \ (H * (d \ - a))\<^sup>\ * H * a * top" using star.left_plus_below_circ mult_left_isotone by simp have 52: "(H * (d \ - a) \ H * (d \ - a)\<^sup>T) * (H * (d \ - a))\<^sup>\ * H * a * top = H * (d \ - a) * (H * (d \ - a))\<^sup>\ * H * a * top \ H * (d \ - a)\<^sup>T * (H * (d \ - a))\<^sup>\ * H * a * top" using mult_right_dist_sup by auto hence "... \ (H * (d \ - a))\<^sup>\ * H * a * top \ H * (d \ - a)\<^sup>T * (H * (d \ - a))\<^sup>\ * H * a * top" using star.left_plus_below_circ mult_left_isotone sup_left_isotone by auto thus ?thesis using 4 51 52 mult_assoc by auto qed hence "(H * (d \ - a) \ H * (d \ - a)\<^sup>T)\<^sup>\ * H * a * top \ ?lhs" proof - have "(H * (d \ - a) \ H * (d \ - a)\<^sup>T)\<^sup>\ * (H * (d \ - a))\<^sup>\ * H * a * top \ ?lhs" using 5 star_left_induct_mult_iff mult_assoc by auto thus ?thesis using star.circ_decompose_11 star_decompose_1 by auto qed hence 6: "(H * ((d \ - a) \ (d \ - a)\<^sup>T))\<^sup>\ * H * a * top \ ?lhs" using mult_left_dist_sup by auto have 7: "(H * (d \ - a))\<^sup>\ * H * a * top \ (H * ((d \ - a) \ (d \ - a)\<^sup>T))\<^sup>\ * H * a * top" by (simp add: mult_left_isotone semiring.distrib_left star_isotone) thus ?thesis using 6 7 by (simp add: mult_assoc) qed end -context stone_relation_algebra -begin - -text \ -A \vector_classes\ corresponds to one or more equivalence classes and a \unique_vector_class\ corresponds to a single equivalence class. -\ - -definition vector_classes :: "'a \ 'a \ bool" where "vector_classes x v \ regular x \ regular v \ equivalence x \ vector v \ x * v \ v \ v \ bot" -definition unique_vector_class :: "'a \ 'a \ bool" where "unique_vector_class x v \ vector_classes x v \ v * v\<^sup>T \ x" - -end - subsection \An operation to select components\ text \ -We introduce the operation \choose_component\. -\begin{itemize} - \item Axiom \component_in_v\ expresses that the result of \choose_component\ is contained in the set of vertices, $v$, we are selecting from, ignoring the weights. - \item Axiom \component_is_vector\ states that the result of \choose_component\ is a vector. - \item Axiom \component_is_regular\ states that the result of \choose_component\ is regular. - \item Axiom \component_is_connected\ states that any two vertices from the result of \choose_component\ are connected in $e$. - \item Axiom \component_single\ states that the result of \choose_component\ is closed under being connected in $e$. - \item Finally, axiom \component_not_bot_when_v_bot_bot\ expresses that the operation \choose_component\ returns a non-empty component if the input satisfies the given criteria. -\end{itemize} -\ - -class choose_component = - fixes choose_component :: "'a \ 'a \ 'a" - -class choose_component_algebra = choose_component + stone_relation_algebra + - assumes component_is_vector: "vector (choose_component e v)" - assumes component_is_regular: "regular (choose_component e v)" - assumes component_in_v: "choose_component e v \ --v" - assumes component_is_connected: "choose_component e v * (choose_component e v)\<^sup>T \ e" - assumes component_single: "e * choose_component e v \ choose_component e v" - assumes component_not_bot_when_v_bot_bot: "vector_classes e v \ choose_component e v \ bot" - -text \ -Every \m_kleene_algebra\ is an instance of \choose_component_algebra\ when the \choose_component\ operation is defined as follows: +This section has been moved to theories \Stone_Relation_Algebras.Choose_Component\ and \Aggregation_Algebras.M_Choose_Component\. \ -context m_kleene_algebra -begin - -definition "m_choose_component e v \ - if vector_classes e v then - e * minarc(v) * top - else - bot" - -sublocale m_choose_component_algebra: choose_component_algebra where choose_component = m_choose_component -proof - fix e v - show "m_choose_component e v \ -- v" - proof (cases "vector_classes e v") - case True - hence "m_choose_component e v = e * minarc(v) * top" - by (simp add: m_choose_component_def) - also have "... \ e * --v * top" - by (simp add: comp_isotone minarc_below) - also have "... = e * v * top" - using True vector_classes_def by auto - also have "... \ v * top" - using True vector_classes_def mult_assoc by auto - finally show ?thesis - using True vector_classes_def by auto - next - case False - hence "m_choose_component e v = bot" - using False m_choose_component_def by auto - thus ?thesis - by simp - qed -next - fix e v - show "vector (m_choose_component e v)" - proof (cases "vector_classes e v") - case True - thus ?thesis - by (simp add: mult_assoc m_choose_component_def) - next - case False - thus ?thesis - by (simp add: m_choose_component_def) - qed -next - fix e v - show "regular (m_choose_component e v)" - using minarc_regular regular_mult_closed vector_classes_def m_choose_component_def by auto -next - fix e v - show "m_choose_component e v * (m_choose_component e v)\<^sup>T \ e" - proof (cases "vector_classes e v") - case True - assume 1: "vector_classes e v" - hence "m_choose_component e v * (m_choose_component e v)\<^sup>T = e * minarc(v) * top * (e * minarc(v) * top)\<^sup>T" - by (simp add: m_choose_component_def) - also have "... = e * minarc(v) * top * top\<^sup>T * minarc(v)\<^sup>T * e\<^sup>T" - by (metis comp_associative conv_dist_comp) - also have "... = e * minarc(v) * top * top * minarc(v)\<^sup>T * e" - using True vector_classes_def by auto - also have "... = e * minarc(v) * top * minarc(v)\<^sup>T * e" - by (simp add: comp_associative) - also have "... \ e" - proof (cases "v = bot") - case True - thus ?thesis - by (simp add: True minarc_bot) - next - case False - assume 3: "v \ bot" - hence "e * minarc(v) * top * minarc(v)\<^sup>T \ e * 1" - using 3 minarc_arc arc_expanded comp_associative mult_right_isotone by fastforce - hence "e * minarc(v) * top * minarc(v)\<^sup>T * e \ e * 1 * e" - using mult_left_isotone by auto - also have "... = e" - using True preorder_idempotent vector_classes_def by auto - thus ?thesis - using calculation by auto - qed - thus ?thesis - by (simp add: calculation) - next - case False - thus ?thesis - by (simp add: m_choose_component_def) - qed -next - fix e v - show "e * m_choose_component e v \ m_choose_component e v" - proof (cases "vector_classes e v") - case True - thus ?thesis - using comp_right_one dual_order.eq_iff mult_isotone vector_classes_def m_choose_component_def mult_assoc by metis - next - case False - thus ?thesis - by (simp add: m_choose_component_def) - qed -next - fix e v - show "vector_classes e v \ m_choose_component e v \ bot" - proof (cases "vector_classes e v") - case True - hence "m_choose_component e v \ minarc(v) * top" - using vector_classes_def m_choose_component_def comp_associative minarc_arc shunt_bijective by fastforce - also have "... \ minarc(v)" - using calculation dual_order.trans top_right_mult_increasing by blast - thus ?thesis - using le_bot minarc_bot_iff vector_classes_def by fastforce - next - case False - thus ?thesis - by blast - qed -qed - -end - subsection \m-k-Stone-Kleene relation algebras\ text \ $m$-$k$-Stone-Kleene relation algebras are an extension of $m$-Kleene algebras where the \choose_component\ operation has been added. \ -class m_kleene_algebra_choose_component = - m_kleene_algebra - + choose_component_algebra +context m_kleene_algebra_choose_component begin text \ A \selected_edge\ is a minimum-weight edge whose source is in a component, with respect to $h$, $j$ and $g$, and whose target is not in that component. \ abbreviation "selected_edge h j g \ minarc (choose_component (forest_components h) j * - choose_component (forest_components h) j\<^sup>T \ g)" text \ A \path\ is any sequence of edges in the forest, $f$, of the graph, $g$, backwards from the target of the \selected_edge\ to a root in $f$. \ abbreviation "path f h j g \ top * selected_edge h j g * (f \ - selected_edge h j g\<^sup>T)\<^sup>T\<^sup>\" definition "boruvka_outer_invariant f g \ symmetric g \ forest f \ f \ --g \ regular f \ (\w . minimum_spanning_forest w g \ f \ w \ w\<^sup>T)" definition "boruvka_inner_invariant j f h g d \ boruvka_outer_invariant f g \ g \ bot \ regular d \ regular j \ vector j \ regular h \ forest h \ forest_components h * j = j \ forest_modulo_equivalence (forest_components h) d \ d * top \ - j \ f \ f\<^sup>T = h \ h\<^sup>T \ d \ d\<^sup>T \ (\ a b . forest_modulo_equivalence_path a b (forest_components h) d \ a \ -(forest_components h) \ -- g \ b \ d \ sum(b \ g) \ sum(a \ g))" lemma F_is_H_and_d: assumes "f \ f\<^sup>T = h \ h\<^sup>T \ d \ d\<^sup>T" and "injective f" and "injective h" shows "forest_components f = (forest_components h * (d \ d\<^sup>T))\<^sup>\ * forest_components h" proof - have "forest_components f = (f \ f\<^sup>T)\<^sup>\" using assms(2) cancel_separate_1 by simp also have "... = (h \ h\<^sup>T \ d \ d\<^sup>T)\<^sup>\" using assms(1) by auto also have "... = ((h \ h\<^sup>T)\<^sup>\ * (d \ d\<^sup>T))\<^sup>\ * (h \ h\<^sup>T)\<^sup>\" using star.circ_sup_9 sup_assoc by metis also have "... = (forest_components h * (d \ d\<^sup>T))\<^sup>\ * forest_components h" using assms(3) forest_components_wcc by simp finally show ?thesis by simp qed lemma H_below_F: assumes "f \ f\<^sup>T = h \ h\<^sup>T \ d \ d\<^sup>T" and "injective f" and "injective h" shows "forest_components h \ forest_components f" using assms(1, 2, 3) cancel_separate_1 dual_order.trans star.circ_sub_dist by metis lemma H_below_regular_g: assumes "f \ f\<^sup>T = h \ h\<^sup>T \ d \ d\<^sup>T" and "f \ --g" and "symmetric g" shows "h \ --g" proof - have "h \ f \ f\<^sup>T" using assms(1) sup_assoc by simp also have "... \ --g" using assms(2, 3) conv_complement conv_isotone by fastforce finally show ?thesis using order_trans by blast qed lemma expression_equivalent_without_e_complement: assumes "selected_edge h j g \ - forest_components f" shows "f \ - (selected_edge h j g)\<^sup>T \ - (path f h j g) \ (f \ - (selected_edge h j g)\<^sup>T \ (path f h j g))\<^sup>T \ (selected_edge h j g) = f \ - (path f h j g) \ (f \ (path f h j g))\<^sup>T \ (selected_edge h j g)" proof - let ?p = "path f h j g" let ?e = "selected_edge h j g" let ?F = "forest_components f" have 1: "?e \ - ?F" by (simp add: assms) have "f\<^sup>T \ ?F" by (metis conv_dist_comp conv_involutive conv_order conv_star_commute forest_components_increasing) hence "- ?F \ - f\<^sup>T" using p_antitone by auto hence "?e \ - f\<^sup>T" using 1 dual_order.trans by blast hence "f\<^sup>T \ - ?e" by (simp add: p_antitone_iff) hence "f\<^sup>T\<^sup>T \ - ?e\<^sup>T" by (metis conv_complement conv_dist_inf inf.orderE inf.orderI) hence "f \ - ?e\<^sup>T" by auto hence "f = f \ - ?e\<^sup>T" using inf.orderE by blast hence "f \ - ?e\<^sup>T \ - ?p \ (f \ - ?e\<^sup>T \ ?p)\<^sup>T \ ?e = f \ - ?p \ (f \ ?p)\<^sup>T \ ?e" by auto thus ?thesis by auto qed text \ The source of the \selected_edge\ is contained in $j$, the vector describing those vertices still to be processed in the inner loop of Bor\r{u}vka's algorithm. \ lemma et_below_j: assumes "vector j" and "regular j" and "j \ bot" shows "selected_edge h j g * top \ j" proof - let ?e = "selected_edge h j g" let ?c = "choose_component (forest_components h) j" have "?e * top \ --(?c * -?c\<^sup>T \ g) * top" using comp_isotone minarc_below by blast also have "... = (--(?c * -?c\<^sup>T) \ --g) * top" by simp also have "... = (?c * -?c\<^sup>T \ --g) * top" using component_is_regular regular_mult_closed by auto also have "... = (?c \ -?c\<^sup>T \ --g) * top" by (metis component_is_vector conv_complement vector_complement_closed vector_covector) also have "... \ ?c * top" using inf.cobounded1 mult_left_isotone order_trans by blast also have "... \ j * top" by (metis assms(2) comp_inf.star.circ_sup_2 comp_isotone component_in_v) also have "... = j" by (simp add: assms(1)) finally show ?thesis by simp qed subsubsection \Components of forests and forests modulo an equivalence\ text \ We prove a number of properties about \forest_modulo_equivalence\ and \forest_components\. \ -lemma component_single_eq: - assumes "equivalence x" - shows "choose_component x v = x * choose_component x v" -proof - - have 1: "choose_component x v \ x * choose_component x v" - by (meson component_is_connected ex231c mult_isotone order_lesseq_imp) - thus ?thesis - by (simp add: component_single order.antisym) -qed - lemma fc_j_eq_j_inv: assumes "forest h" and "forest_components h * j = j" shows "forest_components h * (j \ - choose_component (forest_components h) j) = j \ - choose_component (forest_components h) j" proof - let ?c = "choose_component (forest_components h) j" let ?H = "forest_components h" have 1:"equivalence ?H" by (simp add: assms(1) forest_components_equivalence) have "?H * (j \ - ?c) = ?H * j \ ?H * - ?c" using 1 by (metis assms(2) equivalence_comp_dist_inf inf.sup_monoid.add_commute) hence 2: "?H * (j \ - ?c) = j \ ?H * - ?c" by (simp add: assms(2)) have 3: "j \ - ?c \ ?H * - ?c" using 1 by (metis assms(2) dedekind_1 dual_order.trans equivalence_comp_dist_inf inf.cobounded2) have "?H * ?c \ ?c" using component_single by auto hence "?H\<^sup>T * ?c \ ?c" using 1 by simp hence "?H * - ?c \ - ?c" using component_is_regular schroeder_3_p by force hence "j \ ?H * - ?c \ j \ - ?c" using inf.sup_right_isotone by auto thus ?thesis using 2 3 order.antisym by simp qed text \ There is a path in the \forest_modulo_equivalence\ between edges $a$ and $b$ if and only if there is either a path in the \forest_modulo_equivalence\ from $a$ to $b$ or one from $a$ to $c$ and one from $c$ to $b$. \ lemma forest_modulo_equivalence_path_split_disj: assumes "equivalence H" and "arc c" and "regular a \ regular b \ regular c \ regular d \ regular H" shows "forest_modulo_equivalence_path a b H (d \ c) \ forest_modulo_equivalence_path a b H d \ (forest_modulo_equivalence_path a c H d \ forest_modulo_equivalence_path c b H d)" proof - have 1: "forest_modulo_equivalence_path a b H (d \ c) \ forest_modulo_equivalence_path a b H d \ (forest_modulo_equivalence_path a c H d \ forest_modulo_equivalence_path c b H d)" proof (rule impI) assume 11: "forest_modulo_equivalence_path a b H (d \ c)" hence "a\<^sup>T * top \ (H * (d \ c))\<^sup>\ * H * b * top" by (simp add: forest_modulo_equivalence_path_def) also have "... = ((H * d)\<^sup>\ \ (H * d)\<^sup>\ * H * c * (H * d)\<^sup>\) * H * b * top" using assms(1, 2) path_through_components by simp also have "... = (H * d)\<^sup>\ * H * b * top \ (H * d)\<^sup>\ * H * c * (H * d)\<^sup>\ * H * b * top" by (simp add: mult_right_dist_sup) finally have 12:"a\<^sup>T * top \ (H * d)\<^sup>\ * H * b * top \ (H * d)\<^sup>\ * H * c * (H * d)\<^sup>\ * H * b * top" by simp have 13: "a\<^sup>T * top \ (H * d)\<^sup>\ * H * b * top \ a\<^sup>T * top \ (H * d)\<^sup>\ * H * c * (H * d)\<^sup>\ * H * b * top" proof (rule point_in_vector_sup) show "point (a\<^sup>T * top)" using 11 forest_modulo_equivalence_path_def mult_assoc by auto next show "vector ((H * d)\<^sup>\ * H * b * top)" using vector_mult_closed by simp next show "regular ((H * d)\<^sup>\ * H * b * top)" using assms(3) pp_dist_comp pp_dist_star by auto next show "a\<^sup>T * top \ (H * d)\<^sup>\ * H * b * top \ (H * d)\<^sup>\ * H * c * (H * d)\<^sup>\ * H * b * top" using 12 by simp qed thus "forest_modulo_equivalence_path a b H d \ (forest_modulo_equivalence_path a c H d \ forest_modulo_equivalence_path c b H d)" proof (cases "a\<^sup>T * top \ (H * d)\<^sup>\ * H * b * top") case True assume "a\<^sup>T * top \ (H * d)\<^sup>\ * H * b * top" hence "forest_modulo_equivalence_path a b H d" using 11 forest_modulo_equivalence_path_def by auto thus ?thesis by simp next case False have 14: "a\<^sup>T * top \ (H * d)\<^sup>\ * H * c * (H * d)\<^sup>\ * H * b * top" using 13 by (simp add: False) hence 15: "a\<^sup>T * top \ (H * d)\<^sup>\ * H * c * top" by (metis mult_right_isotone order_lesseq_imp top_greatest mult_assoc) have "c\<^sup>T * top \ (H * d)\<^sup>\ * H * b * top" proof (rule ccontr) assume "\ c\<^sup>T * top \ (H * d)\<^sup>\ * H * b * top" hence "c\<^sup>T * top \ -((H * d)\<^sup>\ * H * b * top)" by (meson assms(2, 3) point_in_vector_or_complement regular_closed_star regular_closed_top regular_mult_closed vector_mult_closed vector_top_closed) hence "c * (H * d)\<^sup>\ * H * b * top \ bot" using schroeder_3_p mult_assoc by auto thus "False" using 13 False sup.absorb_iff1 mult_assoc by auto qed hence "forest_modulo_equivalence_path a c H d \ forest_modulo_equivalence_path c b H d" using 11 15 assms(2) forest_modulo_equivalence_path_def by auto thus ?thesis by simp qed qed have 2: "forest_modulo_equivalence_path a b H d \ (forest_modulo_equivalence_path a c H d \ forest_modulo_equivalence_path c b H d) \ forest_modulo_equivalence_path a b H (d \ c)" proof - have 21: "forest_modulo_equivalence_path a b H d \ forest_modulo_equivalence_path a b H (d \ c)" proof (rule impI) assume 22:"forest_modulo_equivalence_path a b H d" hence "a\<^sup>T * top \ (H * d)\<^sup>\ * H * b * top" using forest_modulo_equivalence_path_def by blast hence "a\<^sup>T * top \ (H * (d \ c))\<^sup>\ * H * b * top" by (simp add: mult_left_isotone mult_right_dist_sup mult_right_isotone order.trans star_isotone star_slide) thus "forest_modulo_equivalence_path a b H (d \ c)" using 22 forest_modulo_equivalence_path_def by blast qed have "forest_modulo_equivalence_path a c H d \ forest_modulo_equivalence_path c b H d \ forest_modulo_equivalence_path a b H (d \ c)" proof (rule impI) assume 23: "forest_modulo_equivalence_path a c H d \ forest_modulo_equivalence_path c b H d" hence "a\<^sup>T * top \ (H * d)\<^sup>\ * H * c * top" using forest_modulo_equivalence_path_def by blast also have "... \ (H * d)\<^sup>\ * H * c * c\<^sup>T * c * top" by (metis ex231c comp_inf.star.circ_sup_2 mult_isotone mult_right_isotone mult_assoc) also have "... \ (H * d)\<^sup>\ * H * c * c\<^sup>T * top" by (simp add: mult_right_isotone mult_assoc) also have "... \ (H * d)\<^sup>\ * H * c * (H * d)\<^sup>\ * H * b * top" using 23 mult_right_isotone mult_assoc by (simp add: forest_modulo_equivalence_path_def) also have "... \ (H * d)\<^sup>\ * H * b * top \ (H * d)\<^sup>\ * H * c * (H * d)\<^sup>\ * H * b * top" by (simp add: forest_modulo_equivalence_path_def) finally have "a\<^sup>T * top \ (H * (d \ c))\<^sup>\ * H * b * top" using assms(1, 2) path_through_components mult_right_dist_sup by simp thus "forest_modulo_equivalence_path a b H (d \ c)" using 23 forest_modulo_equivalence_path_def by blast qed thus ?thesis using 21 by auto qed thus ?thesis using 1 2 by blast qed lemma dT_He_eq_bot: assumes "vector j" and "regular j" and "d * top \ - j" and "forest_components h * j = j" and "j \ bot" shows "d\<^sup>T * forest_components h * selected_edge h j g \ bot" proof - let ?e = "selected_edge h j g" let ?H = "forest_components h" have 1: "?e * top \ j" using assms(1, 2, 5) et_below_j by auto have "d\<^sup>T * ?H * ?e \ (d * top)\<^sup>T * ?H * (?e * top)" by (simp add: comp_isotone conv_isotone top_right_mult_increasing) also have "... \ (d * top)\<^sup>T * ?H * j" using 1 mult_right_isotone by auto also have "... \ (- j)\<^sup>T * ?H * j" by (simp add: assms(3) conv_isotone mult_left_isotone) also have "... = (- j)\<^sup>T * j" using assms(4) comp_associative by auto also have "... = bot" by (simp add: assms(1) conv_complement covector_vector_comp) finally show ?thesis using coreflexive_bot_closed le_bot by blast qed lemma forest_modulo_equivalence_d_U_e: assumes "forest f" and "vector j" and "regular j" and "forest h" and "forest_modulo_equivalence (forest_components h) d" and "d * top \ - j" and "forest_components h * j = j" and "f \ f\<^sup>T = h \ h\<^sup>T \ d \ d\<^sup>T" and "selected_edge h j g \ - forest_components f" and "j \ bot" shows "forest_modulo_equivalence (forest_components h) (d \ selected_edge h j g)" proof (cases "selected_edge h j g = bot") let ?e = "selected_edge h j g" case True assume "?e = bot" thus ?thesis by (simp add: True assms(5)) next let ?H = "forest_components h" let ?F = "forest_components f" let ?e = "selected_edge h j g" let ?d' = "d \ ?e" case False assume e_not_bot: "?e \ bot" have "forest_modulo_equivalence (forest_components h) (d \ selected_edge h j g)" proof (unfold forest_modulo_equivalence_def, intro conjI) show 01: "reflexive ?H" by (simp add: assms(4) forest_components_equivalence) show 02: "transitive ?H" by (simp add: assms(4) forest_components_equivalence) show 03: "symmetric ?H" by (simp add: assms(4) forest_components_equivalence) have 04: "equivalence ?H" by (simp add: 01 02 03) show "univalent (?H * ?d')" proof - have "(?H * ?d')\<^sup>T * (?H * ?d') = ?d'\<^sup>T * ?H\<^sup>T * ?H * ?d'" using conv_dist_comp mult_assoc by auto also have "... = ?d'\<^sup>T * ?H * ?H * ?d'" by (simp add: conv_dist_comp conv_star_commute) also have "... = ?d'\<^sup>T * ?H * ?d'" using 01 02 by (metis preorder_idempotent mult_assoc) finally have 21: "univalent (?H * ?d') \ ?e\<^sup>T * ?H * d \ d\<^sup>T * ?H * ?e \ ?e\<^sup>T * ?H * ?e \ d\<^sup>T * ?H * d \ 1" using conv_dist_sup semiring.distrib_left semiring.distrib_right by auto have 22: "?e\<^sup>T * ?H * ?e \ 1" proof - have 221: "?e\<^sup>T * ?H * ?e \ ?e\<^sup>T * top * ?e" by (simp add: mult_left_isotone mult_right_isotone) have "arc ?e" using e_not_bot minarc_arc minarc_bot_iff by blast hence "?e\<^sup>T * top * ?e \ 1" using arc_expanded by blast thus ?thesis using 221 dual_order.trans by blast qed have 24: "d\<^sup>T * ?H * ?e \ 1" by (metis assms(2, 3, 6, 7, 10) dT_He_eq_bot coreflexive_bot_closed le_bot) hence "(d\<^sup>T * ?H * ?e)\<^sup>T \ 1\<^sup>T" using conv_isotone by blast hence "?e\<^sup>T * ?H\<^sup>T * d\<^sup>T\<^sup>T \ 1" by (simp add: conv_dist_comp mult_assoc) hence 25: "?e\<^sup>T * ?H * d \ 1" using assms(4) fch_equivalence by auto have 8: "d\<^sup>T * ?H * d \ 1" using 04 assms(5) dTransHd_le_1 forest_modulo_equivalence_def by blast thus ?thesis using 21 22 24 25 by simp qed show "coreflexive (?H \ ?d' * ?d'\<^sup>T)" proof - have "coreflexive (?H \ ?d' * ?d'\<^sup>T) \ ?H \ (d \ ?e) * (d\<^sup>T \ ?e\<^sup>T) \ 1" by (simp add: conv_dist_sup) also have "... \ ?H \ (d * d\<^sup>T \ d * ?e\<^sup>T \ ?e * d\<^sup>T \ ?e * ?e\<^sup>T) \ 1" by (metis mult_left_dist_sup mult_right_dist_sup sup.left_commute sup_commute) finally have 1: "coreflexive (?H \ ?d' * ?d'\<^sup>T) \ ?H \ d * d\<^sup>T \ ?H \ d * ?e\<^sup>T \ ?H \ ?e * d\<^sup>T \ ?H \ ?e * ?e\<^sup>T \ 1" by (simp add: inf_sup_distrib1) have 31: "?H \ d * d\<^sup>T \ 1" using assms(5) forest_modulo_equivalence_def by blast have 32: "?H \ ?e * d\<^sup>T \ 1" proof - have "?e * d\<^sup>T \ ?e * top * (d * top)\<^sup>T" by (simp add: conv_isotone mult_isotone top_right_mult_increasing) also have "... \ ?e * top * - j\<^sup>T" by (metis assms(6) conv_complement conv_isotone mult_right_isotone) also have "... \ j * - j\<^sup>T" using assms(2, 3, 10) et_below_j mult_left_isotone by auto also have "... \ - ?H" using 03 by (metis assms(2, 3, 7) conv_complement conv_dist_comp equivalence_top_closed mult_left_isotone schroeder_3_p vector_top_closed) finally have "?e * d\<^sup>T \ - ?H" by simp thus ?thesis by (metis inf.coboundedI1 p_antitone_iff p_shunting_swap regular_one_closed) qed have 33: "?H \ d * ?e\<^sup>T \ 1" proof - have 331: "injective h" by (simp add: assms(4)) have "(?H \ ?e * d\<^sup>T)\<^sup>T \ 1" using 32 coreflexive_conv_closed by auto hence "?H \ (?e * d\<^sup>T)\<^sup>T \ 1" using 331 conv_dist_inf forest_components_equivalence by auto thus ?thesis using conv_dist_comp by auto qed have 34: "?H \ ?e * ?e\<^sup>T \ 1" proof - have 341:"arc ?e \ arc (?e\<^sup>T)" using e_not_bot minarc_arc minarc_bot_iff by auto have "?H \ ?e * ?e\<^sup>T \ ?e * ?e\<^sup>T" by auto thus ?thesis using 341 arc_injective le_infI2 by blast qed thus ?thesis using 1 31 32 33 34 by simp qed show 4:"(?H * (d \ ?e))\<^sup>+ \ ?H \ bot" proof - have 40: "(?H * d)\<^sup>+ \ -?H" using assms(5) forest_modulo_equivalence_def bot_unique pseudo_complement by blast have "?e \ - ?F" by (simp add: assms(9)) hence "?F \ - ?e" by (simp add: p_antitone_iff) hence "?F\<^sup>T * ?F \ - ?e" using assms(1) fch_equivalence by fastforce hence "?F\<^sup>T * ?F * ?F\<^sup>T \ - ?e" by (metis assms(1) fch_equivalence forest_components_star star.circ_decompose_9) hence 41: "?F * ?e * ?F \ - ?F" using triple_schroeder_p by blast hence 42:"(?F * ?F)\<^sup>\ * ?F * ?e * (?F * ?F)\<^sup>\ \ - ?F" proof - have 43: "?F * ?F = ?F" using assms(1) forest_components_equivalence preorder_idempotent by auto hence "?F * ?e * ?F = ?F * ?F * ?e * ?F" by simp also have "... = (?F)\<^sup>\ * ?F * ?e * (?F)\<^sup>\" by (simp add: assms(1) forest_components_star) also have "... = (?F * ?F)\<^sup>\ * ?F * ?e * (?F * ?F)\<^sup>\" using 43 by simp finally show ?thesis using 41 by simp qed hence 44: "(?H * d)\<^sup>\ * ?H * ?e * (?H * d)\<^sup>\ \ - ?H" proof - have 45: "?H \ ?F" using assms(1, 4, 8) H_below_F by blast hence 46:"?H * ?e \ ?F * ?e" by (simp add: mult_left_isotone) have "d \ f \ f\<^sup>T" using assms(8) sup.left_commute sup_commute by auto also have "... \ ?F" by (metis forest_components_increasing le_supI2 star.circ_back_loop_fixpoint star.circ_increasing sup.bounded_iff) finally have "d \ ?F" by simp hence "?H * d \ ?F * ?F" using 45 mult_isotone by auto hence 47: "(?H * d)\<^sup>\ \ (?F * ?F)\<^sup>\" by (simp add: star_isotone) hence "(?H * d)\<^sup>\ * ?H * ?e * (?H * d)\<^sup>\ \ (?H * d)\<^sup>\ * ?F * ?e * (?H * d)\<^sup>\" using 46 by (metis mult_left_isotone mult_right_isotone mult_assoc) also have "... \ (?F * ?F)\<^sup>\ * ?F * ?e * (?F * ?F)\<^sup>\" using 47 mult_left_isotone mult_right_isotone by (simp add: comp_isotone) also have "... \ - ?F" using 42 by simp also have "... \ - ?H" using 45 by (simp add: p_antitone) finally show ?thesis by simp qed have "(?H * (d \ ?e))\<^sup>+ = (?H * (d \ ?e))\<^sup>\ * (?H * (d \ ?e))" using star.circ_plus_same by auto also have "... = ((?H * d)\<^sup>\ \ (?H * d)\<^sup>\ * ?H * ?e * (?H * d)\<^sup>\) * (?H * (d \ ?e))" using assms(4) e_not_bot forest_components_equivalence minarc_arc minarc_bot_iff path_through_components by auto also have "... = (?H * d)\<^sup>\ * (?H * (d \ ?e)) \ (?H * d)\<^sup>\ * ?H * ?e * (?H * d)\<^sup>\ * (?H * (d \ ?e))" using mult_right_dist_sup by auto also have "... = (?H * d)\<^sup>\ * (?H * d \ ?H * ?e) \ (?H * d)\<^sup>\ * ?H * ?e * (?H * d)\<^sup>\ * (?H * d \ ?H * ?e)" by (simp add: mult_left_dist_sup) also have "... = (?H * d)\<^sup>\ * ?H * d \ (?H * d)\<^sup>\ * ?H * ?e \ (?H * d)\<^sup>\ * ?H * ?e * (?H * d)\<^sup>\ * (?H * d \ ?H * ?e)" using mult_left_dist_sup mult_assoc by auto also have "... = (?H * d)\<^sup>+ \ (?H * d)\<^sup>\ * ?H * ?e \ (?H * d)\<^sup>\ * ?H * ?e * (?H * d)\<^sup>\ * (?H * d \ ?H * ?e)" by (simp add: star.circ_plus_same mult_assoc) also have "... = (?H * d)\<^sup>+ \ (?H * d)\<^sup>\ * ?H * ?e \ (?H * d)\<^sup>\ * ?H * ?e * (?H * d)\<^sup>\ * ?H * d \ (?H * d)\<^sup>\ * ?H * ?e * (?H * d)\<^sup>\ * ?H * ?e" by (simp add: mult.semigroup_axioms semiring.distrib_left sup.semigroup_axioms semigroup.assoc) also have "... \ (?H * d)\<^sup>+ \ (?H * d)\<^sup>\ * ?H * ?e \ (?H * d)\<^sup>\ * ?H * ?e * (?H * d)\<^sup>\ * ?H * d \ (?H * d)\<^sup>\ * ?H * ?e" proof - have "?e * (?H * d)\<^sup>\ * ?H * ?e \ ?e * top * ?e" by (metis comp_associative comp_inf.coreflexive_idempotent comp_inf.coreflexive_transitive comp_isotone top.extremum) also have "... \ ?e" using e_not_bot arc_top_arc minarc_arc minarc_bot_iff by auto finally have "?e * (?H * d)\<^sup>\ * ?H * ?e \ ?e" by simp hence "(?H * d)\<^sup>\ * ?H * ?e * (?H * d)\<^sup>\ * ?H * ?e \ (?H * d)\<^sup>\ * ?H * ?e" by (simp add: comp_associative comp_isotone) thus ?thesis using sup_right_isotone by blast qed also have "... = (?H * d)\<^sup>+ \ (?H * d)\<^sup>\ * ?H * ?e \ (?H * d)\<^sup>\ * ?H * ?e * (?H * d)\<^sup>\ * ?H * d" by (simp add: order.eq_iff ac_simps) also have "... = (?H * d)\<^sup>+ \ (?H * d)\<^sup>\ * ?H * ?e \ (?H * d)\<^sup>\ * ?H * ?e * (?H * d)\<^sup>+" using star.circ_plus_same mult_assoc by auto also have "... = (?H * d)\<^sup>+ \ (?H * d)\<^sup>\ * ?H * ?e * (1 \ (?H * d)\<^sup>+)" by (simp add: mult_left_dist_sup sup_assoc) also have "... = (?H * d)\<^sup>+ \ (?H * d)\<^sup>\ * ?H * ?e * (?H * d)\<^sup>\" by (simp add: star_left_unfold_equal) also have "... \ - ?H" using 40 44 assms(5) sup.boundedI by blast finally show ?thesis using pseudo_complement by force qed qed thus ?thesis by blast qed subsubsection \Identifying arcs\ text \ The expression $d \sqcap \top e^\top H \sqcap (H d^\top)^* H a^\top \top$ identifies the edge incoming to the component that the \selected_edge\, $e$, is outgoing from and which is on the path from edge $a$ to $e$. Here, we prove this expression is an \arc\. \ lemma shows_arc_x: assumes "forest_modulo_equivalence H d" and "forest_modulo_equivalence_path a e H d" and "H * d * (H * d)\<^sup>\ \ - H" and "\ a\<^sup>T * top \ H * e * top" and "regular a" and "regular e" and "regular H" and "regular d" shows "arc (d \ top * e\<^sup>T * H \ (H * d\<^sup>T)\<^sup>\ * H * a\<^sup>T * top)" proof - let ?x = "d \ top * e\<^sup>T * H \ (H * d\<^sup>T)\<^sup>\ * H * a\<^sup>T * top" have 1:"regular ?x" using assms(5, 6, 7, 8) regular_closed_star regular_conv_closed regular_mult_closed by auto have 2: "a\<^sup>T * top * a \ 1" using arc_expanded assms(2) forest_modulo_equivalence_path_def by auto have 3: "e * top * e\<^sup>T \ 1" using assms(2) forest_modulo_equivalence_path_def arc_expanded by blast have 4: "top * ?x * top = top" proof - have "a\<^sup>T * top \ (H * d)\<^sup>\ * H * e * top" using assms(2) forest_modulo_equivalence_path_def by blast also have "... = H * e * top \ (H * d)\<^sup>\ * H * d * H * e * top" by (metis star.circ_loop_fixpoint star.circ_plus_same sup_commute mult_assoc) finally have "a\<^sup>T * top \ H * e * top \ (H * d)\<^sup>\ * H * d * H * e * top" by simp hence "a\<^sup>T * top \ H * e * top \ a\<^sup>T * top \ (H * d)\<^sup>\ * H * d * H * e * top" using assms(2, 6, 7) point_in_vector_sup forest_modulo_equivalence_path_def regular_mult_closed vector_mult_closed by auto hence "a\<^sup>T * top \ (H * d)\<^sup>\ * H * d * H * e * top" using assms(4) by blast also have "... = (H * d)\<^sup>\ * H * d * (H * e * top \ H * e * top)" by (simp add: mult_assoc) also have "... = (H * d)\<^sup>\ * H * (d \ (H * e * top)\<^sup>T) * H * e * top" by (metis comp_associative covector_inf_comp_3 star.circ_left_top star.circ_top) also have "... = (H * d)\<^sup>\ * H * (d \ top\<^sup>T * e\<^sup>T * H\<^sup>T) * H * e * top" using conv_dist_comp mult_assoc by auto also have "... = (H * d)\<^sup>\ * H * (d \ top * e\<^sup>T * H) * H * e * top" using assms(1) by (simp add: forest_modulo_equivalence_def) finally have 2: "a\<^sup>T * top \ (H * d)\<^sup>\ * H * (d \ top * e\<^sup>T * H) * H * e * top" by simp hence "e * top \ ((H * d)\<^sup>\ * H * (d \ top * e\<^sup>T * H) * H)\<^sup>T * a\<^sup>T * top" proof - have "bijective (e * top) \ bijective (a\<^sup>T * top)" using assms(2) forest_modulo_equivalence_path_def by auto thus ?thesis using 2 by (metis bijective_reverse mult_assoc) qed also have "... = H\<^sup>T * (d \ top * e\<^sup>T * H)\<^sup>T * H\<^sup>T * (H * d)\<^sup>\\<^sup>T * a\<^sup>T * top" by (simp add: conv_dist_comp mult_assoc) also have "... = H * (d \ top * e\<^sup>T * H)\<^sup>T * H * (H * d)\<^sup>\\<^sup>T * a\<^sup>T * top" using assms(1) forest_modulo_equivalence_def by auto also have "... = H * (d \ top * e\<^sup>T * H)\<^sup>T * H * (d\<^sup>T * H)\<^sup>\ * a\<^sup>T * top" using assms(1) forest_modulo_equivalence_def conv_dist_comp conv_star_commute by auto also have "... = H * (d\<^sup>T \ H * e * top) * H * (d\<^sup>T * H)\<^sup>\ * a\<^sup>T * top" using assms(1) conv_dist_comp forest_modulo_equivalence_def comp_associative conv_dist_inf by auto also have "... = H * (d\<^sup>T \ H * e * top) * (H * d\<^sup>T)\<^sup>\ * H * a\<^sup>T * top" by (simp add: comp_associative star_slide) also have "... = H * (d\<^sup>T \ H * e * top) * ((H * d\<^sup>T)\<^sup>\ * H * a\<^sup>T * top \ (H * d\<^sup>T)\<^sup>\ * H * a\<^sup>T * top)" using mult_assoc by auto also have "... = H * (d\<^sup>T \ H * e * top \ ((H * d\<^sup>T)\<^sup>\ * H * a\<^sup>T * top)\<^sup>T) * (H * d\<^sup>T)\<^sup>\ * H * a\<^sup>T * top" by (smt comp_inf_vector covector_comp_inf vector_conv_covector vector_top_closed mult_assoc) also have "... = H * (d\<^sup>T \ (top * e\<^sup>T * H)\<^sup>T \ ((H * d\<^sup>T)\<^sup>\ * H * a\<^sup>T * top)\<^sup>T) * (H * d\<^sup>T)\<^sup>\ * H * a\<^sup>T * top" using assms(1) forest_modulo_equivalence_def conv_dist_comp mult_assoc by auto also have "... = H * (d \ top * e\<^sup>T * H \ (H * d\<^sup>T)\<^sup>\ * H * a\<^sup>T * top)\<^sup>T * (H * d\<^sup>T)\<^sup>\ * H * a\<^sup>T * top" by (simp add: conv_dist_inf) finally have 3: "e * top \ H * ?x\<^sup>T * (H * d\<^sup>T)\<^sup>\ * H * a\<^sup>T * top" by auto have "?x \ bot" proof (rule ccontr) assume "\ ?x \ bot" hence "e * top = bot" using 3 le_bot by auto thus "False" using assms(2, 4) forest_modulo_equivalence_path_def mult_assoc semiring.mult_zero_right by auto qed thus ?thesis using 1 using tarski by blast qed have 5: "?x * top * ?x\<^sup>T \ 1" proof - have 51: "H * (d * H)\<^sup>\ \ d * H * d\<^sup>T \ 1" proof - have 511: "d * (H * d)\<^sup>\ \ - H" using assms(1, 3) forest_modulo_equivalence_def preorder_idempotent schroeder_4_p triple_schroeder_p by fastforce hence "(d * H)\<^sup>\ * d \ - H" using star_slide by auto hence "H * (d\<^sup>T * H)\<^sup>\ \ - d" by (smt assms(1) forest_modulo_equivalence_def conv_dist_comp conv_star_commute schroeder_4_p star_slide) hence "H * (d * H)\<^sup>\ \ - d\<^sup>T" using 511 by (metis assms(1) forest_modulo_equivalence_def schroeder_5_p star_slide) hence "H * (d * H)\<^sup>\ \ - (H * d\<^sup>T)" by (metis assms(3) p_antitone_iff schroeder_4_p star_slide mult_assoc) hence "H * (d * H)\<^sup>\ \ H * d\<^sup>T \ bot" by (simp add: bot_unique pseudo_complement) hence "H * d * (H * (d * H)\<^sup>\ \ H * d\<^sup>T) \ 1" by (simp add: bot_unique) hence 512: "H * d * H * (d * H)\<^sup>\ \ H * d * H * d\<^sup>T \ 1" using univalent_comp_left_dist_inf assms(1) forest_modulo_equivalence_def mult_assoc by fastforce hence 513: "H * d * H * (d * H)\<^sup>\ \ d * H * d\<^sup>T \ 1" proof - have "d * H * d\<^sup>T \ H * d * H * d\<^sup>T" by (metis assms(1) forest_modulo_equivalence_def conv_dist_comp conv_involutive mult_1_right mult_left_isotone) thus ?thesis using 512 by (smt dual_order.trans p_antitone p_shunting_swap regular_one_closed) qed have "d\<^sup>T * H * d \ 1 \ - H" using assms(1) forest_modulo_equivalence_def dTransHd_le_1 le_supI1 by blast hence "(- 1 \ H) * d\<^sup>T * H \ - d\<^sup>T" by (metis assms(1) forest_modulo_equivalence_def dTransHd_le_1 inf.sup_monoid.add_commute le_infI2 p_antitone_iff regular_one_closed schroeder_4_p mult_assoc) hence "d * (- 1 \ H) * d\<^sup>T \ - H" by (metis assms(1) forest_modulo_equivalence_def conv_dist_comp schroeder_3_p triple_schroeder_p) hence "H \ d * (- 1 \ H) * d\<^sup>T \ 1" by (metis inf.coboundedI1 p_antitone_iff p_shunting_swap regular_one_closed) hence "H \ d * d\<^sup>T \ H \ d * (- 1 \ H) * d\<^sup>T \ 1" using assms(1) forest_modulo_equivalence_def le_supI by blast hence "H \ (d * 1 * d\<^sup>T \ d * (- 1 \ H) * d\<^sup>T) \ 1" using comp_inf.semiring.distrib_left by auto hence "H \ (d * (1 \ (- 1 \ H)) * d\<^sup>T) \ 1" by (simp add: mult_left_dist_sup mult_right_dist_sup) hence 514: "H \ d * H * d\<^sup>T \ 1" by (metis assms(1) forest_modulo_equivalence_def comp_inf.semiring.distrib_left inf.le_iff_sup inf.sup_monoid.add_commute inf_top_right regular_one_closed stone) thus ?thesis proof - have "H \ d * H * d\<^sup>T \ H * d * H * (d * H)\<^sup>\ \ d * H * d\<^sup>T \ 1" using 513 514 by simp hence "d * H * d\<^sup>T \ (H \ H * d * H * (d * H)\<^sup>\) \ 1" by (simp add: comp_inf.semiring.distrib_left inf.sup_monoid.add_commute) hence "d * H * d\<^sup>T \ H * (1 \ d * H * (d * H)\<^sup>\) \ 1" by (simp add: mult_left_dist_sup mult_assoc) thus ?thesis by (simp add: inf.sup_monoid.add_commute star_left_unfold_equal) qed qed have "?x * top * ?x\<^sup>T = (d \ top * e\<^sup>T * H \ (H * d\<^sup>T)\<^sup>\ * H * a\<^sup>T * top) * top * (d\<^sup>T \ H\<^sup>T * e\<^sup>T\<^sup>T * top\<^sup>T \ top\<^sup>T * a\<^sup>T\<^sup>T * H\<^sup>T * (d\<^sup>T\<^sup>T * H\<^sup>T)\<^sup>\)" by (simp add: conv_dist_comp conv_dist_inf conv_star_commute mult_assoc) also have "... = (d \ top * e\<^sup>T * H \ (H * d\<^sup>T)\<^sup>\ * H * a\<^sup>T * top) * top * (d\<^sup>T \ H * e * top \ top * a * H * (d * H)\<^sup>\)" using assms(1) forest_modulo_equivalence_def by auto also have "... = (H * d\<^sup>T)\<^sup>\ * H * a\<^sup>T * top \ (d \ top * e\<^sup>T * H) * top * (d\<^sup>T \ H * e * top \ top * a * H * (d * H)\<^sup>\)" by (metis inf_vector_comp vector_export_comp) also have "... = (H * d\<^sup>T)\<^sup>\ * H * a\<^sup>T * top \ (d \ top * e\<^sup>T * H) * top * top * (d\<^sup>T \ H * e * top \ top * a * H * (d * H)\<^sup>\)" by (simp add: vector_mult_closed) also have "... = (H * d\<^sup>T)\<^sup>\ * H * a\<^sup>T * top \ d * ((top * e\<^sup>T * H)\<^sup>T \ top) * top * (d\<^sup>T \ H * e * top \ top * a * H * (d * H)\<^sup>\)" by (simp add: covector_comp_inf_1 covector_mult_closed) also have "... = (H * d\<^sup>T)\<^sup>\ * H * a\<^sup>T * top \ d * ((top * e\<^sup>T * H)\<^sup>T \ (H * e * top)\<^sup>T) * d\<^sup>T \ top * a * H * (d * H)\<^sup>\" by (smt comp_associative comp_inf.star_star_absorb comp_inf_vector conv_star_commute covector_comp_inf covector_conv_vector fc_top star.circ_top total_conv_surjective vector_conv_covector vector_inf_comp) also have "... = (H * d\<^sup>T)\<^sup>\ * H * a\<^sup>T * top \ top * a * H * (d * H)\<^sup>\ \ d * ((top * e\<^sup>T * H)\<^sup>T \ (H * e * top)\<^sup>T) * d\<^sup>T" using inf.sup_monoid.add_assoc inf.sup_monoid.add_commute by auto also have "... = (H * d\<^sup>T)\<^sup>\ * H * a\<^sup>T * top * top * a * H * (d * H)\<^sup>\ \ d * ((top * e\<^sup>T * H)\<^sup>T \ (H * e * top)\<^sup>T) * d\<^sup>T" by (smt comp_inf.star.circ_decompose_9 comp_inf.star_star_absorb comp_inf_covector fc_top star.circ_decompose_11 star.circ_top vector_export_comp) also have "... = (H * d\<^sup>T)\<^sup>\ * H * a\<^sup>T * top * a * H * (d * H)\<^sup>\ \ d * (H * e * top \ top * e\<^sup>T * H) * d\<^sup>T" using assms(1) forest_modulo_equivalence_def conv_dist_comp mult_assoc by auto also have "... = (H * d\<^sup>T)\<^sup>\ * H * a\<^sup>T * top * a * H * (d * H)\<^sup>\ \ d * H * e * top * e\<^sup>T * H * d\<^sup>T" by (metis comp_inf_covector inf_top.left_neutral mult_assoc) also have "... \ (H * d\<^sup>T)\<^sup>\ * (H * d)\<^sup>\ * H \ d * H * e * top * e\<^sup>T * H * d\<^sup>T" proof - have "(H * d\<^sup>T)\<^sup>\ * H * a\<^sup>T * top * a * H * (d * H)\<^sup>\ \ (H * d\<^sup>T)\<^sup>\ * H * 1 * H * (d * H)\<^sup>\" using 2 by (metis comp_associative comp_isotone mult_left_isotone mult_semi_associative star.circ_transitive_equal) also have "... = (H * d\<^sup>T)\<^sup>\ * H * (d * H)\<^sup>\" using assms(1) forest_modulo_equivalence_def mult.semigroup_axioms preorder_idempotent semigroup.assoc by fastforce also have "... = (H * d\<^sup>T)\<^sup>\ * (H * d)\<^sup>\ * H" by (metis star_slide mult_assoc) finally show ?thesis using inf.sup_left_isotone by auto qed also have "... \ (H * d\<^sup>T)\<^sup>\ * (H * d)\<^sup>\ * H \ d * H * d\<^sup>T" proof - have "d * H * e * top * e\<^sup>T * H * d\<^sup>T \ d * H * 1 * H * d\<^sup>T" using 3 by (metis comp_isotone idempotent_one_closed mult_left_isotone mult_sub_right_one mult_assoc) also have "... \ d * H * d\<^sup>T" by (metis assms(1) forest_modulo_equivalence_def mult_left_isotone mult_one_associative mult_semi_associative preorder_idempotent) finally show ?thesis using inf.sup_right_isotone by auto qed also have "... = H * (d\<^sup>T * H)\<^sup>\ * (H * d)\<^sup>\ * H \ d * H * d\<^sup>T" by (metis assms(1) forest_modulo_equivalence_def comp_associative preorder_idempotent star_slide) also have "... = H * ((d\<^sup>T * H)\<^sup>\ \ (H * d)\<^sup>\) * H \ d * H * d\<^sup>T" by (simp add: assms(1) expand_forest_modulo_equivalence mult.semigroup_axioms semigroup.assoc) also have "... = (H * (d\<^sup>T * H)\<^sup>\ * H \ H * (H * d)\<^sup>\ * H) \ d * H * d\<^sup>T" by (simp add: mult_left_dist_sup mult_right_dist_sup) also have "... = (H * d\<^sup>T)\<^sup>\ * H \ d * H * d\<^sup>T \ H * (d * H)\<^sup>\ \ d * H * d\<^sup>T" by (smt assms(1) forest_modulo_equivalence_def inf_sup_distrib2 mult.semigroup_axioms preorder_idempotent star_slide semigroup.assoc) also have "... \ (H * d\<^sup>T)\<^sup>\ * H \ d * H * d\<^sup>T \ 1" using 51 comp_inf.semiring.add_left_mono by blast finally have "?x * top * ?x\<^sup>T \ 1" using 51 by (smt assms(1) forest_modulo_equivalence_def conv_dist_comp conv_dist_inf conv_dist_sup conv_involutive conv_star_commute equivalence_one_closed mult.semigroup_axioms sup.absorb2 semigroup.assoc conv_isotone conv_order) thus ?thesis by simp qed have 6: "?x\<^sup>T * top * ?x \ 1" proof - have "?x\<^sup>T * top * ?x = (d\<^sup>T \ H\<^sup>T * e\<^sup>T\<^sup>T * top\<^sup>T \ top\<^sup>T * a\<^sup>T\<^sup>T * H\<^sup>T * (d\<^sup>T\<^sup>T * H\<^sup>T)\<^sup>\) * top * (d \ top * e\<^sup>T * H \ (H * d\<^sup>T)\<^sup>\ * H * a\<^sup>T * top)" by (simp add: conv_dist_comp conv_dist_inf conv_star_commute mult_assoc) also have "... = (d\<^sup>T \ H * e * top \ top * a * H * (d * H)\<^sup>\) * top * (d \ top * e\<^sup>T * H \ (H * d\<^sup>T)\<^sup>\ * H * a\<^sup>T * top)" using assms(1) forest_modulo_equivalence_def by auto also have "... = H * e * top \ (d\<^sup>T \ top * a * H * (d * H)\<^sup>\) * top * (d \ top * e\<^sup>T * H \ (H * d\<^sup>T)\<^sup>\ * H * a\<^sup>T * top)" by (smt comp_associative inf.sup_monoid.add_assoc inf.sup_monoid.add_commute star.circ_left_top star.circ_top vector_inf_comp) also have "... = H * e * top \ d\<^sup>T * ((top * a * H * (d * H)\<^sup>\)\<^sup>T \ top) * (d \ top * e\<^sup>T * H \ (H * d\<^sup>T)\<^sup>\ * H * a\<^sup>T * top)" by (simp add: covector_comp_inf_1 covector_mult_closed) also have "... = H * e * top \ d\<^sup>T * (d * H)\<^sup>\\<^sup>T * H * a\<^sup>T * top * (d \ top * e\<^sup>T * H \ (H * d\<^sup>T)\<^sup>\ * H * a\<^sup>T * top)" using assms(1) forest_modulo_equivalence_def comp_associative conv_dist_comp by auto also have "... = H * e * top \ d\<^sup>T * (d * H)\<^sup>\\<^sup>T * H * a\<^sup>T * top * (d \ (H * d\<^sup>T)\<^sup>\ * H * a\<^sup>T * top) \ top * e\<^sup>T * H" by (smt comp_associative comp_inf_covector inf.sup_monoid.add_assoc inf.sup_monoid.add_commute) also have "... = H * e * top \ d\<^sup>T * (d * H)\<^sup>\\<^sup>T * H * a\<^sup>T * (top \ ((H * d\<^sup>T)\<^sup>\ * H * a\<^sup>T * top)\<^sup>T) * d \ top * e\<^sup>T * H" by (metis comp_associative comp_inf_vector vector_conv_covector vector_top_closed) also have "... = H * e * top \ (H * e * top)\<^sup>T \ d\<^sup>T * (d * H)\<^sup>\\<^sup>T * H * a\<^sup>T * ((H * d\<^sup>T)\<^sup>\ * H * a\<^sup>T * top)\<^sup>T * d" by (smt assms(1) forest_modulo_equivalence_def conv_dist_comp inf.left_commute inf.sup_monoid.add_commute symmetric_top_closed mult_assoc inf_top.left_neutral) also have "... = H * e * top * (H * e * top)\<^sup>T \ d\<^sup>T * (d * H)\<^sup>\\<^sup>T * H * a\<^sup>T * ((H * d\<^sup>T)\<^sup>\ * H * a\<^sup>T * top)\<^sup>T * d" using vector_covector vector_mult_closed by auto also have "... = H * e * top * top\<^sup>T * e\<^sup>T * H\<^sup>T \ d\<^sup>T * (d * H)\<^sup>\\<^sup>T * H * a\<^sup>T * top\<^sup>T * a\<^sup>T\<^sup>T * H\<^sup>T * (H * d\<^sup>T)\<^sup>\\<^sup>T * d" by (smt conv_dist_comp mult.semigroup_axioms symmetric_top_closed semigroup.assoc) also have "... = H * e * top * top * e\<^sup>T * H \ d\<^sup>T * (H * d\<^sup>T)\<^sup>\ * H * a\<^sup>T * top * a * H * (d * H)\<^sup>\ * d" using assms(1) forest_modulo_equivalence_def conv_dist_comp conv_star_commute by auto also have "... = H * e * top * e\<^sup>T * H \ d\<^sup>T * (H * d\<^sup>T)\<^sup>\ * H * a\<^sup>T * top * a * H * (d * H)\<^sup>\ * d" using vector_top_closed mult_assoc by auto also have "... \ H \ d\<^sup>T * (H * d\<^sup>T)\<^sup>\ * H * (d * H)\<^sup>\ * d" proof - have "H * e * top * e\<^sup>T * H \ H * 1 * H" using 3 by (metis comp_associative mult_left_isotone mult_right_isotone) also have "... = H" using assms(1) forest_modulo_equivalence_def preorder_idempotent by auto finally have 611: "H * e * top * e\<^sup>T * H \ H" by simp have "d\<^sup>T * (H * d\<^sup>T)\<^sup>\ * H * a\<^sup>T * top * a * H * (d * H)\<^sup>\ * d \ d\<^sup>T * (H * d\<^sup>T)\<^sup>\ * H * 1 * H * (d * H)\<^sup>\ * d" using 2 by (metis comp_associative mult_left_isotone mult_right_isotone) also have "... = d\<^sup>T * (H * d\<^sup>T)\<^sup>\ * H * (d * H)\<^sup>\ * d" using assms(1) forest_modulo_equivalence_def mult.semigroup_axioms preorder_idempotent semigroup.assoc by fastforce finally have "d\<^sup>T * (H * d\<^sup>T)\<^sup>\ * H * a\<^sup>T * top * a * H * (d * H)\<^sup>\ * d \ d\<^sup>T * (H * d\<^sup>T)\<^sup>\ * H * (d * H)\<^sup>\ * d" by simp thus ?thesis using 611 comp_inf.comp_isotone by blast qed also have "... = H \ (d\<^sup>T * H)\<^sup>\ * d\<^sup>T * H * d * (H * d)\<^sup>\" using star_slide mult_assoc by auto also have "... \ H \ (d\<^sup>T * H)\<^sup>\ * (H * d)\<^sup>\" proof - have "(d\<^sup>T * H)\<^sup>\ * d\<^sup>T * H * d * (H * d)\<^sup>\ \ (d\<^sup>T * H)\<^sup>\ * 1 * (H * d)\<^sup>\" by (smt assms(1) forest_modulo_equivalence_def conv_dist_comp mult_left_isotone mult_right_isotone preorder_idempotent mult_assoc) also have "... = (d\<^sup>T * H)\<^sup>\ * (H * d)\<^sup>\" by simp finally show ?thesis using inf.sup_right_isotone by blast qed also have "... = H \ ((d\<^sup>T * H)\<^sup>\ \ (H * d)\<^sup>\)" by (simp add: assms(1) expand_forest_modulo_equivalence) also have "... = H \ (d\<^sup>T * H)\<^sup>\ \ H \ (H * d)\<^sup>\" by (simp add: comp_inf.semiring.distrib_left) also have "... = 1 \ H \ (d\<^sup>T * H)\<^sup>+ \ H \ (H * d)\<^sup>+" proof - have 612: "H \ (H * d)\<^sup>\ = 1 \ H \ (H * d)\<^sup>+" using assms(1) forest_modulo_equivalence_def reflexive_inf_star by blast have "H \ (d\<^sup>T * H)\<^sup>\ = 1 \ H \ (d\<^sup>T * H)\<^sup>+" using assms(1) forest_modulo_equivalence_def reflexive_inf_star by auto thus ?thesis using 612 sup_assoc sup_commute by auto qed also have "... \ 1" proof - have 613: "H \ (H * d)\<^sup>+ \ 1" by (metis assms(3) inf.coboundedI1 p_antitone_iff p_shunting_swap regular_one_closed) hence "H \ (d\<^sup>T * H)\<^sup>+ \ 1" by (metis assms(1) forest_modulo_equivalence_def conv_dist_comp conv_dist_inf conv_plus_commute coreflexive_symmetric) thus ?thesis by (simp add: 613) qed finally show ?thesis by simp qed have 7:"bijective (?x * top)" using 4 5 6 arc_expanded by blast have "bijective (?x\<^sup>T * top)" using 4 5 6 arc_expanded by blast thus ?thesis using 7 by simp qed text \ To maintain that $f$ can be extended to a minimum spanning forest we identify an edge, $i = v \sqcap \overline{F}e\top \sqcap \top e^\top F$, that may be exchanged with the \selected_edge\, $e$. Here, we show that $i$ is an \arc\. \ lemma boruvka_edge_arc: assumes "equivalence F" and "forest v" and "arc e" and "regular F" and "F \ forest_components (F \ v)" and "regular v" and "v * e\<^sup>T = bot" and "e * F * e = bot" and "e\<^sup>T \ v\<^sup>\" and "e \ bot" shows "arc (v \ -F * e * top \ top * e\<^sup>T * F)" proof - let ?i = "v \ -F * e * top \ top * e\<^sup>T * F" have 1: "?i\<^sup>T * top * ?i \ 1" proof - have "?i\<^sup>T * top * ?i = (v\<^sup>T \ top * e\<^sup>T * -F \ F * e * top) * top * (v \ -F * e * top \ top * e\<^sup>T * F)" using assms(1) conv_complement conv_dist_comp conv_dist_inf mult.semigroup_axioms semigroup.assoc by fastforce also have "... = F * e * top \ (v\<^sup>T \ top * e\<^sup>T * -F) * top * (v \ -F * e * top) \ top * e\<^sup>T * F" by (smt covector_comp_inf covector_mult_closed inf_vector_comp vector_export_comp vector_top_closed) also have "... = F * e * top \ (v\<^sup>T \ top * e\<^sup>T * -F) * top * top * (v \ -F * e * top) \ top * e\<^sup>T * F" by (simp add: comp_associative) also have "... = F * e * top \ v\<^sup>T * (top \ (top * e\<^sup>T * -F)\<^sup>T) * top * (v \ -F * e * top) \ top * e\<^sup>T * F" using comp_associative comp_inf_vector_1 by auto also have "... = F * e * top \ v\<^sup>T * (top \ (top * e\<^sup>T * -F)\<^sup>T) * (top \ (-F * e * top)\<^sup>T) * v \ top * e\<^sup>T * F" by (smt comp_inf_vector conv_dist_comp mult.semigroup_axioms symmetric_top_closed semigroup.assoc) also have "... = F * e * top \ v\<^sup>T * (top * e\<^sup>T * -F)\<^sup>T * (-F * e * top)\<^sup>T * v \ top * e\<^sup>T * F" by simp also have "... = F * e * top \ v\<^sup>T * -F\<^sup>T * e\<^sup>T\<^sup>T * top\<^sup>T * top\<^sup>T * e\<^sup>T * -F\<^sup>T * v \ top * e\<^sup>T * F" by (metis comp_associative conv_complement conv_dist_comp) also have "... = F * e * top \ v\<^sup>T * -F * e * top * top * e\<^sup>T * -F * v \ top * e\<^sup>T * F" by (simp add: assms(1)) also have "... = F * e * top \ v\<^sup>T * -F * e * top \ top * e\<^sup>T * -F * v \ top * e\<^sup>T * F" by (metis comp_associative comp_inf_covector inf.sup_monoid.add_assoc inf_top.left_neutral vector_top_closed) also have "... = (F \ v\<^sup>T * -F) * e * top \ top * e\<^sup>T * -F * v \ top * e\<^sup>T * F" using assms(3) injective_comp_right_dist_inf mult_assoc by auto also have "... = (F \ v\<^sup>T * -F) * e * top \ top * e\<^sup>T * (F \ -F * v)" using assms(3) conv_dist_comp inf.sup_monoid.add_assoc inf.sup_monoid.add_commute mult.semigroup_axioms univalent_comp_left_dist_inf semigroup.assoc by fastforce also have "... = (F \ v\<^sup>T * -F) * e * top * top * e\<^sup>T * (F \ -F * v)" by (metis comp_associative comp_inf_covector inf_top.left_neutral vector_top_closed) also have "... = (F \ v\<^sup>T * -F) * e * top * e\<^sup>T * (F \ -F * v)" by (simp add: comp_associative) also have "... \ (F \ v\<^sup>T * -F) * (F \ -F * v)" by (smt assms(3) conv_dist_comp mult_left_isotone shunt_bijective symmetric_top_closed top_right_mult_increasing mult_assoc) also have "... \ (F \ v\<^sup>T * -F) * (F \ -F * v) \ F" by (metis assms(1) inf.absorb1 inf.cobounded1 mult_isotone preorder_idempotent) also have "... \ (F \ v\<^sup>T * -F) * (F \ -F * v) \ (F \ v)\<^sup>T\<^sup>\ * (F \ v)\<^sup>\" using assms(5) comp_inf.mult_right_isotone by auto also have "... \ (-F \ v\<^sup>T) * -F * -F * (-F \ v) \ (F \ v)\<^sup>T\<^sup>\ * (F \ v)\<^sup>\" proof - have "F \ v\<^sup>T * -F \ (v\<^sup>T \ F * -F\<^sup>T) * -F" by (metis conv_complement dedekind_2 inf_commute) also have "... = (v\<^sup>T \ -F\<^sup>T) * -F" using assms(1) equivalence_comp_left_complement by simp finally have "F \ v\<^sup>T * -F \ F \ (v\<^sup>T \ -F) * -F" using assms(1) by auto hence 11: "F \ v\<^sup>T * -F = F \ (-F \ v\<^sup>T) * -F" by (metis inf.antisym_conv inf.sup_monoid.add_commute comp_left_subdist_inf inf.boundedE inf.sup_right_isotone) hence "F\<^sup>T \ -F\<^sup>T * v\<^sup>T\<^sup>T = F\<^sup>T \ -F\<^sup>T * (-F\<^sup>T \ v\<^sup>T\<^sup>T)" by (metis (full_types) assms(1) conv_complement conv_dist_comp conv_dist_inf) hence 12: "F \ -F * v = F \ -F * (-F \ v)" using assms(1) by (simp add: abel_semigroup.commute inf.abel_semigroup_axioms) have "(F \ v\<^sup>T * -F) * (F \ -F * v) = (F \ (-F \ v\<^sup>T) * -F) * (F \ -F * (-F \ v))" using 11 12 by auto also have "... \ (-F \ v\<^sup>T) * -F * -F * (-F \ v)" by (metis comp_associative comp_isotone inf.cobounded2) finally show ?thesis using comp_inf.mult_left_isotone by blast qed also have "... = ((-F \ v\<^sup>T) * -F * -F * (-F \ v) \ (F \ v)\<^sup>T * (F \ v)\<^sup>T\<^sup>\ * (F \ v)\<^sup>\) \ ((-F \ v\<^sup>T) * -F * -F * (-F \ v) \ (F \ v)\<^sup>\)" by (metis comp_associative inf_sup_distrib1 star.circ_loop_fixpoint) also have "... = ((-F \ v\<^sup>T) * -F * -F * (-F \ v) \ (F \ v\<^sup>T) * (F \ v)\<^sup>T\<^sup>\ * (F \ v)\<^sup>\) \ ((-F \ v\<^sup>T) * -F * -F * (-F \ v) \ (F \ v)\<^sup>\)" using assms(1) conv_dist_inf by auto also have "... = bot \ ((-F \ v\<^sup>T) * -F * -F * (-F \ v) \ (F \ v)\<^sup>\)" proof - have "(-F \ v\<^sup>T) * -F * -F * (-F \ v) \ (F \ v\<^sup>T) * (F \ v)\<^sup>T\<^sup>\ * (F \ v)\<^sup>\ \ bot" using assms(1, 2) forests_bot_2 by (simp add: comp_associative) thus ?thesis using le_bot by blast qed also have "... = (-F \ v\<^sup>T) * -F * -F * (-F \ v) \ (1 \ (F \ v)\<^sup>\ * (F \ v))" by (simp add: star.circ_plus_same star_left_unfold_equal) also have "... = ((-F \ v\<^sup>T) * -F * -F * (-F \ v) \ 1) \ ((-F \ v\<^sup>T) * -F * -F * (-F \ v) \ (F \ v)\<^sup>\ * (F \ v))" by (simp add: comp_inf.semiring.distrib_left) also have "... \ 1 \ ((-F \ v\<^sup>T) * -F * -F * (-F \ v) \ (F \ v)\<^sup>\ * (F \ v))" using sup_left_isotone by auto also have "... \ 1 \ bot" using assms(1, 2) forests_bot_3 comp_inf.semiring.add_left_mono by simp finally show ?thesis by simp qed have 2: "?i * top * ?i\<^sup>T \ 1" proof - have "?i * top * ?i\<^sup>T = (v \ -F * e * top \ top * e\<^sup>T * F) * top * (v\<^sup>T \ (-F * e * top)\<^sup>T \ (top * e\<^sup>T * F)\<^sup>T)" by (simp add: conv_dist_inf) also have "... = (v \ -F * e * top \ top * e\<^sup>T * F) * top * (v\<^sup>T \ top\<^sup>T * e\<^sup>T * -F\<^sup>T \ F\<^sup>T * e\<^sup>T\<^sup>T * top\<^sup>T)" by (simp add: conv_complement conv_dist_comp mult_assoc) also have "... = (v \ -F * e * top \ top * e\<^sup>T * F) * top * (v\<^sup>T \ top * e\<^sup>T * -F \ F * e * top)" by (simp add: assms(1)) also have "... = -F * e * top \ (v \ top * e\<^sup>T * F) * top * (v\<^sup>T \ top * e\<^sup>T * -F \ F * e * top)" by (smt inf.left_commute inf.sup_monoid.add_assoc vector_export_comp) also have "... = -F * e * top \ (v \ top * e\<^sup>T * F) * top * (v\<^sup>T \ F * e * top) \ top * e\<^sup>T * -F" by (smt comp_inf_covector inf.sup_monoid.add_assoc inf.sup_monoid.add_commute mult_assoc) also have "... = -F * e * top \ (v \ top * e\<^sup>T * F) * top * top * (v\<^sup>T \ F * e * top) \ top * e\<^sup>T * -F" by (simp add: mult_assoc) also have "... = -F * e * top \ v * ((top * e\<^sup>T * F)\<^sup>T \ top) * top * (v\<^sup>T \ F * e * top) \ top * e\<^sup>T * -F" by (simp add: comp_inf_vector_1 mult.semigroup_axioms semigroup.assoc) also have "... = -F * e * top \ v * ((top * e\<^sup>T * F)\<^sup>T \ top) * (top \ (F * e * top)\<^sup>T) * v\<^sup>T \ top * e\<^sup>T * -F" by (smt comp_inf_vector covector_comp_inf vector_conv_covector vector_mult_closed vector_top_closed) also have "... = -F * e * top \ v * (top * e\<^sup>T * F)\<^sup>T * (F * e * top)\<^sup>T * v\<^sup>T \ top * e\<^sup>T * -F" by simp also have "... = -F * e * top \ v * F\<^sup>T * e\<^sup>T\<^sup>T * top\<^sup>T * top\<^sup>T * e\<^sup>T * F\<^sup>T * v\<^sup>T \ top * e\<^sup>T * -F" by (metis comp_associative conv_dist_comp) also have "... = -F * e * top \ v * F * e * top * top * e\<^sup>T * F * v\<^sup>T \ top * e\<^sup>T * -F" using assms(1) by auto also have "... = -F * e * top \ v * F * e * top \ top * e\<^sup>T * F * v\<^sup>T \ top * e\<^sup>T * -F" by (smt comp_associative comp_inf_covector inf.sup_monoid.add_assoc inf_top.left_neutral vector_top_closed) also have "... = (-F \ v * F) * e * top \ top * e\<^sup>T * F * v\<^sup>T \ top * e\<^sup>T * -F" using injective_comp_right_dist_inf assms(3) mult.semigroup_axioms semigroup.assoc by fastforce also have "... = (-F \ v * F) * e * top \ top * e\<^sup>T * (F * v\<^sup>T \ -F)" using injective_comp_right_dist_inf assms(3) conv_dist_comp inf.sup_monoid.add_assoc mult.semigroup_axioms univalent_comp_left_dist_inf semigroup.assoc by fastforce also have "... = (-F \ v * F) * e * top * top * e\<^sup>T * (F * v\<^sup>T \ -F)" by (metis inf_top_right vector_export_comp vector_top_closed) also have "... = (-F \ v * F) * e * top * e\<^sup>T * (F * v\<^sup>T \ -F)" by (simp add: comp_associative) also have "... \ (-F \ v * F) * (F * v\<^sup>T \ -F)" by (smt assms(3) conv_dist_comp mult.semigroup_axioms mult_left_isotone shunt_bijective symmetric_top_closed top_right_mult_increasing semigroup.assoc) also have "... = (-F \ v * F) * ((v * F)\<^sup>T \ -F)" by (simp add: assms(1) conv_dist_comp) also have "... = (-F \ v * F) * (-F \ v * F)\<^sup>T" using assms(1) conv_complement conv_dist_inf by (simp add: inf.sup_monoid.add_commute) also have "... \ (-F \ v) * (F \ v)\<^sup>\ * (F \ v)\<^sup>T\<^sup>\ * (-F \ v)\<^sup>T" proof - let ?Fv = "F \ v" have "-F \ v * F \ -F \ v * (F \ v)\<^sup>T\<^sup>\ * (F \ v)\<^sup>\" using assms(5) inf.sup_right_isotone mult_right_isotone comp_associative by auto also have "... \ -F \ v * (F \ v)\<^sup>\" proof - have "v * v\<^sup>T \ 1" by (simp add: assms(2)) hence "v * v\<^sup>T * F \ F" using assms(1) dual_order.trans mult_left_isotone by blast hence "v * v\<^sup>T * F\<^sup>T\<^sup>\ * F\<^sup>\ \ F" by (metis assms(1) mult_1_right preorder_idempotent star.circ_sup_one_right_unfold star.circ_transitive_equal star_one star_simulation_right_equal mult_assoc) hence "v * (F \ v)\<^sup>T * F\<^sup>T\<^sup>\ * F\<^sup>\ \ F" by (meson conv_isotone dual_order.trans inf.cobounded2 inf.sup_monoid.add_commute mult_left_isotone mult_right_isotone) hence "v * (F \ v)\<^sup>T * (F \ v)\<^sup>T\<^sup>\ * (F \ v)\<^sup>\ \ F" by (meson conv_isotone dual_order.trans inf.cobounded2 inf.sup_monoid.add_commute mult_left_isotone mult_right_isotone comp_isotone conv_dist_inf inf.cobounded1 star_isotone) hence "-F \ v * (F \ v)\<^sup>T * (F \ v)\<^sup>T\<^sup>\ * (F \ v)\<^sup>\ \ bot" using order.eq_iff p_antitone pseudo_complement by auto hence "(-F \ v * (F \ v)\<^sup>T * (F \ v)\<^sup>T\<^sup>\ * (F \ v)\<^sup>\) \ v * (v \ F)\<^sup>\ \ v * (v \ F)\<^sup>\" using bot_least le_bot by fastforce hence "(-F \ v * (v \ F)\<^sup>\) \ (v * (F \ v)\<^sup>T * (F \ v)\<^sup>T\<^sup>\ * (F \ v)\<^sup>\ \ v * (v \ F)\<^sup>\) \ v * (v \ F)\<^sup>\" by (simp add: sup_inf_distrib2) hence "(-F \ v * (v \ F)\<^sup>\) \ v * ((F \ v)\<^sup>T * (F \ v)\<^sup>T\<^sup>\ \ 1) * (v \ F)\<^sup>\ \ v * (v \ F)\<^sup>\" by (simp add: inf.sup_monoid.add_commute mult.semigroup_axioms mult_left_dist_sup mult_right_dist_sup semigroup.assoc) hence "(-F \ v * (v \ F)\<^sup>\) \ v * (F \ v)\<^sup>T\<^sup>\ * (v \ F)\<^sup>\ \ v * (v \ F)\<^sup>\" by (simp add: star_left_unfold_equal sup_commute) hence "-F \ v * (F \ v)\<^sup>T\<^sup>\ * (v \ F)\<^sup>\ \ v * (v \ F)\<^sup>\" using comp_inf.mult_right_sub_dist_sup_left inf.order_lesseq_imp by blast thus ?thesis by (simp add: inf.sup_monoid.add_commute) qed also have "... \ (v \ -F * (F \ v)\<^sup>T\<^sup>\) * (F \ v)\<^sup>\" by (metis dedekind_2 conv_star_commute inf.sup_monoid.add_commute) also have "... \ (v \ -F * F\<^sup>T\<^sup>\) * (F \ v)\<^sup>\" using conv_isotone inf.sup_right_isotone mult_left_isotone mult_right_isotone star_isotone by auto also have "... = (v \ -F * F) * (F \ v)\<^sup>\" by (metis assms(1) equivalence_comp_right_complement mult_left_one star_one star_simulation_right_equal) also have "... = (-F \ v) * (F \ v)\<^sup>\" using assms(1) equivalence_comp_right_complement inf.sup_monoid.add_commute by auto finally have "-F \ v * F \ (-F \ v) * (F \ v)\<^sup>\" by simp hence "(-F \ v * F) * (-F \ v * F)\<^sup>T \ (-F \ v) * (F \ v)\<^sup>\ * ((-F \ v) * (F \ v)\<^sup>\)\<^sup>T" by (simp add: comp_isotone conv_isotone) also have "... = (-F \ v) * (F \ v)\<^sup>\ * (F \ v)\<^sup>T\<^sup>\ * (-F \ v)\<^sup>T" by (simp add: comp_associative conv_dist_comp conv_star_commute) finally show ?thesis by simp qed also have "... \ (-F \ v) * ((F \ v\<^sup>\) \ (F \ v\<^sup>T\<^sup>\)) * (-F \ v)\<^sup>T" proof - have "(F \ v)\<^sup>\ * (F \ v)\<^sup>T\<^sup>\ \ F\<^sup>\ * F\<^sup>T\<^sup>\" using fc_isotone by auto also have "... \ F * F" by (metis assms(1) preorder_idempotent star.circ_sup_one_left_unfold star.circ_transitive_equal star_right_induct_mult) finally have 21: "(F \ v)\<^sup>\ * (F \ v)\<^sup>T\<^sup>\ \ F" using assms(1) dual_order.trans by blast have "(F \ v)\<^sup>\ * (F \ v)\<^sup>T\<^sup>\ \ v\<^sup>\ * v\<^sup>T\<^sup>\" by (simp add: fc_isotone) hence "(F \ v)\<^sup>\ * (F \ v)\<^sup>T\<^sup>\ \ F \ v\<^sup>\ * v\<^sup>T\<^sup>\" using 21 by simp also have "... = F \ (v\<^sup>\ \ v\<^sup>T\<^sup>\)" by (simp add: assms(2) cancel_separate_eq) finally show ?thesis by (metis assms(4, 6) comp_associative comp_inf.semiring.distrib_left comp_isotone inf_pp_semi_commute mult_left_isotone regular_closed_inf) qed also have "... \ (-F \ v) * (F \ v\<^sup>T\<^sup>\) * (-F \ v)\<^sup>T \ (-F \ v) * (F \ v\<^sup>\) * (-F \ v)\<^sup>T" by (simp add: mult_left_dist_sup mult_right_dist_sup) also have "... \ (-F \ v) * (-F \ v)\<^sup>T \ (-F \ v) * (-F \ v)\<^sup>T" proof - have "(-F \ v) * (F \ v\<^sup>T\<^sup>\) \ (-F \ v) * ((F \ v)\<^sup>T\<^sup>\ * (F \ v)\<^sup>\ \ v\<^sup>T\<^sup>\)" by (simp add: assms(5) inf.coboundedI1 mult_right_isotone) also have "... = (-F \ v) * ((F \ v)\<^sup>T * (F \ v)\<^sup>T\<^sup>\ * (F \ v)\<^sup>\ \ v\<^sup>T\<^sup>\) \ (-F \ v) * ((F \ v)\<^sup>\ \ v\<^sup>T\<^sup>\)" by (metis comp_associative comp_inf.mult_right_dist_sup mult_left_dist_sup star.circ_loop_fixpoint) also have "... \ (-F \ v) * (F \ v)\<^sup>T * top \ (-F \ v) * ((F \ v)\<^sup>\ \ v\<^sup>T\<^sup>\)" by (simp add: comp_associative comp_isotone inf.coboundedI2 inf.sup_monoid.add_commute le_supI1) also have "... \ (-F \ v) * (F \ v)\<^sup>T * top \ (-F \ v) * (v\<^sup>\ \ v\<^sup>T\<^sup>\)" by (smt comp_inf.mult_right_isotone comp_inf.semiring.add_mono order.eq_iff inf.cobounded2 inf.sup_monoid.add_commute mult_right_isotone star_isotone) also have "... \ bot \ (-F \ v) * (v\<^sup>\ \ v\<^sup>T\<^sup>\)" by (metis assms(1, 2) forests_bot_1 comp_associative comp_inf.semiring.add_right_mono mult_semi_associative vector_bot_closed) also have "... \ -F \ v" by (simp add: assms(2) acyclic_star_inf_conv) finally have 22: "(-F \ v) * (F \ v\<^sup>T\<^sup>\) \ -F \ v" by simp have "((-F \ v) * (F \ v\<^sup>T\<^sup>\))\<^sup>T = (F \ v\<^sup>\) * (-F \ v)\<^sup>T" by (simp add: assms(1) conv_dist_inf conv_star_commute conv_dist_comp) hence "(F \ v\<^sup>\) * (-F \ v)\<^sup>T \ (-F \ v)\<^sup>T" using 22 conv_isotone by fastforce thus ?thesis using 22 by (metis assms(4, 6) comp_associative comp_inf.pp_comp_semi_commute comp_inf.semiring.add_mono comp_isotone inf_pp_commute mult_left_isotone) qed also have "... = (-F \ v) * (-F \ v)\<^sup>T" by simp also have "... \ v * v\<^sup>T" by (simp add: comp_isotone conv_isotone) also have "... \ 1" by (simp add: assms(2)) thus ?thesis using calculation dual_order.trans by blast qed have 3: "top * ?i * top = top" proof - have 31: "regular (e\<^sup>T * -F * v * F * e)" using assms(3, 4, 6) arc_regular regular_mult_closed by auto have 32: "bijective ((top * e\<^sup>T)\<^sup>T)" using assms(3) by (simp add: conv_dist_comp) have "top * ?i * top = top * (v \ -F * e * top) * ((top * e\<^sup>T * F)\<^sup>T \ top)" by (simp add: comp_associative comp_inf_vector_1) also have "... = (top \ (-F * e * top)\<^sup>T) * v * ((top * e\<^sup>T * F)\<^sup>T \ top)" using comp_inf_vector conv_dist_comp by auto also have "... = (-F * e * top)\<^sup>T * v * (top * e\<^sup>T * F)\<^sup>T" by simp also have "... = top\<^sup>T * e\<^sup>T * -F\<^sup>T * v * F\<^sup>T * e\<^sup>T\<^sup>T * top\<^sup>T" by (simp add: comp_associative conv_complement conv_dist_comp) finally have 33: "top * ?i * top = top * e\<^sup>T * -F * v * F * e * top" by (simp add: assms(1)) have "top * ?i * top \ bot" proof (rule ccontr) assume "\ top * (v \ - F * e * top \ top * e\<^sup>T * F) * top \ bot" hence "top * e\<^sup>T * -F * v * F * e * top = bot" using 33 by auto hence "e\<^sup>T * -F * v * F * e = bot" using 31 tarski comp_associative le_bot by fastforce hence "top * (-F * v * F * e)\<^sup>T \ -(e\<^sup>T)" by (metis comp_associative conv_complement_sub_leq conv_involutive p_bot schroeder_5_p) hence "top * e\<^sup>T * F\<^sup>T * v\<^sup>T * -F\<^sup>T \ -(e\<^sup>T)" by (simp add: comp_associative conv_complement conv_dist_comp) hence "v * F * e * top * e\<^sup>T \ F" by (metis assms(1, 4) comp_associative conv_dist_comp schroeder_3_p symmetric_top_closed) hence "v * F * e * top * top * e\<^sup>T \ F" by (simp add: comp_associative) hence "v * F * e * top \ F * (top * e\<^sup>T)\<^sup>T" using 32 by (metis shunt_bijective comp_associative conv_involutive) hence "v * F * e * top \ F * e * top" using comp_associative conv_dist_comp by auto hence "v\<^sup>\ * F * e * top \ F * e * top" using comp_associative star_left_induct_mult_iff by auto hence "e\<^sup>T * F * e * top \ F * e * top" by (meson assms(9) mult_left_isotone order_trans) hence "e\<^sup>T * F * e * top * (e * top)\<^sup>T \ F" using 32 shunt_bijective assms(3) mult_assoc by auto hence 34: "e\<^sup>T * F * e * top * top * e\<^sup>T \ F" by (metis conv_dist_comp mult.semigroup_axioms symmetric_top_closed semigroup.assoc) hence "e\<^sup>T \ F" proof - have "e\<^sup>T \ e\<^sup>T * e * e\<^sup>T" by (metis conv_involutive ex231c) also have "... \ e\<^sup>T * F * e * e\<^sup>T" using assms(1) comp_associative mult_left_isotone mult_right_isotone by fastforce also have "... \ e\<^sup>T * F * e * top * top * e\<^sup>T" by (simp add: mult_left_isotone top_right_mult_increasing vector_mult_closed) finally show ?thesis using 34 by simp qed hence 35: "e \ F" using assms(1) conv_order by fastforce have "top * (F * e)\<^sup>T \ - e" using assms(8) comp_associative schroeder_4_p by auto hence "top * e\<^sup>T * F \ - e" by (simp add: assms(1) comp_associative conv_dist_comp) hence "(top * e\<^sup>T)\<^sup>T * e \ - F" using schroeder_3_p by auto hence "e * top * e \ - F" by (simp add: conv_dist_comp) hence "e \ - F" by (simp add: assms(3) arc_top_arc) hence "e \ F \ - F" using 35 inf.boundedI by blast hence "e = bot" using bot_unique by auto thus "False" using assms(10) by auto qed thus ?thesis by (metis assms(3, 4, 6) arc_regular regular_closed_inf regular_closed_top regular_conv_closed regular_mult_closed semiring.mult_not_zero tarski) qed have "bijective (?i * top) \ bijective (?i\<^sup>T * top)" using 1 2 3 arc_expanded by blast thus ?thesis by blast qed subsubsection \Comparison of edge weights\ text \ In this section we compare the weight of the \selected_edge\ with other edges of interest. For example, Theorem \e_leq_c_c_complement_transpose_general\ is used to show that the \selected_edge\ has its source inside and its target outside the component it is chosen for. \ lemma e_leq_c_c_complement_transpose_general: assumes "e = minarc (v * -(v)\<^sup>T \ g)" and "regular v" shows "e \ v * -(v)\<^sup>T" proof - have "e \ -- (v * - v\<^sup>T \ g)" using assms(1) minarc_below order_trans by blast also have "... \ -- (v * - v\<^sup>T)" using order_lesseq_imp pp_isotone_inf by blast also have "... = v * - v\<^sup>T" using assms(2) regular_mult_closed by auto finally show ?thesis by simp qed lemma x_leq_c_transpose_general: assumes "vector_classes x v" and "a\<^sup>T * top \ x * e * top" and "e \ v * -v\<^sup>T" shows "a \ v\<^sup>T" proof - have 1: "equivalence x" using assms(1) using vector_classes_def by blast have "a \ top * a" using top_left_mult_increasing by blast also have "... \ (x * e * top)\<^sup>T" using assms(2) conv_dist_comp conv_isotone by fastforce also have "... = top * e\<^sup>T * x" using 1 by (simp add: conv_dist_comp mult_assoc) also have "... \ top * (v * - v\<^sup>T)\<^sup>T * x" by (metis assms(3) conv_dist_comp conv_isotone mult_left_isotone symmetric_top_closed) also have "... = top * (- v * v\<^sup>T) * x" by (simp add: conv_complement conv_dist_comp) also have "... \ top * v\<^sup>T * x" by (metis mult_left_isotone top.extremum mult_assoc) also have "... = v\<^sup>T * x" using assms(1) vector_classes_def vector_conv_covector by auto also have "... = v\<^sup>T" by (metis assms(1) order.antisym conv_dist_comp conv_order dual_order.trans mult_right_isotone mult_sub_right_one vector_classes_def) finally show ?thesis by simp qed lemma x_leq_c_complement_general: assumes "vector v" and "v * v\<^sup>T \ x" and "a \ v\<^sup>T" and "a \ -x" shows "a \ -v" proof - have "a \ -x \ v\<^sup>T" using assms(3, 4) by auto also have "... \ - v" proof - have "v \ v\<^sup>T \ x" using assms(1, 2) vector_covector by auto hence "-x \ v \ v\<^sup>T \ bot" using inf.sup_monoid.add_assoc p_antitone pseudo_complement by fastforce thus ?thesis using le_bot p_shunting_swap pseudo_complement by blast qed finally show ?thesis by simp qed lemma sum_e_below_sum_a_when_outgoing_same_component_general: assumes "e = minarc (v * -(v)\<^sup>T \ g)" and "symmetric g" and "arc a" and "a \ -x \ -- g" and "a\<^sup>T * top \ x * e * top" and "unique_vector_class x v" shows "sum (e \ g) \ sum (a \ g)" proof - have 1:"e \ v * - v\<^sup>T" using assms(1, 6) e_leq_c_c_complement_transpose_general unique_vector_class_def vector_classes_def by auto have 2: "a \ v\<^sup>T" using 1 assms(5) assms(6) x_leq_c_transpose_general unique_vector_class_def by blast hence "a \ -v" using assms(4, 6) inf.boundedE unique_vector_class_def vector_classes_def x_leq_c_complement_general by meson hence "a \ - v \ v\<^sup>T" using 2 by simp hence "a \ - v * v\<^sup>T" using assms(6) vector_complement_closed vector_covector unique_vector_class_def vector_classes_def by metis hence "a\<^sup>T \ v\<^sup>T\<^sup>T * - v\<^sup>T" using conv_complement conv_dist_comp conv_isotone by metis hence 3: "a\<^sup>T \ v * - v\<^sup>T" by simp hence "a \ -- g" using assms(4) by auto hence "a\<^sup>T \ -- g" using assms(2) conv_complement conv_isotone by fastforce hence "a\<^sup>T \ v * - v\<^sup>T \ -- g \ bot" using 3 assms(3, 6) inf.orderE semiring.mult_not_zero unique_vector_class_def vector_classes_def by metis hence "a\<^sup>T \ v * - v\<^sup>T \ g \ bot" using inf.sup_monoid.add_commute pp_inf_bot_iff by auto hence "sum (minarc (v * - v\<^sup>T \ g) \ (v * - v\<^sup>T \ g)) \ sum (a\<^sup>T \ v * - v\<^sup>T \ g)" using assms(3) minarc_min inf.sup_monoid.add_assoc by simp hence "sum (e \ v * - v\<^sup>T \ g) \ sum (a\<^sup>T \ v * - v\<^sup>T \ g)" using assms(1, 6) inf.sup_monoid.add_assoc by simp hence "sum (e \ g) \ sum (a\<^sup>T \ g)" using 1 3 by (metis inf.orderE) hence "sum (e \ g) \ sum (a \ g)" by (simp add: assms(2) sum_symmetric) thus ?thesis by simp qed lemma sum_e_below_sum_x_when_outgoing_same_component: assumes "symmetric g" and "vector j" and "forest h" and "regular h" and "x \ - forest_components h \ -- g" and "x\<^sup>T * top \ forest_components h * selected_edge h j g * top" and "j \ bot" and "arc x" shows "sum (selected_edge h j g \ g) \ sum (x \ g)" proof - let ?e = "selected_edge h j g" let ?c = "choose_component (forest_components h) j" let ?H = "forest_components h" show ?thesis proof (rule sum_e_below_sum_a_when_outgoing_same_component_general) next show "?e = minarc (?c * - ?c\<^sup>T \ g)" by simp next show "symmetric g" by (simp add: assms(1)) next show "arc x" by (simp add: assms(8)) next show "x \ -?H \ -- g" using assms(5) by auto next show "x\<^sup>T * top \ ?H * ?e * top" by (simp add: assms(6)) next show "unique_vector_class ?H ?c" proof (unfold unique_vector_class_def, unfold vector_classes_def, intro conjI) next show "regular ?H" by (metis assms(4) conv_complement pp_dist_star regular_mult_closed) next show "regular ?c" using component_is_regular by auto next show "reflexive ?H" using assms(3) forest_components_equivalence by blast next show "transitive ?H" using assms(3) fch_equivalence by blast next show "symmetric ?H" by (simp add: assms(3) fch_equivalence) next show "vector ?c" by (simp add: assms(2, 6) component_is_vector) next show "?H * ?c \ ?c" using component_single by auto next show "?c \ bot" using assms(2, 6 , 7, 8) inf_bot_left le_bot minarc_bot mult_left_zero mult_right_zero by fastforce next show "?c * ?c\<^sup>T \ ?H" by (simp add: component_is_connected) qed qed qed text \ If there is a path in the \forest_modulo_equivalence\ from an edge between components, $a$, to the \selected_edge\, $e$, then the weight of $e$ is no greater than the weight of $a$. This is because either, \begin{itemize} \item the edges $a$ and $e$ are adjacent the same component so that we can use \sum_e_below_sum_x_when_outgoing_same_component\, or \item there is at least one edge between $a$ and $e$, namely $x$, the edge incoming to the component that $e$ is outgoing from. The path from $a$ to $e$ is split on $x$ using \forest_modulo_equivalence_path_split_disj\. We show that the weight of $e$ is no greater than the weight of $x$ by making use of lemma \sum_e_below_sum_x_when_outgoing_same_component\. We define $x$ in a way that we can show that the weight of $x$ is no greater than the weight of $a$ using the invariant. Then, it follows that the weight of $e$ is no greater than the weight of $a$ owing to transitivity. \end{itemize} \ lemma a_to_e_in_forest_modulo_equivalence: assumes "symmetric g" and "f \ --g" and "vector j" and "forest h" and "forest_modulo_equivalence (forest_components h) d" and "f \ f\<^sup>T = h \ h\<^sup>T \ d \ d\<^sup>T" and "(\ a b . forest_modulo_equivalence_path a b (forest_components h) d \ a \ -(forest_components h) \ -- g \ b \ d \ sum(b \ g) \ sum(a \ g))" and "regular d" and "j \ bot" and "b = selected_edge h j g" and "arc a" and "forest_modulo_equivalence_path a b (forest_components h) (d \ selected_edge h j g)" and "a \ - forest_components h \ -- g" and "regular h" shows "sum (b \ g) \ sum (a \ g)" proof - let ?p = "path f h j g" let ?e = "selected_edge h j g" let ?F = "forest_components f" let ?H = "forest_components h" have "sum (b \ g) \ sum (a \ g)" proof (cases "a\<^sup>T * top \ ?H * ?e * top") case True show "a\<^sup>T * top \ ?H * ?e * top \ sum (b \ g) \ sum (a \ g)" proof- have "sum (?e \ g) \ sum (a \ g)" proof (rule sum_e_below_sum_x_when_outgoing_same_component) show "symmetric g" using assms(1) by auto next show "vector j" using assms(3) by blast next show "forest h" by (simp add: assms(4)) next show "a \ - ?H \ -- g" using assms(13) by auto next show "a\<^sup>T * top \ ?H * ?e * top" using True by auto next show "j \ bot" by (simp add: assms(9)) next show "arc a" by (simp add: assms(11)) next show "regular h" using assms(14) by auto qed thus ?thesis using assms(10) by auto qed next case False show "\ a\<^sup>T * top \ ?H * ?e * top \ sum (b \ g) \ sum (a \ g)" proof - let ?d' = "d \ ?e" let ?x = "d \ top * ?e\<^sup>T * ?H \ (?H * d\<^sup>T)\<^sup>\ * ?H * a\<^sup>T * top" have 61: "arc (?x)" proof (rule shows_arc_x) show "forest_modulo_equivalence ?H d" by (simp add: assms(5)) next show "forest_modulo_equivalence_path a ?e ?H d" proof - have 611: "forest_modulo_equivalence_path a b ?H (d \ b)" using assms(10, 12) by auto have 616: "regular h" using assms(14) by auto have "regular a" using 611 forest_modulo_equivalence_path_def arc_regular by fastforce thus ?thesis using 616 by (smt forest_modulo_equivalence_path_split_disj assms(4, 8, 10, 12) forest_modulo_equivalence_path_def fch_equivalence minarc_regular regular_closed_star regular_conv_closed regular_mult_closed) qed next show "(?H * d)\<^sup>+ \ - ?H" using assms(5) forest_modulo_equivalence_def le_bot pseudo_complement by blast next show "\ a\<^sup>T * top \ ?H * ?e * top" by (simp add: False) next show "regular a" using assms(12) forest_modulo_equivalence_path_def arc_regular by auto next show "regular ?e" using minarc_regular by auto next show "regular ?H" using assms(14) pp_dist_star regular_conv_closed regular_mult_closed by auto next show "regular d" using assms(8) by auto qed have 62: "bijective (a\<^sup>T * top)" by (simp add: assms(11)) have 63: "bijective (?x * top)" using 61 by simp have 64: "?x \ (?H * d\<^sup>T)\<^sup>\ * ?H * a\<^sup>T * top" by simp hence "?x * top \ (?H * d\<^sup>T)\<^sup>\ * ?H * a\<^sup>T * top" using mult_left_isotone inf_vector_comp by auto hence "a\<^sup>T * top \ ((?H * d\<^sup>T)\<^sup>\ * ?H)\<^sup>T * ?x * top" using 62 63 64 by (smt bijective_reverse mult_assoc) also have "... = ?H * (d * ?H)\<^sup>\ * ?x * top" using conv_dist_comp conv_star_commute by auto also have "... = (?H * d)\<^sup>\ * ?H * ?x * top" by (simp add: star_slide) finally have "a\<^sup>T * top \ (?H * d)\<^sup>\ * ?H * ?x * top" by simp hence 65: "forest_modulo_equivalence_path a ?x ?H d" using 61 assms(12) forest_modulo_equivalence_path_def by blast have 66: "?x \ d" by (simp add: inf.sup_monoid.add_assoc) hence x_below_a: "sum (?x \ g) \ sum (a \ g)" using 65 forest_modulo_equivalence_path_def assms(7, 13) by blast have "sum (?e \ g) \ sum (?x \ g)" proof (rule sum_e_below_sum_x_when_outgoing_same_component) show "symmetric g" using assms(1) by auto next show "vector j" using assms(3) by blast next show "forest h" by (simp add: assms(4)) next show "?x \ - ?H \ -- g" proof - have 67: "?x \ - ?H" proof - have "?x \ d" using 66 by blast also have "... \ ?H * d" using dual_order.trans star.circ_loop_fixpoint sup.cobounded2 mult_assoc by metis also have "... \ (?H * d)\<^sup>+" using star.circ_mult_increasing by blast also have "... \ - ?H" using assms(5) bot_unique pseudo_complement forest_modulo_equivalence_def by blast thus ?thesis using calculation inf.order_trans by blast qed have "?x \ d" by (simp add: conv_isotone inf.sup_monoid.add_assoc) also have "... \ f \ f\<^sup>T" proof - have "h \ h\<^sup>T \ d \ d\<^sup>T = f \ f\<^sup>T" by (simp add: assms(6)) thus ?thesis by (metis (no_types) le_supE sup.absorb_iff2 sup.idem) qed also have "... \ -- g" using assms(1, 2) conv_complement conv_isotone by fastforce finally have "?x \ -- g" by simp thus ?thesis by (simp add: 67) qed next show "?x\<^sup>T * top \ ?H * ?e * top" proof - have "?x \ top * ?e\<^sup>T * ?H" using inf.coboundedI1 by auto hence "?x\<^sup>T \ ?H * ?e * top" using conv_dist_comp conv_dist_inf conv_star_commute inf.orderI inf.sup_monoid.add_assoc inf.sup_monoid.add_commute mult_assoc by auto hence "?x\<^sup>T * top \ ?H * ?e * top * top" by (simp add: mult_left_isotone) thus ?thesis by (simp add: mult_assoc) qed next show "j \ bot" by (simp add: assms(9)) next show "arc (?x)" using 61 by blast next show "regular h" using assms(14) by auto qed hence "sum (?e \ g) \ sum (a \ g)" using x_below_a order.trans by blast thus ?thesis by (simp add: assms(10)) qed qed thus ?thesis by simp qed subsubsection \Maintenance of algorithm invariants\ text \ In this section, most of the work is done to maintain the invariants of the inner and outer loops of the algorithm. In particular, we use \exists_a_w\ to maintain that $f$ can be extended to a minimum spanning forest. \ lemma boruvka_exchange_spanning_inv: assumes "forest v" and "v\<^sup>\ * e\<^sup>T = e\<^sup>T" and "i \ v \ top * e\<^sup>T * w\<^sup>T\<^sup>\" and "arc i" and "arc e" and "v \ --g" and "w \ --g" and "e \ --g" and "components g \ forest_components v" shows "i \ (v \ -i)\<^sup>T\<^sup>\ * e\<^sup>T * top" proof - have 1: "(v \ -i \ -i\<^sup>T) * (v\<^sup>T \ -i \ -i\<^sup>T) \ 1" using assms(1) comp_isotone order.trans inf.cobounded1 by blast have 2: "bijective (i * top) \ bijective (e\<^sup>T * top)" using assms(4, 5) mult_assoc by auto have "i \ v * (top * e\<^sup>T * w\<^sup>T\<^sup>\)\<^sup>T" using assms(3) covector_mult_closed covector_restrict_comp_conv order_lesseq_imp vector_top_closed by blast also have "... \ v * w\<^sup>T\<^sup>\\<^sup>T * e\<^sup>T\<^sup>T * top\<^sup>T" by (simp add: comp_associative conv_dist_comp) also have "... \ v * w\<^sup>\ * e * top" by (simp add: conv_star_commute) also have "... = v * w\<^sup>\ * e * e\<^sup>T * e * top" using assms(5) arc_eq_1 by (simp add: comp_associative) also have "... \ v * w\<^sup>\ * e * e\<^sup>T * top" by (simp add: comp_associative mult_right_isotone) also have "... \ (--g) * (--g)\<^sup>\ * (--g) * e\<^sup>T * top" using assms(6, 7, 8) by (simp add: comp_isotone star_isotone) also have "... \ (--g)\<^sup>\ * e\<^sup>T * top" by (metis comp_isotone mult_left_isotone star.circ_increasing star.circ_transitive_equal) also have "... \ v\<^sup>T\<^sup>\ * v\<^sup>\ * e\<^sup>T * top" by (simp add: assms(9) mult_left_isotone) also have "... \ v\<^sup>T\<^sup>\ * e\<^sup>T * top" by (simp add: assms(2) comp_associative) finally have "i \ v\<^sup>T\<^sup>\ * e\<^sup>T * top" by simp hence "i * top \ v\<^sup>T\<^sup>\ * e\<^sup>T * top" by (metis comp_associative mult_left_isotone vector_top_closed) hence "e\<^sup>T * top \ v\<^sup>T\<^sup>\\<^sup>T * i * top" using 2 by (metis bijective_reverse mult_assoc) also have "... = v\<^sup>\ * i * top" by (simp add: conv_star_commute) also have "... \ (v \ -i \ -i\<^sup>T)\<^sup>\ * i * top" proof - have 3: "i * top \ (v \ -i \ -i\<^sup>T)\<^sup>\ * i * top" using star.circ_loop_fixpoint sup_right_divisibility mult_assoc by auto have "(v \ i) * (v \ -i \ -i\<^sup>T)\<^sup>\ * i * top \ i * top * i * top" by (metis comp_isotone inf.cobounded1 inf.sup_monoid.add_commute mult_left_isotone top.extremum) also have "... \ i * top" by simp finally have 4: "(v \ i) * (v \ -i \ -i\<^sup>T)\<^sup>\ * i * top \ (v \ -i \ -i\<^sup>T)\<^sup>\ * i * top" using 3 dual_order.trans by blast have 5: "(v \ -i \ -i\<^sup>T) * (v \ -i \ -i\<^sup>T)\<^sup>\ * i * top \ (v \ -i \ -i\<^sup>T)\<^sup>\ * i * top" by (metis mult_left_isotone star.circ_increasing star.left_plus_circ) have "v\<^sup>+ \ -1" by (simp add: assms(1)) hence "v * v \ -1" by (metis mult_left_isotone order_trans star.circ_increasing star.circ_plus_same) hence "v * 1 \ -v\<^sup>T" by (simp add: schroeder_5_p) hence "v \ -v\<^sup>T" by simp hence "v \ v\<^sup>T \ bot" by (simp add: bot_unique pseudo_complement) hence 7: "v \ i\<^sup>T \ bot" by (metis assms(3) comp_inf.mult_right_isotone conv_dist_inf inf.boundedE inf.le_iff_sup le_bot) hence "(v \ i\<^sup>T) * (v \ -i \ -i\<^sup>T)\<^sup>\ * i * top \ bot" using le_bot semiring.mult_zero_left by fastforce hence 6: "(v \ i\<^sup>T) * (v \ -i \ -i\<^sup>T)\<^sup>\ * i * top \ (v \ -i \ -i\<^sup>T)\<^sup>\ * i * top" using bot_least le_bot by blast have 8: "v = (v \ i) \ (v \ i\<^sup>T) \ (v \ -i \ -i\<^sup>T)" proof - have 81: "regular i" by (simp add: assms(4) arc_regular) have "(v \ i\<^sup>T) \ (v \ -i \ -i\<^sup>T) = (v \ -i)" using 7 by (metis comp_inf.coreflexive_comp_inf_complement inf_import_p inf_p le_bot maddux_3_11_pp top.extremum) hence "(v \ i) \ (v \ i\<^sup>T) \ (v \ -i \ -i\<^sup>T) = (v \ i) \ (v \ -i)" by (simp add: sup.semigroup_axioms semigroup.assoc) also have "... = v" using 81 by (metis maddux_3_11_pp) finally show ?thesis by simp qed have "(v \ i) * (v \ -i \ -i\<^sup>T)\<^sup>\ * i * top \ (v \ i\<^sup>T) * (v \ -i \ -i\<^sup>T)\<^sup>\ * i * top \ (v \ -i \ -i\<^sup>T) * (v \ -i \ -i\<^sup>T)\<^sup>\ * i * top \ (v \ -i \ -i\<^sup>T)\<^sup>\ * i * top" using 4 5 6 by simp hence "((v \ i) \ (v \ i\<^sup>T) \ (v \ -i \ -i\<^sup>T)) * (v \ -i \ -i\<^sup>T)\<^sup>\ * i * top \ (v \ -i \ -i\<^sup>T)\<^sup>\ * i * top" by (simp add: mult_right_dist_sup) hence "v * (v \ -i \ -i\<^sup>T)\<^sup>\ * i * top \ (v \ -i \ -i\<^sup>T)\<^sup>\ * i * top" using 8 by auto hence "i * top \ v * (v \ -i \ -i\<^sup>T)\<^sup>\ * i * top \ (v \ -i \ -i\<^sup>T)\<^sup>\ * i * top" using 3 by auto hence 9:"v\<^sup>\ * (v \ -i \ -i\<^sup>T)\<^sup>\ * i * top \ (v \ -i \ -i\<^sup>T)\<^sup>\ * i * top" by (simp add: star_left_induct_mult mult_assoc) have "v\<^sup>\ * i * top \ v\<^sup>\ * (v \ -i \ -i\<^sup>T)\<^sup>\ * i * top" using 3 mult_right_isotone mult_assoc by auto thus ?thesis using 9 order.trans by blast qed finally have "e\<^sup>T * top \ (v \ -i \ -i\<^sup>T)\<^sup>\ * i * top" by simp hence "i * top \ (v \ -i \ -i\<^sup>T)\<^sup>\\<^sup>T * e\<^sup>T * top" using 2 by (metis bijective_reverse mult_assoc) also have "... = (v\<^sup>T \ -i \ -i\<^sup>T)\<^sup>\ * e\<^sup>T * top" using comp_inf.inf_vector_comp conv_complement conv_dist_inf conv_star_commute inf.sup_monoid.add_commute by auto also have "... \ ((v \ -i \ -i\<^sup>T) \ (v\<^sup>T \ -i \ -i\<^sup>T))\<^sup>\ * e\<^sup>T * top" by (simp add: mult_left_isotone star_isotone) finally have "i \ ((v\<^sup>T \ -i \ -i\<^sup>T) \ (v \ -i \ -i\<^sup>T))\<^sup>\ * e\<^sup>T * top" using dual_order.trans top_right_mult_increasing sup_commute by auto also have "... = (v\<^sup>T \ -i \ -i\<^sup>T)\<^sup>\ * (v \ -i \ -i\<^sup>T)\<^sup>\ * e\<^sup>T * top" using 1 cancel_separate_1 by (simp add: sup_commute) also have "... \ (v\<^sup>T \ -i \ -i\<^sup>T)\<^sup>\ * v\<^sup>\ * e\<^sup>T * top" by (simp add: inf_assoc mult_left_isotone mult_right_isotone star_isotone) also have "... = (v\<^sup>T \ -i \ -i\<^sup>T)\<^sup>\ * e\<^sup>T * top" using assms(2) mult_assoc by simp also have "... \ (v\<^sup>T \ -i\<^sup>T)\<^sup>\ * e\<^sup>T * top" by (metis mult_left_isotone star_isotone inf.cobounded2 inf.left_commute inf.sup_monoid.add_commute) also have "... = (v \ -i)\<^sup>T\<^sup>\ * e\<^sup>T * top" using conv_complement conv_dist_inf by auto finally show ?thesis by simp qed lemma exists_a_w: assumes "symmetric g" and "forest f" and "f \ --g" and "regular f" and "(\w . minimum_spanning_forest w g \ f \ w \ w\<^sup>T)" and "vector j" and "regular j" and "forest h" and "forest_modulo_equivalence (forest_components h) d" and "d * top \ - j" and "forest_components h * j = j" and "f \ f\<^sup>T = h \ h\<^sup>T \ d \ d\<^sup>T" and "(\ a b . forest_modulo_equivalence_path a b (forest_components h) d \ a \ -(forest_components h) \ -- g \ b \ d \ sum(b \ g) \ sum(a \ g))" and "regular d" and "selected_edge h j g \ - forest_components f" and "selected_edge h j g \ bot" and "j \ bot" and "regular h" and "h \ --g" shows "\w. minimum_spanning_forest w g \ f \ - (selected_edge h j g)\<^sup>T \ - (path f h j g) \ (f \ - (selected_edge h j g)\<^sup>T \ (path f h j g))\<^sup>T \ (selected_edge h j g) \ w \ w\<^sup>T" proof - let ?p = "path f h j g" let ?e = "selected_edge h j g" let ?f = "(f \ -?e\<^sup>T \ -?p) \ (f \ -?e\<^sup>T \ ?p)\<^sup>T \ ?e" let ?F = "forest_components f" let ?H = "forest_components h" let ?ec = "choose_component (forest_components h) j * - choose_component (forest_components h) j\<^sup>T \ g" from assms(4) obtain w where 2: "minimum_spanning_forest w g \ f \ w \ w\<^sup>T" using assms(5) by blast hence 3: "regular w \ regular f \ regular ?e" by (metis assms(4) minarc_regular minimum_spanning_forest_def spanning_forest_def) have 5: "equivalence ?F" using assms(2) forest_components_equivalence by auto have "?e\<^sup>T * top * ?e\<^sup>T = ?e\<^sup>T" by (metis arc_conv_closed arc_top_arc coreflexive_bot_closed coreflexive_symmetric minarc_arc minarc_bot_iff semiring.mult_not_zero) hence "?e\<^sup>T * top * ?e\<^sup>T \ -?F" using 5 assms(15) conv_complement conv_isotone by fastforce hence 6: "?e * ?F * ?e = bot" using assms(2) le_bot triple_schroeder_p by simp let ?q = "w \ top * ?e * w\<^sup>T\<^sup>\" let ?v = "(w \ -(top * ?e * w\<^sup>T\<^sup>\)) \ ?q\<^sup>T" have 7: "regular ?q" using 3 regular_closed_star regular_conv_closed regular_mult_closed by auto have 8: "injective ?v" proof (rule kruskal_exchange_injective_inv_1) show "injective w" using 2 minimum_spanning_forest_def spanning_forest_def by blast next show "covector (top * ?e * w\<^sup>T\<^sup>\)" by (simp add: covector_mult_closed) next show "top * ?e * w\<^sup>T\<^sup>\ * w\<^sup>T \ top * ?e * w\<^sup>T\<^sup>\" by (simp add: mult_right_isotone star.right_plus_below_circ mult_assoc) next show "coreflexive ((top * ?e * w\<^sup>T\<^sup>\)\<^sup>T * (top * ?e * w\<^sup>T\<^sup>\) \ w\<^sup>T * w)" using 2 by (metis comp_inf.semiring.mult_not_zero forest_bot kruskal_injective_inv_3 minarc_arc minarc_bot_iff minimum_spanning_forest_def semiring.mult_not_zero spanning_forest_def) qed have 9: "components g \ forest_components ?v" proof (rule kruskal_exchange_spanning_inv_1) show "injective (w \ - (top *?e * w\<^sup>T\<^sup>\) \ (w \ top * ?e * w\<^sup>T\<^sup>\)\<^sup>T)" using 8 by simp next show "regular (w \ top * ?e * w\<^sup>T\<^sup>\)" using 7 by simp next show "components g \ forest_components w" using 2 minimum_spanning_forest_def spanning_forest_def by blast qed have 10: "spanning_forest ?v g" proof (unfold spanning_forest_def, intro conjI) show "injective ?v" using 8 by auto next show "acyclic ?v" proof (rule kruskal_exchange_acyclic_inv_1) show "pd_kleene_allegory_class.acyclic w" using 2 minimum_spanning_forest_def spanning_forest_def by blast next show "covector (top * ?e * w\<^sup>T\<^sup>\)" by (simp add: covector_mult_closed) qed next show "?v \ --g" proof (rule sup_least) show "w \ - (top * ?e * w\<^sup>T\<^sup>\) \ - - g" using 7 inf.coboundedI1 minimum_spanning_forest_def spanning_forest_def 2 by blast next show "(w \ top * ?e * w\<^sup>T\<^sup>\)\<^sup>T \ - - g" using 2 by (metis assms(1) conv_complement conv_isotone inf.coboundedI1 minimum_spanning_forest_def spanning_forest_def) qed next show "components g \ forest_components ?v" using 9 by simp next show "regular ?v" using 3 regular_closed_star regular_conv_closed regular_mult_closed by auto qed have 11: "sum (?v \ g) = sum (w \ g)" proof - have "sum (?v \ g) = sum (w \ -(top * ?e * w\<^sup>T\<^sup>\) \ g) + sum (?q\<^sup>T \ g)" using 2 by (smt conv_complement conv_top epm_8 inf_import_p inf_top_right regular_closed_top vector_top_closed minimum_spanning_forest_def spanning_forest_def sum_disjoint) also have "... = sum (w \ -(top * ?e * w\<^sup>T\<^sup>\) \ g) + sum (?q \ g)" by (simp add: assms(1) sum_symmetric) also have "... = sum (((w \ -(top * ?e * w\<^sup>T\<^sup>\)) \ ?q) \ g)" using inf_commute inf_left_commute sum_disjoint by simp also have "... = sum (w \ g)" using 3 7 8 maddux_3_11_pp by auto finally show ?thesis by simp qed have 12: "?v \ ?v\<^sup>T = w \ w\<^sup>T" proof - have "?v \ ?v\<^sup>T = (w \ -?q) \ ?q\<^sup>T \ (w\<^sup>T \ -?q\<^sup>T) \ ?q" using conv_complement conv_dist_inf conv_dist_sup inf_import_p sup_assoc by simp also have "... = w \ w\<^sup>T" using 3 7 conv_complement conv_dist_inf inf_import_p maddux_3_11_pp sup_monoid.add_assoc sup_monoid.add_commute by auto finally show ?thesis by simp qed have 13: "?v * ?e\<^sup>T = bot" proof (rule kruskal_reroot_edge) show "injective (?e\<^sup>T * top)" using assms(16) minarc_arc minarc_bot_iff by blast next show "pd_kleene_allegory_class.acyclic w" using 2 minimum_spanning_forest_def spanning_forest_def by simp qed have "?v \ ?e \ ?v \ top * ?e" using inf.sup_right_isotone top_left_mult_increasing by simp also have "... \ ?v * (top * ?e)\<^sup>T" using covector_restrict_comp_conv covector_mult_closed vector_top_closed by simp finally have 14: "?v \ ?e = bot" using 13 by (metis conv_dist_comp mult_assoc le_bot mult_left_zero) let ?i = "?v \ (- ?F) * ?e * top \ top * ?e\<^sup>T * ?F" let ?w = "(?v \ -?i) \ ?e" have 15: "regular ?i" using 3 regular_closed_star regular_conv_closed regular_mult_closed by simp have 16: "?F \ -?i" proof - have 161: "bijective (?e * top)" using assms(16) minarc_arc minarc_bot_iff by auto have "?i \ - ?F * ?e * top" using inf.cobounded2 inf.coboundedI1 by blast also have "... = - (?F * ?e * top)" using 161 comp_bijective_complement by (simp add: mult_assoc) finally have "?i \ - (?F * ?e * top)" by blast hence 162: "?i \ ?F \ - (?F * ?e * top)" using inf.coboundedI1 by blast have "?i \ ?F \ ?F \ (top * ?e\<^sup>T * ?F)" by (meson inf_le1 inf_le2 le_infI order_trans) also have "... \ ?F * (top * ?e\<^sup>T * ?F)\<^sup>T" by (simp add: covector_mult_closed covector_restrict_comp_conv) also have "... = ?F * ?F\<^sup>T * ?e\<^sup>T\<^sup>T * top\<^sup>T" by (simp add: conv_dist_comp mult_assoc) also have "... = ?F * ?F * ?e * top" by (simp add: conv_dist_comp conv_star_commute) also have "... = ?F * ?e * top" by (simp add: 5 preorder_idempotent) finally have "?i \ ?F \ ?F * ?e * top" by simp hence "?i \ ?F \ ?F * ?e * top \ - (?F * ?e * top)" using 162 inf.bounded_iff by blast also have "... = bot" by simp finally show ?thesis using le_bot p_antitone_iff pseudo_complement by blast qed have 17: "?i \ top * ?e\<^sup>T * (?F \ ?v \ -?i)\<^sup>T\<^sup>\" proof - have "?i \ ?v \ - ?F * ?e * top \ top * ?e\<^sup>T * (?F \ ?v)\<^sup>T\<^sup>\ * (?F \ ?v)\<^sup>\" using 2 8 12 by (smt inf.sup_right_isotone kruskal_forest_components_inf mult_right_isotone mult_assoc) also have "... = ?v \ - ?F * ?e * top \ top * ?e\<^sup>T * (?F \ ?v)\<^sup>T\<^sup>\ * (1 \ (?F \ ?v)\<^sup>\ * (?F \ ?v))" using star_left_unfold_equal star.circ_right_unfold_1 by auto also have "... = ?v \ - ?F * ?e * top \ (top * ?e\<^sup>T * (?F \ ?v)\<^sup>T\<^sup>\ \ top * ?e\<^sup>T * (?F \ ?v)\<^sup>T\<^sup>\ * (?F \ ?v)\<^sup>\ * (?F \ ?v))" by (simp add: mult_left_dist_sup mult_assoc) also have "... = (?v \ - ?F * ?e * top \ top * ?e\<^sup>T * (?F \ ?v)\<^sup>T\<^sup>\) \ (?v \ - ?F * ?e * top \ top * ?e\<^sup>T * (?F \ ?v)\<^sup>T\<^sup>\ * (?F \ ?v)\<^sup>\ * (?F \ ?v))" using comp_inf.semiring.distrib_left by blast also have "... \ top * ?e\<^sup>T * (?F \ ?v)\<^sup>T\<^sup>\ \ (?v \ - ?F * ?e * top \ top * ?e\<^sup>T * (?F \ ?v)\<^sup>T\<^sup>\ * (?F \ ?v)\<^sup>\ * (?F \ ?v))" using comp_inf.semiring.add_right_mono inf_le2 by blast also have "... \ top * ?e\<^sup>T * (?F \ ?v)\<^sup>T\<^sup>\ \ (?v \ - ?F * ?e * top \ top * ?e\<^sup>T * (?F\<^sup>T \ ?v\<^sup>T)\<^sup>\ * (?F \ ?v)\<^sup>\ * (?F \ ?v))" by (simp add: conv_dist_inf) also have "... \ top * ?e\<^sup>T * (?F \ ?v)\<^sup>T\<^sup>\ \ (?v \ - ?F * ?e * top \ top * ?e\<^sup>T * ?F\<^sup>T\<^sup>\ * ?F\<^sup>\ * (?F \ ?v))" proof - have "top * ?e\<^sup>T * (?F\<^sup>T \ ?v\<^sup>T)\<^sup>\ * (?F \ ?v)\<^sup>\ * (?F \ ?v) \ top * ?e\<^sup>T * ?F\<^sup>T\<^sup>\ * ?F\<^sup>\ * (?F \ ?v)" using star_isotone by (simp add: comp_isotone) hence "?v \ - ?F * ?e * top \ top * ?e\<^sup>T * (?F\<^sup>T \ ?v\<^sup>T)\<^sup>\ * (?F \ ?v)\<^sup>\ * (?F \ ?v) \ ?v \ - ?F * ?e * top \ top * ?e\<^sup>T * ?F\<^sup>T\<^sup>\ * ?F\<^sup>\ * (?F \ ?v)" using inf.sup_right_isotone by blast thus ?thesis using sup_right_isotone by blast qed also have "... = top * ?e\<^sup>T * (?F \ ?v)\<^sup>T\<^sup>\ \ (?v \ - ?F * ?e * top \ top * ?e\<^sup>T * ?F\<^sup>\ * ?F\<^sup>\ * (?F \ ?v))" using 5 by auto also have "... = top * ?e\<^sup>T * (?F \ ?v)\<^sup>T\<^sup>\ \ (?v \ - ?F * ?e * top \ top * ?e\<^sup>T * ?F * ?F * (?F \ ?v))" by (simp add: assms(2) forest_components_star) also have "... = top * ?e\<^sup>T * (?F \ ?v)\<^sup>T\<^sup>\ \ (?v \ - ?F * ?e * top \ top * ?e\<^sup>T * ?F * (?F \ ?v))" using 5 mult.semigroup_axioms preorder_idempotent semigroup.assoc by fastforce also have "... = top * ?e\<^sup>T * (?F \ ?v)\<^sup>T\<^sup>\" proof - have "?e * top * ?e\<^sup>T \ 1" using assms(16) arc_expanded minarc_arc minarc_bot_iff by auto hence "?F * ?e * top * ?e\<^sup>T \ ?F * 1" by (metis comp_associative comp_isotone mult_semi_associative star.circ_transitive_equal) hence "?v * ?v\<^sup>T * ?F * ?e * top * ?e\<^sup>T \ 1 * ?F * 1" using 8 by (smt comp_isotone mult_assoc) hence 171: "?v * ?v\<^sup>T * ?F * ?e * top * ?e\<^sup>T \ ?F" by simp hence "?v * (?F \ ?v)\<^sup>T * ?F * ?e * top * ?e\<^sup>T \ ?F" proof - have "?v * (?F \ ?v)\<^sup>T * ?F * ?e * top * ?e\<^sup>T \ ?v * ?v\<^sup>T * ?F * ?e * top * ?e\<^sup>T" by (simp add: conv_dist_inf mult_left_isotone mult_right_isotone) thus ?thesis using 171 order_trans by blast qed hence 172: "-?F * ((?F \ ?v)\<^sup>T * ?F * ?e * top * ?e\<^sup>T)\<^sup>T \ -?v" by (smt schroeder_4_p comp_associative order_lesseq_imp pp_increasing) have "-?F * ((?F \ ?v)\<^sup>T * ?F * ?e * top * ?e\<^sup>T)\<^sup>T = -?F * ?e\<^sup>T\<^sup>T * top\<^sup>T * ?e\<^sup>T * ?F\<^sup>T * (?F \ ?v)\<^sup>T\<^sup>T" by (simp add: comp_associative conv_dist_comp) also have "... = -?F * ?e * top * ?e\<^sup>T * ?F * (?F \ ?v)" using 5 by auto also have "... = -?F * ?e * top * top * ?e\<^sup>T * ?F * (?F \ ?v)" using comp_associative by auto also have "... = -?F * ?e * top \ top * ?e\<^sup>T * ?F * (?F \ ?v)" by (smt comp_associative comp_inf.star.circ_decompose_9 comp_inf.star_star_absorb comp_inf_covector inf_vector_comp vector_top_closed) finally have "-?F * ((?F \ ?v)\<^sup>T * ?F * ?e * top * ?e\<^sup>T)\<^sup>T = -?F * ?e * top \ top * ?e\<^sup>T * ?F * (?F \ ?v)" by simp hence "-?F * ?e * top \ top * ?e\<^sup>T * ?F * (?F \ ?v) \ -?v" using 172 by auto hence "?v \ -?F * ?e * top \ top * ?e\<^sup>T * ?F * (?F \ ?v) \ bot" by (smt bot_unique inf.sup_monoid.add_commute p_shunting_swap pseudo_complement) thus ?thesis using le_bot sup_monoid.add_0_right by blast qed also have "... = top * ?e\<^sup>T * (?F \ ?v \ -?i)\<^sup>T\<^sup>\" using 16 by (smt comp_inf.coreflexive_comp_inf_complement inf_top_right p_bot pseudo_complement top.extremum) finally show ?thesis by blast qed have 18: "?i \ top * ?e\<^sup>T * ?w\<^sup>T\<^sup>\" proof - have "?i \ top * ?e\<^sup>T * (?F \ ?v \ -?i)\<^sup>T\<^sup>\" using 17 by simp also have "... \ top * ?e\<^sup>T * (?v \ -?i)\<^sup>T\<^sup>\" using mult_right_isotone conv_isotone star_isotone inf.cobounded2 inf.sup_monoid.add_assoc by (simp add: inf.sup_monoid.add_assoc order.eq_iff inf.sup_monoid.add_commute) also have "... \ top * ?e\<^sup>T * ((?v \ -?i) \ ?e)\<^sup>T\<^sup>\" using mult_right_isotone conv_isotone star_isotone sup_ge1 by simp finally show ?thesis by blast qed have 19: "?i \ top * ?e\<^sup>T * ?v\<^sup>T\<^sup>\" proof - have "?i \ top * ?e\<^sup>T * (?F \ ?v \ -?i)\<^sup>T\<^sup>\" using 17 by simp also have "... \ top * ?e\<^sup>T * (?v \ -?i)\<^sup>T\<^sup>\" using mult_right_isotone conv_isotone star_isotone inf.cobounded2 inf.sup_monoid.add_assoc by (simp add: inf.sup_monoid.add_assoc order.eq_iff inf.sup_monoid.add_commute) also have "... \ top * ?e\<^sup>T * (?v)\<^sup>T\<^sup>\" using mult_right_isotone conv_isotone star_isotone by auto finally show ?thesis by blast qed have 20: "f \ f\<^sup>T \ (?v \ -?i \ -?i\<^sup>T) \ (?v\<^sup>T \ -?i \ -?i\<^sup>T)" proof (rule kruskal_edge_between_components_2) show "?F \ - ?i" using 16 by simp next show "injective f" by (simp add: assms(2)) next show "f \ f\<^sup>T \ w \ - (top * ?e * w\<^sup>T\<^sup>\) \ (w \ top * ?e * w\<^sup>T\<^sup>\)\<^sup>T \ (w \ - (top * ?e * w\<^sup>T\<^sup>\) \ (w \ top * ?e * w\<^sup>T\<^sup>\)\<^sup>T)\<^sup>T" using 2 12 by (metis conv_dist_sup conv_involutive conv_isotone le_supI sup_commute) qed have "minimum_spanning_forest ?w g \ ?f \ ?w \ ?w\<^sup>T" proof (intro conjI) have 211: "?e\<^sup>T \ ?v\<^sup>\" proof (rule kruskal_edge_arc_1[where g=g and h="?ec"]) show "?e \ -- ?ec" using minarc_below by blast next show "?ec \ g" using assms(4) inf.cobounded2 by (simp add: boruvka_inner_invariant_def boruvka_outer_invariant_def conv_dist_inf) next show "symmetric g" by (meson assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def) next show "components g \ forest_components (w \ - (top * ?e * w\<^sup>T\<^sup>\) \ (w \ top * ?e * w\<^sup>T\<^sup>\)\<^sup>T)" using 9 by simp next show "(w \ - (top * ?e * w\<^sup>T\<^sup>\) \ (w \ top * ?e * w\<^sup>T\<^sup>\)\<^sup>T) * ?e\<^sup>T = bot" using 13 by blast qed have 212: "arc ?i" proof (rule boruvka_edge_arc) show "equivalence ?F" by (simp add: 5) next show "forest ?v" using 10 spanning_forest_def by blast next show "arc ?e" using assms(16) minarc_arc minarc_bot_iff by blast next show "regular ?F" using 3 regular_closed_star regular_conv_closed regular_mult_closed by auto next show "?F \ forest_components (?F \ ?v)" by (simp add: 12 2 8 kruskal_forest_components_inf) next show "regular ?v" using 10 spanning_forest_def by blast next show "?v * ?e\<^sup>T = bot" using 13 by auto next show "?e * ?F * ?e = bot" by (simp add: 6) next show "?e\<^sup>T \ ?v\<^sup>\" using 211 by auto next show "?e \ bot" by (simp add: assms(16)) qed show "minimum_spanning_forest ?w g" proof (unfold minimum_spanning_forest_def, intro conjI) have "(?v \ -?i) * ?e\<^sup>T \ ?v * ?e\<^sup>T" using inf_le1 mult_left_isotone by simp hence "(?v \ -?i) * ?e\<^sup>T = bot" using 13 le_bot by simp hence 221: "?e * (?v \ -?i)\<^sup>T = bot" using conv_dist_comp conv_involutive conv_bot by force have 222: "injective ?w" proof (rule injective_sup) show "injective (?v \ -?i)" using 8 by (simp add: injective_inf_closed) next show "coreflexive (?e * (?v \ -?i)\<^sup>T)" using 221 by simp next show "injective ?e" by (metis arc_injective minarc_arc coreflexive_bot_closed coreflexive_injective minarc_bot_iff) qed show "spanning_forest ?w g" proof (unfold spanning_forest_def, intro conjI) show "injective ?w" using 222 by simp next show "acyclic ?w" proof (rule kruskal_exchange_acyclic_inv_2) show "acyclic ?v" using 10 spanning_forest_def by blast next show "injective ?v" using 8 by simp next show "?i \?v" using inf.coboundedI1 by simp next show "bijective (?i\<^sup>T * top)" using 212 by simp next show "bijective (?e * top)" using 14 212 by (smt assms(4) comp_inf.idempotent_bot_closed conv_complement minarc_arc minarc_bot_iff p_bot regular_closed_bot semiring.mult_not_zero symmetric_top_closed) next show "?i \ top * ?e\<^sup>T *?v\<^sup>T\<^sup>\" using 19 by simp next show "?v * ?e\<^sup>T * top = bot" using 13 by simp qed next have "?w \ ?v \ ?e" using inf_le1 sup_left_isotone by simp also have "... \ --g \ ?e" using 10 sup_left_isotone spanning_forest_def by blast also have "... \ --g \ --h" proof - have 1: "--g \ --g \ --h" by simp have 2: "?e \ --g \ --h" by (metis inf.coboundedI1 inf.sup_monoid.add_commute minarc_below order.trans p_dist_inf p_dist_sup sup.cobounded1) thus ?thesis using 1 2 by simp qed also have "... \ --g" using assms(18, 19) by auto finally show "?w \ --g" by simp next have 223: "?i \ (?v \ -?i)\<^sup>T\<^sup>\ * ?e\<^sup>T * top" proof (rule boruvka_exchange_spanning_inv) show "forest ?v" using 10 spanning_forest_def by blast next show "?v\<^sup>\ * ?e\<^sup>T = ?e\<^sup>T" using 13 by (smt conv_complement conv_dist_comp conv_involutive conv_star_commute dense_pp fc_top regular_closed_top star_absorb) next show "?i \ ?v \ top * ?e\<^sup>T * ?w\<^sup>T\<^sup>\" using 18 inf.sup_monoid.add_assoc by auto next show "arc ?i" using 212 by blast next show "arc ?e" using assms(16) minarc_arc minarc_bot_iff by auto next show "?v \ --g" using 10 spanning_forest_def by blast next show "?w \ --g" proof - have 2231: "?e \ --g" by (metis inf.boundedE minarc_below pp_dist_inf) have "?w \ ?v \ ?e" using inf_le1 sup_left_isotone by simp also have "... \ --g" using 2231 10 spanning_forest_def sup_least by blast finally show ?thesis by blast qed next show "?e \ --g" by (metis inf.boundedE minarc_below pp_dist_inf) next show "components g \ forest_components ?v" by (simp add: 9) qed have "components g \ forest_components ?v" using 10 spanning_forest_def by auto also have "... \ forest_components ?w" proof (rule kruskal_exchange_forest_components_inv) next show "injective ((?v \ -?i) \ ?e)" using 222 by simp next show "regular ?i" using 15 by simp next show "?e * top * ?e = ?e" by (metis arc_top_arc minarc_arc minarc_bot_iff semiring.mult_not_zero) next show "?i \ top * ?e\<^sup>T * ?v\<^sup>T\<^sup>\" using 19 by blast next show "?v * ?e\<^sup>T * top = bot" using 13 by simp next show "injective ?v" using 8 by simp next show "?i \ ?v" by (simp add: le_infI1) next show "?i \ (?v \ -?i)\<^sup>T\<^sup>\ * ?e\<^sup>T * top" using 223 by blast qed finally show "components g \ forest_components ?w" by simp next show "regular ?w" using 3 7 regular_conv_closed by simp qed next have 224: "?e \ g \ bot" using assms(16) inf.left_commute inf_bot_right minarc_meet_bot by fastforce have 225: "sum (?e \ g) \ sum (?i \ g)" proof (rule a_to_e_in_forest_modulo_equivalence) show "symmetric g" using assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def by auto next show "j \ bot" by (simp add: assms(17)) next show "f \ -- g" by (simp add: assms(3)) next show "vector j" using assms(6) boruvka_inner_invariant_def by blast next show "forest h" by (simp add: assms(8)) next show " forest_modulo_equivalence (forest_components h) d" by (simp add: assms(9)) next show "f \ f\<^sup>T = h \ h\<^sup>T \ d \ d\<^sup>T" by (simp add: assms(12)) next show "\a b. forest_modulo_equivalence_path a b (?H) d \ a \ - ?H \ - - g \ b \ d \ sum (b \ g) \ sum (a \ g)" by (simp add: assms(13)) next show "regular d" using assms(14) by auto next show "?e = ?e" by simp next show "arc ?i" using 212 by blast next show "forest_modulo_equivalence_path ?i ?e ?H (d \ ?e)" proof - have "d\<^sup>T * ?H * ?e = bot" using assms(6, 7, 10 ,11, 17) dT_He_eq_bot le_bot by blast hence 251: "d\<^sup>T * ?H * ?e \ (?H * d)\<^sup>\ * ?H * ?e" by simp hence "d\<^sup>T * ?H * ?H * ?e \ (?H * d)\<^sup>\ * ?H * ?e" by (metis assms(8) forest_components_star star.circ_decompose_9 mult_assoc) hence "d\<^sup>T * (?H * d)\<^sup>\ * ?H * ?e \ (?H * d)\<^sup>\ * ?H * ?e" proof - have "d\<^sup>T * ?H * d \ 1" using assms(9) forest_modulo_equivalence_def dTransHd_le_1 by blast hence "d\<^sup>T * ?H * d * (?H * d)\<^sup>\ * ?H * ?e \ (?H * d)\<^sup>\ * ?H * ?e" by (metis mult_left_isotone star.circ_circ_mult star_involutive star_one) hence "d\<^sup>T * ?H * ?e \ d\<^sup>T * ?H * d * (?H * d)\<^sup>\ * ?H * ?e \ (?H * d)\<^sup>\ * ?H * ?e" using 251 by simp hence "d\<^sup>T * (1 \ ?H * d * (?H * d)\<^sup>\) * ?H * ?e \ (?H * d)\<^sup>\ * ?H * ?e" by (simp add: comp_associative comp_left_dist_sup semiring.distrib_right) thus ?thesis by (simp add: star_left_unfold_equal) qed hence "?H * d\<^sup>T * (?H * d)\<^sup>\ * ?H * ?e \ ?H * (?H * d)\<^sup>\ * ?H * ?e" by (simp add: mult_right_isotone mult_assoc) hence "?H * d\<^sup>T * (?H * d)\<^sup>\ * ?H * ?e \ ?H * ?H * (d * ?H)\<^sup>\ * ?e" by (smt star_slide mult_assoc) hence "?H * d\<^sup>T * (?H * d)\<^sup>\ * ?H * ?e \ ?H * (d * ?H)\<^sup>\ * ?e" by (metis assms(8) forest_components_star star.circ_decompose_9) hence "?H * d\<^sup>T * (?H * d)\<^sup>\ * ?H * ?e \ (?H * d)\<^sup>\ * ?H * ?e" using star_slide by auto hence "?H * d * (?H * d)\<^sup>\ * ?H * ?e \ ?H * d\<^sup>T * (?H * d)\<^sup>\ * ?H * ?e \ (?H * d)\<^sup>\ * ?H * ?e" by (smt le_supI star.circ_loop_fixpoint sup.cobounded2 sup_commute mult_assoc) hence "(?H * (d \ d\<^sup>T)) * (?H * d)\<^sup>\ * ?H * ?e \ (?H * d)\<^sup>\ * ?H * ?e" by (simp add: semiring.distrib_left semiring.distrib_right) hence "(?H * (d \ d\<^sup>T))\<^sup>\ * (?H * d)\<^sup>\ * ?H * ?e \ (?H * d)\<^sup>\ * ?H * ?e" by (simp add: star_left_induct_mult mult_assoc) hence 252: "(?H * (d \ d\<^sup>T))\<^sup>\ * ?H * ?e \ (?H * d)\<^sup>\ * ?H * ?e" by (smt mult_left_dist_sup star.circ_transitive_equal star_slide star_sup_1 mult_assoc) have "?i \ top * ?e\<^sup>T * ?F" by auto hence "?i\<^sup>T \ ?F\<^sup>T * ?e\<^sup>T\<^sup>T * top\<^sup>T" by (simp add: conv_dist_comp conv_dist_inf mult_assoc) hence "?i\<^sup>T * top \ ?F\<^sup>T * ?e\<^sup>T\<^sup>T * top\<^sup>T * top" using comp_isotone by blast also have "... = ?F\<^sup>T * ?e\<^sup>T\<^sup>T * top\<^sup>T" by (simp add: vector_mult_closed) also have "... = ?F * ?e\<^sup>T\<^sup>T * top\<^sup>T" by (simp add: conv_dist_comp conv_star_commute) also have "... = ?F * ?e * top" by simp also have "... = (?H * (d \ d\<^sup>T))\<^sup>\ * ?H * ?e * top" using assms(2, 8, 12) F_is_H_and_d by simp also have "... \ (?H * d)\<^sup>\ * ?H * ?e * top" by (simp add: 252 comp_isotone) also have "... \ (?H * (d \ ?e))\<^sup>\ * ?H * ?e * top" by (simp add: comp_isotone star_isotone) finally have "?i\<^sup>T * top \ (?H * (d \ ?e))\<^sup>\ * ?H * ?e * top" by blast thus ?thesis using 212 assms(16) forest_modulo_equivalence_path_def minarc_arc minarc_bot_iff by blast qed next show "?i \ - ?H \ -- g" proof - have "forest_components h \ forest_components f" using assms(2, 8, 12) H_below_F by blast then have 241: "?i \ -?H" using 16 assms(9) inf.order_lesseq_imp p_antitone_iff by blast have "?i \ -- g" using 10 inf.coboundedI1 spanning_forest_def by blast thus ?thesis using 241 inf_greatest by blast qed next show "regular h" using assms(18) by auto qed have "?v \ ?e \ -?i = bot" using 14 by simp hence "sum (?w \ g) = sum (?v \ -?i \ g) + sum (?e \ g)" using sum_disjoint inf_commute inf_assoc by simp also have "... \ sum (?v \ -?i \ g) + sum (?i \ g)" using 224 225 sum_plus_right_isotone by simp also have "... = sum (((?v \ -?i) \ ?i) \ g)" using sum_disjoint inf_le2 pseudo_complement by simp also have "... = sum ((?v \ ?i) \ (-?i \ ?i) \ g)" by (simp add: sup_inf_distrib2) also have "... = sum ((?v \ ?i) \ g)" using 15 by (metis inf_top_right stone) also have "... = sum (?v \ g)" by (simp add: inf.sup_monoid.add_assoc) finally have "sum (?w \ g) \ sum (?v \ g)" by simp thus "\u . spanning_forest u g \ sum (?w \ g) \ sum (u \ g)" using 2 11 minimum_spanning_forest_def by auto qed next have "?f \ f \ f\<^sup>T \ ?e" by (smt conv_dist_inf inf_le1 sup_left_isotone sup_mono inf.order_lesseq_imp) also have "... \ (?v \ -?i \ -?i\<^sup>T) \ (?v\<^sup>T \ -?i \ -?i\<^sup>T) \ ?e" using 20 sup_left_isotone by simp also have "... \ (?v \ -?i) \ (?v\<^sup>T \ -?i \ -?i\<^sup>T) \ ?e" by (metis inf.cobounded1 sup_inf_distrib2) also have "... = ?w \ (?v\<^sup>T \ -?i \ -?i\<^sup>T)" by (simp add: sup_assoc sup_commute) also have "... \ ?w \ (?v\<^sup>T \ -?i\<^sup>T)" using inf.sup_right_isotone inf_assoc sup_right_isotone by simp also have "... \ ?w \ ?w\<^sup>T" using conv_complement conv_dist_inf conv_dist_sup sup_right_isotone by simp finally show "?f \ ?w \ ?w\<^sup>T" by simp qed thus ?thesis by auto qed lemma boruvka_outer_invariant_when_e_not_bot: assumes "boruvka_inner_invariant j f h g d" and "j \ bot" and "selected_edge h j g \ - forest_components f" and "selected_edge h j g \ bot" shows "boruvka_outer_invariant (f \ - selected_edge h j g\<^sup>T \ - path f h j g \ (f \ - selected_edge h j g\<^sup>T \ path f h j g)\<^sup>T \ selected_edge h j g) g" proof - let ?c = "choose_component (forest_components h) j" let ?p = "path f h j g" let ?F = "forest_components f" let ?H = "forest_components h" let ?e = "selected_edge h j g" let ?f' = "f \ -?e\<^sup>T \ -?p \ (f \ -?e\<^sup>T \ ?p)\<^sup>T \ ?e" let ?d' = "d \ ?e" let ?j' = "j \ -?c" show "boruvka_outer_invariant ?f' g" proof (unfold boruvka_outer_invariant_def, intro conjI) show "symmetric g" by (meson assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def) next show "injective ?f'" proof (rule kruskal_injective_inv) show "injective (f \ - ?e\<^sup>T)" by (meson assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def injective_inf_closed) show "covector (?p)" using covector_mult_closed by simp show "?p * (f \ - ?e\<^sup>T)\<^sup>T \ ?p" by (simp add: mult_right_isotone star.left_plus_below_circ star_plus mult_assoc) show "?e \ ?p" by (meson mult_left_isotone order.trans star_outer_increasing top.extremum) show "?p * (f \ - ?e\<^sup>T)\<^sup>T \ - ?e" proof - have "?p * (f \ - ?e\<^sup>T)\<^sup>T \ ?p * f\<^sup>T" by (simp add: conv_dist_inf mult_right_isotone) also have "... \ top * ?e * (f)\<^sup>T\<^sup>\ * f\<^sup>T" using conv_dist_inf star_isotone comp_isotone by simp also have "... \ - ?e" using assms(1, 4) boruvka_inner_invariant_def boruvka_outer_invariant_def kruskal_injective_inv_2 minarc_arc minarc_bot_iff by auto finally show ?thesis . qed show "injective (?e)" by (metis arc_injective coreflexive_bot_closed minarc_arc minarc_bot_iff semiring.mult_not_zero) show "coreflexive (?p\<^sup>T * ?p \ (f \ - ?e\<^sup>T)\<^sup>T * (f \ - ?e\<^sup>T))" proof - have "(?p\<^sup>T * ?p \ (f \ - ?e\<^sup>T)\<^sup>T * (f \ - ?e\<^sup>T)) \ ?p\<^sup>T * ?p \ f\<^sup>T * f" using conv_dist_inf inf.sup_right_isotone mult_isotone by simp also have "... \ (top * ?e * f\<^sup>T\<^sup>\)\<^sup>T * (top * ?e * f\<^sup>T\<^sup>\) \ f\<^sup>T * f" by (metis comp_associative comp_inf.coreflexive_transitive comp_inf.mult_right_isotone comp_isotone conv_isotone inf.cobounded1 inf.idem inf.sup_monoid.add_commute star_isotone top.extremum) also have "... \ 1" using assms(1, 4) boruvka_inner_invariant_def boruvka_outer_invariant_def kruskal_injective_inv_3 minarc_arc minarc_bot_iff by auto finally show ?thesis by simp qed qed next show "acyclic ?f'" proof (rule kruskal_acyclic_inv) show "acyclic (f \ - ?e\<^sup>T)" proof - have f_intersect_below: "(f \ - ?e\<^sup>T) \ f" by simp have "acyclic f" by (meson assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def) thus ?thesis using comp_isotone dual_order.trans star_isotone f_intersect_below by blast qed next show "covector ?p" by (metis comp_associative vector_top_closed) next show "(f \ - ?e\<^sup>T \ ?p)\<^sup>T * (f \ - ?e\<^sup>T)\<^sup>\ * ?e = bot" proof - have "?e \ - (f\<^sup>T\<^sup>\ * f\<^sup>\)" by (simp add: assms(3)) hence "?e * top * ?e \ - (f\<^sup>T\<^sup>\ * f\<^sup>\)" by (metis arc_top_arc minarc_arc minarc_bot_iff semiring.mult_not_zero) hence "?e\<^sup>T * top * ?e\<^sup>T \ - (f\<^sup>T\<^sup>\ * f\<^sup>\)\<^sup>T" by (metis comp_associative conv_complement conv_dist_comp conv_isotone symmetric_top_closed) hence "?e\<^sup>T * top * ?e\<^sup>T \ - (f\<^sup>T\<^sup>\ * f\<^sup>\)" by (simp add: conv_dist_comp conv_star_commute) hence "?e * (f\<^sup>T\<^sup>\ * f\<^sup>\) * ?e \ bot" using triple_schroeder_p by auto hence 1: "?e * f\<^sup>T\<^sup>\ * f\<^sup>\ * ?e \ bot" using mult_assoc by auto have 2: "(f \ - ?e\<^sup>T)\<^sup>T\<^sup>\ \ f\<^sup>T\<^sup>\" by (simp add: conv_dist_inf star_isotone) have "(f \ - ?e\<^sup>T \ ?p)\<^sup>T * (f \ - ?e\<^sup>T)\<^sup>\ * ?e \ (f \ ?p)\<^sup>T * (f \ - ?e\<^sup>T)\<^sup>\ * ?e" by (simp add: comp_isotone conv_dist_inf inf.orderI inf.sup_monoid.add_assoc) also have "... \ (f \ ?p)\<^sup>T * f\<^sup>\ * ?e" by (simp add: comp_isotone star_isotone) also have "... \ (f \ top * ?e * (f)\<^sup>T\<^sup>\)\<^sup>T * f\<^sup>\ * ?e" using 2 by (metis comp_inf.comp_isotone comp_inf.coreflexive_transitive comp_isotone conv_isotone inf.idem top.extremum) also have "... = (f\<^sup>T \ (top * ?e * f\<^sup>T\<^sup>\)\<^sup>T) * f\<^sup>\ * ?e" by (simp add: conv_dist_inf) also have "... \ top * (f\<^sup>T \ (top * ?e * f\<^sup>T\<^sup>\)\<^sup>T) * f\<^sup>\ * ?e" using top_left_mult_increasing mult_assoc by auto also have "... = (top \ top * ?e * f\<^sup>T\<^sup>\) * f\<^sup>T * f\<^sup>\ * ?e" by (smt covector_comp_inf_1 covector_mult_closed order.eq_iff inf.sup_monoid.add_commute vector_top_closed) also have "... = top * ?e * f\<^sup>T\<^sup>\ * f\<^sup>T * f\<^sup>\ * ?e" by simp also have "... \ top * ?e * f\<^sup>T\<^sup>\ * f\<^sup>\ * ?e" by (smt conv_dist_comp conv_isotone conv_star_commute mult_left_isotone mult_right_isotone star.left_plus_below_circ mult_assoc) also have "... \ bot" using 1 covector_bot_closed le_bot mult_assoc by fastforce finally show ?thesis using le_bot by auto qed next show "?e * (f \ - ?e\<^sup>T)\<^sup>\ * ?e = bot" proof - have 1: "?e \ - ?F" by (simp add: assms(3)) have 2: "injective f" by (meson assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def) have 3: "equivalence ?F" using 2 forest_components_equivalence by simp hence 4: "?e\<^sup>T = ?e\<^sup>T * top * ?e\<^sup>T" by (smt arc_conv_closed arc_top_arc covector_complement_closed covector_conv_vector ex231e minarc_arc minarc_bot_iff pp_surjective regular_closed_top vector_mult_closed vector_top_closed) also have "... \ - ?F" using 1 3 conv_isotone conv_complement calculation by fastforce finally have 5: "?e * ?F * ?e = bot" using 4 by (smt triple_schroeder_p le_bot pp_total regular_closed_top vector_top_closed) have "(f \ - ?e\<^sup>T)\<^sup>\ \ f\<^sup>\" by (simp add: star_isotone) hence "?e * (f \ - ?e\<^sup>T)\<^sup>\ * ?e \ ?e * f\<^sup>\ * ?e" using mult_left_isotone mult_right_isotone by blast also have "... \ ?e * ?F * ?e" by (metis conv_star_commute forest_components_increasing mult_left_isotone mult_right_isotone star_involutive) also have 6: "... = bot" using 5 by simp finally show ?thesis using 6 le_bot by blast qed next show "forest_components (f \ -?e\<^sup>T) \ - ?e" proof - have 1: "?e \ - ?F" by (simp add: assms(3)) have "f \ - ?e\<^sup>T \ f" by simp hence "forest_components (f \ - ?e\<^sup>T) \ ?F" using forest_components_isotone by blast thus ?thesis using 1 order_lesseq_imp p_antitone_iff by blast qed qed next show "?f' \ --g" proof - have 1: "(f \ - ?e\<^sup>T \ - ?p) \ --g" by (meson assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def inf.coboundedI1) have 2: "(f \ - ?e\<^sup>T \ ?p)\<^sup>T \ --g" proof - have "(f \ - ?e\<^sup>T \ ?p)\<^sup>T \ f\<^sup>T" by (simp add: conv_isotone inf.sup_monoid.add_assoc) also have "... \ --g" by (metis assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def conv_complement conv_isotone) finally show ?thesis by simp qed have 3: "?e \ --g" by (metis inf.boundedE minarc_below pp_dist_inf) show ?thesis using 1 2 3 by simp qed next show "regular ?f'" using assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def minarc_regular regular_closed_star regular_conv_closed regular_mult_closed by auto next show "\w. minimum_spanning_forest w g \ ?f' \ w \ w\<^sup>T" proof (rule exists_a_w) show "symmetric g" using assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def by auto next show "forest f" using assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def by auto next show "f \ --g" using assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def by auto next show "regular f" using assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def by auto next show "(\w . minimum_spanning_forest w g \ f \ w \ w\<^sup>T)" using assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def by auto next show "vector j" using assms(1) boruvka_inner_invariant_def by blast next show "regular j" using assms(1) boruvka_inner_invariant_def by blast next show "forest h" using assms(1) boruvka_inner_invariant_def by blast next show "forest_modulo_equivalence (forest_components h) d" using assms(1) boruvka_inner_invariant_def by blast next show "d * top \ - j" using assms(1) boruvka_inner_invariant_def by blast next show "forest_components h * j = j" using assms(1) boruvka_inner_invariant_def by blast next show "f \ f\<^sup>T = h \ h\<^sup>T \ d \ d\<^sup>T" using assms(1) boruvka_inner_invariant_def by blast next show "(\ a b . forest_modulo_equivalence_path a b (forest_components h) d \ a \ -(forest_components h) \ -- g \ b \ d \ sum(b \ g) \ sum(a \ g))" using assms(1) boruvka_inner_invariant_def by blast next show "regular d" using assms(1) boruvka_inner_invariant_def by blast next show "selected_edge h j g \ - forest_components f" by (simp add: assms(3)) next show "selected_edge h j g \ bot" by (simp add: assms(4)) next show "j \ bot" by (simp add: assms(2)) next show "regular h" using assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def by auto next show "h \ --g" using H_below_regular_g assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def by auto qed qed qed lemma second_inner_invariant_when_e_not_bot: assumes "boruvka_inner_invariant j f h g d" and "j \ bot" and "selected_edge h j g \ - forest_components f" and "selected_edge h j g \ bot" shows "boruvka_inner_invariant (j \ - choose_component (forest_components h) j) (f \ - selected_edge h j g\<^sup>T \ - path f h j g \ (f \ - selected_edge h j g\<^sup>T \ path f h j g)\<^sup>T \ selected_edge h j g) h g (d \ selected_edge h j g)" proof - let ?c = "choose_component (forest_components h) j" let ?p = "path f h j g" let ?F = "forest_components f" let ?H = "forest_components h" let ?e = "selected_edge h j g" let ?f' = "f \ -?e\<^sup>T \ -?p \ (f \ -?e\<^sup>T \ ?p)\<^sup>T \ ?e" let ?d' = "d \ ?e" let ?j' = "j \ -?c" show "boruvka_inner_invariant ?j' ?f' h g ?d'" proof (unfold boruvka_inner_invariant_def, intro conjI) show 1: "boruvka_outer_invariant ?f' g" using assms(1, 2, 3, 4) boruvka_outer_invariant_when_e_not_bot by blast next show "g \ bot" using assms(1) boruvka_inner_invariant_def by force next show "regular ?d'" using assms(1) boruvka_inner_invariant_def minarc_regular by auto next show "regular ?j'" using assms(1) boruvka_inner_invariant_def by auto next show "vector ?j'" using assms(1, 2) boruvka_inner_invariant_def component_is_vector vector_complement_closed vector_inf_closed by simp next show "regular h" by (meson assms(1) boruvka_inner_invariant_def) next show "injective h" by (meson assms(1) boruvka_inner_invariant_def) next show "pd_kleene_allegory_class.acyclic h" by (meson assms(1) boruvka_inner_invariant_def) next show "?H * ?j' = ?j'" using fc_j_eq_j_inv assms(1) boruvka_inner_invariant_def by blast next show "forest_modulo_equivalence ?H ?d'" using assms(1, 2, 3) forest_modulo_equivalence_d_U_e boruvka_inner_invariant_def boruvka_outer_invariant_def by auto next show "?d' * top \ -?j'" proof - have 31: "?d' * top = d * top \ ?e * top" by (simp add: mult_right_dist_sup) have 32: "d * top \ -?j'" by (meson assms(1) boruvka_inner_invariant_def inf.coboundedI1 p_antitone_iff) have "regular (?c * - ?c\<^sup>T)" using assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def component_is_regular regular_conv_closed regular_mult_closed by presburger then have "minarc(?c * - ?c\<^sup>T \ g) = minarc(?c \ - ?c\<^sup>T \ g)" by (metis component_is_vector covector_comp_inf inf_top.left_neutral vector_conv_compl) also have "... \ -- (?c \ - ?c\<^sup>T \ g)" using minarc_below by blast also have "... \ -- ?c" by (simp add: inf.sup_monoid.add_assoc) also have "... = ?c" using component_is_regular by auto finally have "?e \ ?c" by simp then have "?e * top \ ?c" by (metis component_is_vector mult_left_isotone) also have "... \ -j \ ?c" by simp also have "... = - (j \ - ?c)" using component_is_regular by auto finally have 33: "?e * top \ - (j \ - ?c)" by simp show ?thesis using 31 32 33 by auto qed next show "?f' \ ?f'\<^sup>T = h \ h\<^sup>T \ ?d' \ ?d'\<^sup>T" proof - have "?f' \ ?f'\<^sup>T = f \ - ?e\<^sup>T \ - ?p \ (f \ - ?e\<^sup>T \ ?p)\<^sup>T \ ?e \ (f \ - ?e\<^sup>T \ - ?p)\<^sup>T \ (f \ - ?e\<^sup>T \ ?p) \ ?e\<^sup>T" by (simp add: conv_dist_sup sup_monoid.add_assoc) also have "... = (f \ - ?e\<^sup>T \ - ?p) \ (f \ - ?e\<^sup>T \ ?p) \ (f \ - ?e\<^sup>T \ ?p)\<^sup>T \ (f \ - ?e\<^sup>T \ - ?p)\<^sup>T \ ?e\<^sup>T \ ?e" by (simp add: sup.left_commute sup_commute) also have "... = f \ f\<^sup>T \ ?e \ ?e\<^sup>T" proof (rule simplify_f) show "regular ?p" using assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def minarc_regular regular_closed_star regular_conv_closed regular_mult_closed by auto next show "regular ?e" using minarc_regular by blast qed also have "... = h \ h\<^sup>T \ d \ d\<^sup>T \ ?e \ ?e\<^sup>T" using assms(1) boruvka_inner_invariant_def by auto finally show ?thesis by (smt conv_dist_sup sup.left_commute sup_commute) qed next show "\ a b . forest_modulo_equivalence_path a b ?H ?d' \ a \ - ?H \ -- g \ b \ ?d' \ sum (b \ g) \ sum (a \ g)" proof (intro allI, rule impI, unfold forest_modulo_equivalence_path_def) fix a b assume 1: "(arc a \ arc b \ a\<^sup>T * top \ (?H * ?d')\<^sup>\ * ?H * b * top) \ a \ - ?H \ -- g \ b \ ?d'" thus "sum (b \ g) \ sum (a \ g)" proof (cases "b = ?e") case b_equals_e: True thus ?thesis proof (cases "a = ?e") case True thus ?thesis using b_equals_e by auto next case a_ne_e: False have "sum (b \ g) \ sum (a \ g)" proof (rule a_to_e_in_forest_modulo_equivalence) show "symmetric g" using assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def by auto next show "j \ bot" by (simp add: assms(2)) next show "f \ -- g" using assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def by auto next show "vector j" using assms(1) boruvka_inner_invariant_def by blast next show "forest h" using assms(1) boruvka_inner_invariant_def by blast next show "forest_modulo_equivalence (forest_components h) d" using assms(1) boruvka_inner_invariant_def by blast next show "f \ f\<^sup>T = h \ h\<^sup>T \ d \ d\<^sup>T" using assms(1) boruvka_inner_invariant_def by blast next show "\a b. forest_modulo_equivalence_path a b (?H) d \ a \ - ?H \ - - g \ b \ d \ sum (b \ g) \ sum (a \ g)" using assms(1) boruvka_inner_invariant_def by blast next show "regular d" using assms(1) boruvka_inner_invariant_def by blast next show "b = ?e" using b_equals_e by simp next show "arc a" using 1 by simp next show "forest_modulo_equivalence_path a b ?H ?d'" using 1 forest_modulo_equivalence_path_def by simp next show "a \ - ?H \ -- g" using 1 by simp next show "regular h" using assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def by auto qed thus ?thesis by simp qed next case b_not_equal_e: False then have b_below_d: "b \ d" using 1 assms(4) different_arc_in_sup_arc minarc_arc minarc_bot_iff by metis thus ?thesis proof (cases "?e \ d") case True then have "forest_modulo_equivalence_path a b ?H d \ b \ d" using 1 forest_modulo_equivalence_path_def sup.absorb1 by auto thus ?thesis using 1 assms(1) boruvka_inner_invariant_def by blast next case e_not_less_than_d: False have 71:"equivalence ?H" using assms(1) fch_equivalence boruvka_inner_invariant_def by auto then have 72: "forest_modulo_equivalence_path a b ?H ?d' \ forest_modulo_equivalence_path a b ?H d \ (forest_modulo_equivalence_path a ?e ?H d \ forest_modulo_equivalence_path ?e b ?H d)" proof (rule forest_modulo_equivalence_path_split_disj) show "arc ?e" using assms(4) minarc_arc minarc_bot_iff by blast next show "regular a \ regular b \ regular ?e \ regular d \ regular ?H" using assms(1) 1 boruvka_inner_invariant_def boruvka_outer_invariant_def arc_regular minarc_regular regular_closed_star regular_conv_closed regular_mult_closed by auto qed thus ?thesis proof (cases "forest_modulo_equivalence_path a b ?H d") case True have "forest_modulo_equivalence_path a b ?H d \ b \ d" using 1 True forest_modulo_equivalence_path_def sup.absorb1 by (metis assms(4) b_not_equal_e minarc_arc minarc_bot_iff different_arc_in_sup_arc) thus ?thesis using 1 assms(1) b_below_d boruvka_inner_invariant_def by auto next case False have 73:"forest_modulo_equivalence_path a ?e ?H d \ forest_modulo_equivalence_path ?e b ?H d" using 1 72 False forest_modulo_equivalence_path_def by blast have 74: "?e \ --g" by (metis inf.boundedE minarc_below pp_dist_inf) have "?H \ ?F" using assms(1) H_below_F boruvka_inner_invariant_def boruvka_outer_invariant_def by blast then have "?e \ - ?H" using assms(3) order_trans p_antitone by blast then have "?e \ - ?H \ --g" using 74 by simp then have 75: "sum (b \ g) \ sum (?e \ g)" using assms(1) b_below_d 73 boruvka_inner_invariant_def by blast have 76: "forest_modulo_equivalence_path a ?e ?H ?d'" by (meson 73 forest_modulo_equivalence_path_split_disj assms(1) forest_modulo_equivalence_path_def boruvka_inner_invariant_def boruvka_outer_invariant_def fch_equivalence arc_regular regular_closed_star regular_conv_closed regular_mult_closed) have 77: "sum (?e \ g) \ sum (a \ g)" proof (rule a_to_e_in_forest_modulo_equivalence) show "symmetric g" using assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def by auto next show "j \ bot" by (simp add: assms(2)) next show "f \ -- g" using assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def by auto next show "vector j" using assms(1) boruvka_inner_invariant_def by blast next show "forest h" using assms(1) boruvka_inner_invariant_def by blast next show "forest_modulo_equivalence (forest_components h) d" using assms(1) boruvka_inner_invariant_def by blast next show "f \ f\<^sup>T = h \ h\<^sup>T \ d \ d\<^sup>T" using assms(1) boruvka_inner_invariant_def by blast next show "\a b. forest_modulo_equivalence_path a b (?H) d \ a \ - ?H \ - - g \ b \ d \ sum (b \ g) \ sum (a \ g)" using assms(1) boruvka_inner_invariant_def by blast next show "regular d" using assms(1) boruvka_inner_invariant_def by blast next show "?e = ?e" by simp next show "arc a" using 1 by simp next show "forest_modulo_equivalence_path a ?e ?H ?d'" by (simp add: "76") next show "a \ - ?H \ --g" using 1 by simp next show "regular h" using assms(1) boruvka_inner_invariant_def boruvka_outer_invariant_def by auto qed thus ?thesis using 75 order.trans by blast qed qed qed qed qed qed lemma second_inner_invariant_when_e_bot: assumes "selected_edge h j g = bot" and "selected_edge h j g \ - forest_components f" and "boruvka_inner_invariant j f h g d" shows "boruvka_inner_invariant (j \ - choose_component (forest_components h) j) (f \ - selected_edge h j g\<^sup>T \ - path f h j g \ (f \ - selected_edge h j g\<^sup>T \ path f h j g)\<^sup>T \ selected_edge h j g) h g (d \ selected_edge h j g)" proof - let ?c = "choose_component (forest_components h) j" let ?p = "path f h j g" let ?F = "forest_components f" let ?H = "forest_components h" let ?e = "selected_edge h j g" let ?f' = "f \ -?e\<^sup>T \ -?p \ (f \ -?e\<^sup>T \ ?p)\<^sup>T \ ?e" let ?d' = "d \ ?e" let ?j' = "j \ -?c" show "boruvka_inner_invariant ?j' ?f' h g ?d'" proof (unfold boruvka_inner_invariant_def, intro conjI) next show "boruvka_outer_invariant ?f' g" using assms(1) assms(3) boruvka_inner_invariant_def by auto next show "g \ bot" using assms(3) boruvka_inner_invariant_def by blast next show "regular ?d'" using assms(1) assms(3) boruvka_inner_invariant_def by auto next show "regular ?j'" using assms(3) boruvka_inner_invariant_def by auto next show "vector ?j'" by (metis assms(3) boruvka_inner_invariant_def component_is_vector vector_complement_closed vector_inf_closed) next show "regular h" using assms(3) boruvka_inner_invariant_def by blast next show "injective h" using assms(3) boruvka_inner_invariant_def by blast next show "pd_kleene_allegory_class.acyclic h" using assms(3) boruvka_inner_invariant_def by blast next show "?H * ?j' = ?j'" using assms(3) fc_j_eq_j_inv boruvka_inner_invariant_def by blast next show "forest_modulo_equivalence ?H ?d'" using assms(1) assms(3) boruvka_inner_invariant_def by auto next show "?d' * top \ -?j'" using assms(1) assms(3) boruvka_inner_invariant_def by (metis order.trans p_antitone_inf sup_monoid.add_0_right) next show "?f' \ ?f'\<^sup>T = h \ h\<^sup>T \ ?d' \ ?d'\<^sup>T" using assms(1, 3) boruvka_inner_invariant_def by auto next show "\a b. forest_modulo_equivalence_path a b ?H ?d' \ a \ -?H \ --g \ b \ ?d' \ sum(b \ g) \ sum(a \ g)" using assms(1, 3) boruvka_inner_invariant_def by auto qed qed subsection \Formalization and correctness proof\ text \ The following result shows that Bor\r{u}vka's algorithm constructs a minimum spanning forest. We have the same postcondition as the proof of Kruskal's minimum spanning tree algorithm. We show only partial correctness. \ theorem boruvka_mst: "VARS f j h c e d { symmetric g } f := bot; WHILE -(forest_components f) \ g \ bot INV { boruvka_outer_invariant f g } DO j := top; h := f; d := bot; WHILE j \ bot INV { boruvka_inner_invariant j f h g d } DO c := choose_component (forest_components h) j; e := minarc(c * -c\<^sup>T \ g); IF e \ -(forest_components f) THEN f := f \ -e\<^sup>T; f := (f \ -(top * e * f\<^sup>T\<^sup>\)) \ (f \ top * e * f\<^sup>T\<^sup>\)\<^sup>T \ e; d := d \ e ELSE SKIP FI; j := j \ -c OD OD { minimum_spanning_forest f g }" proof vcg_simp assume 1: "symmetric g" show "boruvka_outer_invariant bot g" using 1 boruvka_outer_invariant_def kruskal_exists_minimal_spanning by auto next fix f let ?F = "forest_components f" assume 1: "boruvka_outer_invariant f g \ - ?F \ g \ bot" have 2: "equivalence ?F" using 1 boruvka_outer_invariant_def forest_components_equivalence by auto show "boruvka_inner_invariant top f f g bot" proof (unfold boruvka_inner_invariant_def, intro conjI) show "boruvka_outer_invariant f g" by (simp add: 1) next show "g \ bot" using 1 by auto next show "surjective top" by simp next show "regular top" by simp next show "regular bot" by auto next show "regular f" using 1 boruvka_outer_invariant_def by blast next show "injective f" using 1 boruvka_outer_invariant_def by blast next show "pd_kleene_allegory_class.acyclic f" using 1 boruvka_outer_invariant_def by blast next show "forest_modulo_equivalence ?F bot" by (simp add: 2 forest_modulo_equivalence_def) next show "bot * top \ - top" by simp next show "times_top_class.total (?F)" by (simp add: star.circ_right_top mult_assoc) next show "f \ f\<^sup>T = f \ f\<^sup>T \ bot \ bot\<^sup>T" by simp next show "\ a b. forest_modulo_equivalence_path a b ?F bot \ a \ - ?F \ -- g \ b \ bot \ sum (b \ g) \ sum (a \ g)" by (metis (full_types) forest_modulo_equivalence_path_def bot_unique mult_left_zero mult_right_zero top.extremum) qed next fix f j h d let ?c = "choose_component (forest_components h) j" let ?p = "path f h j g" let ?F = "forest_components f" let ?H = "forest_components h" let ?e = "selected_edge h j g" let ?f' = "f \ -?e\<^sup>T \ -?p \ (f \ -?e\<^sup>T \ ?p)\<^sup>T \ ?e" let ?d' = "d \ ?e" let ?j' = "j \ -?c" assume 1: "boruvka_inner_invariant j f h g d \ j \ bot" show "(?e \ -?F \ boruvka_inner_invariant ?j' ?f' h g ?d') \ (\ ?e \ -?F \ boruvka_inner_invariant ?j' f h g d)" proof (intro conjI) show "?e \ -?F \ boruvka_inner_invariant ?j' ?f' h g ?d'" proof (cases "?e = bot") case True thus ?thesis using 1 second_inner_invariant_when_e_bot by simp next case False thus ?thesis using 1 second_inner_invariant_when_e_not_bot by simp qed next show "\ ?e \ -?F \ boruvka_inner_invariant ?j' f h g d" proof (rule impI, unfold boruvka_inner_invariant_def, intro conjI) show "boruvka_outer_invariant f g" using 1 boruvka_inner_invariant_def by blast next show "g \ bot" using 1 boruvka_inner_invariant_def by blast next show "vector ?j'" using 1 boruvka_inner_invariant_def component_is_vector vector_complement_closed vector_inf_closed by auto next show "regular ?j'" using 1 boruvka_inner_invariant_def by auto next show "regular d" using 1 boruvka_inner_invariant_def by blast next show "regular h" using 1 boruvka_inner_invariant_def by blast next show "injective h" using 1 boruvka_inner_invariant_def by blast next show "pd_kleene_allegory_class.acyclic h" using 1 boruvka_inner_invariant_def by blast next show "forest_modulo_equivalence ?H d" using 1 boruvka_inner_invariant_def by blast next show "d * top \ -?j'" using 1 by (meson boruvka_inner_invariant_def dual_order.trans p_antitone_inf) next show "?H * ?j' = ?j'" using 1 fc_j_eq_j_inv boruvka_inner_invariant_def by blast next show "f \ f\<^sup>T = h \ h\<^sup>T \ d \ d\<^sup>T" using 1 boruvka_inner_invariant_def by blast next show "\ ?e \ -?F \ \a b. forest_modulo_equivalence_path a b ?H d \ a \ -?H \ --g \ b \ d \ sum(b \ g) \ sum(a \ g)" using 1 boruvka_inner_invariant_def by blast qed qed next fix f h d assume "boruvka_inner_invariant bot f h g d" thus "boruvka_outer_invariant f g" by (meson boruvka_inner_invariant_def) next fix f assume 1: "boruvka_outer_invariant f g \ - forest_components f \ g = bot" hence 2:"spanning_forest f g" proof (unfold spanning_forest_def, intro conjI) show "injective f" using 1 boruvka_outer_invariant_def by blast next show "acyclic f" using 1 boruvka_outer_invariant_def by blast next show "f \ --g" using 1 boruvka_outer_invariant_def by blast next show "components g \ forest_components f" proof - let ?F = "forest_components f" have "-?F \ g \ bot" by (simp add: 1) hence "--g \ bot \ --?F" using 1 shunting_p p_antitone pseudo_complement by auto hence "--g \ ?F" using 1 boruvka_outer_invariant_def pp_dist_comp pp_dist_star regular_conv_closed by auto hence "(--g)\<^sup>\ \ ?F\<^sup>\" by (simp add: star_isotone) thus ?thesis using 1 boruvka_outer_invariant_def forest_components_star by auto qed next show "regular f" using 1 boruvka_outer_invariant_def by auto qed from 1 obtain w where 3: "minimum_spanning_forest w g \ f \ w \ w\<^sup>T" using boruvka_outer_invariant_def by blast hence "w = w \ --g" by (simp add: inf.absorb1 minimum_spanning_forest_def spanning_forest_def) also have "... \ w \ components g" by (metis inf.sup_right_isotone star.circ_increasing) also have "... \ w \ f\<^sup>T\<^sup>\ * f\<^sup>\" using 2 spanning_forest_def inf.sup_right_isotone by simp also have "... \ f \ f\<^sup>T" proof (rule cancel_separate_6[where z=w and y="w\<^sup>T"]) show "injective w" using 3 minimum_spanning_forest_def spanning_forest_def by simp next show "f\<^sup>T \ w\<^sup>T \ w" using 3 by (metis conv_dist_inf conv_dist_sup conv_involutive inf.cobounded2 inf.orderE) next show "f \ w\<^sup>T \ w" using 3 by (simp add: sup_commute) next show "injective w" using 3 minimum_spanning_forest_def spanning_forest_def by simp next show "w \ w\<^sup>T\<^sup>\ = bot" using 3 by (metis acyclic_star_below_complement comp_inf.mult_right_isotone inf_p le_bot minimum_spanning_forest_def spanning_forest_def) qed finally have 4: "w \ f \ f\<^sup>T" by simp have "sum (f \ g) = sum ((w \ w\<^sup>T) \ (f \ g))" using 3 by (metis inf_absorb2 inf.assoc) also have "... = sum (w \ (f \ g)) + sum (w\<^sup>T \ (f \ g))" using 3 inf.commute acyclic_asymmetric sum_disjoint minimum_spanning_forest_def spanning_forest_def by simp also have "... = sum (w \ (f \ g)) + sum (w \ (f\<^sup>T \ g\<^sup>T))" by (metis conv_dist_inf conv_involutive sum_conv) also have "... = sum (f \ (w \ g)) + sum (f\<^sup>T \ (w \ g))" proof - have 51:"f\<^sup>T \ (w \ g) = f\<^sup>T \ (w \ g\<^sup>T)" using 1 boruvka_outer_invariant_def by auto have 52:"f \ (w \ g) = w \ (f \ g)" by (simp add: inf.left_commute) thus ?thesis using 51 52 abel_semigroup.left_commute inf.abel_semigroup_axioms by fastforce qed also have "... = sum ((f \ f\<^sup>T) \ (w \ g))" using 2 acyclic_asymmetric inf.sup_monoid.add_commute sum_disjoint spanning_forest_def by simp also have "... = sum (w \ g)" using 4 by (metis inf_absorb2 inf.assoc) finally show "minimum_spanning_forest f g" using 2 3 minimum_spanning_forest_def by simp qed end end diff --git a/thys/Relational_Minimum_Spanning_Trees/ROOT b/thys/Relational_Minimum_Spanning_Trees/ROOT --- a/thys/Relational_Minimum_Spanning_Trees/ROOT +++ b/thys/Relational_Minimum_Spanning_Trees/ROOT @@ -1,19 +1,18 @@ chapter AFP -session Relational_Minimum_Spanning_Trees = Aggregation_Algebras + +session Relational_Minimum_Spanning_Trees = Relational_Disjoint_Set_Forests + options [timeout = 600] sessions "HOL-Hoare" - Relational_Disjoint_Set_Forests theories Kruskal Prim Boruvka document_files "root.tex" "root.bib" diff --git a/thys/Stone_Kleene_Relation_Algebras/Iterings.thy b/thys/Stone_Kleene_Relation_Algebras/Iterings.thy --- a/thys/Stone_Kleene_Relation_Algebras/Iterings.thy +++ b/thys/Stone_Kleene_Relation_Algebras/Iterings.thy @@ -1,858 +1,864 @@ (* Title: Iterings Author: Walter Guttmann Maintainer: Walter Guttmann *) section \Iterings\ text \ This theory introduces algebraic structures with an operation that describes iteration in various relational computation models. An iteration describes the repeated sequential execution of a computation. This is typically modelled by fixpoints, but different computation models use different fixpoints in the refinement order. We therefore look at equational and simulation axioms rather than induction axioms. Our development is based on \<^cite>\"Guttmann2012c"\ and the proposed algebras generalise Kleene algebras. We first consider a variant of Conway semirings \<^cite>\"BloomEsik1993a"\ based on idempotent left semirings. Conway semirings expand semirings by an iteration operation satisfying Conway's sumstar and productstar axioms \<^cite>\"Conway1971"\. Many properties of iteration follow already from these equational axioms. Next we introduce iterings, which use generalised versions of simulation axioms in addition to sumstar and productstar. Unlike the induction axioms of the Kleene star, which hold only in partial-correctness models, the simulation axioms are also valid in total and general correctness models. They are still powerful enough to prove the correctness of complex results such as separation theorems of \<^cite>\"Cohen2000"\ and Back's atomicity refinement theorem \<^cite>\"BackWright1999" and "Wright2004"\. \ theory Iterings imports Stone_Relation_Algebras.Semirings begin subsection \Conway Semirings\ text \ In this section, we consider equational axioms for iteration. The algebraic structures are based on idempotent left semirings, which are expanded by a unary iteration operation. We start with an unfold property, one inequality of the sliding rule and distributivity over joins, which is similar to Conway's sumstar. \ class circ = fixes circ :: "'a \ 'a" ("_\<^sup>\" [100] 100) class left_conway_semiring = idempotent_left_semiring + circ + assumes circ_left_unfold: "1 \ x * x\<^sup>\ = x\<^sup>\" assumes circ_left_slide: "(x * y)\<^sup>\ * x \ x * (y * x)\<^sup>\" assumes circ_sup_1: "(x \ y)\<^sup>\ = x\<^sup>\ * (y * x\<^sup>\)\<^sup>\" begin text \ We obtain one inequality of Conway's productstar, as well as of the other unfold rule. \ lemma circ_mult_sub: "1 \ x * (y * x)\<^sup>\ * y \ (x * y)\<^sup>\" by (metis sup_right_isotone circ_left_slide circ_left_unfold mult_assoc mult_right_isotone) lemma circ_right_unfold_sub: "1 \ x\<^sup>\ * x \ x\<^sup>\" by (metis circ_mult_sub mult_1_left mult_1_right) lemma circ_zero: "bot\<^sup>\ = 1" by (metis sup_monoid.add_0_right circ_left_unfold mult_left_zero) lemma circ_increasing: "x \ x\<^sup>\" by (metis le_supI2 circ_left_unfold circ_right_unfold_sub mult_1_left mult_right_sub_dist_sup_left order_trans) lemma circ_reflexive: "1 \ x\<^sup>\" by (metis sup_left_divisibility circ_left_unfold) lemma circ_mult_increasing: "x \ x * x\<^sup>\" by (metis circ_reflexive mult_right_isotone mult_1_right) lemma circ_mult_increasing_2: "x \ x\<^sup>\ * x" by (metis circ_reflexive mult_left_isotone mult_1_left) lemma circ_transitive_equal: "x\<^sup>\ * x\<^sup>\ = x\<^sup>\" by (metis sup_idem circ_sup_1 circ_left_unfold mult_assoc) text \ While iteration is not idempotent, a fixpoint is reached after applying this operation twice. Iteration is idempotent for the unit. \ lemma circ_circ_circ: "x\<^sup>\\<^sup>\\<^sup>\ = x\<^sup>\\<^sup>\" by (metis sup_idem circ_sup_1 circ_increasing circ_transitive_equal le_iff_sup) lemma circ_one: "1\<^sup>\ = 1\<^sup>\\<^sup>\" by (metis circ_circ_circ circ_zero) lemma circ_sup_sub: "(x\<^sup>\ * y)\<^sup>\ * x\<^sup>\ \ (x \ y)\<^sup>\" by (metis circ_sup_1 circ_left_slide) lemma circ_plus_one: "x\<^sup>\ = 1 \ x\<^sup>\" by (metis le_iff_sup circ_reflexive) text \ Iteration satisfies a characteristic property of reflexive transitive closures. \ lemma circ_rtc_2: "1 \ x \ x\<^sup>\ * x\<^sup>\ = x\<^sup>\" by (metis sup_assoc circ_increasing circ_plus_one circ_transitive_equal le_iff_sup) lemma mult_zero_circ: "(x * bot)\<^sup>\ = 1 \ x * bot" by (metis circ_left_unfold mult_assoc mult_left_zero) lemma mult_zero_sup_circ: "(x \ y * bot)\<^sup>\ = x\<^sup>\ * (y * bot)\<^sup>\" by (metis circ_sup_1 mult_assoc mult_left_zero) lemma circ_plus_sub: "x\<^sup>\ * x \ x * x\<^sup>\" by (metis circ_left_slide mult_1_left mult_1_right) lemma circ_loop_fixpoint: "y * (y\<^sup>\ * z) \ z = y\<^sup>\ * z" by (metis sup_commute circ_left_unfold mult_assoc mult_1_left mult_right_dist_sup) lemma left_plus_below_circ: "x * x\<^sup>\ \ x\<^sup>\" by (metis sup.cobounded2 circ_left_unfold) lemma right_plus_below_circ: "x\<^sup>\ * x \ x\<^sup>\" using circ_right_unfold_sub by auto lemma circ_sup_upper_bound: "x \ z\<^sup>\ \ y \ z\<^sup>\ \ x \ y \ z\<^sup>\" by simp lemma circ_mult_upper_bound: "x \ z\<^sup>\ \ y \ z\<^sup>\ \ x * y \ z\<^sup>\" by (metis mult_isotone circ_transitive_equal) lemma circ_sub_dist: "x\<^sup>\ \ (x \ y)\<^sup>\" by (metis circ_sup_sub circ_plus_one mult_1_left mult_right_sub_dist_sup_left order_trans) lemma circ_sub_dist_1: "x \ (x \ y)\<^sup>\" using circ_increasing le_supE by blast lemma circ_sub_dist_2: "x * y \ (x \ y)\<^sup>\" by (metis sup_commute circ_mult_upper_bound circ_sub_dist_1) lemma circ_sub_dist_3: "x\<^sup>\ * y\<^sup>\ \ (x \ y)\<^sup>\" by (metis sup_commute circ_mult_upper_bound circ_sub_dist) lemma circ_isotone: "x \ y \ x\<^sup>\ \ y\<^sup>\" by (metis circ_sub_dist le_iff_sup) lemma circ_sup_2: "(x \ y)\<^sup>\ \ (x\<^sup>\ * y\<^sup>\)\<^sup>\" by (metis sup.bounded_iff circ_increasing circ_isotone circ_reflexive mult_isotone mult_1_left mult_1_right) lemma circ_sup_one_left_unfold: "1 \ x \ x * x\<^sup>\ = x\<^sup>\" by (metis order.antisym le_iff_sup mult_1_left mult_right_sub_dist_sup_left left_plus_below_circ) lemma circ_sup_one_right_unfold: "1 \ x \ x\<^sup>\ * x = x\<^sup>\" by (metis order.antisym le_iff_sup mult_left_sub_dist_sup_left mult_1_right right_plus_below_circ) lemma circ_decompose_4: "(x\<^sup>\ * y\<^sup>\)\<^sup>\ = x\<^sup>\ * (y\<^sup>\ * x\<^sup>\)\<^sup>\" by (metis sup_assoc sup_commute circ_sup_1 circ_loop_fixpoint circ_plus_one circ_rtc_2 circ_transitive_equal mult_assoc) lemma circ_decompose_5: "(x\<^sup>\ * y\<^sup>\)\<^sup>\ = (y\<^sup>\ * x\<^sup>\)\<^sup>\" by (metis circ_decompose_4 circ_loop_fixpoint order.antisym mult_right_sub_dist_sup_right mult_assoc) lemma circ_decompose_6: "x\<^sup>\ * (y * x\<^sup>\)\<^sup>\ = y\<^sup>\ * (x * y\<^sup>\)\<^sup>\" by (metis sup_commute circ_sup_1) lemma circ_decompose_7: "(x \ y)\<^sup>\ = x\<^sup>\ * y\<^sup>\ * (x \ y)\<^sup>\" by (metis circ_sup_1 circ_decompose_6 circ_transitive_equal mult_assoc) lemma circ_decompose_8: "(x \ y)\<^sup>\ = (x \ y)\<^sup>\ * x\<^sup>\ * y\<^sup>\" by (metis order.antisym eq_refl mult_assoc mult_isotone mult_1_right circ_mult_upper_bound circ_reflexive circ_sub_dist_3) lemma circ_decompose_9: "(x\<^sup>\ * y\<^sup>\)\<^sup>\ = x\<^sup>\ * y\<^sup>\ * (x\<^sup>\ * y\<^sup>\)\<^sup>\" by (metis circ_decompose_4 mult_assoc) lemma circ_decompose_10: "(x\<^sup>\ * y\<^sup>\)\<^sup>\ = (x\<^sup>\ * y\<^sup>\)\<^sup>\ * x\<^sup>\ * y\<^sup>\" by (metis sup_ge2 circ_loop_fixpoint circ_reflexive circ_sup_one_right_unfold mult_assoc order_trans) lemma circ_back_loop_prefixpoint: "(z * y\<^sup>\) * y \ z \ z * y\<^sup>\" by (metis sup.bounded_iff circ_left_unfold mult_assoc mult_left_sub_dist_sup_left mult_right_isotone mult_1_right right_plus_below_circ) text \ We obtain the fixpoint and prefixpoint properties of iteration, but not least or greatest fixpoint properties. \ lemma circ_loop_is_fixpoint: "is_fixpoint (\x . y * x \ z) (y\<^sup>\ * z)" by (metis circ_loop_fixpoint is_fixpoint_def) lemma circ_back_loop_is_prefixpoint: "is_prefixpoint (\x . x * y \ z) (z * y\<^sup>\)" by (metis circ_back_loop_prefixpoint is_prefixpoint_def) lemma circ_circ_sup: "(1 \ x)\<^sup>\ = x\<^sup>\\<^sup>\" by (metis sup_commute circ_sup_1 circ_decompose_4 circ_zero mult_1_right) lemma circ_circ_mult_sub: "x\<^sup>\ * 1\<^sup>\ \ x\<^sup>\\<^sup>\" by (metis circ_increasing circ_isotone circ_mult_upper_bound circ_reflexive) lemma left_plus_circ: "(x * x\<^sup>\)\<^sup>\ = x\<^sup>\" by (metis circ_left_unfold circ_sup_1 mult_1_right mult_sub_right_one sup.absorb1 mult_assoc) lemma right_plus_circ: "(x\<^sup>\ * x)\<^sup>\ = x\<^sup>\" by (metis sup_commute circ_isotone circ_loop_fixpoint circ_plus_sub circ_sub_dist order.eq_iff left_plus_circ) lemma circ_square: "(x * x)\<^sup>\ \ x\<^sup>\" by (metis circ_increasing circ_isotone left_plus_circ mult_right_isotone) lemma circ_mult_sub_sup: "(x * y)\<^sup>\ \ (x \ y)\<^sup>\" by (metis sup_ge1 sup_ge2 circ_isotone circ_square mult_isotone order_trans) lemma circ_sup_mult_zero: "x\<^sup>\ * y = (x \ y * bot)\<^sup>\ * y" proof - have "(x \ y * bot)\<^sup>\ * y = x\<^sup>\ * (1 \ y * bot) * y" by (metis mult_zero_sup_circ mult_zero_circ) also have "... = x\<^sup>\ * (y \ y * bot)" by (metis mult_assoc mult_1_left mult_left_zero mult_right_dist_sup) also have "... = x\<^sup>\ * y" by (metis sup_commute le_iff_sup zero_right_mult_decreasing) finally show ?thesis by simp qed lemma troeger_1: "(x \ y)\<^sup>\ = x\<^sup>\ * (1 \ y * (x \ y)\<^sup>\)" by (metis circ_sup_1 circ_left_unfold mult_assoc) lemma troeger_2: "(x \ y)\<^sup>\ * z = x\<^sup>\ * (y * (x \ y)\<^sup>\ * z \ z)" by (metis circ_sup_1 circ_loop_fixpoint mult_assoc) lemma troeger_3: "(x \ y * bot)\<^sup>\ = x\<^sup>\ * (1 \ y * bot)" by (metis mult_zero_sup_circ mult_zero_circ) lemma circ_sup_sub_sup_one_1: "x \ y \ x\<^sup>\ * (1 \ y)" by (metis circ_increasing circ_left_unfold mult_1_left mult_1_right mult_left_sub_dist_sup mult_right_sub_dist_sup_left order_trans sup_mono) lemma circ_sup_sub_sup_one_2: "x\<^sup>\ * (x \ y) \ x\<^sup>\ * (1 \ y)" by (metis circ_sup_sub_sup_one_1 circ_transitive_equal mult_assoc mult_right_isotone) lemma circ_sup_sub_sup_one: "x * x\<^sup>\ * (x \ y) \ x * x\<^sup>\ * (1 \ y)" by (metis circ_sup_sub_sup_one_2 mult_assoc mult_right_isotone) lemma circ_square_2: "(x * x)\<^sup>\ * (x \ 1) \ x\<^sup>\" by (metis sup.bounded_iff circ_increasing circ_mult_upper_bound circ_reflexive circ_square) lemma circ_extra_circ: "(y * x\<^sup>\)\<^sup>\ = (y * y\<^sup>\ * x\<^sup>\)\<^sup>\" by (metis circ_decompose_6 circ_transitive_equal left_plus_circ mult_assoc) lemma circ_circ_sub_mult: "1\<^sup>\ * x\<^sup>\ \ x\<^sup>\\<^sup>\" by (metis circ_increasing circ_isotone circ_mult_upper_bound circ_reflexive) lemma circ_decompose_11: "(x\<^sup>\ * y\<^sup>\)\<^sup>\ = (x\<^sup>\ * y\<^sup>\)\<^sup>\ * x\<^sup>\" by (metis circ_decompose_10 circ_decompose_4 circ_decompose_5 circ_decompose_9 left_plus_circ) lemma circ_mult_below_circ_circ: "(x * y)\<^sup>\ \ (x\<^sup>\ * y)\<^sup>\ * x\<^sup>\" by (metis circ_increasing circ_isotone circ_reflexive dual_order.trans mult_left_isotone mult_right_isotone mult_1_right) +lemma power_below_circ: + "power x i \ x\<^sup>\" + apply (induct rule: nat.induct) + apply (simp add: circ_reflexive) + by (simp add: circ_increasing circ_mult_upper_bound) + (* lemma circ_right_unfold: "1 \ x\<^sup>\ * x = x\<^sup>\" nitpick [expect=genuine] oops lemma circ_mult: "1 \ x * (y * x)\<^sup>\ * y = (x * y)\<^sup>\" nitpick [expect=genuine] oops lemma circ_slide: "(x * y)\<^sup>\ * x = x * (y * x)\<^sup>\" nitpick [expect=genuine] oops lemma circ_plus_same: "x\<^sup>\ * x = x * x\<^sup>\" nitpick [expect=genuine] oops lemma "1\<^sup>\ * x\<^sup>\ \ x\<^sup>\ * 1\<^sup>\" nitpick [expect=genuine,card=7] oops lemma circ_circ_mult_1: "x\<^sup>\ * 1\<^sup>\ = x\<^sup>\\<^sup>\" nitpick [expect=genuine,card=7] oops lemma "x\<^sup>\ * 1\<^sup>\ \ 1\<^sup>\ * x\<^sup>\" nitpick [expect=genuine,card=7] oops lemma circ_circ_mult: "1\<^sup>\ * x\<^sup>\ = x\<^sup>\\<^sup>\" nitpick [expect=genuine,card=7] oops lemma circ_sup: "(x\<^sup>\ * y)\<^sup>\ * x\<^sup>\ = (x \ y)\<^sup>\" nitpick [expect=genuine,card=8] oops lemma circ_unfold_sum: "(x \ y)\<^sup>\ = x\<^sup>\ \ x\<^sup>\ * y * (x \ y)\<^sup>\" nitpick [expect=genuine,card=7] oops lemma mult_zero_sup_circ_2: "(x \ y * bot)\<^sup>\ = x\<^sup>\ \ x\<^sup>\ * y * bot" nitpick [expect=genuine,card=7] oops lemma sub_mult_one_circ: "x * 1\<^sup>\ \ 1\<^sup>\ * x" nitpick [expect=genuine] oops lemma circ_back_loop_fixpoint: "(z * y\<^sup>\) * y \ z = z * y\<^sup>\" nitpick [expect=genuine] oops lemma circ_back_loop_is_fixpoint: "is_fixpoint (\x . x * y \ z) (z * y\<^sup>\)" nitpick [expect=genuine] oops lemma "x\<^sup>\ * y\<^sup>\ \ (x\<^sup>\ * y)\<^sup>\ * x\<^sup>\" nitpick [expect=genuine,card=7] oops *) end text \ The next class considers the interaction of iteration with a greatest element. \ class bounded_left_conway_semiring = bounded_idempotent_left_semiring + left_conway_semiring begin lemma circ_top: "top\<^sup>\ = top" by (simp add: order.antisym circ_increasing) lemma circ_right_top: "x\<^sup>\ * top = top" by (metis sup_right_top circ_loop_fixpoint) lemma circ_left_top: "top * x\<^sup>\ = top" by (metis circ_right_top circ_top circ_decompose_11) lemma mult_top_circ: "(x * top)\<^sup>\ = 1 \ x * top" by (metis circ_left_top circ_left_unfold mult_assoc) end class left_zero_conway_semiring = idempotent_left_zero_semiring + left_conway_semiring begin lemma mult_zero_sup_circ_2: "(x \ y * bot)\<^sup>\ = x\<^sup>\ \ x\<^sup>\ * y * bot" by (metis mult_assoc mult_left_dist_sup mult_1_right troeger_3) lemma circ_unfold_sum: "(x \ y)\<^sup>\ = x\<^sup>\ \ x\<^sup>\ * y * (x \ y)\<^sup>\" by (metis mult_assoc mult_left_dist_sup mult_1_right troeger_1) end text \ The next class assumes the full sliding equation. \ class left_conway_semiring_1 = left_conway_semiring + assumes circ_right_slide: "x * (y * x)\<^sup>\ \ (x * y)\<^sup>\ * x" begin lemma circ_slide_1: "x * (y * x)\<^sup>\ = (x * y)\<^sup>\ * x" by (metis order.antisym circ_left_slide circ_right_slide) text \ This implies the full unfold rules and Conway's productstar. \ lemma circ_right_unfold_1: "1 \ x\<^sup>\ * x = x\<^sup>\" by (metis circ_left_unfold circ_slide_1 mult_1_left mult_1_right) lemma circ_mult_1: "(x * y)\<^sup>\ = 1 \ x * (y * x)\<^sup>\ * y" by (metis circ_left_unfold circ_slide_1 mult_assoc) lemma circ_sup_9: "(x \ y)\<^sup>\ = (x\<^sup>\ * y)\<^sup>\ * x\<^sup>\" by (metis circ_sup_1 circ_slide_1) lemma circ_plus_same: "x\<^sup>\ * x = x * x\<^sup>\" by (metis circ_slide_1 mult_1_left mult_1_right) lemma circ_decompose_12: "x\<^sup>\ * y\<^sup>\ \ (x\<^sup>\ * y)\<^sup>\ * x\<^sup>\" by (metis circ_sup_9 circ_sub_dist_3) end class left_zero_conway_semiring_1 = left_zero_conway_semiring + left_conway_semiring_1 begin lemma circ_back_loop_fixpoint: "(z * y\<^sup>\) * y \ z = z * y\<^sup>\" by (metis sup_commute circ_left_unfold circ_plus_same mult_assoc mult_left_dist_sup mult_1_right) lemma circ_back_loop_is_fixpoint: "is_fixpoint (\x . x * y \ z) (z * y\<^sup>\)" by (metis circ_back_loop_fixpoint is_fixpoint_def) lemma circ_elimination: "x * y = bot \ x * y\<^sup>\ \ x" by (metis sup_monoid.add_0_left circ_back_loop_fixpoint circ_plus_same mult_assoc mult_left_zero order_refl) end subsection \Iterings\ text \ This section adds simulation axioms to Conway semirings. We consider several classes with increasingly general simulation axioms. \ class itering_1 = left_conway_semiring_1 + assumes circ_simulate: "z * x \ y * z \ z * x\<^sup>\ \ y\<^sup>\ * z" begin lemma circ_circ_mult: "1\<^sup>\ * x\<^sup>\ = x\<^sup>\\<^sup>\" by (metis order.antisym circ_circ_sup circ_reflexive circ_simulate circ_sub_dist_3 circ_sup_one_left_unfold circ_transitive_equal mult_1_left order_refl) lemma sub_mult_one_circ: "x * 1\<^sup>\ \ 1\<^sup>\ * x" by (metis circ_simulate mult_1_left mult_1_right order_refl) text \ The left simulation axioms is enough to prove a basic import property of tests. \ lemma circ_import: assumes "p \ p * p" and "p \ 1" and "p * x \ x * p" shows "p * x\<^sup>\ = p * (p * x)\<^sup>\" proof - have "p * x \ p * (p * x * p) * p" by (metis assms coreflexive_transitive order.eq_iff test_preserves_equation mult_assoc) hence "p * x\<^sup>\ \ p * (p * x)\<^sup>\" by (metis (no_types) assms circ_simulate circ_slide_1 test_preserves_equation) thus ?thesis by (metis assms(2) circ_isotone mult_left_isotone mult_1_left mult_right_isotone order.antisym) qed end text \ Including generalisations of both simulation axioms allows us to prove separation rules. \ class itering_2 = left_conway_semiring_1 + assumes circ_simulate_right: "z * x \ y * z \ w \ z * x\<^sup>\ \ y\<^sup>\ * (z \ w * x\<^sup>\)" assumes circ_simulate_left: "x * z \ z * y \ w \ x\<^sup>\ * z \ (z \ x\<^sup>\ * w) * y\<^sup>\" begin subclass itering_1 apply unfold_locales by (metis sup_monoid.add_0_right circ_simulate_right mult_left_zero) lemma circ_simulate_left_1: "x * z \ z * y \ x\<^sup>\ * z \ z * y\<^sup>\ \ x\<^sup>\ * bot" by (metis sup_monoid.add_0_right circ_simulate_left mult_assoc mult_left_zero mult_right_dist_sup) lemma circ_separate_1: assumes "y * x \ x * y" shows "(x \ y)\<^sup>\ = x\<^sup>\ * y\<^sup>\" proof - have "y\<^sup>\ * x \ x * y\<^sup>\ \ y\<^sup>\ * bot" by (metis assms circ_simulate_left_1) hence "y\<^sup>\ * x * y\<^sup>\ \ x * y\<^sup>\ * y\<^sup>\ \ y\<^sup>\ * bot * y\<^sup>\" by (metis mult_assoc mult_left_isotone mult_right_dist_sup) also have "... = x * y\<^sup>\ \ y\<^sup>\ * bot" by (metis circ_transitive_equal mult_assoc mult_left_zero) finally have "y\<^sup>\ * (x * y\<^sup>\)\<^sup>\ \ x\<^sup>\ * (y\<^sup>\ \ y\<^sup>\ * bot)" using circ_simulate_right mult_assoc by fastforce also have "... = x\<^sup>\ * y\<^sup>\" by (simp add: sup_absorb1 zero_right_mult_decreasing) finally have "(x \ y)\<^sup>\ \ x\<^sup>\ * y\<^sup>\" by (simp add: circ_decompose_6 circ_sup_1) thus ?thesis by (simp add: order.antisym circ_sub_dist_3) qed lemma circ_circ_mult_1: "x\<^sup>\ * 1\<^sup>\ = x\<^sup>\\<^sup>\" by (metis sup_commute circ_circ_sup circ_separate_1 mult_1_left mult_1_right order_refl) end text \ With distributivity, we also get Back's atomicity refinement theorem. \ class itering_3 = itering_2 + left_zero_conway_semiring_1 begin lemma circ_simulate_1: assumes "y * x \ x * y" shows "y\<^sup>\ * x\<^sup>\ \ x\<^sup>\ * y\<^sup>\" proof - have "y * x\<^sup>\ \ x\<^sup>\ * y" by (metis assms circ_simulate) hence "y\<^sup>\ * x\<^sup>\ \ x\<^sup>\ * y\<^sup>\ \ y\<^sup>\ * bot" by (metis circ_simulate_left_1) thus ?thesis by (metis sup_assoc sup_monoid.add_0_right circ_loop_fixpoint mult_assoc mult_left_zero mult_zero_sup_circ_2) qed lemma atomicity_refinement: assumes "s = s * q" and "x = q * x" and "q * b = bot" and "r * b \ b * r" and "r * l \ l * r" and "x * l \ l * x" and "b * l \ l * b" and "q * l \ l * q" and "r\<^sup>\ * q \ q * r\<^sup>\" and "q \ 1" shows "s * (x \ b \ r \ l)\<^sup>\ * q \ s * (x * b\<^sup>\ * q \ r \ l)\<^sup>\" proof - have "(x \ b \ r) * l \ l * (x \ b \ r)" using assms(5-7) mult_left_dist_sup mult_right_dist_sup semiring.add_mono by presburger hence "s * (x \ b \ r \ l)\<^sup>\ * q = s * l\<^sup>\ * (x \ b \ r)\<^sup>\ * q" by (metis sup_commute circ_separate_1 mult_assoc) also have "... = s * l\<^sup>\ * b\<^sup>\ * r\<^sup>\ * q * (x * b\<^sup>\ * r\<^sup>\ * q)\<^sup>\" proof - have "(b \ r)\<^sup>\ = b\<^sup>\ * r\<^sup>\" by (simp add: assms(4) circ_separate_1) hence "b\<^sup>\ * r\<^sup>\ * (q * (x * b\<^sup>\ * r\<^sup>\))\<^sup>\ = (x \ b \ r)\<^sup>\" by (metis (full_types) assms(2) circ_sup_1 sup_assoc sup_commute mult_assoc) thus ?thesis by (metis circ_slide_1 mult_assoc) qed also have "... \ s * l\<^sup>\ * b\<^sup>\ * r\<^sup>\ * q * (x * b\<^sup>\ * q * r\<^sup>\)\<^sup>\" by (metis assms(9) circ_isotone mult_assoc mult_right_isotone) also have "... \ s * q * l\<^sup>\ * b\<^sup>\ * r\<^sup>\ * (x * b\<^sup>\ * q * r\<^sup>\)\<^sup>\" by (metis assms(1,10) mult_left_isotone mult_right_isotone mult_1_right) also have "... \ s * l\<^sup>\ * q * b\<^sup>\ * r\<^sup>\ * (x * b\<^sup>\ * q * r\<^sup>\)\<^sup>\" by (metis assms(1,8) circ_simulate mult_assoc mult_left_isotone mult_right_isotone) also have "... \ s * l\<^sup>\ * r\<^sup>\ * (x * b\<^sup>\ * q * r\<^sup>\)\<^sup>\" by (metis assms(3,10) sup_monoid.add_0_left circ_back_loop_fixpoint circ_plus_same mult_assoc mult_left_zero mult_left_isotone mult_right_isotone mult_1_right) also have "... \ s * (x * b\<^sup>\ * q \ r \ l)\<^sup>\" by (metis sup_commute circ_sup_1 circ_sub_dist_3 mult_assoc mult_right_isotone) finally show ?thesis . qed end text \ The following class contains the most general simulation axioms we consider. They allow us to prove further separation properties. \ class itering = idempotent_left_zero_semiring + circ + assumes circ_sup: "(x \ y)\<^sup>\ = (x\<^sup>\ * y)\<^sup>\ * x\<^sup>\" assumes circ_mult: "(x * y)\<^sup>\ = 1 \ x * (y * x)\<^sup>\ * y" assumes circ_simulate_right_plus: "z * x \ y * y\<^sup>\ * z \ w \ z * x\<^sup>\ \ y\<^sup>\ * (z \ w * x\<^sup>\)" assumes circ_simulate_left_plus: "x * z \ z * y\<^sup>\ \ w \ x\<^sup>\ * z \ (z \ x\<^sup>\ * w) * y\<^sup>\" begin lemma circ_right_unfold: "1 \ x\<^sup>\ * x = x\<^sup>\" by (metis circ_mult mult_1_left mult_1_right) lemma circ_slide: "x * (y * x)\<^sup>\ = (x * y)\<^sup>\ * x" proof - have "x * (y * x)\<^sup>\ = Rf x (y * 1 \ y * (x * (y * x)\<^sup>\ * y)) * x" by (metis (no_types) circ_mult mult_1_left mult_1_right mult_left_dist_sup mult_right_dist_sup mult_assoc) thus ?thesis by (metis (no_types) circ_mult mult_1_right mult_left_dist_sup mult_assoc) qed subclass itering_3 apply unfold_locales apply (metis circ_mult mult_1_left mult_1_right) apply (metis circ_slide order_refl) apply (metis circ_sup circ_slide) apply (metis circ_slide order_refl) apply (metis sup_left_isotone circ_right_unfold mult_left_isotone mult_left_sub_dist_sup_left mult_1_right order_trans circ_simulate_right_plus) by (metis sup_commute sup_ge1 sup_right_isotone circ_mult mult_right_isotone mult_1_right order_trans circ_simulate_left_plus) lemma circ_simulate_right_plus_1: "z * x \ y * y\<^sup>\ * z \ z * x\<^sup>\ \ y\<^sup>\ * z" by (metis sup_monoid.add_0_right circ_simulate_right_plus mult_left_zero) lemma circ_simulate_left_plus_1: "x * z \ z * y\<^sup>\ \ x\<^sup>\ * z \ z * y\<^sup>\ \ x\<^sup>\ * bot" by (metis sup_monoid.add_0_right circ_simulate_left_plus mult_assoc mult_left_zero mult_right_dist_sup) lemma circ_simulate_2: "y * x\<^sup>\ \ x\<^sup>\ * y\<^sup>\ \ y\<^sup>\ * x\<^sup>\ \ x\<^sup>\ * y\<^sup>\" apply (rule iffI) apply (metis sup_assoc sup_monoid.add_0_right circ_loop_fixpoint circ_simulate_left_plus_1 mult_assoc mult_left_zero mult_zero_sup_circ_2) by (metis circ_increasing mult_left_isotone order_trans) lemma circ_simulate_absorb: "y * x \ x \ y\<^sup>\ * x \ x \ y\<^sup>\ * bot" by (metis circ_simulate_left_plus_1 circ_zero mult_1_right) lemma circ_simulate_3: "y * x\<^sup>\ \ x\<^sup>\ \ y\<^sup>\ * x\<^sup>\ \ x\<^sup>\ * y\<^sup>\" by (metis sup.bounded_iff circ_reflexive circ_simulate_2 le_iff_sup mult_right_isotone mult_1_right) lemma circ_separate_mult_1: "y * x \ x * y \ (x * y)\<^sup>\ \ x\<^sup>\ * y\<^sup>\" by (metis circ_mult_sub_sup circ_separate_1) lemma circ_separate_unfold: "(y * x\<^sup>\)\<^sup>\ = y\<^sup>\ \ y\<^sup>\ * y * x * x\<^sup>\ * (y * x\<^sup>\)\<^sup>\" by (metis circ_back_loop_fixpoint circ_plus_same circ_unfold_sum sup_commute mult_assoc) lemma separation: assumes "y * x \ x * y\<^sup>\" shows "(x \ y)\<^sup>\ = x\<^sup>\ * y\<^sup>\" proof - have "y\<^sup>\ * x * y\<^sup>\ \ x * y\<^sup>\ \ y\<^sup>\ * bot" by (metis assms circ_simulate_left_plus_1 circ_transitive_equal mult_assoc mult_left_isotone) thus ?thesis by (metis sup_commute circ_sup_1 circ_simulate_right circ_sub_dist_3 le_iff_sup mult_assoc mult_left_zero zero_right_mult_decreasing) qed lemma simulation: "y * x \ x * y\<^sup>\ \ y\<^sup>\ * x\<^sup>\ \ x\<^sup>\ * y\<^sup>\" by (metis sup_ge2 circ_isotone circ_mult_upper_bound circ_sub_dist separation) lemma circ_simulate_4: assumes "y * x \ x * x\<^sup>\ * (1 \ y)" shows "y\<^sup>\ * x\<^sup>\ \ x\<^sup>\ * y\<^sup>\" proof - have "x \ (x * x\<^sup>\ * x * x \ x * x) = x * x\<^sup>\" by (metis (no_types) circ_back_loop_fixpoint mult_right_dist_sup sup_commute) hence "x \ x * x\<^sup>\ * 1 \ x * x\<^sup>\ * y" by (metis mult_1_right sup_assoc sup_ge1) hence "(1 \ y) * x \ x * x\<^sup>\ * (1 \ y)" using assms mult_left_dist_sup mult_right_dist_sup by force hence "y * x\<^sup>\ \ x\<^sup>\ * y\<^sup>\" by (metis circ_sup_upper_bound circ_increasing circ_reflexive circ_simulate_right_plus_1 mult_right_isotone mult_right_sub_dist_sup_right order_trans) thus ?thesis by (metis circ_simulate_2) qed lemma circ_simulate_5: "y * x \ x * x\<^sup>\ * (x \ y) \ y\<^sup>\ * x\<^sup>\ \ x\<^sup>\ * y\<^sup>\" by (metis circ_sup_sub_sup_one circ_simulate_4 order_trans) lemma circ_simulate_6: "y * x \ x * (x \ y) \ y\<^sup>\ * x\<^sup>\ \ x\<^sup>\ * y\<^sup>\" by (metis sup_commute circ_back_loop_fixpoint circ_simulate_5 mult_right_sub_dist_sup_left order_trans) lemma circ_separate_4: assumes "y * x \ x * x\<^sup>\ * (1 \ y)" shows "(x \ y)\<^sup>\ = x\<^sup>\ * y\<^sup>\" proof - have "y * x * x\<^sup>\ \ x * x\<^sup>\ * (1 \ y) * x\<^sup>\" by (simp add: assms mult_left_isotone) also have "... = x * x\<^sup>\ \ x * x\<^sup>\ * y * x\<^sup>\" by (simp add: circ_transitive_equal mult_left_dist_sup mult_right_dist_sup mult_assoc) also have "... \ x * x\<^sup>\ \ x * x\<^sup>\ * x\<^sup>\ * y\<^sup>\" by (metis assms sup_right_isotone circ_simulate_2 circ_simulate_4 mult_assoc mult_right_isotone) finally have "y * x * x\<^sup>\ \ x * x\<^sup>\ * y\<^sup>\" by (metis circ_reflexive circ_transitive_equal le_iff_sup mult_assoc mult_right_isotone mult_1_right) thus ?thesis by (metis circ_sup_1 left_plus_circ mult_assoc separation) qed lemma circ_separate_5: "y * x \ x * x\<^sup>\ * (x \ y) \ (x \ y)\<^sup>\ = x\<^sup>\ * y\<^sup>\" by (metis circ_sup_sub_sup_one circ_separate_4 order_trans) lemma circ_separate_6: "y * x \ x * (x \ y) \ (x \ y)\<^sup>\ = x\<^sup>\ * y\<^sup>\" by (metis sup_commute circ_back_loop_fixpoint circ_separate_5 mult_right_sub_dist_sup_left order_trans) end class bounded_itering = bounded_idempotent_left_zero_semiring + itering begin subclass bounded_left_conway_semiring .. (* lemma "1 = x\<^sup>\" nitpick [expect=genuine] oops lemma "x = x\<^sup>\" nitpick [expect=genuine] oops lemma "x = x * x\<^sup>\" nitpick [expect=genuine] oops lemma "x * x\<^sup>\ = x\<^sup>\" nitpick [expect=genuine] oops lemma "x\<^sup>\ = x\<^sup>\\<^sup>\" nitpick [expect=genuine] oops lemma "(x * y)\<^sup>\ = (x \ y)\<^sup>\" nitpick [expect=genuine] oops lemma "x\<^sup>\ * y\<^sup>\ = (x \ y)\<^sup>\" nitpick [expect=genuine,card=6] oops lemma "(x \ y)\<^sup>\ = (x\<^sup>\ * y\<^sup>\)\<^sup>\" nitpick [expect=genuine] oops lemma "1 = 1\<^sup>\" nitpick [expect=genuine] oops lemma "1 = (x * bot)\<^sup>\" nitpick [expect=genuine] oops lemma "1 \ x * bot = x\<^sup>\" nitpick [expect=genuine] oops lemma "x\<^sup>\ = x\<^sup>\ * 1\<^sup>\" nitpick [expect=genuine] oops lemma "z \ y * x = x \ y\<^sup>\ * z \ x" nitpick [expect=genuine] oops lemma "y * x = x \ y\<^sup>\ * x \ x" nitpick [expect=genuine] oops lemma "z \ x * y = x \ z * y\<^sup>\ \ x" nitpick [expect=genuine] oops lemma "x * y = x \ x * y\<^sup>\ \ x" nitpick [expect=genuine] oops lemma "x = z \ y * x \ x \ y\<^sup>\ * z" nitpick [expect=genuine] oops lemma "x = y * x \ x \ y\<^sup>\" nitpick [expect=genuine] oops lemma "x * z = z * y \ x\<^sup>\ * z \ z * y\<^sup>\" nitpick [expect=genuine] oops lemma "x\<^sup>\ = (x * x)\<^sup>\ * (x \ 1)" oops lemma "y\<^sup>\ * x\<^sup>\ \ x\<^sup>\ * y\<^sup>\ \ (x \ y)\<^sup>\ = x\<^sup>\ * y\<^sup>\" oops lemma "y * x \ (1 \ x) * y\<^sup>\ \ (x \ y)\<^sup>\ = x\<^sup>\ * y\<^sup>\" oops lemma "y * x \ x \ y\<^sup>\ * x \ 1\<^sup>\ * x" oops *) end text \ We finally expand Conway semirings and iterings by an element that corresponds to the endless loop. \ class L = fixes L :: "'a" class left_conway_semiring_L = left_conway_semiring + L + assumes one_circ_mult_split: "1\<^sup>\ * x = L \ x" assumes L_split_sup: "x * (y \ L) \ x * y \ L" begin lemma L_def: "L = 1\<^sup>\ * bot" by (metis sup_monoid.add_0_right one_circ_mult_split) lemma one_circ_split: "1\<^sup>\ = L \ 1" by (metis mult_1_right one_circ_mult_split) lemma one_circ_circ_split: "1\<^sup>\\<^sup>\ = L \ 1" by (metis circ_one one_circ_split) lemma sub_mult_one_circ: "x * 1\<^sup>\ \ 1\<^sup>\ * x" by (metis L_split_sup sup_commute mult_1_right one_circ_mult_split) lemma one_circ_mult_split_2: "1\<^sup>\ * x = x * 1\<^sup>\ \ L" proof - have 1: "x * 1\<^sup>\ \ L \ x" using one_circ_mult_split sub_mult_one_circ by presburger have "x \ x * 1\<^sup>\ = x * 1\<^sup>\" by (meson circ_back_loop_prefixpoint le_iff_sup sup.boundedE) thus ?thesis using 1 by (simp add: le_iff_sup one_circ_mult_split sup_assoc sup_commute) qed lemma sub_mult_one_circ_split: "x * 1\<^sup>\ \ x \ L" by (metis sup_commute one_circ_mult_split sub_mult_one_circ) lemma sub_mult_one_circ_split_2: "x * 1\<^sup>\ \ x \ 1\<^sup>\" by (metis L_def sup_right_isotone order_trans sub_mult_one_circ_split zero_right_mult_decreasing) lemma L_split: "x * L \ x * bot \ L" by (metis L_split_sup sup_monoid.add_0_left) lemma L_left_zero: "L * x = L" by (metis L_def mult_assoc mult_left_zero) lemma one_circ_L: "1\<^sup>\ * L = L" by (metis L_def circ_transitive_equal mult_assoc) lemma mult_L_circ: "(x * L)\<^sup>\ = 1 \ x * L" by (metis L_left_zero circ_left_unfold mult_assoc) lemma mult_L_circ_mult: "(x * L)\<^sup>\ * y = y \ x * L" by (metis L_left_zero mult_L_circ mult_assoc mult_1_left mult_right_dist_sup) lemma circ_L: "L\<^sup>\ = L \ 1" by (metis L_left_zero sup_commute circ_left_unfold) lemma L_below_one_circ: "L \ 1\<^sup>\" by (metis L_def zero_right_mult_decreasing) lemma circ_circ_mult_1: "x\<^sup>\ * 1\<^sup>\ = x\<^sup>\\<^sup>\" by (metis L_left_zero sup_commute circ_sup_1 circ_circ_sup mult_zero_circ one_circ_split) lemma circ_circ_mult: "1\<^sup>\ * x\<^sup>\ = x\<^sup>\\<^sup>\" by (metis order.antisym circ_circ_mult_1 circ_circ_sub_mult sub_mult_one_circ) lemma circ_circ_split: "x\<^sup>\\<^sup>\ = L \ x\<^sup>\" by (metis circ_circ_mult one_circ_mult_split) lemma circ_sup_6: "L \ (x \ y)\<^sup>\ = (x\<^sup>\ * y\<^sup>\)\<^sup>\" by (metis sup_assoc sup_commute circ_sup_1 circ_circ_sup circ_circ_split circ_decompose_4) end class itering_L = itering + L + assumes L_def: "L = 1\<^sup>\ * bot" begin lemma one_circ_split: "1\<^sup>\ = L \ 1" by (metis L_def sup_commute order.antisym circ_sup_upper_bound circ_reflexive circ_simulate_absorb mult_1_right order_refl zero_right_mult_decreasing) lemma one_circ_mult_split: "1\<^sup>\ * x = L \ x" by (metis L_def sup_commute circ_loop_fixpoint mult_assoc mult_left_zero mult_zero_circ one_circ_split) lemma sub_mult_one_circ_split: "x * 1\<^sup>\ \ x \ L" by (metis sup_commute one_circ_mult_split sub_mult_one_circ) lemma sub_mult_one_circ_split_2: "x * 1\<^sup>\ \ x \ 1\<^sup>\" by (metis L_def sup_right_isotone order_trans sub_mult_one_circ_split zero_right_mult_decreasing) lemma L_split: "x * L \ x * bot \ L" by (metis L_def mult_assoc mult_left_isotone mult_right_dist_sup sub_mult_one_circ_split_2) subclass left_conway_semiring_L apply unfold_locales apply (metis L_def sup_commute circ_loop_fixpoint mult_assoc mult_left_zero mult_zero_circ one_circ_split) by (metis sup_commute mult_assoc mult_left_isotone one_circ_mult_split sub_mult_one_circ) lemma circ_left_induct_mult_L: "L \ x \ x * y \ x \ x * y\<^sup>\ \ x" by (metis circ_one circ_simulate le_iff_sup one_circ_mult_split) lemma circ_left_induct_mult_iff_L: "L \ x \ x * y \ x \ x * y\<^sup>\ \ x" by (metis sup.bounded_iff circ_back_loop_fixpoint circ_left_induct_mult_L le_iff_sup) lemma circ_left_induct_L: "L \ x \ x * y \ z \ x \ z * y\<^sup>\ \ x" by (metis sup.bounded_iff circ_left_induct_mult_L le_iff_sup mult_right_dist_sup) end end diff --git a/thys/Stone_Kleene_Relation_Algebras/Kleene_Relation_Algebras.thy b/thys/Stone_Kleene_Relation_Algebras/Kleene_Relation_Algebras.thy --- a/thys/Stone_Kleene_Relation_Algebras/Kleene_Relation_Algebras.thy +++ b/thys/Stone_Kleene_Relation_Algebras/Kleene_Relation_Algebras.thy @@ -1,3300 +1,3309 @@ (* 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, opaque_lifting) 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 +class linorder_stone_kleene_relation_algebra_tarski_consistent_expansion = linorder_stone_kleene_relation_algebra_expansion + non_trivial_bounded_order +begin + +subclass stone_kleene_relation_algebra_tarski_consistent + apply unfold_locales + by (simp_all add: bot_not_top) + end +end + diff --git a/thys/Stone_Kleene_Relation_Algebras/Matrix_Kleene_Algebras.thy b/thys/Stone_Kleene_Relation_Algebras/Matrix_Kleene_Algebras.thy --- a/thys/Stone_Kleene_Relation_Algebras/Matrix_Kleene_Algebras.thy +++ b/thys/Stone_Kleene_Relation_Algebras/Matrix_Kleene_Algebras.thy @@ -1,1313 +1,1322 @@ (* Title: Matrix Kleene Algebras Author: Walter Guttmann Maintainer: Walter Guttmann *) section \Matrix Kleene Algebras\ text \ This theory gives a matrix model of Stone-Kleene relation algebras. The main result is that matrices over Kleene algebras form Kleene algebras. The automata-based construction is due to Conway \<^cite>\"Conway1971"\. An implementation of the construction in Isabelle/HOL that extends \<^cite>\"ArmstrongGomesStruthWeber2016"\ was given in \<^cite>\"Asplund2014"\ without a correctness proof. For specifying the size of matrices, Isabelle/HOL's type system requires the use of types, not sets. This creates two issues when trying to implement Conway's recursive construction directly. First, the matrix size changes for recursive calls, which requires dependent types. Second, some submatrices used in the construction are not square, which requires typed Kleene algebras \<^cite>\"Kozen1998"\, that is, categories of Kleene algebras. Because these instruments are not available in Isabelle/HOL, we use square matrices with a constant size given by the argument of the Kleene star operation. Smaller, possibly rectangular submatrices are identified by two lists of indices: one for the rows to include and one for the columns to include. Lists are used to make recursive calls deterministic; otherwise sets would be sufficient. \ theory Matrix_Kleene_Algebras imports Stone_Relation_Algebras.Matrix_Relation_Algebras Kleene_Relation_Algebras begin subsection \Matrix Restrictions\ text \ In this section we develop a calculus of matrix restrictions. The restriction of a matrix to specific row and column indices is implemented by the following function, which keeps the size of the matrix and sets all unused entries to \bot\. \ definition restrict_matrix :: "'a list \ ('a,'b::bot) square \ 'a list \ ('a,'b) square" ("_ \_\ _" [90,41,90] 91) where "restrict_matrix as f bs = (\(i,j) . if List.member as i \ List.member bs j then f (i,j) else bot)" text \ The following function captures Conway's automata-based construction of the Kleene star of a matrix. An index \k\ is chosen and \s\ contains all other indices. The matrix is split into four submatrices \a\, \b\, \c\, \d\ including/not including row/column \k\. Four matrices are computed containing the entries given by Conway's construction. These four matrices are added to obtain the result. All matrices involved in the function have the same size, but matrix restriction is used to set irrelevant entries to \bot\. \ primrec star_matrix' :: "'a list \ ('a,'b::{star,times,bounded_semilattice_sup_bot}) square \ ('a,'b) square" where "star_matrix' Nil g = mbot" | "star_matrix' (k#s) g = ( let r = [k] in let a = r\g\r in let b = r\g\s in let c = s\g\r in let d = s\g\s in let as = r\star o a\r in let ds = star_matrix' s d in let e = a \ b \ ds \ c in let es = r\star o e\r in let f = d \ c \ as \ b in let fs = star_matrix' s f in es \ as \ b \ fs \ ds \ c \ es \ fs )" text \ The Kleene star of the whole matrix is obtained by taking as indices all elements of the underlying type \'a\. This is conveniently supplied by the \enum\ class. \ fun star_matrix :: "('a::enum,'b::{star,times,bounded_semilattice_sup_bot}) square \ ('a,'b) square" ("_\<^sup>\" [100] 100) where "star_matrix f = star_matrix' (enum_class.enum::'a list) f" text \ The following lemmas deconstruct matrices with non-empty restrictions. \ lemma restrict_empty_left: "[]\f\ls = mbot" by (unfold restrict_matrix_def List.member_def bot_matrix_def) auto lemma restrict_empty_right: "ks\f\[] = mbot" by (unfold restrict_matrix_def List.member_def bot_matrix_def) auto lemma restrict_nonempty_left: fixes f :: "('a,'b::bounded_semilattice_sup_bot) square" shows "(k#ks)\f\ls = [k]\f\ls \ ks\f\ls" by (unfold restrict_matrix_def List.member_def sup_matrix_def) auto lemma restrict_nonempty_right: fixes f :: "('a,'b::bounded_semilattice_sup_bot) square" shows "ks\f\(l#ls) = ks\f\[l] \ ks\f\ls" by (unfold restrict_matrix_def List.member_def sup_matrix_def) auto lemma restrict_nonempty: fixes f :: "('a,'b::bounded_semilattice_sup_bot) square" shows "(k#ks)\f\(l#ls) = [k]\f\[l] \ [k]\f\ls \ ks\f\[l] \ ks\f\ls" by (unfold restrict_matrix_def List.member_def sup_matrix_def) auto text \ The following predicate captures that two index sets are disjoint. This has consequences for composition and the unit matrix. \ abbreviation "disjoint ks ls \ \(\x . List.member ks x \ List.member ls x)" lemma times_disjoint: fixes f g :: "('a,'b::idempotent_semiring) square" assumes "disjoint ls ms" shows "ks\f\ls \ ms\g\ns = mbot" proof (rule ext, rule prod_cases) fix i j have "(ks\f\ls \ ms\g\ns) (i,j) = (\\<^sub>k (ks\f\ls) (i,k) * (ms\g\ns) (k,j))" by (simp add: times_matrix_def) also have "... = (\\<^sub>k (if List.member ks i \ List.member ls k then f (i,k) else bot) * (if List.member ms k \ List.member ns j then g (k,j) else bot))" by (simp add: restrict_matrix_def) also have "... = (\\<^sub>k if List.member ms k \ List.member ns j then bot * g (k,j) else (if List.member ks i \ List.member ls k then f (i,k) else bot) * bot)" using assms by (auto intro: sup_monoid.sum.cong) also have "... = (\\<^sub>(k::'a) bot)" by (simp add: sup_monoid.sum.neutral) also have "... = bot" by (simp add: eq_iff le_funI) also have "... = mbot (i,j)" by (simp add: bot_matrix_def) finally show "(ks\f\ls \ ms\g\ns) (i,j) = mbot (i,j)" . qed lemma one_disjoint: assumes "disjoint ks ls" shows "ks\(mone::('a,'b::idempotent_semiring) square)\ls = mbot" proof (rule ext, rule prod_cases) let ?o = "mone::('a,'b) square" fix i j have "(ks\?o\ls) (i,j) = (if List.member ks i \ List.member ls j then if i = j then 1 else bot else bot)" by (simp add: restrict_matrix_def one_matrix_def) also have "... = bot" using assms by auto also have "... = mbot (i,j)" by (simp add: bot_matrix_def) finally show "(ks\?o\ls) (i,j) = mbot (i,j)" . qed text \ The following predicate captures that an index set is a subset of another index set. This has consequences for repeated restrictions. \ abbreviation "is_sublist ks ls \ \x . List.member ks x \ List.member ls x" lemma restrict_sublist: assumes "is_sublist ls ks" and "is_sublist ms ns" shows "ls\ks\f\ns\ms = ls\f\ms" proof (rule ext, rule prod_cases) fix i j show "(ls\ks\f\ns\ms) (i,j) = (ls\f\ms) (i,j)" proof (cases "List.member ls i \ List.member ms j") case True thus ?thesis by (simp add: assms restrict_matrix_def) next case False thus ?thesis by (unfold restrict_matrix_def) auto qed qed lemma restrict_superlist: assumes "is_sublist ls ks" and "is_sublist ms ns" shows "ks\ls\f\ms\ns = ls\f\ms" proof (rule ext, rule prod_cases) fix i j show "(ks\ls\f\ms\ns) (i,j) = (ls\f\ms) (i,j)" proof (cases "List.member ls i \ List.member ms j") case True thus ?thesis by (simp add: assms restrict_matrix_def) next case False thus ?thesis by (unfold restrict_matrix_def) auto qed qed text \ The following lemmas give the sizes of the results of some matrix operations. \ lemma restrict_sup: fixes f g :: "('a,'b::bounded_semilattice_sup_bot) square" shows "ks\f \ g\ls = ks\f\ls \ ks\g\ls" by (unfold restrict_matrix_def sup_matrix_def) auto lemma restrict_times: fixes f g :: "('a,'b::idempotent_semiring) square" shows "ks\ks\f\ls \ ls\g\ms\ms = ks\f\ls \ ls\g\ms" proof (rule ext, rule prod_cases) fix i j have "(ks\(ks\f\ls \ ls\g\ms)\ms) (i,j) = (if List.member ks i \ List.member ms j then (\\<^sub>k (ks\f\ls) (i,k) * (ls\g\ms) (k,j)) else bot)" by (simp add: times_matrix_def restrict_matrix_def) also have "... = (if List.member ks i \ List.member ms j then (\\<^sub>k (if List.member ks i \ List.member ls k then f (i,k) else bot) * (if List.member ls k \ List.member ms j then g (k,j) else bot)) else bot)" by (simp add: restrict_matrix_def) also have "... = (if List.member ks i \ List.member ms j then (\\<^sub>k if List.member ls k then f (i,k) * g (k,j) else bot) else bot)" by (auto intro: sup_monoid.sum.cong) also have "... = (\\<^sub>k if List.member ks i \ List.member ms j then (if List.member ls k then f (i,k) * g (k,j) else bot) else bot)" by auto also have "... = (\\<^sub>k (if List.member ks i \ List.member ls k then f (i,k) else bot) * (if List.member ls k \ List.member ms j then g (k,j) else bot))" by (auto intro: sup_monoid.sum.cong) also have "... = (\\<^sub>k (ks\f\ls) (i,k) * (ls\g\ms) (k,j))" by (simp add: restrict_matrix_def) also have "... = (ks\f\ls \ ls\g\ms) (i,j)" by (simp add: times_matrix_def) finally show "(ks\(ks\f\ls \ ls\g\ms)\ms) (i,j) = (ks\f\ls \ ls\g\ms) (i,j)" . qed lemma restrict_star: fixes g :: "('a,'b::kleene_algebra) square" shows "t\star_matrix' t g\t = star_matrix' t g" proof (induct arbitrary: g rule: list.induct) case Nil show ?case by (simp add: restrict_empty_left) next case (Cons k s) let ?t = "k#s" assume "\g::('a,'b) square . s\star_matrix' s g\s = star_matrix' s g" hence 1: "\g::('a,'b) square . ?t\star_matrix' s g\?t = star_matrix' s g" by (metis member_rec(1) restrict_superlist) show "?t\star_matrix' ?t g\?t = star_matrix' ?t g" proof - let ?r = "[k]" let ?a = "?r\g\?r" let ?b = "?r\g\s" let ?c = "s\g\?r" let ?d = "s\g\s" let ?as = "?r\star o ?a\?r" let ?ds = "star_matrix' s ?d" let ?e = "?a \ ?b \ ?ds \ ?c" let ?es = "?r\star o ?e\?r" let ?f = "?d \ ?c \ ?as \ ?b" let ?fs = "star_matrix' s ?f" have 2: "?t\?as\?t = ?as \ ?t\?b\?t = ?b \ ?t\?c\?t = ?c \ ?t\?es\?t = ?es" by (simp add: restrict_superlist member_def) have 3: "?t\?ds\?t = ?ds \ ?t\?fs\?t = ?fs" using 1 by simp have 4: "?t\?t\?as\?t \ ?t\?b\?t \ ?t\?fs\?t\?t = ?t\?as\?t \ ?t\?b\?t \ ?t\?fs\?t" by (metis (no_types) restrict_times) have 5: "?t\?t\?ds\?t \ ?t\?c\?t \ ?t\?es\?t\?t = ?t\?ds\?t \ ?t\?c\?t \ ?t\?es\?t" by (metis (no_types) restrict_times) have "?t\star_matrix' ?t g\?t = ?t\?es \ ?as \ ?b \ ?fs \ ?ds \ ?c \ ?es \ ?fs\?t" by (metis star_matrix'.simps(2)) also have "... = ?t\?es\?t \ ?t\?as \ ?b \ ?fs\?t \ ?t\?ds \ ?c \ ?es\?t \ ?t\?fs\?t" by (simp add: restrict_sup) also have "... = ?es \ ?as \ ?b \ ?fs \ ?ds \ ?c \ ?es \ ?fs" using 2 3 4 5 by simp also have "... = star_matrix' ?t g" by (metis star_matrix'.simps(2)) finally show ?thesis . qed qed lemma restrict_one: assumes "\ List.member ks k" shows "(k#ks)\(mone::('a,'b::idempotent_semiring) square)\(k#ks) = [k]\mone\[k] \ ks\mone\ks" by (subst restrict_nonempty) (simp add: assms member_rec one_disjoint) lemma restrict_one_left_unit: "ks\(mone::('a::finite,'b::idempotent_semiring) square)\ks \ ks\f\ls = ks\f\ls" proof (rule ext, rule prod_cases) let ?o = "mone::('a,'b::idempotent_semiring) square" fix i j have "(ks\?o\ks \ ks\f\ls) (i,j) = (\\<^sub>k (ks\?o\ks) (i,k) * (ks\f\ls) (k,j))" by (simp add: times_matrix_def) also have "... = (\\<^sub>k (if List.member ks i \ List.member ks k then ?o (i,k) else bot) * (if List.member ks k \ List.member ls j then f (k,j) else bot))" by (simp add: restrict_matrix_def) also have "... = (\\<^sub>k (if List.member ks i \ List.member ks k then (if i = k then 1 else bot) else bot) * (if List.member ks k \ List.member ls j then f (k,j) else bot))" by (unfold one_matrix_def) auto also have "... = (\\<^sub>k (if i = k then (if List.member ks i then 1 else bot) else bot) * (if List.member ks k \ List.member ls j then f (k,j) else bot))" by (auto intro: sup_monoid.sum.cong) also have "... = (\\<^sub>k if i = k then (if List.member ks i then 1 else bot) * (if List.member ks i \ List.member ls j then f (i,j) else bot) else bot)" by (rule sup_monoid.sum.cong) simp_all also have "... = (if List.member ks i then 1 else bot) * (if List.member ks i \ List.member ls j then f (i,j) else bot)" by simp also have "... = (if List.member ks i \ List.member ls j then f (i,j) else bot)" by simp also have "... = (ks\f\ls) (i,j)" by (simp add: restrict_matrix_def) finally show "(ks\?o\ks \ ks\f\ls) (i,j) = (ks\f\ls) (i,j)" . qed text \ The following lemmas consider restrictions to singleton index sets. \ lemma restrict_singleton: "([k]\f\[l]) (i,j) = (if i = k \ j = l then f (i,j) else bot)" by (simp add: restrict_matrix_def List.member_def) lemma restrict_singleton_list: "([k]\f\ls) (i,j) = (if i = k \ List.member ls j then f (i,j) else bot)" by (simp add: restrict_matrix_def List.member_def) lemma restrict_list_singleton: "(ks\f\[l]) (i,j) = (if List.member ks i \ j = l then f (i,j) else bot)" by (simp add: restrict_matrix_def List.member_def) lemma restrict_singleton_product: fixes f g :: "('a::finite,'b::kleene_algebra) square" shows "([k]\f\[l] \ [m]\g\[n]) (i,j) = (if i = k \ l = m \ j = n then f (i,l) * g (m,j) else bot)" proof - have "([k]\f\[l] \ [m]\g\[n]) (i,j) = (\\<^sub>h ([k]\f\[l]) (i,h) * ([m]\g\[n]) (h,j))" by (simp add: times_matrix_def) also have "... = (\\<^sub>h (if i = k \ h = l then f (i,h) else bot) * (if h = m \ j = n then g (h,j) else bot))" by (simp add: restrict_singleton) also have "... = (\\<^sub>h if h = l then (if i = k then f (i,h) else bot) * (if h = m \ j = n then g (h,j) else bot) else bot)" by (rule sup_monoid.sum.cong) auto also have "... = (if i = k then f (i,l) else bot) * (if l = m \ j = n then g (l,j) else bot)" by simp also have "... = (if i = k \ l = m \ j = n then f (i,l) * g (m,j) else bot)" by simp finally show ?thesis . qed text \ The Kleene star unfold law holds for matrices with a single entry on the diagonal. \ lemma restrict_star_unfold: "[l]\(mone::('a::finite,'b::kleene_algebra) square)\[l] \ [l]\f\[l] \ [l]\star o f\[l] = [l]\star o f\[l]" proof (rule ext, rule prod_cases) let ?o = "mone::('a,'b::kleene_algebra) square" fix i j have "([l]\?o\[l] \ [l]\f\[l] \ [l]\star o f\[l]) (i,j) = ([l]\?o\[l]) (i,j) \ ([l]\f\[l] \ [l]\star o f\[l]) (i,j)" by (simp add: sup_matrix_def) also have "... = ([l]\?o\[l]) (i,j) \ (\\<^sub>k ([l]\f\[l]) (i,k) * ([l]\star o f\[l]) (k,j))" by (simp add: times_matrix_def) also have "... = ([l]\?o\[l]) (i,j) \ (\\<^sub>k (if i = l \ k = l then f (i,k) else bot) * (if k = l \ j = l then (f (k,j))\<^sup>\ else bot))" by (simp add: restrict_singleton o_def) also have "... = ([l]\?o\[l]) (i,j) \ (\\<^sub>k if k = l then (if i = l then f (i,k) else bot) * (if j = l then (f (k,j))\<^sup>\ else bot) else bot)" apply (rule arg_cong2[where f=sup]) apply simp by (rule sup_monoid.sum.cong) auto also have "... = ([l]\?o\[l]) (i,j) \ (if i = l then f (i,l) else bot) * (if j = l then (f (l,j))\<^sup>\ else bot)" by simp also have "... = (if i = l \ j = l then 1 \ f (l,l) * (f (l,l))\<^sup>\ else bot)" by (simp add: restrict_singleton one_matrix_def) also have "... = (if i = l \ j = l then (f (l,l))\<^sup>\ else bot)" by (simp add: star_left_unfold_equal) also have "... = ([l]\star o f\[l]) (i,j)" by (simp add: restrict_singleton o_def) finally show "([l]\?o\[l] \ [l]\f\[l] \ [l]\star o f\[l]) (i,j) = ([l]\star o f\[l]) (i,j)" . qed lemma restrict_all: "enum_class.enum\f\enum_class.enum = f" by (simp add: restrict_matrix_def List.member_def enum_UNIV) text \ The following shows the various components of a matrix product. It is essentially a recursive implementation of the product. \ lemma restrict_nonempty_product: fixes f g :: "('a::finite,'b::idempotent_semiring) square" assumes "\ List.member ls l" shows "(k#ks)\f\(l#ls) \ (l#ls)\g\(m#ms) = ([k]\f\[l] \ [l]\g\[m] \ [k]\f\ls \ ls\g\[m]) \ ([k]\f\[l] \ [l]\g\ms \ [k]\f\ls \ ls\g\ms) \ (ks\f\[l] \ [l]\g\[m] \ ks\f\ls \ ls\g\[m]) \ (ks\f\[l] \ [l]\g\ms \ ks\f\ls \ ls\g\ms)" proof - have "(k#ks)\f\(l#ls) \ (l#ls)\g\(m#ms) = ([k]\f\[l] \ [k]\f\ls \ ks\f\[l] \ ks\f\ls) \ ([l]\g\[m] \ [l]\g\ms \ ls\g\[m] \ ls\g\ms)" by (metis restrict_nonempty) also have "... = [k]\f\[l] \ ([l]\g\[m] \ [l]\g\ms \ ls\g\[m] \ ls\g\ms) \ [k]\f\ls \ ([l]\g\[m] \ [l]\g\ms \ ls\g\[m] \ ls\g\ms) \ ks\f\[l] \ ([l]\g\[m] \ [l]\g\ms \ ls\g\[m] \ ls\g\ms) \ ks\f\ls \ ([l]\g\[m] \ [l]\g\ms \ ls\g\[m] \ ls\g\ms)" by (simp add: matrix_idempotent_semiring.mult_right_dist_sup) also have "... = ([k]\f\[l] \ [l]\g\[m] \ [k]\f\[l] \ [l]\g\ms \ [k]\f\[l] \ ls\g\[m] \ [k]\f\[l] \ ls\g\ms) \ ([k]\f\ls \ [l]\g\[m] \ [k]\f\ls \ [l]\g\ms \ [k]\f\ls \ ls\g\[m] \ [k]\f\ls \ ls\g\ms) \ (ks\f\[l] \ [l]\g\[m] \ ks\f\[l] \ [l]\g\ms \ ks\f\[l] \ ls\g\[m] \ ks\f\[l] \ ls\g\ms) \ (ks\f\ls \ [l]\g\[m] \ ks\f\ls \ [l]\g\ms \ ks\f\ls \ ls\g\[m] \ ks\f\ls \ ls\g\ms)" by (simp add: matrix_idempotent_semiring.mult_left_dist_sup) also have "... = ([k]\f\[l] \ [l]\g\[m] \ [k]\f\[l] \ [l]\g\ms) \ ([k]\f\ls \ ls\g\[m] \ [k]\f\ls \ ls\g\ms) \ (ks\f\[l] \ [l]\g\[m] \ ks\f\[l] \ [l]\g\ms) \ (ks\f\ls \ ls\g\[m] \ ks\f\ls \ ls\g\ms)" using assms by (simp add: List.member_def times_disjoint) also have "... = ([k]\f\[l] \ [l]\g\[m] \ [k]\f\ls \ ls\g\[m]) \ ([k]\f\[l] \ [l]\g\ms \ [k]\f\ls \ ls\g\ms) \ (ks\f\[l] \ [l]\g\[m] \ ks\f\ls \ ls\g\[m]) \ (ks\f\[l] \ [l]\g\ms \ ks\f\ls \ ls\g\ms)" by (simp add: matrix_bounded_semilattice_sup_bot.sup_monoid.add_assoc matrix_semilattice_sup.sup_left_commute) finally show ?thesis . qed text \ Equality of matrices is componentwise. \ lemma restrict_nonempty_eq: "(k#ks)\f\(l#ls) = (k#ks)\g\(l#ls) \ [k]\f\[l] = [k]\g\[l] \ [k]\f\ls = [k]\g\ls \ ks\f\[l] = ks\g\[l] \ ks\f\ls = ks\g\ls" proof assume 1: "(k#ks)\f\(l#ls) = (k#ks)\g\(l#ls)" have 2: "is_sublist [k] (k#ks) \ is_sublist ks (k#ks) \ is_sublist [l] (l#ls) \ is_sublist ls (l#ls)" by (simp add: member_rec) hence "[k]\f\[l] = [k]\(k#ks)\f\(l#ls)\[l] \ [k]\f\ls = [k]\(k#ks)\f\(l#ls)\ls \ ks\f\[l] = ks\(k#ks)\f\(l#ls)\[l] \ ks\f\ls = ks\(k#ks)\f\(l#ls)\ls" by (simp add: restrict_sublist) thus "[k]\f\[l] = [k]\g\[l] \ [k]\f\ls = [k]\g\ls \ ks\f\[l] = ks\g\[l] \ ks\f\ls = ks\g\ls" using 1 2 by (simp add: restrict_sublist) next assume 3: "[k]\f\[l] = [k]\g\[l] \ [k]\f\ls = [k]\g\ls \ ks\f\[l] = ks\g\[l] \ ks\f\ls = ks\g\ls" show "(k#ks)\f\(l#ls) = (k#ks)\g\(l#ls)" proof (rule ext, rule prod_cases) fix i j have 4: "f (k,l) = g (k,l)" using 3 by (metis restrict_singleton) have 5: "List.member ls j \ f (k,j) = g (k,j)" using 3 by (metis restrict_singleton_list) have 6: "List.member ks i \ f (i,l) = g (i,l)" using 3 by (metis restrict_list_singleton) have "(ks\f\ls) (i,j) = (ks\g\ls) (i,j)" using 3 by simp hence 7: "List.member ks i \ List.member ls j \ f (i,j) = g (i,j)" by (simp add: restrict_matrix_def) have "((k#ks)\f\(l#ls)) (i,j) = (if (i = k \ List.member ks i) \ (j = l \ List.member ls j) then f (i,j) else bot)" by (simp add: restrict_matrix_def List.member_def) also have "... = (if i = k \ j = l then f (i,j) else if i = k \ List.member ls j then f (i,j) else if List.member ks i \ j = l then f (i,j) else if List.member ks i \ List.member ls j then f (i,j) else bot)" by auto also have "... = (if i = k \ j = l then g (i,j) else if i = k \ List.member ls j then g (i,j) else if List.member ks i \ j = l then g (i,j) else if List.member ks i \ List.member ls j then g (i,j) else bot)" using 4 5 6 7 by simp also have "... = (if (i = k \ List.member ks i) \ (j = l \ List.member ls j) then g (i,j) else bot)" by auto also have "... = ((k#ks)\g\(l#ls)) (i,j)" by (simp add: restrict_matrix_def List.member_def) finally show "((k#ks)\f\(l#ls)) (i,j) = ((k#ks)\g\(l#ls)) (i,j)" . qed qed text \ Inequality of matrices is componentwise. \ lemma restrict_nonempty_less_eq: fixes f g :: "('a,'b::idempotent_semiring) square" shows "(k#ks)\f\(l#ls) \ (k#ks)\g\(l#ls) \ [k]\f\[l] \ [k]\g\[l] \ [k]\f\ls \ [k]\g\ls \ ks\f\[l] \ ks\g\[l] \ ks\f\ls \ ks\g\ls" by (unfold matrix_semilattice_sup.sup.order_iff) (metis (no_types, lifting) restrict_nonempty_eq restrict_sup) text \ The following lemmas treat repeated restrictions to disjoint index sets. \ lemma restrict_disjoint_left: assumes "disjoint ks ms" shows "ms\ks\f\ls\ns = mbot" proof (rule ext, rule prod_cases) fix i j have "(ms\ks\f\ls\ns) (i,j) = (if List.member ms i \ List.member ns j then if List.member ks i \ List.member ls j then f (i,j) else bot else bot)" by (simp add: restrict_matrix_def) thus "(ms\ks\f\ls\ns) (i,j) = mbot (i,j)" using assms by (simp add: bot_matrix_def) qed lemma restrict_disjoint_right: assumes "disjoint ls ns" shows "ms\ks\f\ls\ns = mbot" proof (rule ext, rule prod_cases) fix i j have "(ms\ks\f\ls\ns) (i,j) = (if List.member ms i \ List.member ns j then if List.member ks i \ List.member ls j then f (i,j) else bot else bot)" by (simp add: restrict_matrix_def) thus "(ms\ks\f\ls\ns) (i,j) = mbot (i,j)" using assms by (simp add: bot_matrix_def) qed text \ The following lemma expresses the equality of a matrix and a product of two matrices componentwise. \ lemma restrict_nonempty_product_eq: fixes f g h :: "('a::finite,'b::idempotent_semiring) square" assumes "\ List.member ks k" and "\ List.member ls l" and "\ List.member ms m" shows "(k#ks)\f\(l#ls) \ (l#ls)\g\(m#ms) = (k#ks)\h\(m#ms) \ [k]\f\[l] \ [l]\g\[m] \ [k]\f\ls \ ls\g\[m] = [k]\h\[m] \ [k]\f\[l] \ [l]\g\ms \ [k]\f\ls \ ls\g\ms = [k]\h\ms \ ks\f\[l] \ [l]\g\[m] \ ks\f\ls \ ls\g\[m] = ks\h\[m] \ ks\f\[l] \ [l]\g\ms \ ks\f\ls \ ls\g\ms = ks\h\ms" proof - have 1: "disjoint [k] ks \ disjoint [m] ms" by (simp add: assms(1,3) member_rec) have 2: "[k]\(k#ks)\f\(l#ls) \ (l#ls)\g\(m#ms)\[m] = [k]\f\[l] \ [l]\g\[m] \ [k]\f\ls \ ls\g\[m]" proof - have "[k]\(k#ks)\f\(l#ls) \ (l#ls)\g\(m#ms)\[m] = [k]\([k]\f\[l] \ [l]\g\[m] \ [k]\f\ls \ ls\g\[m]) \ ([k]\f\[l] \ [l]\g\ms \ [k]\f\ls \ ls\g\ms) \ (ks\f\[l] \ [l]\g\[m] \ ks\f\ls \ ls\g\[m]) \ (ks\f\[l] \ [l]\g\ms \ ks\f\ls \ ls\g\ms)\[m]" by (simp add: assms(2) restrict_nonempty_product) also have "... = [k]\[k]\f\[l] \ [l]\g\[m]\[m] \ [k]\[k]\f\ls \ ls\g\[m]\[m] \ [k]\[k]\f\[l] \ [l]\g\ms\[m] \ [k]\[k]\f\ls \ ls\g\ms\[m] \ [k]\ks\f\[l] \ [l]\g\[m]\[m] \ [k]\ks\f\ls \ ls\g\[m]\[m] \ [k]\ks\f\[l] \ [l]\g\ms\[m] \ [k]\ks\f\ls \ ls\g\ms\[m]" by (simp add: matrix_bounded_semilattice_sup_bot.sup_monoid.add_assoc restrict_sup) also have "... = [k]\f\[l] \ [l]\g\[m] \ [k]\f\ls \ ls\g\[m] \ [k]\[k]\[k]\f\[l] \ [l]\g\ms\ms\[m] \ [k]\[k]\[k]\f\ls \ ls\g\ms\ms\[m] \ [k]\ks\ks\f\[l] \ [l]\g\[m]\[m]\[m] \ [k]\ks\ks\f\ls \ ls\g\[m]\[m]\[m] \ [k]\ks\ks\f\[l] \ [l]\g\ms\ms\[m] \ [k]\ks\ks\f\ls \ ls\g\ms\ms\[m]" by (simp add: restrict_times) also have "... = [k]\f\[l] \ [l]\g\[m] \ [k]\f\ls \ ls\g\[m]" using 1 by (metis restrict_disjoint_left restrict_disjoint_right matrix_bounded_semilattice_sup_bot.sup_monoid.add_0_right) finally show ?thesis . qed have 3: "[k]\(k#ks)\f\(l#ls) \ (l#ls)\g\(m#ms)\ms = [k]\f\[l] \ [l]\g\ms \ [k]\f\ls \ ls\g\ms" proof - have "[k]\(k#ks)\f\(l#ls) \ (l#ls)\g\(m#ms)\ms = [k]\([k]\f\[l] \ [l]\g\[m] \ [k]\f\ls \ ls\g\[m]) \ ([k]\f\[l] \ [l]\g\ms \ [k]\f\ls \ ls\g\ms) \ (ks\f\[l] \ [l]\g\[m] \ ks\f\ls \ ls\g\[m]) \ (ks\f\[l] \ [l]\g\ms \ ks\f\ls \ ls\g\ms)\ms" by (simp add: assms(2) restrict_nonempty_product) also have "... = [k]\[k]\f\[l] \ [l]\g\[m]\ms \ [k]\[k]\f\ls \ ls\g\[m]\ms \ [k]\[k]\f\[l] \ [l]\g\ms\ms \ [k]\[k]\f\ls \ ls\g\ms\ms \ [k]\ks\f\[l] \ [l]\g\[m]\ms \ [k]\ks\f\ls \ ls\g\[m]\ms \ [k]\ks\f\[l] \ [l]\g\ms\ms \ [k]\ks\f\ls \ ls\g\ms\ms" by (simp add: matrix_bounded_semilattice_sup_bot.sup_monoid.add_assoc restrict_sup) also have "... = [k]\[k]\[k]\f\[l] \ [l]\g\[m]\[m]\ms \ [k]\[k]\[k]\f\ls \ ls\g\[m]\[m]\ms \ [k]\f\[l] \ [l]\g\ms \ [k]\f\ls \ ls\g\ms \ [k]\ks\ks\f\[l] \ [l]\g\[m]\[m]\ms \ [k]\ks\ks\f\ls \ ls\g\[m]\[m]\ms \ [k]\ks\ks\f\[l] \ [l]\g\ms\ms\ms \ [k]\ks\ks\f\ls \ ls\g\ms\ms\ms" by (simp add: restrict_times) also have "... = [k]\f\[l] \ [l]\g\ms \ [k]\f\ls \ ls\g\ms" using 1 by (metis restrict_disjoint_left restrict_disjoint_right matrix_bounded_semilattice_sup_bot.sup_monoid.add_0_right matrix_bounded_semilattice_sup_bot.sup_monoid.add_0_left) finally show ?thesis . qed have 4: "ks\(k#ks)\f\(l#ls) \ (l#ls)\g\(m#ms)\[m] = ks\f\[l] \ [l]\g\[m] \ ks\f\ls \ ls\g\[m]" proof - have "ks\(k#ks)\f\(l#ls) \ (l#ls)\g\(m#ms)\[m] = ks\([k]\f\[l] \ [l]\g\[m] \ [k]\f\ls \ ls\g\[m]) \ ([k]\f\[l] \ [l]\g\ms \ [k]\f\ls \ ls\g\ms) \ (ks\f\[l] \ [l]\g\[m] \ ks\f\ls \ ls\g\[m]) \ (ks\f\[l] \ [l]\g\ms \ ks\f\ls \ ls\g\ms)\[m]" by (simp add: assms(2) restrict_nonempty_product) also have "... = ks\[k]\f\[l] \ [l]\g\[m]\[m] \ ks\[k]\f\ls \ ls\g\[m]\[m] \ ks\[k]\f\[l] \ [l]\g\ms\[m] \ ks\[k]\f\ls \ ls\g\ms\[m] \ ks\ks\f\[l] \ [l]\g\[m]\[m] \ ks\ks\f\ls \ ls\g\[m]\[m] \ ks\ks\f\[l] \ [l]\g\ms\[m] \ ks\ks\f\ls \ ls\g\ms\[m]" by (simp add: matrix_bounded_semilattice_sup_bot.sup_monoid.add_assoc restrict_sup) also have "... = ks\[k]\[k]\f\[l] \ [l]\g\[m]\[m]\[m] \ ks\[k]\[k]\f\ls \ ls\g\[m]\[m]\[m] \ ks\[k]\[k]\f\[l] \ [l]\g\ms\ms\[m] \ ks\[k]\[k]\f\ls \ ls\g\ms\ms\[m] \ ks\f\[l] \ [l]\g\[m] \ ks\f\ls \ ls\g\[m] \ ks\ks\ks\f\[l] \ [l]\g\ms\ms\[m] \ ks\ks\ks\f\ls \ ls\g\ms\ms\[m]" by (simp add: restrict_times) also have "... = ks\f\[l] \ [l]\g\[m] \ ks\f\ls \ ls\g\[m]" using 1 by (metis restrict_disjoint_left restrict_disjoint_right matrix_bounded_semilattice_sup_bot.sup_monoid.add_0_right matrix_bounded_semilattice_sup_bot.sup_monoid.add_0_left) finally show ?thesis . qed have 5: "ks\(k#ks)\f\(l#ls) \ (l#ls)\g\(m#ms)\ms = ks\f\[l] \ [l]\g\ms \ ks\f\ls \ ls\g\ms" proof - have "ks\(k#ks)\f\(l#ls) \ (l#ls)\g\(m#ms)\ms = ks\([k]\f\[l] \ [l]\g\[m] \ [k]\f\ls \ ls\g\[m]) \ ([k]\f\[l] \ [l]\g\ms \ [k]\f\ls \ ls\g\ms) \ (ks\f\[l] \ [l]\g\[m] \ ks\f\ls \ ls\g\[m]) \ (ks\f\[l] \ [l]\g\ms \ ks\f\ls \ ls\g\ms)\ms" by (simp add: assms(2) restrict_nonempty_product) also have "... = ks\[k]\f\[l] \ [l]\g\[m]\ms \ ks\[k]\f\ls \ ls\g\[m]\ms \ ks\[k]\f\[l] \ [l]\g\ms\ms \ ks\[k]\f\ls \ ls\g\ms\ms \ ks\ks\f\[l] \ [l]\g\[m]\ms \ ks\ks\f\ls \ ls\g\[m]\ms \ ks\ks\f\[l] \ [l]\g\ms\ms \ ks\ks\f\ls \ ls\g\ms\ms" by (simp add: matrix_bounded_semilattice_sup_bot.sup_monoid.add_assoc restrict_sup) also have "... = ks\[k]\[k]\f\[l] \ [l]\g\[m]\[m]\ms \ ks\[k]\[k]\f\ls \ ls\g\[m]\[m]\ms \ ks\[k]\[k]\f\[l] \ [l]\g\ms\ms\ms \ ks\[k]\[k]\f\ls \ ls\g\ms\ms\ms \ ks\ks\ks\f\[l] \ [l]\g\[m]\[m]\ms \ ks\ks\ks\f\ls \ ls\g\[m]\[m]\ms \ ks\f\[l] \ [l]\g\ms \ ks\f\ls \ ls\g\ms" by (simp add: restrict_times) also have "... = ks\f\[l] \ [l]\g\ms \ ks\f\ls \ ls\g\ms" using 1 by (metis restrict_disjoint_left restrict_disjoint_right matrix_bounded_semilattice_sup_bot.sup_monoid.add_0_left) finally show ?thesis . qed have "(k#ks)\f\(l#ls) \ (l#ls)\g\(m#ms) = (k#ks)\h\(m#ms) \ (k#ks)\(k#ks)\f\(l#ls) \ (l#ls)\g\(m#ms)\(m#ms) = (k#ks)\h\(m#ms)" by (simp add: restrict_times) also have "... \ [k]\(k#ks)\f\(l#ls) \ (l#ls)\g\(m#ms)\[m] = [k]\h\[m] \ [k]\(k#ks)\f\(l#ls) \ (l#ls)\g\(m#ms)\ms = [k]\h\ms \ ks\(k#ks)\f\(l#ls) \ (l#ls)\g\(m#ms)\[m] = ks\h\[m] \ ks\(k#ks)\f\(l#ls) \ (l#ls)\g\(m#ms)\ms = ks\h\ms" by (meson restrict_nonempty_eq) also have "... \ [k]\f\[l] \ [l]\g\[m] \ [k]\f\ls \ ls\g\[m] = [k]\h\[m] \ [k]\f\[l] \ [l]\g\ms \ [k]\f\ls \ ls\g\ms = [k]\h\ms \ ks\f\[l] \ [l]\g\[m] \ ks\f\ls \ ls\g\[m] = ks\h\[m] \ ks\f\[l] \ [l]\g\ms \ ks\f\ls \ ls\g\ms = ks\h\ms" using 2 3 4 5 by simp finally show ?thesis by simp qed text \ The following lemma gives a componentwise characterisation of the inequality of a matrix and a product of two matrices. \ lemma restrict_nonempty_product_less_eq: fixes f g h :: "('a::finite,'b::idempotent_semiring) square" assumes "\ List.member ks k" and "\ List.member ls l" and "\ List.member ms m" shows "(k#ks)\f\(l#ls) \ (l#ls)\g\(m#ms) \ (k#ks)\h\(m#ms) \ [k]\f\[l] \ [l]\g\[m] \ [k]\f\ls \ ls\g\[m] \ [k]\h\[m] \ [k]\f\[l] \ [l]\g\ms \ [k]\f\ls \ ls\g\ms \ [k]\h\ms \ ks\f\[l] \ [l]\g\[m] \ ks\f\ls \ ls\g\[m] \ ks\h\[m] \ ks\f\[l] \ [l]\g\ms \ ks\f\ls \ ls\g\ms \ ks\h\ms" proof - have 1: "[k]\(k#ks)\f\(l#ls) \ (l#ls)\g\(m#ms)\[m] = [k]\f\[l] \ [l]\g\[m] \ [k]\f\ls \ ls\g\[m]" by (metis assms restrict_nonempty_product_eq restrict_times) have 2: "[k]\(k#ks)\f\(l#ls) \ (l#ls)\g\(m#ms)\ms = [k]\f\[l] \ [l]\g\ms \ [k]\f\ls \ ls\g\ms" by (metis assms restrict_nonempty_product_eq restrict_times) have 3: "ks\(k#ks)\f\(l#ls) \ (l#ls)\g\(m#ms)\[m] = ks\f\[l] \ [l]\g\[m] \ ks\f\ls \ ls\g\[m]" by (metis assms restrict_nonempty_product_eq restrict_times) have 4: "ks\(k#ks)\f\(l#ls) \ (l#ls)\g\(m#ms)\ms = ks\f\[l] \ [l]\g\ms \ ks\f\ls \ ls\g\ms" by (metis assms restrict_nonempty_product_eq restrict_times) have "(k#ks)\f\(l#ls) \ (l#ls)\g\(m#ms) \ (k#ks)\h\(m#ms) \ (k#ks)\(k#ks)\f\(l#ls) \ (l#ls)\g\(m#ms)\(m#ms) \ (k#ks)\h\(m#ms)" by (simp add: restrict_times) also have "... \ [k]\(k#ks)\f\(l#ls) \ (l#ls)\g\(m#ms)\[m] \ [k]\h\[m] \ [k]\(k#ks)\f\(l#ls) \ (l#ls)\g\(m#ms)\ms \ [k]\h\ms \ ks\(k#ks)\f\(l#ls) \ (l#ls)\g\(m#ms)\[m] \ ks\h\[m] \ ks\(k#ks)\f\(l#ls) \ (l#ls)\g\(m#ms)\ms \ ks\h\ms" by (meson restrict_nonempty_less_eq) also have "... \ [k]\f\[l] \ [l]\g\[m] \ [k]\f\ls \ ls\g\[m] \ [k]\h\[m] \ [k]\f\[l] \ [l]\g\ms \ [k]\f\ls \ ls\g\ms \ [k]\h\ms \ ks\f\[l] \ [l]\g\[m] \ ks\f\ls \ ls\g\[m] \ ks\h\[m] \ ks\f\[l] \ [l]\g\ms \ ks\f\ls \ ls\g\ms \ ks\h\ms" using 1 2 3 4 by simp finally show ?thesis by simp qed text \ The Kleene star induction laws hold for matrices with a single entry on the diagonal. The matrix \g\ can actually contain a whole row/colum at the appropriate index. \ lemma restrict_star_left_induct: fixes f g :: "('a::finite,'b::kleene_algebra) square" shows "distinct ms \ [l]\f\[l] \ [l]\g\ms \ [l]\g\ms \ [l]\star o f\[l] \ [l]\g\ms \ [l]\g\ms" proof (induct ms) case Nil thus ?case by (simp add: restrict_empty_right) next case (Cons m ms) assume 1: "distinct ms \ [l]\f\[l] \ [l]\g\ms \ [l]\g\ms \ [l]\star o f\[l] \ [l]\g\ms \ [l]\g\ms" assume 2: "distinct (m#ms)" assume 3: "[l]\f\[l] \ [l]\g\(m#ms) \ [l]\g\(m#ms)" have 4: "[l]\f\[l] \ [l]\g\[m] \ [l]\g\[m] \ [l]\f\[l] \ [l]\g\ms \ [l]\g\ms" using 2 3 by (metis distinct.simps(2) matrix_semilattice_sup.sup.bounded_iff member_def member_rec(2) restrict_nonempty_product_less_eq) hence 5: "[l]\star o f\[l] \ [l]\g\ms \ [l]\g\ms" using 1 2 by simp have "f (l,l) * g (l,m) \ g (l,m)" using 4 by (metis restrict_singleton_product restrict_singleton less_eq_matrix_def) hence 6: "(f (l,l))\<^sup>\ * g (l,m) \ g (l,m)" by (simp add: star_left_induct_mult) have "[l]\star o f\[l] \ [l]\g\[m] \ [l]\g\[m]" proof (unfold less_eq_matrix_def, rule allI, rule prod_cases) fix i j have "([l]\star o f\[l] \ [l]\g\[m]) (i,j) = (\\<^sub>k ([l]\star o f\[l]) (i,k) * ([l]\g\[m]) (k,j))" by (simp add: times_matrix_def) also have "... = (\\<^sub>k (if i = l \ k = l then (f (i,k))\<^sup>\ else bot) * (if k = l \ j = m then g (k,j) else bot))" by (simp add: restrict_singleton o_def) also have "... = (\\<^sub>k if k = l then (if i = l then (f (i,k))\<^sup>\ else bot) * (if j = m then g (k,j) else bot) else bot)" by (rule sup_monoid.sum.cong) auto also have "... = (if i = l then (f (i,l))\<^sup>\ else bot) * (if j = m then g (l,j) else bot)" by simp also have "... = (if i = l \ j = m then (f (l,l))\<^sup>\ * g (l,m) else bot)" by simp also have "... \ ([l]\g\[m]) (i,j)" using 6 by (simp add: restrict_singleton) finally show "([l]\star o f\[l] \ [l]\g\[m]) (i,j) \ ([l]\g\[m]) (i,j)" . qed thus "[l]\star o f\[l] \ [l]\g\(m#ms) \ [l]\g\(m#ms)" using 2 5 by (metis (no_types, opaque_lifting) matrix_idempotent_semiring.mult_left_dist_sup matrix_semilattice_sup.sup.mono restrict_nonempty_right) qed lemma restrict_star_right_induct: fixes f g :: "('a::finite,'b::kleene_algebra) square" shows "distinct ms \ ms\g\[l] \ [l]\f\[l] \ ms\g\[l] \ ms\g\[l] \ [l]\star o f\[l] \ ms\g\[l]" proof (induct ms) case Nil thus ?case by (simp add: restrict_empty_left) next case (Cons m ms) assume 1: "distinct ms \ ms\g\[l] \ [l]\f\[l] \ ms\g\[l] \ ms\g\[l] \ [l]\star o f\[l] \ ms\g\[l]" assume 2: "distinct (m#ms)" assume 3: "(m#ms)\g\[l] \ [l]\f\[l] \ (m#ms)\g\[l]" have 4: "[m]\g\[l] \ [l]\f\[l] \ [m]\g\[l] \ ms\g\[l] \ [l]\f\[l] \ ms\g\[l]" using 2 3 by (metis distinct.simps(2) matrix_semilattice_sup.sup.bounded_iff member_def member_rec(2) restrict_nonempty_product_less_eq) hence 5: "ms\g\[l] \ [l]\star o f\[l] \ ms\g\[l]" using 1 2 by simp have "g (m,l) * f (l,l) \ g (m,l)" using 4 by (metis restrict_singleton_product restrict_singleton less_eq_matrix_def) hence 6: "g (m,l) * (f (l,l))\<^sup>\ \ g (m,l)" by (simp add: star_right_induct_mult) have "[m]\g\[l] \ [l]\star o f\[l] \ [m]\g\[l]" proof (unfold less_eq_matrix_def, rule allI, rule prod_cases) fix i j have "([m]\g\[l] \ [l]\star o f\[l]) (i,j) = (\\<^sub>k ([m]\g\[l]) (i,k) * ([l]\star o f\[l]) (k,j))" by (simp add: times_matrix_def) also have "... = (\\<^sub>k (if i = m \ k = l then g (i,k) else bot) * (if k = l \ j = l then (f (k,j))\<^sup>\ else bot))" by (simp add: restrict_singleton o_def) also have "... = (\\<^sub>k if k = l then (if i = m then g (i,k) else bot) * (if j = l then (f (k,j))\<^sup>\ else bot) else bot)" by (rule sup_monoid.sum.cong) auto also have "... = (if i = m then g (i,l) else bot) * (if j = l then (f (l,j))\<^sup>\ else bot)" by simp also have "... = (if i = m \ j = l then g (m,l) * (f (l,l))\<^sup>\ else bot)" by simp also have "... \ ([m]\g\[l]) (i,j)" using 6 by (simp add: restrict_singleton) finally show "([m]\g\[l] \ [l]\star o f\[l]) (i,j) \ ([m]\g\[l]) (i,j)" . qed thus "(m#ms)\g\[l] \ [l]\star o f\[l] \ (m#ms)\g\[l]" using 2 5 by (metis (no_types, opaque_lifting) matrix_idempotent_semiring.mult_right_dist_sup matrix_semilattice_sup.sup.mono restrict_nonempty_left) qed lemma restrict_pp: fixes f :: "('a,'b::p_algebra) square" shows "ks\\\f\ls = \\(ks\f\ls)" by (unfold restrict_matrix_def uminus_matrix_def) auto lemma pp_star_commute: fixes f :: "('a,'b::stone_kleene_relation_algebra) square" shows "\\(star o f) = star o \\f" by (simp add: uminus_matrix_def o_def pp_dist_star) subsection \Matrices form a Kleene Algebra\ text \ Matrices over Kleene algebras form a Kleene algebra using Conway's construction. It remains to prove one unfold and two induction axioms of the Kleene star. Each proof is by induction over the size of the matrix represented by an index list. \ interpretation matrix_kleene_algebra: kleene_algebra_var where sup = sup_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix::('a::enum,'b::kleene_algebra) square" and one = one_matrix and times = times_matrix and star = star_matrix proof fix y :: "('a,'b) square" let ?e = "enum_class.enum::'a list" let ?o = "mone :: ('a,'b) square" have "\g :: ('a,'b) square . distinct ?e \ (?e\?o\?e \ ?e\g\?e \ star_matrix' ?e g) = (star_matrix' ?e g)" proof (induct rule: list.induct) case Nil thus ?case by (simp add: restrict_empty_left) next case (Cons k s) let ?t = "k#s" assume 1: "\g :: ('a,'b) square . distinct s \ (s\?o\s \ s\g\s \ star_matrix' s g) = (star_matrix' s g)" show "\g :: ('a,'b) square . distinct ?t \ (?t\?o\?t \ ?t\g\?t \ star_matrix' ?t g) = (star_matrix' ?t g)" proof (rule allI, rule impI) fix g :: "('a,'b) square" assume 2: "distinct ?t" let ?r = "[k]" let ?a = "?r\g\?r" let ?b = "?r\g\s" let ?c = "s\g\?r" let ?d = "s\g\s" let ?as = "?r\star o ?a\?r" let ?ds = "star_matrix' s ?d" let ?e = "?a \ ?b \ ?ds \ ?c" let ?es = "?r\star o ?e\?r" let ?f = "?d \ ?c \ ?as \ ?b" let ?fs = "star_matrix' s ?f" have "s\?ds\s = ?ds \ s\?fs\s = ?fs" by (simp add: restrict_star) hence 3: "?r\?e\?r = ?e \ s\?f\s = ?f" by (metis (no_types, lifting) restrict_one_left_unit restrict_sup restrict_times) have 4: "disjoint s ?r \ disjoint ?r s" using 2 by (simp add: in_set_member member_rec) hence 5: "?t\?o\?t = ?r\?o\?r \ s\?o\s" by (meson member_rec(1) restrict_one) have 6: "?t\g\?t \ ?es = ?a \ ?es \ ?c \ ?es" proof - have "?t\g\?t \ ?es = (?a \ ?b \ ?c \ ?d) \ ?es" by (metis restrict_nonempty) also have "... = ?a \ ?es \ ?b \ ?es \ ?c \ ?es \ ?d \ ?es" by (simp add: matrix_idempotent_semiring.mult_right_dist_sup) also have "... = ?a \ ?es \ ?c \ ?es" using 4 by (simp add: times_disjoint) finally show ?thesis . qed have 7: "?t\g\?t \ ?as \ ?b \ ?fs = ?a \ ?as \ ?b \ ?fs \ ?c \ ?as \ ?b \ ?fs" proof - have "?t\g\?t \ ?as \ ?b \ ?fs = (?a \ ?b \ ?c \ ?d) \ ?as \ ?b \ ?fs" by (metis restrict_nonempty) also have "... = ?a \ ?as \ ?b \ ?fs \ ?b \ ?as \ ?b \ ?fs \ ?c \ ?as \ ?b \ ?fs \ ?d \ ?as \ ?b \ ?fs" by (simp add: matrix_idempotent_semiring.mult_right_dist_sup) also have "... = ?a \ ?as \ ?b \ ?fs \ ?c \ ?as \ ?b \ ?fs" using 4 by (simp add: times_disjoint) finally show ?thesis . qed have 8: "?t\g\?t \ ?ds \ ?c \ ?es = ?b \ ?ds \ ?c \ ?es \ ?d \ ?ds \ ?c \ ?es" proof - have "?t\g\?t \ ?ds \ ?c \ ?es = (?a \ ?b \ ?c \ ?d) \ ?ds \ ?c \ ?es" by (metis restrict_nonempty) also have "... = ?a \ ?ds \ ?c \ ?es \ ?b \ ?ds \ ?c \ ?es \ ?c \ ?ds \ ?c \ ?es \ ?d \ ?ds \ ?c \ ?es" by (simp add: matrix_idempotent_semiring.mult_right_dist_sup) also have "... = ?b \ ?ds \ ?c \ ?es \ ?d \ ?ds \ ?c \ ?es" using 4 by (metis (no_types, lifting) times_disjoint matrix_idempotent_semiring.mult_left_zero restrict_star matrix_bounded_semilattice_sup_bot.sup_monoid.add_0_right matrix_bounded_semilattice_sup_bot.sup_monoid.add_0_left) finally show ?thesis . qed have 9: "?t\g\?t \ ?fs = ?b \ ?fs \ ?d \ ?fs" proof - have "?t\g\?t \ ?fs = (?a \ ?b \ ?c \ ?d) \ ?fs" by (metis restrict_nonempty) also have "... = ?a \ ?fs \ ?b \ ?fs \ ?c \ ?fs \ ?d \ ?fs" by (simp add: matrix_idempotent_semiring.mult_right_dist_sup) also have "... = ?b \ ?fs \ ?d \ ?fs" using 4 by (metis (no_types, lifting) times_disjoint restrict_star matrix_bounded_semilattice_sup_bot.sup_monoid.add_0_right matrix_bounded_semilattice_sup_bot.sup_monoid.add_0_left) finally show ?thesis . qed have "?t\?o\?t \ ?t\g\?t \ star_matrix' ?t g = ?t\?o\?t \ ?t\g\?t \ (?es \ ?as \ ?b \ ?fs \ ?ds \ ?c \ ?es \ ?fs)" by (metis star_matrix'.simps(2)) also have "... = ?t\?o\?t \ ?t\g\?t \ ?es \ ?t\g\?t \ ?as \ ?b \ ?fs \ ?t\g\?t \ ?ds \ ?c \ ?es \ ?t\g\?t \ ?fs" by (simp add: matrix_idempotent_semiring.mult_left_dist_sup matrix_monoid.mult_assoc matrix_semilattice_sup.sup_assoc) also have "... = ?r\?o\?r \ s\?o\s \ ?a \ ?es \ ?c \ ?es \ ?a \ ?as \ ?b \ ?fs \ ?c \ ?as \ ?b \ ?fs \ ?b \ ?ds \ ?c \ ?es \ ?d \ ?ds \ ?c \ ?es \ ?b \ ?fs \ ?d \ ?fs" using 5 6 7 8 9 by (simp add: matrix_semilattice_sup.sup.assoc) also have "... = (?r\?o\?r \ (?a \ ?es \ ?b \ ?ds \ ?c \ ?es)) \ (?b \ ?fs \ ?a \ ?as \ ?b \ ?fs) \ (?c \ ?es \ ?d \ ?ds \ ?c \ ?es) \ (s\?o\s \ (?d \ ?fs \ ?c \ ?as \ ?b \ ?fs))" by (simp only: matrix_semilattice_sup.sup_assoc matrix_semilattice_sup.sup_commute matrix_semilattice_sup.sup_left_commute) also have "... = (?r\?o\?r \ (?a \ ?es \ ?b \ ?ds \ ?c \ ?es)) \ (?r\?o\?r \ ?b \ ?fs \ ?a \ ?as \ ?b \ ?fs) \ (s\?o\s \ ?c \ ?es \ ?d \ ?ds \ ?c \ ?es) \ (s\?o\s \ (?d \ ?fs \ ?c \ ?as \ ?b \ ?fs))" by (simp add: restrict_one_left_unit) also have "... = (?r\?o\?r \ ?e \ ?es) \ ((?r\?o\?r \ ?a \ ?as) \ ?b \ ?fs) \ ((s\?o\s \ ?d \ ?ds) \ ?c \ ?es) \ (s\?o\s \ ?f \ ?fs)" by (simp add: matrix_idempotent_semiring.mult_right_dist_sup) also have "... = (?r\?o\?r \ ?e \ ?es) \ ((?r\?o\?r \ ?a \ ?as) \ ?b \ ?fs) \ ((s\?o\s \ ?d \ ?ds) \ ?c \ ?es) \ ?fs" using 1 2 3 by (metis distinct.simps(2)) also have "... = (?r\?o\?r \ ?e \ ?es) \ ((?r\?o\?r \ ?a \ ?as) \ ?b \ ?fs) \ (?ds \ ?c \ ?es) \ ?fs" using 1 2 by (metis (no_types, lifting) distinct.simps(2) restrict_superlist) also have "... = ?es \ ((?r\?o\?r \ ?a \ ?as) \ ?b \ ?fs) \ (?ds \ ?c \ ?es) \ ?fs" using 3 by (metis restrict_star_unfold) also have "... = ?es \ ?as \ ?b \ ?fs \ ?ds \ ?c \ ?es \ ?fs" by (metis (no_types, lifting) restrict_one_left_unit restrict_star_unfold restrict_times) also have "... = star_matrix' ?t g" by (metis star_matrix'.simps(2)) finally show "?t\?o\?t \ ?t\g\?t \ star_matrix' ?t g = star_matrix' ?t g" . qed qed thus "?o \ y \ y\<^sup>\ \ y\<^sup>\" by (simp add: enum_distinct restrict_all) next fix x y z :: "('a,'b) square" let ?e = "enum_class.enum::'a list" have "\g h :: ('a,'b) square . \zs . distinct ?e \ distinct zs \ (?e\g\?e \ ?e\h\zs \ ?e\h\zs \ star_matrix' ?e g \ ?e\h\zs \ ?e\h\zs)" proof (induct rule: list.induct) case Nil thus ?case by (simp add: restrict_empty_left) case (Cons k s) let ?t = "k#s" assume 1: "\g h :: ('a,'b) square . \zs . distinct s \ distinct zs \ (s\g\s \ s\h\zs \ s\h\zs \ star_matrix' s g \ s\h\zs \ s\h\zs)" show "\g h :: ('a,'b) square . \zs . distinct ?t \ distinct zs \ (?t\g\?t \ ?t\h\zs \ ?t\h\zs \ star_matrix' ?t g \ ?t\h\zs \ ?t\h\zs)" proof (intro allI) fix g h :: "('a,'b) square" fix zs :: "'a list" show "distinct ?t \ distinct zs \ (?t\g\?t \ ?t\h\zs \ ?t\h\zs \ star_matrix' ?t g \ ?t\h\zs \ ?t\h\zs)" proof (cases zs) case Nil thus ?thesis by (metis restrict_empty_right restrict_star restrict_times) next case (Cons y ys) assume 2: "zs = y#ys" show "distinct ?t \ distinct zs \ (?t\g\?t \ ?t\h\zs \ ?t\h\zs \ star_matrix' ?t g \ ?t\h\zs \ ?t\h\zs)" proof (intro impI) let ?y = "[y]" assume 3: "distinct ?t \ distinct zs" hence 4: "distinct s \ distinct ys \ \ List.member s k \ \ List.member ys y" using 2 by (simp add: List.member_def) let ?r = "[k]" let ?a = "?r\g\?r" let ?b = "?r\g\s" let ?c = "s\g\?r" let ?d = "s\g\s" let ?as = "?r\star o ?a\?r" let ?ds = "star_matrix' s ?d" let ?e = "?a \ ?b \ ?ds \ ?c" let ?es = "?r\star o ?e\?r" let ?f = "?d \ ?c \ ?as \ ?b" let ?fs = "star_matrix' s ?f" let ?ha = "?r\h\?y" let ?hb = "?r\h\ys" let ?hc = "s\h\?y" let ?hd = "s\h\ys" assume "?t\g\?t \ ?t\h\zs \ ?t\h\zs" hence 5: "?a \ ?ha \ ?b \ ?hc \ ?ha \ ?a \ ?hb \ ?b \ ?hd \ ?hb \ ?c \ ?ha \ ?d \ ?hc \ ?hc \ ?c \ ?hb \ ?d \ ?hd \ ?hd" using 2 3 4 by (simp add: restrict_nonempty_product_less_eq) have 6: "s\?ds\s = ?ds \ s\?fs\s = ?fs" by (simp add: restrict_star) hence 7: "?r\?e\?r = ?e \ s\?f\s = ?f" by (metis (no_types, lifting) restrict_one_left_unit restrict_sup restrict_times) have 8: "disjoint s ?r \ disjoint ?r s" using 3 by (simp add: in_set_member member_rec(1) member_rec(2)) have 9: "?es \ ?t\h\zs = ?es \ ?ha \ ?es \ ?hb" proof - have "?es \ ?t\h\zs = ?es \ (?ha \ ?hb \ ?hc \ ?hd)" using 2 by (metis restrict_nonempty) also have "... = ?es \ ?ha \ ?es \ ?hb \ ?es \ ?hc \ ?es \ ?hd" by (simp add: matrix_idempotent_semiring.mult_left_dist_sup) also have "... = ?es \ ?ha \ ?es \ ?hb" using 8 by (simp add: times_disjoint) finally show ?thesis . qed have 10: "?as \ ?b \ ?fs \ ?t\h\zs = ?as \ ?b \ ?fs \ ?hc \ ?as \ ?b \ ?fs \ ?hd" proof - have "?as \ ?b \ ?fs \ ?t\h\zs = ?as \ ?b \ ?fs \ (?ha \ ?hb \ ?hc \ ?hd)" using 2 by (metis restrict_nonempty) also have "... = ?as \ ?b \ ?fs \ ?ha \ ?as \ ?b \ ?fs \ ?hb \ ?as \ ?b \ ?fs \ ?hc \ ?as \ ?b \ ?fs \ ?hd" by (simp add: matrix_idempotent_semiring.mult_left_dist_sup) also have "... = ?as \ ?b \ (?fs \ ?ha) \ ?as \ ?b \ (?fs \ ?hb) \ ?as \ ?b \ ?fs \ ?hc \ ?as \ ?b \ ?fs \ ?hd" by (simp add: matrix_monoid.mult_assoc) also have "... = ?as \ ?b \ mbot \ ?as \ ?b \ mbot \ ?as \ ?b \ ?fs \ ?hc \ ?as \ ?b \ ?fs \ ?hd" using 6 8 by (metis (no_types) times_disjoint) also have "... = ?as \ ?b \ ?fs \ ?hc \ ?as \ ?b \ ?fs \ ?hd" by simp finally show ?thesis . qed have 11: "?ds \ ?c \ ?es \ ?t\h\zs = ?ds \ ?c \ ?es \ ?ha \ ?ds \ ?c \ ?es \ ?hb" proof - have "?ds \ ?c \ ?es \ ?t\h\zs = ?ds \ ?c \ ?es \ (?ha \ ?hb \ ?hc \ ?hd)" using 2 by (metis restrict_nonempty) also have "... = ?ds \ ?c \ ?es \ ?ha \ ?ds \ ?c \ ?es \ ?hb \ ?ds \ ?c \ ?es \ ?hc \ ?ds \ ?c \ ?es \ ?hd" by (simp add: matrix_idempotent_semiring.mult_left_dist_sup) also have "... = ?ds \ ?c \ ?es \ ?ha \ ?ds \ ?c \ ?es \ ?hb \ ?ds \ ?c \ (?es \ ?hc) \ ?ds \ ?c \ (?es \ ?hd)" by (simp add: matrix_monoid.mult_assoc) also have "... = ?ds \ ?c \ ?es \ ?ha \ ?ds \ ?c \ ?es \ ?hb \ ?ds \ ?c \ mbot \ ?ds \ ?c \ mbot" using 8 by (metis times_disjoint) also have "... = ?ds \ ?c \ ?es \ ?ha \ ?ds \ ?c \ ?es \ ?hb" by simp finally show ?thesis . qed have 12: "?fs \ ?t\h\zs = ?fs \ ?hc \ ?fs \ ?hd" proof - have "?fs \ ?t\h\zs = ?fs \ (?ha \ ?hb \ ?hc \ ?hd)" using 2 by (metis restrict_nonempty) also have "... = ?fs \ ?ha \ ?fs \ ?hb \ ?fs \ ?hc \ ?fs \ ?hd" by (simp add: matrix_idempotent_semiring.mult_left_dist_sup) also have "... = ?fs \ ?hc \ ?fs \ ?hd" using 6 8 by (metis (no_types) times_disjoint matrix_bounded_semilattice_sup_bot.sup_monoid.add_0_left) finally show ?thesis . qed have 13: "?es \ ?ha \ ?ha" proof - have "?b \ ?ds \ ?c \ ?ha \ ?b \ ?ds \ ?hc" using 5 by (simp add: matrix_idempotent_semiring.mult_right_isotone matrix_monoid.mult_assoc) also have "... \ ?b \ ?hc" using 1 3 5 by (simp add: matrix_idempotent_semiring.mult_right_isotone matrix_monoid.mult_assoc member_rec(2) restrict_sublist) also have "... \ ?ha" using 5 by simp finally have "?e \ ?ha \ ?ha" using 5 by (simp add: matrix_idempotent_semiring.mult_right_dist_sup) thus ?thesis using 7 by (simp add: restrict_star_left_induct) qed have 14: "?es \ ?hb \ ?hb" proof - have "?b \ ?ds \ ?c \ ?hb \ ?b \ ?ds \ ?hd" using 5 by (simp add: matrix_idempotent_semiring.mult_right_isotone matrix_monoid.mult_assoc) also have "... \ ?b \ ?hd" using 1 4 5 by (simp add: matrix_idempotent_semiring.mult_right_isotone matrix_monoid.mult_assoc restrict_sublist) also have "... \ ?hb" using 5 by simp finally have "?e \ ?hb \ ?hb" using 5 by (simp add: matrix_idempotent_semiring.mult_right_dist_sup) thus ?thesis using 4 7 by (simp add: restrict_star_left_induct) qed have 15: "?fs \ ?hc \ ?hc" proof - have "?c \ ?as \ ?b \ ?hc \ ?c \ ?as \ ?ha" using 5 by (simp add: matrix_idempotent_semiring.mult_right_isotone matrix_monoid.mult_assoc) also have "... \ ?c \ ?ha" using 5 by (simp add: matrix_idempotent_semiring.mult_right_isotone matrix_monoid.mult_assoc restrict_star_left_induct restrict_sublist) also have "... \ ?hc" using 5 by simp finally have "?f \ ?hc \ ?hc" using 5 by (simp add: matrix_idempotent_semiring.mult_right_dist_sup) thus ?thesis using 1 3 7 by simp qed have 16: "?fs \ ?hd \ ?hd" proof - have "?c \ ?as \ ?b \ ?hd \ ?c \ ?as \ ?hb" using 5 by (simp add: matrix_idempotent_semiring.mult_right_isotone matrix_monoid.mult_assoc) also have "... \ ?c \ ?hb" using 4 5 by (simp add: matrix_idempotent_semiring.mult_right_isotone matrix_monoid.mult_assoc restrict_star_left_induct restrict_sublist) also have "... \ ?hd" using 5 by simp finally have "?f \ ?hd \ ?hd" using 5 by (simp add: matrix_idempotent_semiring.mult_right_dist_sup) thus ?thesis using 1 4 7 by simp qed have 17: "?as \ ?b \ ?fs \ ?hc \ ?ha" proof - have "?as \ ?b \ ?fs \ ?hc \ ?as \ ?b \ ?hc" using 15 by (simp add: matrix_idempotent_semiring.mult_right_isotone matrix_monoid.mult_assoc) also have "... \ ?as \ ?ha" using 5 by (simp add: matrix_idempotent_semiring.mult_right_isotone matrix_monoid.mult_assoc) also have "... \ ?ha" using 5 by (simp add: restrict_star_left_induct restrict_sublist) finally show ?thesis . qed have 18: "?as \ ?b \ ?fs \ ?hd \ ?hb" proof - have "?as \ ?b \ ?fs \ ?hd \ ?as \ ?b \ ?hd" using 16 by (simp add: matrix_idempotent_semiring.mult_right_isotone matrix_monoid.mult_assoc) also have "... \ ?as \ ?hb" using 5 by (simp add: matrix_idempotent_semiring.mult_right_isotone matrix_monoid.mult_assoc) also have "... \ ?hb" using 4 5 by (simp add: restrict_star_left_induct restrict_sublist) finally show ?thesis . qed have 19: "?ds \ ?c \ ?es \ ?ha \ ?hc" proof - have "?ds \ ?c \ ?es \ ?ha \ ?ds \ ?c \ ?ha" using 13 by (simp add: matrix_idempotent_semiring.mult_right_isotone matrix_monoid.mult_assoc) also have "... \ ?ds \ ?hc" using 5 by (simp add: matrix_idempotent_semiring.mult_right_isotone matrix_monoid.mult_assoc) also have "... \ ?hc" using 1 3 5 by (simp add: restrict_sublist) finally show ?thesis . qed have 20: "?ds \ ?c \ ?es \ ?hb \ ?hd" proof - have "?ds \ ?c \ ?es \ ?hb \ ?ds \ ?c \ ?hb" using 14 by (simp add: matrix_idempotent_semiring.mult_right_isotone matrix_monoid.mult_assoc) also have "... \ ?ds \ ?hd" using 5 by (simp add: matrix_idempotent_semiring.mult_right_isotone matrix_monoid.mult_assoc) also have "... \ ?hd" using 1 4 5 by (simp add: restrict_sublist) finally show ?thesis . qed have 21: "?es \ ?ha \ ?as \ ?b \ ?fs \ ?hc \ ?ha" using 13 17 matrix_semilattice_sup.le_supI by blast have 22: "?es \ ?hb \ ?as \ ?b \ ?fs \ ?hd \ ?hb" using 14 18 matrix_semilattice_sup.le_supI by blast have 23: "?ds \ ?c \ ?es \ ?ha \ ?fs \ ?hc \ ?hc" using 15 19 matrix_semilattice_sup.le_supI by blast have 24: "?ds \ ?c \ ?es \ ?hb \ ?fs \ ?hd \ ?hd" using 16 20 matrix_semilattice_sup.le_supI by blast have "star_matrix' ?t g \ ?t\h\zs = (?es \ ?as \ ?b \ ?fs \ ?ds \ ?c \ ?es \ ?fs) \ ?t\h\zs" by (metis star_matrix'.simps(2)) also have "... = ?es \ ?t\h\zs \ ?as \ ?b \ ?fs \ ?t\h\zs \ ?ds \ ?c \ ?es \ ?t\h\zs \ ?fs \ ?t\h\zs" by (simp add: matrix_idempotent_semiring.mult_right_dist_sup) also have "... = ?es \ ?ha \ ?es \ ?hb \ ?as \ ?b \ ?fs \ ?hc \ ?as \ ?b \ ?fs \ ?hd \ ?ds \ ?c \ ?es \ ?ha \ ?ds \ ?c \ ?es \ ?hb \ ?fs \ ?hc \ ?fs \ ?hd" using 9 10 11 12 by (simp only: matrix_semilattice_sup.sup_assoc) also have "... = (?es \ ?ha \ ?as \ ?b \ ?fs \ ?hc) \ (?es \ ?hb \ ?as \ ?b \ ?fs \ ?hd) \ (?ds \ ?c \ ?es \ ?ha \ ?fs \ ?hc) \ (?ds \ ?c \ ?es \ ?hb \ ?fs \ ?hd)" by (simp only: matrix_semilattice_sup.sup_assoc matrix_semilattice_sup.sup_commute matrix_semilattice_sup.sup_left_commute) also have "... \ ?ha \ ?hb \ ?hc \ ?hd" using 21 22 23 24 matrix_semilattice_sup.sup.mono by blast also have "... = ?t\h\zs" using 2 by (metis restrict_nonempty) finally show "star_matrix' ?t g \ ?t\h\zs \ ?t\h\zs" . qed qed qed qed hence "\zs . distinct zs \ (y \ ?e\x\zs \ ?e\x\zs \ y\<^sup>\ \ ?e\x\zs \ ?e\x\zs)" by (simp add: enum_distinct restrict_all) thus "y \ x \ x \ y\<^sup>\ \ x \ x" by (metis restrict_all enum_distinct) next fix x y z :: "('a,'b) square" let ?e = "enum_class.enum::'a list" have "\g h :: ('a,'b) square . \zs . distinct ?e \ distinct zs \ (zs\h\?e \ ?e\g\?e \ zs\h\?e \ zs\h\?e \ star_matrix' ?e g \ zs\h\?e)" proof (induct rule:list.induct) case Nil thus ?case by (simp add: restrict_empty_left) case (Cons k s) let ?t = "k#s" assume 1: "\g h :: ('a,'b) square . \zs . distinct s \ distinct zs \ (zs\h\s \ s\g\s \ zs\h\s \ zs\h\s \ star_matrix' s g \ zs\h\s)" show "\g h :: ('a,'b) square . \zs . distinct ?t \ distinct zs \ (zs\h\?t \ ?t\g\?t \ zs\h\?t \ zs\h\?t \ star_matrix' ?t g \ zs\h\?t)" proof (intro allI) fix g h :: "('a,'b) square" fix zs :: "'a list" show "distinct ?t \ distinct zs \ (zs\h\?t \ ?t\g\?t \ zs\h\?t \ zs\h\?t \ star_matrix' ?t g \ zs\h\?t)" proof (cases zs) case Nil thus ?thesis by (metis restrict_empty_left restrict_star restrict_times) next case (Cons y ys) assume 2: "zs = y#ys" show "distinct ?t \ distinct zs \ (zs\h\?t \ ?t\g\?t \ zs\h\?t \ zs\h\?t \ star_matrix' ?t g \ zs\h\?t)" proof (intro impI) let ?y = "[y]" assume 3: "distinct ?t \ distinct zs" hence 4: "distinct s \ distinct ys \ \ List.member s k \ \ List.member ys y" using 2 by (simp add: List.member_def) let ?r = "[k]" let ?a = "?r\g\?r" let ?b = "?r\g\s" let ?c = "s\g\?r" let ?d = "s\g\s" let ?as = "?r\star o ?a\?r" let ?ds = "star_matrix' s ?d" let ?e = "?a \ ?b \ ?ds \ ?c" let ?es = "?r\star o ?e\?r" let ?f = "?d \ ?c \ ?as \ ?b" let ?fs = "star_matrix' s ?f" let ?ha = "?y\h\?r" let ?hb = "?y\h\s" let ?hc = "ys\h\?r" let ?hd = "ys\h\s" assume "zs\h\?t \ ?t\g\?t \ zs\h\?t" hence 5: "?ha \ ?a \ ?hb \ ?c \ ?ha \ ?ha \ ?b \ ?hb \ ?d \ ?hb \ ?hc \ ?a \ ?hd \ ?c \ ?hc \ ?hc \ ?b \ ?hd \ ?d \ ?hd" using 2 3 4 by (simp add: restrict_nonempty_product_less_eq) have 6: "s\?ds\s = ?ds \ s\?fs\s = ?fs" by (simp add: restrict_star) hence 7: "?r\?e\?r = ?e \ s\?f\s = ?f" by (metis (no_types, lifting) restrict_one_left_unit restrict_sup restrict_times) have 8: "disjoint s ?r \ disjoint ?r s" using 3 by (simp add: in_set_member member_rec) have 9: "zs\h\?t \ ?es = ?ha \ ?es \ ?hc \ ?es" proof - have "zs\h\?t \ ?es = (?ha \ ?hb \ ?hc \ ?hd) \ ?es" using 2 by (metis restrict_nonempty) also have "... = ?ha \ ?es \ ?hb \ ?es \ ?hc \ ?es \ ?hd \ ?es" by (simp add: matrix_idempotent_semiring.mult_right_dist_sup) also have "... = ?ha \ ?es \ ?hc \ ?es" using 8 by (simp add: times_disjoint) finally show ?thesis . qed have 10: "zs\h\?t \ ?as \ ?b \ ?fs = ?ha \ ?as \ ?b \ ?fs \ ?hc \ ?as \ ?b \ ?fs" proof - have "zs\h\?t \ ?as \ ?b \ ?fs = (?ha \ ?hb \ ?hc \ ?hd) \ ?as \ ?b \ ?fs" using 2 by (metis restrict_nonempty) also have "... = ?ha \ ?as \ ?b \ ?fs \ ?hb \ ?as \ ?b \ ?fs \ ?hc \ ?as \ ?b \ ?fs \ ?hd \ ?as \ ?b \ ?fs" by (simp add: matrix_idempotent_semiring.mult_right_dist_sup) also have "... = ?ha \ ?as \ ?b \ ?fs \ mbot \ ?b \ ?fs \ ?hc \ ?as \ ?b \ ?fs \ mbot \ ?b \ ?fs" using 8 by (metis (no_types) times_disjoint) also have "... = ?ha \ ?as \ ?b \ ?fs \ ?hc \ ?as \ ?b \ ?fs" by simp finally show ?thesis . qed have 11: "zs\h\?t \ ?ds \ ?c \ ?es = ?hb \ ?ds \ ?c \ ?es \ ?hd \ ?ds \ ?c \ ?es" proof - have "zs\h\?t \ ?ds \ ?c \ ?es = (?ha \ ?hb \ ?hc \ ?hd) \ ?ds \ ?c \ ?es" using 2 by (metis restrict_nonempty) also have "... = ?ha \ ?ds \ ?c \ ?es \ ?hb \ ?ds \ ?c \ ?es \ ?hc \ ?ds \ ?c \ ?es \ ?hd \ ?ds \ ?c \ ?es" by (simp add: matrix_idempotent_semiring.mult_right_dist_sup) also have "... = mbot \ ?c \ ?es \ ?hb \ ?ds \ ?c \ ?es \ mbot \ ?c \ ?es \ ?hd \ ?ds \ ?c \ ?es" using 6 8 by (metis (no_types) times_disjoint) also have "... = ?hb \ ?ds \ ?c \ ?es \ ?hd \ ?ds \ ?c \ ?es" by simp finally show ?thesis . qed have 12: "zs\h\?t \ ?fs = ?hb \ ?fs \ ?hd \ ?fs" proof - have "zs\h\?t \ ?fs = (?ha \ ?hb \ ?hc \ ?hd) \ ?fs" using 2 by (metis restrict_nonempty) also have "... = ?ha \ ?fs \ ?hb \ ?fs \ ?hc \ ?fs \ ?hd \ ?fs" by (simp add: matrix_idempotent_semiring.mult_right_dist_sup) also have "... = ?hb \ ?fs \ ?hd \ ?fs" using 6 8 by (metis (no_types) times_disjoint matrix_bounded_semilattice_sup_bot.sup_monoid.add_0_right matrix_bounded_semilattice_sup_bot.sup_monoid.add_0_left) finally show ?thesis . qed have 13: "?ha \ ?es \ ?ha" proof - have "?ha \ ?b \ ?ds \ ?c \ ?hb \ ?ds \ ?c" using 5 by (simp add: matrix_idempotent_semiring.mult_left_isotone) also have "... \ ?hb \ ?c" using 1 4 5 by (simp add: matrix_idempotent_semiring.mult_left_isotone restrict_sublist) also have "... \ ?ha" using 5 by simp finally have "?ha \ ?e \ ?ha" using 5 by (simp add: matrix_idempotent_semiring.mult_left_dist_sup matrix_monoid.mult_assoc) thus ?thesis using 7 by (simp add: restrict_star_right_induct) qed have 14: "?hb \ ?fs \ ?hb" proof - have "?hb \ ?c \ ?as \ ?b \ ?ha \ ?as \ ?b" using 5 by (metis matrix_semilattice_sup.le_supE matrix_idempotent_semiring.mult_left_isotone) also have "... \ ?ha \ ?b" using 5 by (simp add: matrix_idempotent_semiring.mult_left_isotone restrict_star_right_induct restrict_sublist) also have "... \ ?hb" using 5 by simp finally have "?hb \ ?f \ ?hb" using 5 by (simp add: matrix_idempotent_semiring.mult_left_dist_sup matrix_monoid.mult_assoc) thus ?thesis using 1 3 7 by simp qed have 15: "?hc \ ?es \ ?hc" proof - have "?hc \ ?b \ ?ds \ ?c \ ?hd \ ?ds \ ?c" using 5 by (simp add: matrix_idempotent_semiring.mult_left_isotone) also have "... \ ?hd \ ?c" using 1 4 5 by (simp add: matrix_idempotent_semiring.mult_left_isotone restrict_sublist) also have "... \ ?hc" using 5 by simp finally have "?hc \ ?e \ ?hc" using 5 by (simp add: matrix_idempotent_semiring.mult_left_dist_sup matrix_monoid.mult_assoc) thus ?thesis using 4 7 by (simp add: restrict_star_right_induct) qed have 16: "?hd \ ?fs \ ?hd" proof - have "?hd \ ?c \ ?as \ ?b \ ?hc \ ?as \ ?b" using 5 by (simp add: matrix_idempotent_semiring.mult_left_isotone) also have "... \ ?hc \ ?b" using 4 5 by (simp add: matrix_idempotent_semiring.mult_left_isotone restrict_star_right_induct restrict_sublist) also have "... \ ?hd" using 5 by simp finally have "?hd \ ?f \ ?hd" using 5 by (simp add: matrix_idempotent_semiring.mult_left_dist_sup matrix_monoid.mult_assoc) thus ?thesis using 1 4 7 by simp qed have 17: "?hb \ ?ds \ ?c \ ?es \ ?ha" proof - have "?hb \ ?ds \ ?c \ ?es \ ?hb \ ?c \ ?es" using 1 4 5 by (simp add: matrix_idempotent_semiring.mult_left_isotone restrict_sublist) also have "... \ ?ha \ ?es" using 5 by (simp add: matrix_idempotent_semiring.mult_left_isotone) also have "... \ ?ha" using 13 by simp finally show ?thesis . qed have 18: "?ha \ ?as \ ?b \ ?fs \ ?hb" proof - have "?ha \ ?as \ ?b \ ?fs \ ?ha \ ?b \ ?fs" using 5 by (simp add: matrix_idempotent_semiring.mult_left_isotone restrict_star_right_induct restrict_sublist) also have "... \ ?hb \ ?fs" using 5 by (simp add: matrix_idempotent_semiring.mult_left_isotone) also have "... \ ?hb" using 14 by simp finally show ?thesis by simp qed have 19: "?hd \ ?ds \ ?c \ ?es \ ?hc" proof - have "?hd \ ?ds \ ?c \ ?es \ ?hd \ ?c \ ?es" using 1 4 5 by (simp add: matrix_idempotent_semiring.mult_left_isotone restrict_sublist) also have "... \ ?hc \ ?es" using 5 by (simp add: matrix_idempotent_semiring.mult_left_isotone) also have "... \ ?hc" using 15 by simp finally show ?thesis by simp qed have 20: "?hc \ ?as \ ?b \ ?fs \ ?hd" proof - have "?hc \ ?as \ ?b \ ?fs \ ?hc \ ?b \ ?fs" using 4 5 by (simp add: matrix_idempotent_semiring.mult_left_isotone restrict_star_right_induct restrict_sublist) also have "... \ ?hd \ ?fs" using 5 by (simp add: matrix_idempotent_semiring.mult_left_isotone) also have "... \ ?hd" using 16 by simp finally show ?thesis by simp qed have 21: "?ha \ ?es \ ?hb \ ?ds \ ?c \ ?es \ ?ha" using 13 17 matrix_semilattice_sup.le_supI by blast have 22: "?ha \ ?as \ ?b \ ?fs \ ?hb \ ?fs \ ?hb" using 14 18 matrix_semilattice_sup.le_supI by blast have 23: "?hc \ ?es \ ?hd \ ?ds \ ?c \ ?es \ ?hc" using 15 19 matrix_semilattice_sup.le_supI by blast have 24: "?hc \ ?as \ ?b \ ?fs \ ?hd \ ?fs \ ?hd" using 16 20 matrix_semilattice_sup.le_supI by blast have "zs\h\?t \ star_matrix' ?t g = zs\h\?t \ (?es \ ?as \ ?b \ ?fs \ ?ds \ ?c \ ?es \ ?fs)" by (metis star_matrix'.simps(2)) also have "... = zs\h\?t \ ?es \ zs\h\?t \ ?as \ ?b \ ?fs \ zs\h\?t \ ?ds \ ?c \ ?es \ zs\h\?t \ ?fs" by (simp add: matrix_idempotent_semiring.mult_left_dist_sup matrix_monoid.mult_assoc) also have "... = ?ha \ ?es \ ?hc \ ?es \ ?ha \ ?as \ ?b \ ?fs \ ?hc \ ?as \ ?b \ ?fs \ ?hb \ ?ds \ ?c \ ?es \ ?hd \ ?ds \ ?c \ ?es \ ?hb \ ?fs \ ?hd \ ?fs" using 9 10 11 12 by (simp add: matrix_semilattice_sup.sup_assoc) also have "... = (?ha \ ?es \ ?hb \ ?ds \ ?c \ ?es) \ (?ha \ ?as \ ?b \ ?fs \ ?hb \ ?fs) \ (?hc \ ?es \ ?hd \ ?ds \ ?c \ ?es) \ (?hc \ ?as \ ?b \ ?fs \ ?hd \ ?fs)" using 9 10 11 12 by (simp only: matrix_semilattice_sup.sup_assoc matrix_semilattice_sup.sup_commute matrix_semilattice_sup.sup_left_commute) also have "... \ ?ha \ ?hb \ ?hc \ ?hd" using 21 22 23 24 matrix_semilattice_sup.sup.mono by blast also have "... = zs\h\?t" using 2 by (metis restrict_nonempty) finally show "zs\h\?t \ star_matrix' ?t g \ zs\h\?t" . qed qed qed qed hence "\zs . distinct zs \ (zs\x\?e \ y \ zs\x\?e \ zs\x\?e \ y\<^sup>\ \ zs\x\?e)" by (simp add: enum_distinct restrict_all) thus "x \ y \ x \ x \ y\<^sup>\ \ x" by (metis restrict_all enum_distinct) qed subsection \Matrices form a Stone-Kleene Relation Algebra\ text \ Matrices over Stone-Kleene relation algebras form a Stone-Kleene relation algebra. It remains to prove the axiom about the interaction of Kleene star and double complement. \ interpretation matrix_stone_kleene_relation_algebra: stone_kleene_relation_algebra where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix::('a::enum,'b::stone_kleene_relation_algebra) square" and top = top_matrix and uminus = uminus_matrix and one = one_matrix and times = times_matrix and conv = conv_matrix and star = star_matrix proof fix x :: "('a,'b) square" let ?e = "enum_class.enum::'a list" let ?o = "mone :: ('a,'b) square" show "\\(x\<^sup>\) = (\\x)\<^sup>\" proof (rule matrix_order.order_antisym) have "\g :: ('a,'b) square . distinct ?e \ \\(star_matrix' ?e (\\g)) = star_matrix' ?e (\\g)" proof (induct rule: list.induct) case Nil thus ?case by simp next case (Cons k s) let ?t = "k#s" assume 1: "\g :: ('a,'b) square . distinct s \ \\(star_matrix' s (\\g)) = star_matrix' s (\\g)" show "\g :: ('a,'b) square . distinct ?t \ \\(star_matrix' ?t (\\g)) = star_matrix' ?t (\\g)" proof (rule allI, rule impI) fix g :: "('a,'b) square" assume 2: "distinct ?t" let ?r = "[k]" let ?a = "?r\\\g\?r" let ?b = "?r\\\g\s" let ?c = "s\\\g\?r" let ?d = "s\\\g\s" let ?as = "?r\star o ?a\?r" let ?ds = "star_matrix' s ?d" let ?e = "?a \ ?b \ ?ds \ ?c" let ?es = "?r\star o ?e\?r" let ?f = "?d \ ?c \ ?as \ ?b" let ?fs = "star_matrix' s ?f" have "s\?ds\s = ?ds \ s\?fs\s = ?fs" by (simp add: restrict_star) have 3: "\\?a = ?a \ \\?b = ?b \ \\?c = ?c \ \\?d = ?d" by (metis matrix_p_algebra.regular_closed_p restrict_pp) hence 4: "\\?as = ?as" by (metis pp_star_commute restrict_pp) hence "\\?f = ?f" using 3 by (metis matrix_stone_algebra.regular_closed_sup matrix_stone_relation_algebra.regular_mult_closed) hence 5: "\\?fs = ?fs" using 1 2 by (metis distinct.simps(2)) have 6: "\\?ds = ?ds" using 1 2 by (simp add: restrict_pp) hence "\\?e = ?e" using 3 by (metis matrix_stone_algebra.regular_closed_sup matrix_stone_relation_algebra.regular_mult_closed) hence 7: "\\?es = ?es" by (metis pp_star_commute restrict_pp) have "\\(star_matrix' ?t (\\g)) = \\(?es \ ?as \ ?b \ ?fs \ ?ds \ ?c \ ?es \ ?fs)" by (metis star_matrix'.simps(2)) also have "... = \\?es \ \\?as \ \\?b \ \\?fs \ \\?ds \ \\?c \ \\?es \ \\?fs" by (simp add: matrix_stone_relation_algebra.pp_dist_comp) also have "... = ?es \ ?as \ ?b \ ?fs \ ?ds \ ?c \ ?es \ ?fs" using 3 4 5 6 7 by simp finally show "\\(star_matrix' ?t (\\g)) = star_matrix' ?t (\\g)" by (metis star_matrix'.simps(2)) qed qed hence "(\\x)\<^sup>\ = \\((\\x)\<^sup>\)" by (simp add: enum_distinct restrict_all) thus "\\(x\<^sup>\) \ (\\x)\<^sup>\" by (metis matrix_kleene_algebra.star.circ_isotone matrix_p_algebra.pp_increasing matrix_p_algebra.pp_isotone) next have "?o \ \\x \ \\(x\<^sup>\) \ \\(x\<^sup>\)" by (metis matrix_kleene_algebra.star_left_unfold_equal matrix_p_algebra.sup_pp_semi_commute matrix_stone_relation_algebra.pp_dist_comp) thus "(\\x)\<^sup>\ \ \\(x\<^sup>\)" using matrix_kleene_algebra.star_left_induct by fastforce qed qed +interpretation matrix_stone_kleene_relation_algebra_consistent: stone_kleene_relation_algebra_consistent where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix :: ('a::enum,'b::stone_kleene_relation_algebra_consistent) square" and top = top_matrix and uminus = uminus_matrix and one = one_matrix and times = times_matrix and conv = conv_matrix and star = star_matrix + .. + +interpretation matrix_stone_kleene_relation_algebra_tarski: stone_kleene_relation_algebra_tarski where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix :: ('a::enum,'b::stone_kleene_relation_algebra_tarski) square" and top = top_matrix and uminus = uminus_matrix and one = one_matrix and times = times_matrix and conv = conv_matrix and star = star_matrix + .. + +interpretation matrix_stone_kleene_relation_algebra_tarski_consistent: stone_kleene_relation_algebra_tarski_consistent where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix :: ('a::enum,'b::stone_kleene_relation_algebra_tarski_consistent) square" and top = top_matrix and uminus = uminus_matrix and one = one_matrix and times = times_matrix and conv = conv_matrix and star = star_matrix + .. + end diff --git a/thys/Stone_Relation_Algebras/Choose_Component.thy b/thys/Stone_Relation_Algebras/Choose_Component.thy new file mode 100644 --- /dev/null +++ b/thys/Stone_Relation_Algebras/Choose_Component.thy @@ -0,0 +1,97 @@ +(* Title: An Operation to Select Components + Author: Nicolas Robinson-O'Brien, Walter Guttmann + Maintainer: Walter Guttmann +*) + +section \An Operation to Select Components\ + +text \ +In this theory we axiomatise an operation to select components of a graph. +This is joint work with Nicolas Robinson-O'Brien. +\ + +theory Choose_Component + +imports + Relation_Algebras + +begin + +context stone_relation_algebra +begin + +text \ +A \vector_classes\ corresponds to one or more equivalence classes and a \unique_vector_class\ corresponds to a single equivalence class. +\ + +definition vector_classes :: "'a \ 'a \ bool" where "vector_classes x v \ regular x \ regular v \ equivalence x \ vector v \ x * v \ v \ v \ bot" +definition unique_vector_class :: "'a \ 'a \ bool" where "unique_vector_class x v \ vector_classes x v \ v * v\<^sup>T \ x" + +end + +text \ +We introduce the operation \choose_component\. +\begin{itemize} + \item Axiom \component_in_v\ expresses that the result of \choose_component\ is contained in the set of vertices, $v$, we are selecting from, ignoring the weights. + \item Axiom \component_is_vector\ states that the result of \choose_component\ is a vector. + \item Axiom \component_is_regular\ states that the result of \choose_component\ is regular. + \item Axiom \component_is_connected\ states that any two vertices from the result of \choose_component\ are connected in $e$. + \item Axiom \component_single\ states that the result of \choose_component\ is closed under being connected in $e$. + \item Finally, axiom \component_not_bot_when_v_bot_bot\ expresses that the operation \choose_component\ returns a non-empty component if the input satisfies the given criteria. +\end{itemize} +\ + +class choose_component = + fixes choose_component :: "'a \ 'a \ 'a" + +class choose_component_algebra = choose_component + stone_relation_algebra + + assumes component_is_vector: "vector (choose_component e v)" + assumes component_is_regular: "regular (choose_component e v)" + assumes component_in_v: "choose_component e v \ --v" + assumes component_is_connected: "choose_component e v * (choose_component e v)\<^sup>T \ e" + assumes component_single: "e * choose_component e v \ choose_component e v" + assumes component_not_bot_when_v_bot_bot: "vector_classes e v \ choose_component e v \ bot" +begin + +lemma component_single_eq: + assumes "equivalence x" + shows "choose_component x v = x * choose_component x v" +proof - + have "choose_component x v \ x * choose_component x v" + by (meson component_is_connected ex231c mult_isotone order_lesseq_imp) + thus ?thesis + by (simp add: component_single order.antisym) +qed + +end + +class choose_component_algebra_tarski = choose_component_algebra + stone_relation_algebra_tarski +begin + +definition "choose_component_point x \ choose_component 1 (--x)" + +lemma choose_component_point_point: + assumes "vector x" + and "x \ bot" + shows "point (choose_component_point x)" +proof (intro conjI) + show 1: "vector (choose_component_point x)" + by (simp add: choose_component_point_def component_is_vector) + show "injective (choose_component_point x)" + by (simp add: choose_component_point_def component_is_connected) + have "vector_classes 1 (--x)" + by (metis assms comp_inf.semiring.mult_zero_left coreflexive_symmetric inf.eq_refl mult_1_left pp_one regular_closed_p selection_closed_id vector_classes_def vector_complement_closed) + hence "choose_component_point x \ bot" + by (simp add: choose_component_point_def component_not_bot_when_v_bot_bot) + thus "surjective (choose_component_point x)" + using 1 choose_component_point_def component_is_regular tarski vector_mult_closed by fastforce +qed + +lemma choose_component_point_decreasing: + "choose_component_point x \ --x" + by (metis choose_component_point_def component_in_v regular_closed_p) + +end + +end + diff --git a/thys/Stone_Relation_Algebras/Matrix_Relation_Algebras.thy b/thys/Stone_Relation_Algebras/Matrix_Relation_Algebras.thy --- a/thys/Stone_Relation_Algebras/Matrix_Relation_Algebras.thy +++ b/thys/Stone_Relation_Algebras/Matrix_Relation_Algebras.thy @@ -1,613 +1,651 @@ (* Title: Matrix Relation Algebras Author: Walter Guttmann Maintainer: Walter Guttmann *) section \Matrix Relation Algebras\ text \ This theory gives matrix models of Stone relation algebras and more general structures. We consider only square matrices. The main result is that matrices over Stone relation algebras form a Stone relation algebra. We use the monoid structure underlying semilattices to provide finite sums, which are necessary for defining the composition of two matrices. See \<^cite>\"ArmstrongFosterStruthWeber2016" and "ArmstrongGomesStruthWeber2016"\ for similar liftings to matrices for semirings and relation algebras. A technical difference is that those theories are mostly based on semirings whereas our hierarchy is mostly based on lattices (and our semirings directly inherit from semilattices). Relation algebras have both a semiring and a lattice structure such that semiring addition and lattice join coincide. In particular, finite sums and finite suprema coincide. Isabelle/HOL has separate theories for semirings and lattices, based on separate addition and join operations and different operations for finite sums and finite suprema. Reusing results from both theories is beneficial for relation algebras, but not always easy to realise. \ theory Matrix_Relation_Algebras imports Relation_Algebras begin subsection \Finite Suprema\ text \ We consider finite suprema in idempotent semirings and Stone relation algebras. We mostly use the first of the following notations, which denotes the supremum of expressions \t(x)\ over all \x\ from the type of \x\. For finite types, this is implemented in Isabelle/HOL as the repeated application of binary suprema. \ syntax "_sum_sup_monoid" :: "idt \ 'a::bounded_semilattice_sup_bot \ 'a" ("(\\<^sub>_ _)" [0,10] 10) "_sum_sup_monoid_bounded" :: "idt \ 'b set \ 'a::bounded_semilattice_sup_bot \ 'a" ("(\\<^bsub>_\_\<^esub> _)" [0,51,10] 10) translations "\\<^sub>x t" => "XCONST sup_monoid.sum (\x . t) { x . CONST True }" "\\<^bsub>x\X\<^esub> t" => "XCONST sup_monoid.sum (\x . t) X" context idempotent_semiring begin text \ The following induction principles are useful for comparing two suprema. The first principle works because types are not empty. \ lemma one_sup_induct [case_names one sup]: fixes f g :: "'b::finite \ 'a" assumes one: "\i . P (f i) (g i)" and sup: "\j I . j \ I \ P (\\<^bsub>i\I\<^esub> f i) (\\<^bsub>i\I\<^esub> g i) \ P (f j \ (\\<^bsub>i\I\<^esub> f i)) (g j \ (\\<^bsub>i\I\<^esub> g i))" shows "P (\\<^sub>k f k) (\\<^sub>k g k)" proof - let ?X = "{ k::'b . True }" have "finite ?X" and "?X \ {}" by auto thus ?thesis proof (induct rule: finite_ne_induct) case (singleton i) thus ?case using one by simp next case (insert j I) thus ?case using sup by simp qed qed lemma bot_sup_induct [case_names bot sup]: fixes f g :: "'b::finite \ 'a" assumes bot: "P bot bot" and sup: "\j I . j \ I \ P (\\<^bsub>i\I\<^esub> f i) (\\<^bsub>i\I\<^esub> g i) \ P (f j \ (\\<^bsub>i\I\<^esub> f i)) (g j \ (\\<^bsub>i\I\<^esub> g i))" shows "P (\\<^sub>k f k) (\\<^sub>k g k)" apply (induct rule: one_sup_induct) using bot sup apply fastforce using sup by blast text \ Now many properties of finite suprema follow by simple applications of the above induction rules. In particular, we show distributivity of composition, isotonicity and the upper-bound property. \ lemma comp_right_dist_sum: fixes f :: "'b::finite \ 'a" shows "(\\<^sub>k f k * x) = (\\<^sub>k f k) * x" proof (induct rule: one_sup_induct) case one show ?case by simp next case (sup j I) thus ?case using mult_right_dist_sup by auto qed lemma comp_left_dist_sum: fixes f :: "'b::finite \ 'a" shows "(\\<^sub>k x * f k) = x * (\\<^sub>k f k)" proof (induct rule: one_sup_induct) case one show ?case by simp next case (sup j I) thus ?case by (simp add: mult_left_dist_sup) qed lemma leq_sum: fixes f g :: "'b::finite \ 'a" shows "(\k . f k \ g k) \ (\\<^sub>k f k) \ (\\<^sub>k g k)" proof (induct rule: one_sup_induct) case one thus ?case by simp next case (sup j I) thus ?case using sup_mono by blast qed lemma ub_sum: fixes f :: "'b::finite \ 'a" shows "f i \ (\\<^sub>k f k)" proof - have "i \ { k . True }" by simp thus "f i \ (\\<^sub>k f (k::'b))" by (metis finite_code sup_monoid.sum.insert sup_ge1 mk_disjoint_insert) qed lemma lub_sum: fixes f :: "'b::finite \ 'a" assumes "\k . f k \ x" shows "(\\<^sub>k f k) \ x" proof (induct rule: one_sup_induct) case one show ?case by (simp add: assms) next case (sup j I) thus ?case using assms le_supI by blast qed lemma lub_sum_iff: fixes f :: "'b::finite \ 'a" shows "(\k . f k \ x) \ (\\<^sub>k f k) \ x" using order.trans ub_sum lub_sum by blast +lemma sum_const: + "(\\<^sub>k::'b::finite f) = f" + by (metis lub_sum sup.cobounded1 sup_monoid.add_0_right sup_same_context ub_sum) + end context stone_relation_algebra begin text \ In Stone relation algebras, we can also show that converse, double complement and meet distribute over finite suprema. \ lemma conv_dist_sum: fixes f :: "'b::finite \ 'a" shows "(\\<^sub>k (f k)\<^sup>T) = (\\<^sub>k f k)\<^sup>T" proof (induct rule: one_sup_induct) case one show ?case by simp next case (sup j I) thus ?case by (simp add: conv_dist_sup) qed lemma pp_dist_sum: fixes f :: "'b::finite \ 'a" shows "(\\<^sub>k --f k) = --(\\<^sub>k f k)" proof (induct rule: one_sup_induct) case one show ?case by simp next case (sup j I) thus ?case by simp qed lemma inf_right_dist_sum: fixes f :: "'b::finite \ 'a" shows "(\\<^sub>k f k \ x) = (\\<^sub>k f k) \ x" by (rule comp_inf.comp_right_dist_sum) end subsection \Square Matrices\ text \ Because our semiring and relation algebra type classes only work for homogeneous relations, we only look at square matrices. \ type_synonym ('a,'b) square = "'a \ 'a \ 'b" text \ We use standard matrix operations. The Stone algebra structure is lifted componentwise. Composition is matrix multiplication using given composition and supremum operations. Its unit lifts given zero and one elements into an identity matrix. Converse is matrix transpose with an additional componentwise transpose. \ definition less_eq_matrix :: "('a,'b::ord) square \ ('a,'b) square \ bool" (infix "\" 50) where "f \ g = (\e . f e \ g e)" definition less_matrix :: "('a,'b::ord) square \ ('a,'b) square \ bool" (infix "\" 50) where "f \ g = (f \ g \ \ g \ f)" definition sup_matrix :: "('a,'b::sup) square \ ('a,'b) square \ ('a,'b) square" (infixl "\" 65) where "f \ g = (\e . f e \ g e)" definition inf_matrix :: "('a,'b::inf) square \ ('a,'b) square \ ('a,'b) square" (infixl "\" 67) where "f \ g = (\e . f e \ g e)" definition minus_matrix :: "('a,'b::{uminus,inf}) square \ ('a,'b) square \ ('a,'b) square" (infixl "\" 65) where "f \ g = (\e . f e \ -g e)" definition implies_matrix :: "('a,'b::implies) square \ ('a,'b) square \ ('a,'b) square" (infixl "\" 65) where "f \ g = (\e . f e \ g e)" definition times_matrix :: "('a,'b::{times,bounded_semilattice_sup_bot}) square \ ('a,'b) square \ ('a,'b) square" (infixl "\" 70) where "f \ g = (\(i,j) . \\<^sub>k f (i,k) * g (k,j))" definition uminus_matrix :: "('a,'b::uminus) square \ ('a,'b) square" ("\ _" [80] 80) where "\f = (\e . -f e)" definition conv_matrix :: "('a,'b::conv) square \ ('a,'b) square" ("_\<^sup>t" [100] 100) where "f\<^sup>t = (\(i,j) . (f (j,i))\<^sup>T)" definition bot_matrix :: "('a,'b::bot) square" ("mbot") where "mbot = (\e . bot)" definition top_matrix :: "('a,'b::top) square" ("mtop") where "mtop = (\e . top)" definition one_matrix :: "('a,'b::{one,bot}) square" ("mone") where "mone = (\(i,j) . if i = j then 1 else bot)" subsection \Stone Algebras\ text \ We first lift the Stone algebra structure. Because all operations are componentwise, this also works for infinite matrices. \ interpretation matrix_order: order where less_eq = less_eq_matrix and less = "less_matrix :: ('a,'b::order) square \ ('a,'b) square \ bool" apply unfold_locales apply (simp add: less_matrix_def) apply (simp add: less_eq_matrix_def) apply (meson less_eq_matrix_def order_trans) by (meson less_eq_matrix_def antisym ext) interpretation matrix_semilattice_sup: semilattice_sup where sup = sup_matrix and less_eq = less_eq_matrix and less = "less_matrix :: ('a,'b::semilattice_sup) square \ ('a,'b) square \ bool" apply unfold_locales apply (simp add: sup_matrix_def less_eq_matrix_def) apply (simp add: sup_matrix_def less_eq_matrix_def) by (simp add: sup_matrix_def less_eq_matrix_def) interpretation matrix_semilattice_inf: semilattice_inf where inf = inf_matrix and less_eq = less_eq_matrix and less = "less_matrix :: ('a,'b::semilattice_inf) square \ ('a,'b) square \ bool" apply unfold_locales apply (simp add: inf_matrix_def less_eq_matrix_def) apply (simp add: inf_matrix_def less_eq_matrix_def) by (simp add: inf_matrix_def less_eq_matrix_def) interpretation matrix_bounded_semilattice_sup_bot: bounded_semilattice_sup_bot where sup = sup_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix :: ('a,'b::bounded_semilattice_sup_bot) square" apply unfold_locales by (simp add: bot_matrix_def less_eq_matrix_def) interpretation matrix_bounded_semilattice_inf_top: bounded_semilattice_inf_top where inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and top = "top_matrix :: ('a,'b::bounded_semilattice_inf_top) square" apply unfold_locales by (simp add: less_eq_matrix_def top_matrix_def) interpretation matrix_lattice: lattice where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = "less_matrix :: ('a,'b::lattice) square \ ('a,'b) square \ bool" .. interpretation matrix_distrib_lattice: distrib_lattice where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = "less_matrix :: ('a,'b::distrib_lattice) square \ ('a,'b) square \ bool" apply unfold_locales by (simp add: sup_inf_distrib1 sup_matrix_def inf_matrix_def) interpretation matrix_bounded_lattice: bounded_lattice where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix :: ('a,'b::bounded_lattice) square" and top = top_matrix .. interpretation matrix_bounded_distrib_lattice: bounded_distrib_lattice where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix :: ('a,'b::bounded_distrib_lattice) square" and top = top_matrix .. interpretation matrix_p_algebra: p_algebra where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix :: ('a,'b::p_algebra) square" and top = top_matrix and uminus = uminus_matrix apply unfold_locales apply (unfold inf_matrix_def bot_matrix_def less_eq_matrix_def uminus_matrix_def) by (meson pseudo_complement) interpretation matrix_pd_algebra: pd_algebra where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix :: ('a,'b::pd_algebra) square" and top = top_matrix and uminus = uminus_matrix .. text \ In particular, matrices over Stone algebras form a Stone algebra. \ interpretation matrix_stone_algebra: stone_algebra where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix :: ('a,'b::stone_algebra) square" and top = top_matrix and uminus = uminus_matrix by unfold_locales (simp add: sup_matrix_def uminus_matrix_def top_matrix_def) interpretation matrix_heyting_stone_algebra: heyting_stone_algebra where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix :: ('a,'b::heyting_stone_algebra) square" and top = top_matrix and uminus = uminus_matrix and implies = implies_matrix apply unfold_locales apply (unfold inf_matrix_def sup_matrix_def bot_matrix_def top_matrix_def less_eq_matrix_def uminus_matrix_def implies_matrix_def) apply (simp add: implies_galois) apply (simp add: uminus_eq) by simp interpretation matrix_boolean_algebra: boolean_algebra where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix :: ('a,'b::boolean_algebra) square" and top = top_matrix and uminus = uminus_matrix and minus = minus_matrix apply unfold_locales apply simp apply (simp add: sup_matrix_def uminus_matrix_def top_matrix_def) by (simp add: inf_matrix_def uminus_matrix_def minus_matrix_def) subsection \Semirings\ text \ Next, we lift the semiring structure. Because of composition, this requires a restriction to finite matrices. \ interpretation matrix_monoid: monoid_mult where times = times_matrix and one = "one_matrix :: ('a::finite,'b::idempotent_semiring) square" proof fix f g h :: "('a,'b) square" show "(f \ g) \ h = f \ (g \ h)" proof (rule ext, rule prod_cases) fix i j have "((f \ g) \ h) (i,j) = (\\<^sub>l (f \ g) (i,l) * h (l,j))" by (simp add: times_matrix_def) also have "... = (\\<^sub>l (\\<^sub>k f (i,k) * g (k,l)) * h (l,j))" by (simp add: times_matrix_def) also have "... = (\\<^sub>l \\<^sub>k (f (i,k) * g (k,l)) * h (l,j))" by (metis (no_types) comp_right_dist_sum) also have "... = (\\<^sub>l \\<^sub>k f (i,k) * (g (k,l) * h (l,j)))" by (simp add: mult.assoc) also have "... = (\\<^sub>k \\<^sub>l f (i,k) * (g (k,l) * h (l,j)))" using sup_monoid.sum.swap by auto also have "... = (\\<^sub>k f (i,k) * (\\<^sub>l g (k,l) * h (l,j)))" by (metis (no_types) comp_left_dist_sum) also have "... = (\\<^sub>k f (i,k) * (g \ h) (k,j))" by (simp add: times_matrix_def) also have "... = (f \ (g \ h)) (i,j)" by (simp add: times_matrix_def) finally show "((f \ g) \ h) (i,j) = (f \ (g \ h)) (i,j)" . qed next fix f :: "('a,'b) square" show "mone \ f = f" proof (rule ext, rule prod_cases) fix i j have "(mone \ f) (i,j) = (\\<^sub>k mone (i,k) * f (k,j))" by (simp add: times_matrix_def) also have "... = (\\<^sub>k (if i = k then 1 else bot) * f (k,j))" by (simp add: one_matrix_def) also have "... = (\\<^sub>k if i = k then 1 * f (k,j) else bot * f (k,j))" by (metis (full_types, opaque_lifting)) also have "... = (\\<^sub>k if i = k then f (k,j) else bot)" by (meson mult_left_one mult_left_zero) also have "... = f (i,j)" by simp finally show "(mone \ f) (i,j) = f (i,j)" . qed next fix f :: "('a,'b) square" show "f \ mone = f" proof (rule ext, rule prod_cases) fix i j have "(f \ mone) (i,j) = (\\<^sub>k f (i,k) * mone (k,j))" by (simp add: times_matrix_def) also have "... = (\\<^sub>k f (i,k) * (if k = j then 1 else bot))" by (simp add: one_matrix_def) also have "... = (\\<^sub>k if k = j then f (i,k) * 1 else f (i,k) * bot)" by (metis (full_types, opaque_lifting)) also have "... = (\\<^sub>k if k = j then f (i,k) else bot)" by (meson mult.right_neutral semiring.mult_zero_right) also have "... = f (i,j)" by simp finally show "(f \ mone) (i,j) = f (i,j)" . qed qed interpretation matrix_idempotent_semiring: idempotent_semiring where sup = sup_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix :: ('a::finite,'b::idempotent_semiring) square" and one = one_matrix and times = times_matrix proof fix f g h :: "('a,'b) square" show "f \ g \ f \ h \ f \ (g \ h)" proof (unfold less_eq_matrix_def, rule allI, rule prod_cases) fix i j have "(f \ g \ f \ h) (i,j) = (f \ g) (i,j) \ (f \ h) (i,j)" by (simp add: sup_matrix_def) also have "... = (\\<^sub>k f (i,k) * g (k,j)) \ (\\<^sub>k f (i,k) * h (k,j))" by (simp add: times_matrix_def) also have "... = (\\<^sub>k f (i,k) * g (k,j) \ f (i,k) * h (k,j))" by (simp add: sup_monoid.sum.distrib) also have "... = (\\<^sub>k f (i,k) * (g (k,j) \ h (k,j)))" by (simp add: mult_left_dist_sup) also have "... = (\\<^sub>k f (i,k) * (g \ h) (k,j))" by (simp add: sup_matrix_def) also have "... = (f \ (g \ h)) (i,j)" by (simp add: times_matrix_def) finally show "(f \ g \ f \ h) (i,j) \ (f \ (g \ h)) (i,j)" by simp qed next fix f g h :: "('a,'b) square" show "(f \ g) \ h = f \ h \ g \ h" proof (rule ext, rule prod_cases) fix i j have "((f \ g) \ h) (i,j) = (\\<^sub>k (f \ g) (i,k) * h (k,j))" by (simp add: times_matrix_def) also have "... = (\\<^sub>k (f (i,k) \ g (i,k)) * h (k,j))" by (simp add: sup_matrix_def) also have "... = (\\<^sub>k f (i,k) * h (k,j) \ g (i,k) * h (k,j))" by (meson mult_right_dist_sup) also have "... = (\\<^sub>k f (i,k) * h (k,j)) \ (\\<^sub>k g (i,k) * h (k,j))" by (simp add: sup_monoid.sum.distrib) also have "... = (f \ h) (i,j) \ (g \ h) (i,j)" by (simp add: times_matrix_def) also have "... = (f \ h \ g \ h) (i,j)" by (simp add: sup_matrix_def) finally show "((f \ g) \ h) (i,j) = (f \ h \ g \ h) (i,j)" . qed next fix f :: "('a,'b) square" show "mbot \ f = mbot" proof (rule ext, rule prod_cases) fix i j have "(mbot \ f) (i,j) = (\\<^sub>k mbot (i,k) * f (k,j))" by (simp add: times_matrix_def) also have "... = (\\<^sub>k bot * f (k,j))" by (simp add: bot_matrix_def) also have "... = bot" by simp also have "... = mbot (i,j)" by (simp add: bot_matrix_def) finally show "(mbot \ f) (i,j) = mbot (i,j)" . qed next fix f :: "('a,'b) square" show "mone \ f = f" by simp next fix f :: "('a,'b) square" show "f \ f \ mone" by simp next fix f g h :: "('a,'b) square" show "f \ (g \ h) = f \ g \ f \ h" proof (rule ext, rule prod_cases) fix i j have "(f \ (g \ h)) (i,j) = (\\<^sub>k f (i,k) * (g \ h) (k,j))" by (simp add: times_matrix_def) also have "... = (\\<^sub>k f (i,k) * (g (k,j) \ h (k,j)))" by (simp add: sup_matrix_def) also have "... = (\\<^sub>k f (i,k) * g (k,j) \ f (i,k) * h (k,j))" by (meson mult_left_dist_sup) also have "... = (\\<^sub>k f (i,k) * g (k,j)) \ (\\<^sub>k f (i,k) * h (k,j))" by (simp add: sup_monoid.sum.distrib) also have "... = (f \ g) (i,j) \ (f \ h) (i,j)" by (simp add: times_matrix_def) also have "... = (f \ g \ f \ h) (i,j)" by (simp add: sup_matrix_def) finally show "(f \ (g \ h)) (i,j) = (f \ g \ f \ h) (i,j)" . qed next fix f :: "('a,'b) square" show "f \ mbot = mbot" proof (rule ext, rule prod_cases) fix i j have "(f \ mbot) (i,j) = (\\<^sub>k f (i,k) * mbot (k,j))" by (simp add: times_matrix_def) also have "... = (\\<^sub>k f (i,k) * bot)" by (simp add: bot_matrix_def) also have "... = bot" by simp also have "... = mbot (i,j)" by (simp add: bot_matrix_def) finally show "(f \ mbot) (i,j) = mbot (i,j)" . qed qed interpretation matrix_bounded_idempotent_semiring: bounded_idempotent_semiring where sup = sup_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix :: ('a::finite,'b::bounded_idempotent_semiring) square" and top = top_matrix and one = one_matrix and times = times_matrix proof fix f :: "('a,'b) square" show "f \ mtop = mtop" proof fix e have "(f \ mtop) e = f e \ mtop e" by (simp add: sup_matrix_def) also have "... = f e \ top" by (simp add: top_matrix_def) also have "... = top" by simp also have "... = mtop e" by (simp add: top_matrix_def) finally show "(f \ mtop) e = mtop e" . qed qed subsection \Stone Relation Algebras\ text \ Finally, we show that matrices over Stone relation algebras form a Stone relation algebra. \ interpretation matrix_stone_relation_algebra: stone_relation_algebra where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix :: ('a::finite,'b::stone_relation_algebra) square" and top = top_matrix and uminus = uminus_matrix and one = one_matrix and times = times_matrix and conv = conv_matrix proof fix f g h :: "('a,'b) square" show "(f \ g) \ h = f \ (g \ h)" by (simp add: matrix_monoid.mult_assoc) next fix f g h :: "('a,'b) square" show "(f \ g) \ h = f \ h \ g \ h" by (simp add: matrix_idempotent_semiring.mult_right_dist_sup) next fix f :: "('a,'b) square" show "mbot \ f = mbot" by simp next fix f :: "('a,'b) square" show "mone \ f = f" by simp next fix f :: "('a,'b) square" show "f\<^sup>t\<^sup>t = f" proof (rule ext, rule prod_cases) fix i j have "(f\<^sup>t\<^sup>t) (i,j) = ((f\<^sup>t) (j,i))\<^sup>T" by (simp add: conv_matrix_def) also have "... = f (i,j)" by (simp add: conv_matrix_def) finally show "(f\<^sup>t\<^sup>t) (i,j) = f (i,j)" . qed next fix f g :: "('a,'b) square" show "(f \ g)\<^sup>t = f\<^sup>t \ g\<^sup>t" proof (rule ext, rule prod_cases) fix i j have "((f \ g)\<^sup>t) (i,j) = ((f \ g) (j,i))\<^sup>T" by (simp add: conv_matrix_def) also have "... = (f (j,i) \ g (j,i))\<^sup>T" by (simp add: sup_matrix_def) also have "... = (f\<^sup>t) (i,j) \ (g\<^sup>t) (i,j)" by (simp add: conv_matrix_def conv_dist_sup) also have "... = (f\<^sup>t \ g\<^sup>t) (i,j)" by (simp add: sup_matrix_def) finally show "((f \ g)\<^sup>t) (i,j) = (f\<^sup>t \ g\<^sup>t) (i,j)" . qed next fix f g :: "('a,'b) square" show "(f \ g)\<^sup>t = g\<^sup>t \ f\<^sup>t" proof (rule ext, rule prod_cases) fix i j have "((f \ g)\<^sup>t) (i,j) = ((f \ g) (j,i))\<^sup>T" by (simp add: conv_matrix_def) also have "... = (\\<^sub>k f (j,k) * g (k,i))\<^sup>T" by (simp add: times_matrix_def) also have "... = (\\<^sub>k (f (j,k) * g (k,i))\<^sup>T)" by (metis (no_types) conv_dist_sum) also have "... = (\\<^sub>k (g (k,i))\<^sup>T * (f (j,k))\<^sup>T)" by (simp add: conv_dist_comp) also have "... = (\\<^sub>k (g\<^sup>t) (i,k) * (f\<^sup>t) (k,j))" by (simp add: conv_matrix_def) also have "... = (g\<^sup>t \ f\<^sup>t) (i,j)" by (simp add: times_matrix_def) finally show "((f \ g)\<^sup>t) (i,j) = (g\<^sup>t \ f\<^sup>t) (i,j)" . qed next fix f g h :: "('a,'b) square" show "(f \ g) \ h \ f \ (g \ (f\<^sup>t \ h))" proof (unfold less_eq_matrix_def, rule allI, rule prod_cases) fix i j have "((f \ g) \ h) (i,j) = (f \ g) (i,j) \ h (i,j)" by (simp add: inf_matrix_def) also have "... = (\\<^sub>k f (i,k) * g (k,j)) \ h (i,j)" by (simp add: times_matrix_def) also have "... = (\\<^sub>k f (i,k) * g (k,j) \ h (i,j))" by (metis (no_types) inf_right_dist_sum) also have "... \ (\\<^sub>k f (i,k) * (g (k,j) \ (f (i,k))\<^sup>T * h (i,j)))" by (rule leq_sum, meson dedekind_1) also have "... = (\\<^sub>k f (i,k) * (g (k,j) \ (f\<^sup>t) (k,i) * h (i,j)))" by (simp add: conv_matrix_def) also have "... \ (\\<^sub>k f (i,k) * (g (k,j) \ (\\<^sub>l (f\<^sup>t) (k,l) * h (l,j))))" by (rule leq_sum, rule allI, rule comp_right_isotone, rule inf.sup_right_isotone, rule ub_sum) also have "... = (\\<^sub>k f (i,k) * (g (k,j) \ (f\<^sup>t \ h) (k,j)))" by (simp add: times_matrix_def) also have "... = (\\<^sub>k f (i,k) * (g \ (f\<^sup>t \ h)) (k,j))" by (simp add: inf_matrix_def) also have "... = (f \ (g \ (f\<^sup>t \ h))) (i,j)" by (simp add: times_matrix_def) finally show "((f \ g) \ h) (i,j) \ (f \ (g \ (f\<^sup>t \ h))) (i,j)" . qed next fix f g :: "('a,'b) square" show "\\(f \ g) = \\f \ \\g" proof (rule ext, rule prod_cases) fix i j have "(\\(f \ g)) (i,j) = --((f \ g) (i,j))" by (simp add: uminus_matrix_def) also have "... = --(\\<^sub>k f (i,k) * g (k,j))" by (simp add: times_matrix_def) also have "... = (\\<^sub>k --(f (i,k) * g (k,j)))" by (metis (no_types) pp_dist_sum) also have "... = (\\<^sub>k --(f (i,k)) * --(g (k,j)))" by (meson pp_dist_comp) also have "... = (\\<^sub>k (\\f) (i,k) * (\\g) (k,j))" by (simp add: uminus_matrix_def) also have "... = (\\f \ \\g) (i,j)" by (simp add: times_matrix_def) finally show "(\\(f \ g)) (i,j) = (\\f \ \\g) (i,j)" . qed next let ?o = "mone :: ('a,'b) square" show "\\?o = ?o" proof (rule ext, rule prod_cases) fix i j have "(\\?o) (i,j) = --(?o (i,j))" by (simp add: uminus_matrix_def) also have "... = --(if i = j then 1 else bot)" by (simp add: one_matrix_def) also have "... = (if i = j then --1 else --bot)" by simp also have "... = (if i = j then 1 else bot)" by auto also have "... = ?o (i,j)" by (simp add: one_matrix_def) finally show "(\\?o) (i,j) = ?o (i,j)" . qed qed +interpretation matrix_stone_relation_algebra_consistent: stone_relation_algebra_consistent where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix :: ('a::finite,'b::stone_relation_algebra_consistent) square" and top = top_matrix and uminus = uminus_matrix and one = one_matrix and times = times_matrix and conv = conv_matrix +proof + show "(mbot::('a,'b) square) \ mtop" + by (metis consistent bot_matrix_def top_matrix_def) +qed + +interpretation matrix_stone_relation_algebra_tarski: stone_relation_algebra_tarski where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix :: ('a::finite,'b::stone_relation_algebra_tarski) square" and top = top_matrix and uminus = uminus_matrix and one = one_matrix and times = times_matrix and conv = conv_matrix +proof + fix x :: "('a,'b) square" + assume 1: "matrix_p_algebra.regular x" + assume "x \ mbot" + from this obtain i j where "x (i,j) \ bot" + by (metis bot_matrix_def ext surj_pair) + hence 2: "top * x (i,j) * top = top" + using 1 by (metis tarski uminus_matrix_def) + show "matrix_bounded_idempotent_semiring.total (mtop \ x)" + proof (rule ext, rule prod_cases) + fix k l + have "top * x (i,j) * top \ (\\<^sub>m top * x (m,j)) * top" + using comp_inf.ub_sum comp_isotone by fastforce + also have "... = (mtop \ x) (k,j) * top" + by (simp add: times_matrix_def top_matrix_def) + also have "... \ (\\<^sub>m (mtop \ x) (k,m) * top)" + using comp_inf.ub_sum by force + also have "... = (mtop \ x \ mtop) (k,l)" + by (simp add: times_matrix_def top_matrix_def) + finally show "(mtop \ x \ mtop) (k,l) = mtop (k,l)" + using 2 by (simp add: top_matrix_def inf.bot_unique) + qed +qed + +interpretation matrix_stone_relation_algebra_tarski_consistent: stone_relation_algebra_tarski_consistent where sup = sup_matrix and inf = inf_matrix and less_eq = less_eq_matrix and less = less_matrix and bot = "bot_matrix :: ('a::finite,'b::stone_relation_algebra_tarski_consistent) square" and top = top_matrix and uminus = uminus_matrix and one = one_matrix and times = times_matrix and conv = conv_matrix + .. + end diff --git a/thys/Stone_Relation_Algebras/ROOT b/thys/Stone_Relation_Algebras/ROOT --- a/thys/Stone_Relation_Algebras/ROOT +++ b/thys/Stone_Relation_Algebras/ROOT @@ -1,14 +1,15 @@ chapter AFP session Stone_Relation_Algebras = Stone_Algebras + options [timeout = 600] theories Fixpoints Semirings Relation_Algebras Relation_Subalgebras Matrix_Relation_Algebras Linear_Order_Matrices + Choose_Component document_files "root.tex" "root.bib" diff --git a/thys/Stone_Relation_Algebras/document/root.tex b/thys/Stone_Relation_Algebras/document/root.tex --- a/thys/Stone_Relation_Algebras/document/root.tex +++ b/thys/Stone_Relation_Algebras/document/root.tex @@ -1,71 +1,73 @@ \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{$*$} \begin{document} \title{Stone Relation Algebras} \author{Walter Guttmann} \maketitle \begin{abstract} We develop Stone relation algebras, which generalise relation algebras by replacing the underlying Boolean algebra structure with a Stone algebra. We show that finite matrices over bounded linear orders form an instance. As a consequence, relation-algebraic concepts and methods can be used for reasoning about weighted graphs. We also develop a fixpoint calculus and apply it to compare different definitions of reflexive-transitive closures in semirings. \end{abstract} \tableofcontents \section{Synopsis and Motivation} This document describes the following six theory files: \begin{itemize} \item Fixpoints develops a fixpoint calculus based on partial orders. We also consider least (pre)fixpoints and greatest (post)fixpoints. The derived rules include unfold, square, rolling, fusion, exchange and diagonal rules studied in \cite{AartsBackhouseBoitenDoornbosGasterenGeldropHoogendijkVoermansWoude1995}. Our results are based on the existence of fixpoints instead of completeness of the underlying structure. \item Semirings contains a hierarchy of structures generalising idempotent semirings. In particular, several of these algebras do not assume that multiplication is associative in order to capture models such as multirelations. Even in such a weak setting we can derive several results comparing different definitions of reflexive-transitive closures based on fixpoints. \item Relation Algebras introduces Stone relation algebras, which weaken the Boolean algebra structure of relation algebras to Stone algebras. This is motivated by the wish to represent weighted graphs (matrices over numbers) in addition to unweighted graphs (Boolean matrices) that form relations. Many results of relation algebras can be derived from the weaker axioms and therefore also apply to weighted graphs. Some results hold in Stone relation algebras after small modifications. This allows us to apply relational concepts and methods also to weighted graphs. In particular, we prove a number of properties that have been used to verify graph algorithms. Tarski's relation algebras \cite{Tarski1941} arise as a special case by imposing further axioms. \item Subalgebras of Relation Algebras studies the structures of subsets of elements characterised by a given property. In particular we look at regular elements (which correspond to unweighted graphs), coreflexives (tests), vectors and covectors (which can be used to represent sets). The subsets are turned into Isabelle/HOL types, which are shown to form instances of various algebras. \item Matrix Relation Algebras lifts the Stone algebra hierarchy, the semiring structure and, finally, Stone relation algebras to finite square matrices. These are mostly standard constructions similar to those in \cite{ArmstrongFosterStruthWeber2016,ArmstrongGomesStruthWeber2016} implemented so that they work for many algebraic structures. In particular, they can be instantiated to weighted graphs (see below) and extended to Kleene algebras (not part of this development). \item Matrices over Bounded Linear Orders studies relational properties. In particular, we characterise univalent, injective, total, surjective, mapping, bijective, vector, covector, point, atom, reflexive, coreflexive, irreflexive, symmetric, antisymmetric and asymmetric matrices. Definitions of these properties are taken from relation algebras and their meaning for matrices over bounded linear orders (weighted graphs) is explained by logical formulas in terms of matrix entries. \end{itemize} +Following a refactoring, the selection of components of a graph in Stone relation algebras, which was originally part of Nicolas Robinson-O'Brien's theory \texttt{Relational\_Minimum\_Spanning\_Trees/Boruvka.thy}, has been moved into a new theory in this entry. + The development is based on a theory of Stone algebras \cite{Guttmann2016b} and forms the basis for an extension to Kleene algebras to capture further properties of graphs. We apply Stone relation algebras to verify Prim's minimum spanning tree algorithm in Isabelle/HOL in \cite{Guttmann2016c}. Related libraries for semirings and relation algebras in the Archive of Formal Proofs are \cite{ArmstrongFosterStruthWeber2016,ArmstrongGomesStruthWeber2016}. The theory \texttt{Kleene\_Algebra/Dioid.thy} introduces a number of structures that generalise idempotent semirings, but does not cover most of the semiring structures in the present development. The theory \texttt{Relation\_Algebra/Relation\_Algebra.thy} covers Tarski's relation algebras and hence cannot be reused for the present development as most properties need to be derived from the weaker axioms of Stone relation algebras. The matrix constructions in theories \texttt{Kleene\_Algebra/Inf\_Matrix.thy} and \texttt{Relation\_Algebra/Relation\_Algebra\_Models.thy} are similar, but have strong restrictions on the matrix entry types not appropriate for many algebraic structures in the present development. We also deviate from these hierarchies by basing idempotent semirings directly on the Isabelle/HOL semilattice structures instead of a separate structure; this results in a somewhat smoother integration with the lattice structure of relation algebras. \begin{flushleft} \input{session} \end{flushleft} \bibliographystyle{abbrv} \bibliography{root} \end{document}